Extensions of Gallina¶
Gallina is the kernel language of Coq. We describe here extensions of Gallina’s syntax.
Record types¶
The Record
construction is a macro allowing the definition of
records as is done in many programming languages. Its syntax is
described in the grammar below. In fact, the Record
macro is more general
than the usual record types, since it allows also for “manifest”
expressions. In this sense, the Record
construction allows defining
“signatures”.
-
Command
Record ident binders : sort? := ident? { ident binders : type*; }
¶ The first identifier
ident
is the name of the defined record andsort
is its type. The optional identifier following:=
is the name of its constructor. If it is omitted, the default nameBuild_ident
, whereident
is the record name, is used. Ifsort
is omitted, the default sort is \(\Type\). The identifiers inside the brackets are the names of fields. For a given fieldident
, its type isforall binders, type
. Remark that the type of a particular identifier may depend on a previously-given identifier. Thus the order of the fields is important. Finally,binders
are parameters of the record.
More generally, a record may have explicitly defined (a.k.a. manifest)
fields. For instance, we might have:
Record ident binders : sort := { ident1 : type1 ; ident2 := term2 ; ident3 : type3 }
.
in which case the correctness of type3
may rely on the instance term2
of ident2
and term2
may in turn depend on ident1
.
Example
The set of rational numbers may be defined as:
- Record Rat : Set := mkRat { sign : bool ; top : nat ; bottom : nat ; Rat_bottom_cond : 0 <> bottom ; Rat_irred_cond : forall x y z:nat, (x * y) = top /\ (x * z) = bottom -> x = 1 }.
- Rat is defined sign is defined top is defined bottom is defined Rat_bottom_cond is defined Rat_irred_cond is defined
Note here that the fields Rat_bottom_cond
depends on the field bottom
and Rat_irred_cond
depends on both top
and bottom
.
Let us now see the work done by the Record
macro. First the macro
generates a variant type definition with just one constructor:
Variant ident binders? : sort := ident0 binders?
.
To build an object of type ident
, one should provide the constructor
ident0
with the appropriate number of terms filling the fields of the record.
Example
Let us define the rational \(1/2\):
- Theorem one_two_irred : forall x y z:nat, x * y = 1 /\ x * z = 2 -> x = 1.
- 1 subgoal ============================ forall x y z : nat, x * y = 1 /\ x * z = 2 -> x = 1
- Admitted.
- one_two_irred is declared
- Definition half := mkRat true 1 2 (O_S 1) one_two_irred.
- half is defined
- Check half.
- half : Rat
Alternatively, the following syntax allows creating objects by using named fields, as shown in this grammar. The fields do not have to be in any particular order, nor do they have to be all present if the missing ones can be inferred or prompted for (see Program).
- Definition half' := {| sign := true; Rat_bottom_cond := O_S 1; Rat_irred_cond := one_two_irred |}.
- half' is defined
The following settings let you control the display format for types:
-
Flag
Printing Records
¶ If set, use the record syntax (shown above) as the default display format.
You can override the display format for specified types by adding entries to these tables:
-
Table
Printing Record qualid
¶ Specifies a set of qualids which are displayed as records. Use the
Add @table
andRemove @table
commands to update the set of qualids.
-
Table
Printing Constructor qualid
¶ Specifies a set of qualids which are displayed as constructors. Use the
Add @table
andRemove @table
commands to update the set of qualids.
This syntax can also be used for pattern matching.
- Eval compute in ( match half with | {| sign := true; top := n |} => n | _ => 0 end).
- = 1 : nat
The macro generates also, when it is possible, the projection
functions for destructuring an object of type ident
. These
projection functions are given the names of the corresponding
fields. If a field is named _
then no projection is built
for it. In our example:
- Eval compute in top half.
- = 1 : nat
- Eval compute in bottom half.
- = 2 : nat
- Eval compute in Rat_bottom_cond half.
- = O_S 1 : 0 <> bottom half
An alternative syntax for projections based on a dot notation is available:
- Eval compute in half.(top).
- = 1 : nat
-
Flag
Printing Projections
¶ This flag activates the dot notation for printing.
Example
- Set Printing Projections.
- Check top half.
- half.(top) : nat
The corresponding grammar rules are given in the preceding grammar. When qualid
denotes a projection, the syntax term.(qualid)
is equivalent to qualid term
,
the syntax term.(qualid arg+)
to qualid arg+ term
.
and the syntax term.(@qualid term+)
to @qualid term+ term
.
In each case, term
is the object projected and the
other arguments are the parameters of the inductive type.
Note
Records defined with the Record
keyword are not allowed to be
recursive (references to the record's name in the type of its field
raises an error). To define recursive records, one can use the Inductive
and CoInductive
keywords, resulting in an inductive or co-inductive record.
Definition of mutually inductive or co-inductive records are also allowed, as long
as all of the types in the block are records.
Note
Induction schemes are automatically generated for inductive records.
Automatic generation of induction schemes for non-recursive records
defined with the Record
keyword can be activated with the
Nonrecursive Elimination Schemes
flag (see Generation of induction principles with Scheme).
Note
Structure
is a synonym of the keyword Record
.
-
Warning
ident cannot be defined.
¶ It can happen that the definition of a projection is impossible. This message is followed by an explanation of this impossibility. There may be three reasons:
-
Error
Records declared with the keyword Record or Structure cannot be recursive.
¶ The record name
ident
appears in the type of its fields, but uses the keywordRecord
. Use the keywordInductive
orCoInductive
instead.
-
Error
Cannot handle mutually (co)inductive records.
¶ Records cannot be defined as part of mutually inductive (or co-inductive) definitions, whether with records only or mixed with standard definitions.
During the definition of the one-constructor inductive definition, all the errors of inductive definitions, as described in Section Inductive definitions, may also occur.
See also
Coercions and records in section Classes as Records of the chapter devoted to coercions.
Primitive Projections¶
-
Flag
Primitive Projections
¶ Turns on the use of primitive projections when defining subsequent records (even through the
Inductive
andCoInductive
commands). Primitive projections extended the Calculus of Inductive Constructions with a new binary term constructorr.(p)
representing a primitive projectionp
applied to a record objectr
(i.e., primitive projections are always applied). Even if the record type has parameters, these do not appear in the internal representation of applications of the projection, considerably reducing the sizes of terms when manipulating parameterized records and type checking time. On the user level, primitive projections can be used as a replacement for the usual defined ones, although there are a few notable differences.
-
Flag
Printing Primitive Projection Parameters
¶ This compatibility flag reconstructs internally omitted parameters at printing time (even though they are absent in the actual AST manipulated by the kernel).
Primitive Record Types¶
When the Primitive Projections
flag is on, definitions of
record types change meaning. When a type is declared with primitive
projections, its match
construct is disabled (see Primitive Projections though).
To eliminate the (co-)inductive type, one must use its defined primitive projections.
For compatibility, the parameters still appear to the user when
printing terms even though they are absent in the actual AST
manipulated by the kernel. This can be changed by unsetting the
Printing Primitive Projection Parameters
flag.
There are currently two ways to introduce primitive records types:
- Through the
Record
command, in which case the type has to be non-recursive. The defined type enjoys eta-conversion definitionally, that is the generalized form of surjective pairing for records:r
= Build_
R
(
r
.(
p
\(_{1}\)) …
r
.(
p
\(_{n}\)))
. Eta-conversion allows to define dependent elimination for these types as well. - Through the
Inductive
andCoInductive
commands, when the body of the definition is a record declaration of the formBuild_
R
{
p
\(_{1}\):
t
\(_{1}\); … ;
p
\(_{n}\):
t
\(_{n}\)}
. In this case the types can be recursive and eta-conversion is disallowed. These kind of record types differ from their traditional versions in the sense that dependent elimination is not available for them and only non-dependent case analysis can be defined.
Reduction¶
The basic reduction rule of a primitive projection is
p
\(_{i}\) (Build_
R
t
\(_{1}\) … t
\(_{n}\))
\({\rightarrow_{\iota}}\) t
\(_{i}\).
However, to take the \({\delta}\) flag into
account, projections can be in two states: folded or unfolded. An
unfolded primitive projection application obeys the rule above, while
the folded version delta-reduces to the unfolded version. This allows to
precisely mimic the usual unfolding rules of constants. Projections
obey the usual simpl
flags of the Arguments
command in particular.
There is currently no way to input unfolded primitive projections at the
user-level, and there is no way to display unfolded projections differently
from folded ones.
Compatibility Projections and match
¶
To ease compatibility with ordinary record types, each primitive
projection is also defined as a ordinary constant taking parameters and
an object of the record type as arguments, and whose body is an
application of the unfolded primitive projection of the same name. These
constants are used when elaborating partial applications of the
projection. One can distinguish them from applications of the primitive
projection if the Printing Primitive Projection Parameters
flag
is off: For a primitive projection application, parameters are printed
as underscores while for the compatibility projections they are printed
as usual.
Additionally, user-written match
constructs on primitive records
are desugared into substitution of the projections, they cannot be
printed back as match
constructs.
Variants and extensions of match
¶
Multiple and nested pattern matching¶
The basic version of match
allows pattern matching on simple
patterns. As an extension, multiple nested patterns or disjunction of
patterns are allowed, as in ML-like languages.
The extension just acts as a macro that is expanded during parsing
into a sequence of match on simple patterns. Especially, a
construction defined using the extended match is generally printed
under its expanded form (see Printing Matching
).
See also
Pattern-matching on boolean values: the if expression¶
For inductive types with exactly two constructors and for pattern matching
expressions that do not depend on the arguments of the constructors, it is possible
to use a if … then … else
notation. For instance, the definition
- Definition not (b:bool) := match b with | true => false | false => true end.
- not is defined
can be alternatively written
- Definition not (b:bool) := if b then false else true.
- not is defined
More generally, for an inductive type with constructors C
\(_{1}\) and C
\(_{2}\),
we have the following equivalence
if term [dep_ret_type] then term₁ else term₂ ≡
match term [dep_ret_type] with
| C₁ _ … _ => term₁
| C₂ _ … _ => term₂
end
Example
- Check (fun x (H:{x=0}+{x<>0}) => match H with | left _ => true | right _ => false end).
- fun (x : nat) (H : {x = 0} + {x <> 0}) => if H then true else false : forall x : nat, {x = 0} + {x <> 0} -> bool
Notice that the printing uses the if
syntax because sumbool
is
declared as such (see Controlling pretty-printing of match expressions).
Irrefutable patterns: the destructuring let variants¶
Pattern-matching on terms inhabiting inductive type having only one
constructor can be alternatively written using let … in …
constructions. There are two variants of them.
First destructuring let syntax¶
The expression let (
ident
\(_{1}\), … ,
ident
\(_{n}\)) :=
term
\(_{0}\)in
term
\(_{1}\) performs
case analysis on term
\(_{0}\) which must be in an inductive type with one
constructor having itself \(n\) arguments. Variables ident
\(_{1}\) … ident
\(_{n}\) are
bound to the \(n\) arguments of the constructor in expression term
\(_{1}\). For
instance, the definition
- Definition fst (A B:Set) (H:A * B) := match H with | pair x y => x end.
- fst is defined
can be alternatively written
- Definition fst (A B:Set) (p:A * B) := let (x, _) := p in x.
- fst is defined
Notice that reduction is different from regular let … in …
construction since it happens only if term
\(_{0}\) is in constructor form.
Otherwise, the reduction is blocked.
The pretty-printing of a definition by matching on a irrefutable
pattern can either be done using match
or the let
construction
(see Section Controlling pretty-printing of match expressions).
If term inhabits an inductive type with one constructor C
, we have an
equivalence between
let (ident₁, …, identₙ) [dep_ret_type] := term in term'
and
match term [dep_ret_type] with
C ident₁ … identₙ => term'
end
Second destructuring let syntax¶
Another destructuring let syntax is available for inductive types with one constructor by giving an arbitrary pattern instead of just a tuple for all the arguments. For example, the preceding example can be written:
- Definition fst (A B:Set) (p:A*B) := let 'pair x _ := p in x.
- fst is defined
This is useful to match deeper inside tuples and also to use notations
for the pattern, as the syntax let ’p := t in b
allows arbitrary
patterns to do the deconstruction. For example:
- Definition deep_tuple (A:Set) (x:(A*A)*(A*A)) : A*A*A*A := let '((a,b), (c, d)) := x in (a,b,c,d).
- deep_tuple is defined
- Notation " x 'With' p " := (exist _ x p) (at level 20).
- Identifier 'With' now a keyword
- Definition proj1_sig' (A:Set) (P:A->Prop) (t:{ x:A | P x }) : A := let 'x With p := t in x.
- proj1_sig' is defined
When printing definitions which are written using this construct it takes precedence over let printing directives for the datatype under consideration (see Section Controlling pretty-printing of match expressions).
Controlling pretty-printing of match expressions¶
The following commands give some control over the pretty-printing
of match
expressions.
Printing nested patterns¶
-
Flag
Printing Matching
¶ The Calculus of Inductive Constructions knows pattern matching only over simple patterns. It is however convenient to re-factorize nested pattern matching into a single pattern matching over a nested pattern.
When this flag is on (default), Coq’s printer tries to do such limited re-factorization. Turning it off tells Coq to print only simple pattern matching problems in the same way as the Coq kernel handles them.
Factorization of clauses with same right-hand side¶
-
Flag
Printing Factorizable Match Patterns
¶ When several patterns share the same right-hand side, it is additionally possible to share the clauses using disjunctive patterns. Assuming that the printing matching mode is on, this flag (on by default) tells Coq's printer to try to do this kind of factorization.
Use of a default clause¶
-
Flag
Printing Allow Match Default Clause
¶ When several patterns share the same right-hand side which do not depend on the arguments of the patterns, yet an extra factorization is possible: the disjunction of patterns can be replaced with a
_
default clause. Assuming that the printing matching mode and the factorization mode are on, this flag (on by default) tells Coq's printer to use a default clause when relevant.
Printing of wildcard patterns¶
-
Flag
Printing Wildcard
¶ Some variables in a pattern may not occur in the right-hand side of the pattern matching clause. When this flag is on (default), the variables having no occurrences in the right-hand side of the pattern matching clause are just printed using the wildcard symbol “_”.
Printing of the elimination predicate¶
-
Flag
Printing Synth
¶ In most of the cases, the type of the result of a matched term is mechanically synthesizable. Especially, if the result type does not depend of the matched term. When this flag is on (default), the result type is not printed when Coq knows that it can re- synthesize it.
Printing matching on irrefutable patterns¶
If an inductive type has just one constructor, pattern matching can be written using the first destructuring let syntax.
-
Table
Printing Let qualid
¶ Specifies a set of qualids for which pattern matching is displayed using a let expression. Note that this only applies to pattern matching instances entered with
match
. It doesn't affect pattern matching explicitly entered with a destructuringlet
. Use theAdd @table
andRemove @table
commands to update this set.
Printing matching on booleans¶
If an inductive type is isomorphic to the boolean type, pattern matching
can be written using if
… then
… else
…. This table controls
which types are written this way:
-
Table
Printing If qualid
¶ Specifies a set of qualids for which pattern matching is displayed using
if
…then
…else
…. Use theAdd @table
andRemove @table
commands to update this set.
This example emphasizes what the printing settings offer.
Example
- Definition snd (A B:Set) (H:A * B) := match H with | pair x y => y end.
- snd is defined
- Test Printing Let for prod.
- Cases on elements of prod are printed using a `let' form
- Print snd.
- snd = fun (A B : Set) (H : A * B) => let (_, y) := H in y : forall A B : Set, A * B -> B Arguments snd (_ _)%type_scope
- Remove Printing Let prod.
- Unset Printing Synth.
- Unset Printing Wildcard.
- Print snd.
- snd = fun (A B : Set) (H : A * B) => match H return B with | (x, y) => y end : forall A B : Set, A * B -> B Arguments snd (_ _)%type_scope
Advanced recursive functions¶
The following experimental command is available when the FunInd
library has been loaded via Require Import FunInd
:
-
Command
Function ident binder* { decrease_annot } : type := term
¶ This command can be seen as a generalization of
Fixpoint
. It is actually 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 (seefunction induction
) and its fixpoint equality. The meaning of this declaration is to define a function ident, similarly toFixpoint
. Like inFixpoint
, the decreasing argument must be given (unless the function is not recursive), but it might not necessarily be structurally decreasing. The point of the{ decrease_annot }
annotation is to name the decreasing argument and to describe which kind of decreasing criteria must be used to ensure termination of recursive calls.decrease_annot ::= struct
ident
measureterm
ident
wfterm
ident
The Function
construction also enjoys the with
extension to define
mutually recursive definitions. However, this feature does not work
for non structurally recursive functions.
See the documentation of functional induction (function induction
)
and Functional Scheme
(Generation of induction principles with Functional Scheme) for how to use
the induction principle to easily reason about the function.
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:
- Require Import FunInd.
- [Loading ML file extraction_plugin.cmxs ... done] [Loading ML file recdef_plugin.cmxs ... 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 (decreasing 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:
- 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 (decreasing 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
:
- 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 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 (List.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. Onlyident
is defined; the induction scheme will not be generated. This error happens generally when:- the definition uses pattern matching on dependent types,
which
Function
cannot deal with yet. - the definition is not a pattern matching tree as explained above.
- the definition uses pattern matching on dependent types,
which
-
Warning
Cannot define principle(s) for ident.
¶ The generation of the graph relation (
R_ident
) succeeded but the induction principle could not be built. Onlyident
is defined. Please report.
-
Warning
Cannot build functional inversion principle.
¶ functional inversion
will not be available for the function.
Depending on the {…}
annotation, different definition mechanisms are
used by Function
. A more precise description is given below.
-
Variant
Function ident binder* : type := term
Defines the not recursive function
ident
as if declared withDefinition
. Moreover the following are defined:
-
Variant
Function ident binder* { struct ident } : type := term
Defines the structural recursive function
ident
as if declared withFixpoint
. Moreover the following are defined:
-
Variant
Function ident binder* { measure term ident } : type := term
-
Variant
Function ident binder* { wf term ident } : type := term
Defines a recursive function by well-founded recursion. The module
Recdef
of the standard library must be loaded for this feature. The{}
annotation is mandatory and must be one of the following:{measure term ident }
withident
being the decreasing argument andterm
being a function from type ofident
tonat
for which value on the decreasing argument decreases (for thelt
order onnat
) at each recursive call ofterm
. Parameters of the function are bound interm
;{wf term ident }
withident
being the decreasing argument andterm
an ordering relation on the type ofident
(i.e. of typeT
\(_{\sf ident}\) →T
\(_{\sf ident}\) →Prop
) for which the decreasing argument decreases at each recursive call ofterm
. The order must be well-founded. Parameters of the function are bound interm
.
If the annotation is
measure
orfw
, 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 iswf
) a proof that the ordering relation is well-founded. Once proof obligations are discharged, the following objects are defined:- The same objects as with the struct;
- The lemma
ident
\(_{\sf tcc}\) which collects all proof obligations in one property; - The lemmas
ident
\(_{\sf terminate}\) andident
\(_{\sf F}\) which is needed to be inlined during extraction of ident.
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. Remark: Proof obligations are presented as several subgoals belonging to a Lemma
ident
\(_{\sf tcc}\).
Section mechanism¶
Sections create local contexts which can be shared across multiple definitions.
Example
Sections are opened by the Section
command, and closed by End
.
- Section s1.
Inside a section, local parameters can be introduced using Variable
,
Hypothesis
, or Context
(there are also plural variants for
the first two).
- Variables x y : nat.
- x is declared y is declared
The command Let
introduces section-wide Let-in definitions. These definitions
won't persist when the section is closed, and all persistent definitions which
depend on y'
will be prefixed with let y' := y in
.
- Let y' := y.
- y' is defined
- Definition x' := S x.
- x' is defined
- Definition x'' := x' + y'.
- x'' is defined
- Print x'.
- x' = S x : nat
- Print x''.
- x'' = x' + y' : nat
- End s1.
- Print x'.
- x' = fun x : nat => S x : nat -> nat Arguments x' _%nat_scope
- Print x''.
- x'' = fun x y : nat => let y' := y in x' x + y' : nat -> nat -> nat Arguments x'' (_ _)%nat_scope
Notice the difference between the value of x'
and x''
inside section
s1
and outside.
-
Command
Section ident
¶ This command is used to open a section named
ident
. Section names do not need to be unique.
-
Command
End ident
¶ This command closes the section named
ident
. After closing of the section, the local declarations (variables and local definitions, seeVariable
) get discharged, meaning that they stop being visible and that all global objects defined in the section are generalized with respect to the variables and local definitions they each depended on in the section.-
Error
This is not the last opened section.
¶
-
Error
Note
Most commands, like Hint
, Notation
, option management, … which
appear inside a section are canceled when the section is closed.
-
Command
Variable ident : type
¶ This command links
type
to the nameident
in the context of the current section. When the current section is closed, nameident
will be unknown and every object using this variable will be explicitly parameterized (the variable is discharged).
-
Command
Let ident := term
¶ This command binds the value
term
to the nameident
in the environment of the current section. The nameident
is accessible only within the current section. When the section is closed, all persistent definitions and theorems within it and depending onident
will be prefixed by the let-in definitionlet ident := term in
.-
Variant
Let CoFixpoint ident cofix_body with cofix_body*
¶
-
Variant
-
Command
Context binders
¶ Declare variables in the context of the current section, like
Variable
, but also allowing implicit variables, Implicit generalization, and let-binders.Context {A : Type} (a b : A). Context `{EqDec A}. Context (b' := b).
See also
Section Binders. Section Sections and contexts in chapter Typeclasses.
Module system¶
The module system provides a way of packaging related elements together, as well as a means of massive abstraction.
module_type ::=qualid
module_type
with Definitionqualid
:=term
module_type
with Modulequalid
:=qualid
qualid
qualid
…qualid
!qualid
qualid
…qualid
module_binding ::= ( [Import|Export]ident
…ident
:module_type
) module_bindings ::=module_binding
…module_binding
module_expression ::=qualid
…qualid
!qualid
…qualid
Syntax of modules
In the syntax of module application, the ! prefix indicates that any
Inline
directive in the type of the functor arguments will be ignored
(see the Module Type
command below).
-
Variant
Module ident module_binding*
Starts an interactive functor with parameters given by module_bindings.
-
Variant
Module ident : module_type
Starts an interactive module specifying its module type.
-
Variant
Module ident module_binding* : module_type
Starts an interactive functor with parameters given by the list of
module_bindings
, and output module typemodule_type
.
-
Variant
Module ident <: module_type+<:
- Starts an interactive module satisfying each
module_type
.-
Variant
Module ident module_binding* <: module_type+<:.
Starts an interactive functor with parameters given by the list of
module_binding
. The output module type is verified against eachmodule_type
.
-
Variant
-
Variant
Module ImportExport
Behaves like
Module
, but automatically imports or exports the module.
Reserved commands inside an interactive module¶
-
Command
Include module
¶ Includes the content of module in the current interactive module. Here module can be a module expression or a module type expression. If module is a high-order module or module type expression then the system tries to instantiate module by the current interactive module.
-
Command
Include module+<+
is a shortcut for the commands
Include module
for eachmodule
.
-
Command
End ident
This command closes the interactive module
ident
. If the module type was given the content of the module is matched against it and an error is signaled if the matching fails. If the module is basic (is not a functor) its components (constants, inductive types, submodules etc.) are now available through the dot notation.
-
Command
Module ident := module_expression
This command defines the module identifier
ident
to be equal tomodule_expression
.-
Variant
Module ident module_binding* := module_expression
Defines a functor with parameters given by the list of
module_binding
and bodymodule_expression
.
-
Variant
Module ident module_binding* : module_type := module_expression
Defines a functor with parameters given by the list of
module_binding
(possibly none), and output module typemodule_type
, with bodymodule_expression
.
-
Variant
Module ident module_binding* <: module_type+<: := module_expression
Defines a functor with parameters given by module_bindings (possibly none) with body
module_expression
. The body is checked against eachmodule_typei
.
-
Variant
Module ident module_binding* := module_expression+<+
is equivalent to an interactive module where each
module_expression
is included.
-
Variant
-
Command
Module Type ident
¶ This command is used to start an interactive module type
ident
.-
Variant
Module Type ident module_binding*
Starts an interactive functor type with parameters given by
module_bindings
.
-
Variant
Reserved commands inside an interactive module type:¶
-
Command
Include module
Same as
Include
inside a module.
-
Command
Include module+<+
This is a shortcut for the command
Include module
for eachmodule
.
-
Command
assumption_keyword Inline assums
¶ The instance of this assumption will be automatically expanded at functor application, except when this functor application is prefixed by a
!
annotation.
-
Command
End ident
This command closes the interactive module type
ident
.-
Error
This is not the last opened module type.
¶
-
Error
-
Command
Module Type ident := module_type
Defines a module type
ident
equal tomodule_type
.-
Variant
Module Type ident module_binding* := module_type
Defines a functor type
ident
specifying functors taking argumentsmodule_bindings
and returningmodule_type
.
-
Variant
Module Type ident module_binding* := module_type+<+
is equivalent to an interactive module type were each
module_type
is included.
-
Variant
-
Command
Declare Module ident : module_type
¶ Declares a module
ident
of typemodule_type
.-
Variant
Declare Module ident module_binding* : module_type
Declares a functor with parameters given by the list of
module_binding
and output module typemodule_type
.
-
Variant
Example
Let us define a simple module.
- Module M.
- Interactive Module M started
- Definition T := nat.
- T is defined
- Definition x := 0.
- x is defined
- Definition y : bool.
- 1 subgoal ============================ bool
- exact true.
- No more subgoals.
- Defined.
- End M.
- Module M is defined
Inside a module one can define constants, prove theorems and do any other things that can be done in the toplevel. Components of a closed module can be accessed using the dot notation:
- Print M.x.
- M.x = 0 : nat
A simple module type:
- Module Type SIG.
- Interactive Module Type SIG started
- Parameter T : Set.
- T is declared
- Parameter x : T.
- x is declared
- End SIG.
- Module Type SIG is defined
Now we can create a new module from M, giving it a less precise specification: the y component is dropped as well as the body of x.
- Module N : SIG with Definition T := nat := M.
- Module N is defined
- Print N.T.
- N.T = nat : Set
- Print N.x.
- *** [ N.x : N.T ]
- Fail Print N.y.
- The command has indeed failed with message: N.y not a defined object.
- Module M.
- Interactive Module M started
- Definition T := nat.
- T is defined
- Definition x := 0.
- x is defined
- Definition y : bool.
- 1 subgoal ============================ bool
- exact true.
- No more subgoals.
- Defined.
- End M.
- Module M is defined
- Module Type SIG.
- Interactive Module Type SIG started
- Parameter T : Set.
- T is declared
- Parameter x : T.
- x is declared
- End SIG.
- Module Type SIG is defined
The definition of N
using the module type expression SIG
with
Definition T := nat
is equivalent to the following one:
- Module Type SIG'.
- Interactive Module Type SIG' started
- Definition T : Set := nat.
- T is defined
- Parameter x : T.
- x is declared
- End SIG'.
- Module Type SIG' is defined
- Module N : SIG' := M.
- Module N is defined
If we just want to be sure that our implementation satisfies a given module type without restricting the interface, we can use a transparent constraint
- Module P <: SIG := M.
- Module P is defined
- Print P.y.
- P.y = true : bool
Now let us create a functor, i.e. a parametric module
- Module Two (X Y: SIG).
- Interactive Module Two started
- Definition T := (X.T * Y.T)%type.
- T is defined
- Definition x := (X.x, Y.x).
- x is defined
- End Two.
- Module Two is defined
and apply it to our modules and do some computations:
- Module Q := Two M N.
- Module Q is defined
- Eval compute in (fst Q.x + snd Q.x).
- = N.x : nat
In the end, let us define a module type with two sub-modules, sharing some of the fields and give one of its possible implementations:
- Module Type SIG2.
- Interactive Module Type SIG2 started
- Declare Module M1 : SIG.
- Module M1 is declared
- Module M2 <: SIG.
- Interactive Module M2 started
- Definition T := M1.T.
- T is defined
- Parameter x : T.
- x is declared
- End M2.
- Module M2 is defined
- End SIG2.
- Module Type SIG2 is defined
- Module Mod <: SIG2.
- Interactive Module Mod started
- Module M1.
- Interactive Module M1 started
- Definition T := nat.
- T is defined
- Definition x := 1.
- x is defined
- End M1.
- Module M1 is defined
- Module M2 := M.
- Module M2 is defined
- End Mod.
- Module Mod is defined
Notice that M
is a correct body for the component M2
since its T
component is equal nat
and hence M1.T
as specified.
Note
- Modules and module types can be nested components of each other.
- One can have sections inside a module or a module type, but not a module or a module type inside a section.
- Commands like
Hint
orNotation
can also appear inside modules and module types. Note that in case of a module definition like:
Module N : SIG := M.
or:
Module N : SIG. … End N.
hints and the like valid for N
are not those defined in M
(or the module body) but the ones defined in SIG
.
-
Command
Import qualid
¶ If
qualid
denotes a valid basic module (i.e. its module type is a signature), makes its components available by their short names.Example
- Module Mod.
- Interactive Module Mod started
- Definition T:=nat.
- T is defined
- Check T.
- T : Set
- End Mod.
- Module Mod is defined
- Check Mod.T.
- Mod.T : Set
- Fail Check T.
- The command has indeed failed with message: The reference T was not found in the current environment.
- Import Mod.
- Check T.
- T : Set
Some features defined in modules are activated only when a module is imported. This is for instance the case of notations (see Notations).
Declarations made with the
Local
flag are never imported by theImport
command. Such declarations are only accessible through their fully qualified name.Example
- Module A.
- Interactive Module A started
- Module B.
- Interactive Module B started
- Local Definition T := nat.
- T is defined
- End B.
- Module B is defined
- End A.
- Module A is defined
- Import A.
- Fail Check B.T.
- The command has indeed failed with message: The reference B.T was not found in the current environment.
-
Flag
Short Module Printing
¶ This flag (off by default) disables the printing of the types of fields, leaving only their names, for the commands
Print Module
andPrint Module Type
.
Libraries and qualified names¶
Names of libraries¶
The theories developed in Coq are stored in library files which are
hierarchically classified into libraries and sublibraries. To
express this hierarchy, library names are represented by qualified
identifiers qualid, i.e. as list of identifiers separated by dots (see
Qualified identifiers and simple identifiers). For instance, the library file Mult
of the standard
Coq library Arith
is named Coq.Arith.Mult
. The identifier that starts
the name of a library is called a library root. All library files of
the standard library of Coq have the reserved root Coq but library
filenames based on other roots can be obtained by using Coq commands
(coqc, coqtop, coqdep, …) options -Q
or -R
(see By command line options).
Also, when an interactive Coq session starts, a library of root Top
is
started, unless option -top
or -notop
is set (see By command line options).
Qualified names¶
Library files are modules which possibly contain submodules which
eventually contain constructions (axioms, parameters, definitions,
lemmas, theorems, remarks or facts). The absolute name, or full
name, of a construction in some library file is a qualified
identifier starting with the logical name of the library file,
followed by the sequence of submodules names encapsulating the
construction and ended by the proper name of the construction.
Typically, the absolute name Coq.Init.Logic.eq
denotes Leibniz’
equality defined in the module Logic in the sublibrary Init
of the
standard library of Coq.
The proper name that ends the name of a construction is the short name
(or sometimes base name) of the construction (for instance, the short
name of Coq.Init.Logic.eq
is eq
). Any partial suffix of the absolute
name is a partially qualified name (e.g. Logic.eq
is a partially
qualified name for Coq.Init.Logic.eq
). Especially, the short name of a
construction is its shortest partially qualified name.
Coq does not accept two constructions (definition, theorem, …) with the same absolute name but different constructions can have the same short name (or even same partially qualified names as soon as the full names are different).
Notice that the notion of absolute, partially qualified and short names also applies to library filenames.
Visibility
Coq maintains a table called the name table which maps partially qualified
names of constructions to absolute names. This table is updated by the
commands Require
, Import
and Export
and
also each time a new declaration is added to the context. An absolute
name is called visible from a given short or partially qualified name
when this latter name is enough to denote it. This means that the
short or partially qualified name is mapped to the absolute name in
Coq name table. Definitions flagged as Local are only accessible with
their fully qualified name (see Definitions).
It may happen that a visible name is hidden by the short name or a qualified name of another construction. In this case, the name that has been hidden must be referred to using one more level of qualification. To ensure that a construction always remains accessible, absolute names can never be hidden.
Example
- Check 0.
- 0 : nat
- Definition nat := bool.
- nat is defined
- Check 0.
- 0 : Datatypes.nat
- Check Datatypes.nat.
- Datatypes.nat : Set
- Locate nat.
- Constant Top.nat Inductive Coq.Init.Datatypes.nat (shorter name to refer to it in current context is Datatypes.nat)
See also
Commands Locate
and Locate Library
.
Libraries and filesystem¶
Note
The questions described here have been subject to redesign in Coq 8.5. Former versions of Coq use the same terminology to describe slightly different things.
Compiled files (.vo
and .vio
) store sub-libraries. In order to refer
to them inside Coq, a translation from file-system names to Coq names
is needed. In this translation, names in the file system are called
physical paths while Coq names are contrastingly called logical
names.
A logical prefix Lib can be associated with a physical path using
the command line option -Q
path
Lib
. All subfolders of path are
recursively associated to the logical path Lib
extended with the
corresponding suffix coming from the physical path. For instance, the
folder path/fOO/Bar
maps to Lib.fOO.Bar
. Subdirectories corresponding
to invalid Coq identifiers are skipped, and, by convention,
subdirectories named CVS
or _darcs
are skipped too.
Thanks to this mechanism, .vo
files are made available through the
logical name of the folder they are in, extended with their own
basename. For example, the name associated to the file
path/fOO/Bar/File.vo
is Lib.fOO.Bar.File
. The same caveat applies for
invalid identifiers. When compiling a source file, the .vo
file stores
its logical name, so that an error is issued if it is loaded with the
wrong loadpath afterwards.
Some folders have a special status and are automatically put in the
path. Coq commands associate automatically a logical path to files in
the repository trees rooted at the directory from where the command is
launched, coqlib/user-contrib/
, the directories listed in the
$COQPATH
, ${XDG_DATA_HOME}/coq/
and ${XDG_DATA_DIRS}/coq/
environment variables (see XDG base directory specification)
with the same physical-to-logical translation and with an empty logical prefix.
The command line option -R
is a variant of -Q
which has the strictly
same behavior regarding loadpaths, but which also makes the
corresponding .vo
files available through their short names in a way
not unlike the Import
command (see here). For instance, -R path Lib
associates to the file /path/fOO/Bar/File.vo
the logical name
Lib.fOO.Bar.File
, but allows this file to be accessed through the
short names fOO.Bar.File,Bar.File
and File
. If several files with
identical base name are present in different subdirectories of a
recursive loadpath, which of these files is found first may be system-
dependent and explicit qualification is recommended. The From
argument
of the Require
command can be used to bypass the implicit shortening
by providing an absolute root to the required file (see Compiled files).
There also exists another independent loadpath mechanism attached to
OCaml object files (.cmo
or .cmxs
) rather than Coq object
files as described above. The OCaml loadpath is managed using
the option -I
path
(in the OCaml world, there is neither a
notion of logical name prefix nor a way to access files in
subdirectories of path). See the command Declare ML Module
in
Compiled files to understand the need of the OCaml loadpath.
See By command line options for a more general view over the Coq command line options.
Implicit arguments¶
An implicit argument of a function is an argument which can be inferred from contextual knowledge. There are different kinds of implicit arguments that can be considered implicit in different ways. There are also various commands to control the setting or the inference of implicit arguments.
The different kinds of implicit arguments¶
Implicit arguments inferable from the knowledge of other arguments of a function¶
The first kind of implicit arguments covers the arguments that are inferable from the knowledge of the type of other arguments of the function, or of the type of the surrounding context of the application. Especially, such implicit arguments correspond to parameters dependent in the type of the function. Typical implicit arguments are the type arguments in polymorphic functions. There are several kinds of such implicit arguments.
Strict Implicit Arguments
An implicit argument can be either strict or non strict. An implicit argument is said to be strict if, whatever the other arguments of the function are, it is still inferable from the type of some other argument. Technically, an implicit argument is strict if it corresponds to a parameter which is not applied to a variable which itself is another parameter of the function (since this parameter may erase its arguments), not in the body of a match, and not itself applied or matched against patterns (since the original form of the argument can be lost by reduction).
For instance, the first argument of
cons: forall A:Set, A -> list A -> list A
in module List.v
is strict because list
is an inductive type and A
will always be inferable from the type list A
of the third argument of
cons
. Also, the first argument of cons
is strict with respect to the second one,
since the first argument is exactly the type of the second argument.
On the contrary, the second argument of a term of type
forall P:nat->Prop, forall n:nat, P n -> ex nat P
is implicit but not strict, since it can only be inferred from the
type P n
of the third argument and if P
is, e.g., fun _ => True
, it
reduces to an expression where n
does not occur any longer. The first
argument P
is implicit but not strict either because it can only be
inferred from P n
and P
is not canonically inferable from an arbitrary
n
and the normal form of P n
. Consider, e.g., that n
is \(0\) and the third
argument has type True
, then any P
of the form
fun n => match n with 0 => True | _ => anything end
would be a solution of the inference problem.
Contextual Implicit Arguments
An implicit argument can be contextual or not. An implicit argument is said contextual if it can be inferred only from the knowledge of the type of the context of the current expression. For instance, the only argument of:
nil : forall A:Set, list A`
is contextual. Similarly, both arguments of a term of type:
forall P:nat->Prop, forall n:nat, P n \/ n = 0
are contextual (moreover, n
is strict and P
is not).
Reversible-Pattern Implicit Arguments
There is another class of implicit arguments that can be reinferred
unambiguously if all the types of the remaining arguments are known.
This is the class of implicit arguments occurring in the type of
another argument in position of reversible pattern, which means it is
at the head of an application but applied only to uninstantiated
distinct variables. Such an implicit argument is called reversible-
pattern implicit argument. A typical example is the argument P
of
nat_rec in
nat_rec : forall P : nat -> Set, P 0 ->
(forall n : nat, P n -> P (S n)) -> forall x : nat, P x
(P
is reinferable by abstracting over n
in the type P n
).
See Controlling reversible-pattern implicit arguments for the automatic declaration of reversible-pattern implicit arguments.
Implicit arguments inferable by resolution¶
This corresponds to a class of non-dependent implicit arguments that are solved based on the structure of their type only.
Maximal or non maximal insertion of implicit arguments¶
In case a function is partially applied, and the next argument to be applied is an implicit argument, two disciplines are applicable. In the first case, the function is considered to have no arguments furtherly: one says that the implicit argument is not maximally inserted. In the second case, the function is considered to be implicitly applied to the implicit arguments it is waiting for: one says that the implicit argument is maximally inserted.
Each implicit argument can be declared to have to be inserted maximally or non
maximally. This can be governed argument per argument by the command
Arguments (implicits)
or globally by the Maximal Implicit Insertion
flag.
Casual use of implicit arguments¶
In a given expression, if it is clear that some argument of a function can be inferred from the type of the other arguments, the user can force the given argument to be guessed by replacing it by “_”. If possible, the correct argument will be automatically generated.
-
Error
Cannot infer a term for this placeholder.
¶ Coq was not able to deduce an instantiation of a “_”.
Declaration of implicit arguments¶
In case one wants that some arguments of a given object (constant, inductive types, constructors, assumptions, local or not) are always inferred by Coq, one may declare once and for all which are the expected implicit arguments of this object. There are two ways to do this, a priori and a posteriori.
Implicit Argument Binders¶
In the first setting, one wants to explicitly give the implicit arguments of a declared object as part of its definition. To do this, one has to surround the bindings of implicit arguments by curly braces:
- Definition id {A : Type} (x : A) : A := x.
- id is defined
This automatically declares the argument A of id as a maximally inserted implicit argument. One can then do as-if the argument was absent in every situation but still be able to specify it if needed:
- Definition compose {A B C} (g : B -> C) (f : A -> B) := fun x => g (f x).
- compose is defined
- Goal forall A, compose id id = id (A:=A).
- 1 subgoal ============================ forall A : Type, compose id id = id
The syntax is supported in all top-level definitions:
Definition
, Fixpoint
, Lemma
and so on. For (co-)inductive datatype
declarations, the semantics are the following: an inductive parameter
declared as an implicit argument need not be repeated in the inductive
definition and will become implicit for the inductive type and the constructors.
For example:
- Inductive list {A : Type} : Type := | nil : list | cons : A -> list -> list.
- list is defined list_rect is defined list_ind is defined list_rec is defined list_sind is defined
- Print list.
- Inductive list (A : Type) : Type := nil : list | cons : A -> list -> list Arguments list {A}%type_scope Arguments nil {A}%type_scope Arguments cons {A}%type_scope
One can always specify the parameter if it is not uniform using the usual implicit arguments disambiguation syntax.
Declaring Implicit Arguments¶
-
Command
Arguments qualid [ ident ]{ ident }ident*
¶ This command is used to set implicit arguments a posteriori, where the list of possibly bracketed
ident
is a prefix of the list of arguments ofqualid
where the ones to be declared implicit are surrounded by square brackets and the ones to be declared as maximally inserted implicits are surrounded by curly braces.After the above declaration is issued, implicit arguments can just (and have to) be skipped in any expression involving an application of
qualid
.
-
Variant
Global Arguments qualid [ ident ]{ ident }ident*
This command is used to recompute the implicit arguments of
qualid
after ending of the current section if any, enforcing the implicit arguments known from inside the section to be the ones declared by the command.
-
Variant
Local Arguments qualid [ ident ]{ ident }ident*
When in a module, tell not to activate the implicit arguments of
qualid
declared by this command to contexts that require the module.
-
Variant
GlobalLocal? Arguments qualid [ ident ]{ ident }ident+*,
For names of constants, inductive types, constructors, lemmas which can only be applied to a fixed number of arguments (this excludes for instance constants whose type is polymorphic), multiple implicit arguments declarations can be given. Depending on the number of arguments qualid is applied to in practice, the longest applicable list of implicit arguments is used to select which implicit arguments are inserted. For printing, the omitted arguments are the ones of the longest list of implicit arguments of the sequence.
Example
- Inductive list (A : Type) : Type := | nil : list A | cons : A -> list A -> list A.
- list is defined list_rect is defined list_ind is defined list_rec is defined list_sind is defined
- Check (cons nat 3 (nil nat)).
- cons nat 3 (nil nat) : list nat
- Arguments cons [A] _ _.
- Arguments nil {A}.
- Check (cons 3 nil).
- cons 3 nil : list nat
- Fixpoint map (A B : Type) (f : A -> B) (l : list A) : list B := match l with nil => nil | cons a t => cons (f a) (map A B f t) end.
- map is defined map is recursively defined (decreasing on 4th argument)
- Fixpoint length (A : Type) (l : list A) : nat := match l with nil => 0 | cons _ m => S (length A m) end.
- length is defined length is recursively defined (decreasing on 2nd argument)
- Arguments map [A B] f l.
- Arguments length {A} l.
- Check (fun l:list (list nat) => map length l).
- fun l : list (list nat) => map length l : list (list nat) -> list nat
- Arguments map [A B] f l, [A] B f l, A B f l.
- Check (fun l => map length l = map (list nat) nat length l).
- fun l : list (list nat) => map length l = map length l : list (list nat) -> Prop
Note
To know which are the implicit arguments of an object, use the
command Print Implicit
(see Displaying what the implicit arguments are).
Automatic declaration of implicit arguments¶
-
Command
Arguments qualid : default implicits
¶ This command tells Coq to automatically detect what are the implicit arguments of a defined object.
The auto-detection is governed by flags telling if strict, contextual, or reversible-pattern implicit arguments must be considered or not (see Controlling strict implicit arguments, Controlling strict implicit arguments, Controlling reversible-pattern implicit arguments, and also Controlling the insertion of implicit arguments not followed by explicit arguments).
-
Variant
Global Arguments qualid : default implicits
Tell to recompute the implicit arguments of qualid after ending of the current section if any.
-
Variant
Example
- Inductive list (A:Set) : Set := | nil : list A | cons : A -> list A -> list A.
- list is defined list_rect is defined list_ind is defined list_rec is defined list_sind is defined
- Arguments cons : default implicits.
- Print Implicit cons.
- cons : forall A : Set, A -> list A -> list A Argument A is implicit
- Arguments nil : default implicits.
- Print Implicit nil.
- nil : forall A : Set, list A
- Set Contextual Implicit.
- Arguments nil : default implicits.
- Print Implicit nil.
- nil : forall A : Set, list A Argument A is implicit and maximally inserted
The computation of implicit arguments takes account of the unfolding
of constants. For instance, the variable p
below has type
(Transitivity R)
which is reducible to
forall x,y:U, R x y -> forall z:U, R y z -> R x z
. As the variables x
, y
and z
appear strictly in the body of the type, they are implicit.
- Parameter X : Type.
- X is declared
- Definition Relation := X -> X -> Prop.
- Relation is defined
- Definition Transitivity (R:Relation) := forall x y:X, R x y -> forall z:X, R y z -> R x z.
- Transitivity is defined
- Parameters (R : Relation) (p : Transitivity R).
- R is declared p is declared
- Arguments p : default implicits.
- Print p.
- *** [ p : Transitivity R ] Expanded type for implicit arguments p : forall x y : X, R x y -> forall z : X, R y z -> R x z Arguments p [x y] _ [z]
- Print Implicit p.
- p : forall x y : X, R x y -> forall z : X, R y z -> R x z Arguments x, y, z are implicit
- Parameters (a b c : X) (r1 : R a b) (r2 : R b c).
- a is declared b is declared c is declared r1 is declared r2 is declared
- Check (p r1 r2).
- p r1 r2 : R a c
Mode for automatic declaration of implicit arguments¶
-
Flag
Implicit Arguments
¶ This flag (off by default) allows to systematically declare implicit the arguments detectable as such. Auto-detection of implicit arguments is governed by flags controlling whether strict and contextual implicit arguments have to be considered or not.
Controlling strict implicit arguments¶
-
Flag
Strict Implicit
¶ When the mode for automatic declaration of implicit arguments is on, the default is to automatically set implicit only the strict implicit arguments plus, for historical reasons, a small subset of the non-strict implicit arguments. To relax this constraint and to set implicit all non strict implicit arguments by default, you can turn this flag off.
-
Flag
Strongly Strict Implicit
¶ Use this flag (off by default) to capture exactly the strict implicit arguments and no more than the strict implicit arguments.
Controlling contextual implicit arguments¶
-
Flag
Contextual Implicit
¶ By default, Coq does not automatically set implicit the contextual implicit arguments. You can turn this flag on to tell Coq to also infer contextual implicit argument.
Controlling reversible-pattern implicit arguments¶
-
Flag
Reversible Pattern Implicit
¶ By default, Coq does not automatically set implicit the reversible-pattern implicit arguments. You can turn this flag on to tell Coq to also infer reversible-pattern implicit argument.
Controlling the insertion of implicit arguments not followed by explicit arguments¶
-
Flag
Maximal Implicit Insertion
¶ Assuming the implicit argument mode is on, this flag (off by default) declares implicit arguments to be automatically inserted when a function is partially applied and the next argument of the function is an implicit one.
Explicit applications¶
In presence of non-strict or contextual argument, or in presence of
partial applications, the synthesis of implicit arguments may fail, so
one may have to give explicitly certain implicit arguments of an
application. The syntax for this is (ident := term)
where ident
is the
name of the implicit argument and term is its corresponding explicit
term. Alternatively, one can locally deactivate the hiding of implicit
arguments of a function by using the notation qualid term+
.
This syntax extension is given in the following grammar:
Example: (continued)
- Check (p r1 (z:=c)).
- p r1 (z:=c) : R b c -> R a c
- Check (p (x:=a) (y:=b) r1 (z:=c) r2).
- p r1 r2 : R a c
Renaming implicit arguments¶
-
Command
Arguments qualid name* : rename
¶ This command is used to redefine the names of implicit arguments.
-
Command
Arguments qualid name* : assert
¶ This command is used to assert that a given object has the expected number of arguments and that these arguments are named as expected.
Example: (continued)
- Arguments p [s t] _ [u] _: rename.
- Check (p r1 (u:=c)).
- p r1 (u:=c) : R b c -> R a c
- Check (p (s:=a) (t:=b) r1 (u:=c) r2).
- p r1 r2 : R a c
- Fail Arguments p [s t] _ [w] _ : assert.
- The command has indeed failed with message: Flag "rename" expected to rename u into w.
Displaying what the implicit arguments are¶
Explicit displaying of implicit arguments for pretty-printing¶
-
Flag
Printing Implicit
¶ By default, the basic pretty-printing rules hide the inferable implicit arguments of an application. Turn this flag on to force printing all implicit arguments.
-
Flag
Printing Implicit Defensive
¶ By default, the basic pretty-printing rules display the implicit arguments that are not detected as strict implicit arguments. This “defensive” mode can quickly make the display cumbersome so this can be deactivated by turning this flag off.
See also
Interaction with subtyping¶
When an implicit argument can be inferred from the type of more than one of the other arguments, then only the type of the first of these arguments is taken into account, and not an upper type of all of them. As a consequence, the inference of the implicit argument of “=” fails in
- Fail Check nat = Prop.
- The command has indeed failed with message: The term "Prop" has type "Type" while it is expected to have type "Set" (universe inconsistency).
but succeeds in
- Check Prop = nat.
- Prop = nat : Prop
Deactivation of implicit arguments for parsing¶
-
Flag
Parsing Explicit
¶ Turning this flag on (it is off by default) deactivates the use of implicit arguments.
In this case, all arguments of constants, inductive types, constructors, etc, including the arguments declared as implicit, have to be given as if no arguments were implicit. By symmetry, this also affects printing.
Canonical structures¶
A canonical structure is an instance of a record/structure type that can be used to solve unification problems involving a projection applied to an unknown structure instance (an implicit argument) and a value. The complete documentation of canonical structures can be found in Canonical Structures; here only a simple example is given.
-
Command
Canonical Structure? qualid
¶ This command declares
qualid
as a canonical instance of a structure (a record).Assume that
qualid
denotes an object(Build_struct
c
\(_{1}\) …c
\(_{n}\))
in the structurestruct
of which the fields arex
\(_{1}\), …,x
\(_{n}\). Then, each time an equation of the form(
x
\(_{i}\)_)
=
\(_{\beta\delta\iota\zeta}\)c
\(_{i}\) has to be solved during the type checking process,qualid
is used as a solution. Otherwise said,qualid
is canonically used to extend the fieldc
\(_{i}\) into a complete structure built onc
\(_{i}\).Canonical structures are particularly useful when mixed with coercions and strict implicit arguments.
Example
Here is an example.
- Require Import Relations.
- Require Import EqNat.
- Set Implicit Arguments.
- Unset Strict Implicit.
- Structure Setoid : Type := {Carrier :> Set; Equal : relation Carrier; Prf_equiv : equivalence Carrier Equal}.
- Setoid is defined Carrier is defined Equal is defined Prf_equiv is defined
- Definition is_law (A B:Setoid) (f:A -> B) := forall x y:A, Equal x y -> Equal (f x) (f y).
- is_law is defined
- Axiom eq_nat_equiv : equivalence nat eq_nat.
- eq_nat_equiv is declared
- Definition nat_setoid : Setoid := Build_Setoid eq_nat_equiv.
- nat_setoid is defined
- Canonical nat_setoid.
Thanks to
nat_setoid
declared as canonical, the implicit argumentsA
andB
can be synthesized in the next statement.- Lemma is_law_S : is_law S.
- 1 subgoal ============================ is_law (A:=nat_setoid) (B:=nat_setoid) S
Note
If a same field occurs in several canonical structures, then only the structure declared first as canonical is considered.
Note
To prevent a field from being involved in the inference of canonical instances, its declaration can be annotated with the
#[canonical(false)]
attribute.Example
For instance, when declaring the
Setoid
structure above, thePrf_equiv
field declaration could be written as follows.#[canonical(false)] Prf_equiv : equivalence Carrier EqualSee Canonical Structures for a more realistic example.
-
Command
Print Canonical Projections
¶ This displays the list of global names that are components of some canonical structure. For each of them, the canonical structure of which it is a projection is indicated.
Example
For instance, the above example gives the following output:
- Print Canonical Projections.
- nat <- Carrier ( nat_setoid ) eq_nat <- Equal ( nat_setoid ) eq_nat_equiv <- Prf_equiv ( nat_setoid )
Note
The last line would not show up if the corresponding projection (namely
Prf_equiv
) were annotated as not canonical, as described above.
Implicit types of variables¶
It is possible to bind variable names to a given type (e.g. in a
development using arithmetic, it may be convenient to bind the names n
or m
to the type nat
of natural numbers).
-
Command
Implicit Types ident+ : type
¶ The effect of the command is to automatically set the type of bound variables starting with
ident
(eitherident
itself orident
followed by one or more single quotes, underscore or digits) to betype
(unless the bound variable is already declared with an explicit type in which case, this latter type is considered).
Example
- Require Import List.
- Implicit Types m n : nat.
- Lemma cons_inj_nat : forall m n l, n :: l = m :: l -> n = m.
- 1 subgoal ============================ forall (m n : nat) (l : Datatypes.list nat), n :: l = m :: l -> n = m
- Proof.
- intros m n.
- 1 subgoal m, n : nat ============================ forall l : Datatypes.list nat, n :: l = m :: l -> n = m
- Abort.
- Lemma cons_inj_bool : forall (m n:bool) l, n :: l = m :: l -> n = m.
- 1 subgoal ============================ forall (m n : bool) (l : Datatypes.list bool), n :: l = m :: l -> n = m
- Abort.
Implicit generalization¶
Implicit generalization is an automatic elaboration of a statement with free variables into a closed statement where these variables are quantified explicitly.
It is activated for a binder by prefixing a `, and for terms by surrounding it with `{ } or `( ).
Terms surrounded by `{ } introduce their free variables as maximally inserted implicit arguments, and terms surrounded by `( ) introduce them as explicit arguments.
Generalizing binders always introduce their free variables as maximally inserted implicit arguments. The binder itself introduces its argument as usual.
In the following statement, A
and y
are automatically
generalized, A
is implicit and x
, y
and the anonymous
equality argument are explicit.
- Generalizable All Variables.
- Definition sym `(x:A) : `(x = y -> y = x) := fun _ p => eq_sym p.
- sym is defined
- Print sym.
- sym = fun (A : Type) (x y : A) (p : x = y) => eq_sym p : forall (A : Type) (x y : A), x = y -> y = x Arguments sym {A}%type_scope
Dually to normal binders, the name is optional but the type is required:
- Check (forall `{x = y :> A}, y = x).
- forall (A : Type) (x y : A), x = y -> y = x : Prop
When generalizing a binder whose type is a typeclass, its own class
arguments are omitted from the syntax and are generalized using
automatic names, without instance search. Other arguments are also
generalized unless provided. This produces a fully general statement.
this behaviour may be disabled by prefixing the type with a !
or
by forcing the typeclass name to be an explicit application using
@
(however the later ignores implicit argument information).
- Class Op (A:Type) := op : A -> A -> A.
- Class Commutative (A:Type) `(Op A) := commutative : forall x y, op x y = op y x.
- Instance nat_op : Op nat := plus.
- nat_op is defined
- Set Printing Implicit.
- Check (forall `{Commutative }, True).
- forall (A : Type) (H : Op A), Commutative A H -> True : Prop
- Check (forall `{Commutative nat}, True).
- forall H : Op nat, Commutative nat H -> True : Prop
- Fail Check (forall `{Commutative nat _}, True).
- The command has indeed failed with message: Typeclass does not expect more arguments
- Fail Check (forall `{!Commutative nat}, True).
- The command has indeed failed with message: The term "Commutative nat" has type "Op nat -> Prop" which should be Set, Prop or Type.
- Arguments Commutative _ {_}.
- Check (forall `{!Commutative nat}, True).
- @Commutative nat nat_op -> True : Prop
- Check (forall `{@Commutative nat plus}, True).
- @Commutative nat Nat.add -> True : Prop
Multiple binders can be merged using ,
as a separator:
- Check (forall `{Commutative A, Hnat : !Commutative nat}, True).
- forall (A : Type) (H : Op A), @Commutative A H -> @Commutative nat nat_op -> True : Prop
One can control the set of generalizable identifiers with
the Generalizable
vernacular command to avoid unexpected
generalizations when mistyping identifiers. There are several commands
that specify which variables should be generalizable.
-
Command
Generalizable All Variables
¶ All variables are candidate for generalization if they appear free in the context under a generalization delimiter. This may result in confusing errors in case of typos. In such cases, the context will probably contain some unexpected generalized variable.
-
Command
Generalizable No Variables
¶ Disable implicit generalization entirely. This is the default behavior.
-
Command
Generalizable VariableVariables ident+
¶ Allow generalization of the given identifiers only. Calling this command multiple times adds to the allowed identifiers.
-
Command
Global Generalizable
¶ Allows exporting the choice of generalizable variables.
Coercions¶
Coercions can be used to implicitly inject terms from one class in
which they reside into another one. A class is either a sort
(denoted by the keyword Sortclass
), a product type (denoted by the
keyword Funclass
), or a type constructor (denoted by its name), e.g.
an inductive type or any constant with a type of the form
forall (
x
\(_{1}\) : A
\(_{1}\) ) … (
x
\(_{n}\) : A
\(_{n}\))
, s
where s
is a sort.
Then the user is able to apply an object that is not a function, but
can be coerced to a function, and more generally to consider that a
term of type A
is of type B
provided that there is a declared coercion
between A
and B
.
More details and examples, and a description of the commands related to coercions are provided in Implicit Coercions.
Printing constructions in full¶
-
Flag
Printing All
¶ Coercions, implicit arguments, the type of pattern matching, but also notations (see Syntax extensions and interpretation scopes) can obfuscate the behavior of some tactics (typically the tactics applying to occurrences of subterms are sensitive to the implicit arguments). Turning this flag on deactivates all high-level printing features such as coercions, implicit arguments, returned type of pattern matching, notations and various syntactic sugar for pattern matching or record projections. Otherwise said,
Printing All
includes the effects of the flagsPrinting Implicit
,Printing Coercions
,Printing Synth
,Printing Projections
, andPrinting Notations
. To reactivate the high-level printing features, use the commandUnset Printing All
.
Printing universes¶
-
Flag
Printing Universes
¶ Turn this flag on to activate the display of the actual level of each occurrence of
Type
. See Sorts for details. This wizard flag, in combination withPrinting All
can help to diagnose failures to unify terms apparently identical but internally different in the Calculus of Inductive Constructions.
-
Command
Print Sorted? Universes
¶ This command can be used to print the constraints on the internal level of the occurrences of \(\Type\) (see Sorts).
If the
Sorted
keyword is present, each universe will be made equivalent to a numbered label reflecting its level (with a linear ordering) in the universe hierarchy.
Existential variables¶
Coq terms can include existential variables which represents unknown subterms to eventually be replaced by actual subterms.
Existential variables are generated in place of unsolvable implicit
arguments or “_” placeholders when using commands such as Check
(see
Section Requests to the environment) or when using tactics such as
refine
, as well as in place of unsolvable instances when using
tactics such that eapply
. An existential
variable is defined in a context, which is the context of variables of
the placeholder which generated the existential variable, and a type,
which is the expected type of the placeholder.
As a consequence of typing constraints, existential variables can be duplicated in such a way that they possibly appear in different contexts than their defining context. Thus, any occurrence of a given existential variable comes with an instance of its original context. In the simple case, when an existential variable denotes the placeholder which generated it, or is used in the same context as the one in which it was generated, the context is not displayed and the existential variable is represented by “?” followed by an identifier.
- Parameter identity : forall (X:Set), X -> X.
- identity is declared
- Check identity _ _.
- identity ?X ?y : ?X where ?X : [ |- Set] ?y : [ |- ?X]
- Check identity _ (fun x => _).
- identity (forall x : ?S, ?S0) (fun x : ?S => ?y) : forall x : ?S, ?S0 where ?S : [ |- Set] ?S0 : [x : ?S |- Set] ?y : [x : ?S |- ?S0]
In the general case, when an existential variable ?ident
appears
outside of its context of definition, its instance, written under the
form { ident := term*; }
is appending to its name, indicating
how the variables of its defining context are instantiated.
The variables of the context of the existential variables which are
instantiated by themselves are not written, unless the Printing Existential Instances
flag
is on (see Section Explicit displaying of existential instances for pretty-printing), and this is why an
existential variable used in the same context as its context of definition is written with no instance.
- Check (fun x y => _) 0 1.
- (fun x y : nat => ?y) 0 1 : ?T@{x:=0; y:=1} where ?T : [x : nat y : nat |- Type] ?y : [x : nat y : nat |- ?T]
- Set Printing Existential Instances.
- Check (fun x y => _) 0 1.
- (fun x y : nat => ?y@{x:=x; y:=y}) 0 1 : ?T@{x:=0; y:=1} where ?T : [x : nat y : nat |- Type] ?y : [x : nat y : nat |- ?T@{x:=x; y:=y}]
Existential variables can be named by the user upon creation using
the syntax ?[ident]
. This is useful when the existential
variable needs to be explicitly handled later in the script (e.g.
with a named-goal selector, see Goal selectors).
Explicit displaying of existential instances for pretty-printing¶
-
Flag
Printing Existential Instances
¶ This flag (off by default) activates the full display of how the context of an existential variable is instantiated at each of the occurrences of the existential variable.
Solving existential variables using tactics¶
Instead of letting the unification engine try to solve an existential
variable by itself, one can also provide an explicit hole together
with a tactic to solve it. Using the syntax ltac:(
tacexpr
)
, the user
can put a tactic anywhere a term is expected. The order of resolution
is not specified and is implementation-dependent. The inner tactic may
use any variable defined in its scope, including repeated alternations
between variables introduced by term binding as well as those
introduced by tactic binding. The expression tacexpr
can be any tactic
expression as described in Ltac.
- Definition foo (x : nat) : nat := ltac:(exact x).
- foo is defined
This construction is useful when one wants to define complicated terms using highly automated tactics without resorting to writing the proof-term by means of the interactive proof engine.
Primitive Integers¶
The language of terms features 63-bit machine integers as values. The type of
such a value is axiomatized; it is declared through the following sentence
(excerpt from the Int63
module):
This type is equipped with a few operators, that must be similarly declared.
For instance, equality of two primitive integers can be decided using the Int63.eqb
function,
declared and specified as follows:
The complete set of such operators can be obtained looking at the Int63
module.
These primitive declarations are regular axioms. As such, they must be trusted and are listed by the
Print Assumptions
command, as in the following example.
- From Coq Require Import Int63.
- [Loading ML file newring_plugin.cmxs ... done] [Loading ML file zify_plugin.cmxs ... done] [Loading ML file micromega_plugin.cmxs ... done] [Loading ML file omega_plugin.cmxs ... done] [Loading ML file int63_syntax_plugin.cmxs ... done]
- Lemma one_minus_one_is_zero : (1 - 1 = 0)%int63.
- 1 subgoal ============================ (1 - 1)%int63 = 0%int63
- Proof.
- apply eqb_correct; vm_compute; reflexivity.
- No more subgoals.
- Qed.
- Print Assumptions one_minus_one_is_zero.
- Axioms: sub : int -> int -> int eqb_correct : forall i j : int, (i == j)%int63 = true -> i = j eqb : int -> int -> bool
The reduction machines (vm_compute
, native_compute
) implement
dedicated, efficient, rules to reduce the applications of these primitive
operations.
The extraction of these primitives can be customized similarly to the extraction
of regular axioms (see Extraction of programs in OCaml and Haskell). Nonetheless, the ExtrOCamlInt63
module can be used when extracting to OCaml: it maps the Coq primitives to types
and functions of a Uint63
module. Said OCaml module is not produced by
extraction. Instead, it has to be provided by the user (if they want to compile
or execute the extracted code). For instance, an implementation of this module
can be taken from the kernel of Coq.
Literal values (at type Int63.int
) are extracted to literal OCaml values
wrapped into the Uint63.of_int
(resp. Uint63.of_int64
) constructor on
64-bit (resp. 32-bit) platforms. Currently, this cannot be customized (see the
function Uint63.compile
from the kernel).
Primitive Floats¶
The language of terms features Binary64 floating-point numbers as values.
The type of such a value is axiomatized; it is declared through the
following sentence (excerpt from the PrimFloat
module):
This type is equipped with a few operators, that must be similarly declared.
For instance, the product of two primitive floats can be computed using the
PrimFloat.mul
function, declared and specified as follows:
where Prim2SF
is defined in the FloatOps
module.
The set of such operators is described in section Floats library.
These primitive declarations are regular axioms. As such, they must be trusted, and are listed by the
Print Assumptions
command.
The reduction machines (vm_compute
, native_compute
) implement
dedicated, efficient rules to reduce the applications of these primitive
operations, using the floating-point processor operators that are assumed
to comply with the IEEE 754 standard for floating-point arithmetic.
The extraction of these primitives can be customized similarly to the extraction
of regular axioms (see Extraction of programs in OCaml and Haskell). Nonetheless, the ExtrOCamlFloats
module can be used when extracting to OCaml: it maps the Coq primitives to types
and functions of a Float64
module. Said OCaml module is not produced by
extraction. Instead, it has to be provided by the user (if they want to compile
or execute the extracted code). For instance, an implementation of this module
can be taken from the kernel of Coq.
Literal values (of type Float64.t
) are extracted to literal OCaml
values (of type float
) written in hexadecimal notation and
wrapped into the Float64.of_float
constructor, e.g.:
Float64.of_float (0x1p+0)
.
Bidirectionality hints¶
When type-checking an application, Coq normally does not use information from the context to infer the types of the arguments. It only checks after the fact that the type inferred for the application is coherent with the expected type. Bidirectionality hints make it possible to specify that after type-checking the first arguments of an application, typing information should be propagated from the context to help inferring the types of the remaining arguments.
-
Command
Arguments qualid ident1* & ident2*
¶ This commands tells the typechecking algorithm, when type-checking applications of
qualid
, to first type-check the arguments inident1
and then propagate information from the typing context to type-check the remaining arguments (inident2
).
Example
In a context where a coercion was declared from bool
to nat
:
- Definition b2n (b : bool) := if b then 1 else 0.
- b2n is defined
- Coercion b2n : bool >-> nat.
- b2n is now a coercion
Coq cannot automatically coerce existential statements over bool
to
statements over nat
, because the need for inserting a coercion is known
only from the expected type of a subterm:
- Fail Check (ex_intro _ true _ : exists n : nat, n > 0).
- The command has indeed failed with message: The term "ex_intro ?P true ?y" has type "exists y, ?P y" while it is expected to have type "exists n : nat, n > 0" (cannot unify "bool" and "nat").
However, a suitable bidirectionality hint makes the example work:
- Arguments ex_intro _ _ & _ _.
- Check (ex_intro _ true _ : exists n : nat, n > 0).
- ex_intro (fun n : nat => n > 0) true ?g : exists n : nat, n > 0 : exists n : nat, n > 0 where ?g : [ |- (fun n : nat => n > 0) true]
Coq will attempt to produce a term which uses the arguments you provided, but in some cases involving Program mode the arguments after the bidirectionality starts may be replaced by convertible but syntactically different terms.