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 to be 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 and non-maximal insertion of implicit arguments¶
When a function is partially applied and the next argument to apply is an implicit argument, the application can be interpreted in two ways. If the next argument is declared as maximally inserted, the partial application will include that argument. Otherwise, the argument is non-maximally inserted and the partial application will not include that argument.
Each implicit argument can be declared to be inserted maximally or non maximally. In Coq, maximally inserted implicit arguments are written between curly braces "{ }" and non-maximally inserted implicit arguments are written in square brackets "[ ]".
See also
Trailing Implicit Arguments¶
An implicit argument is considered trailing when all following arguments are implicit. Trailing implicit arguments must be declared as maximally inserted; otherwise they would never be inserted.
-
Error
Argument name is a trailing implicit, so it can't be declared non maximal. Please use { } instead of [ ].
¶ For instance:
- Fail Definition double [n] := n + n.
- The command has indeed failed with message: Argument n is a trailing implicit, so it can't be declared non maximal. Please use { } instead of [ ].
Casual use of implicit arguments¶
If an argument of a function application can be inferred from the type
of the other arguments, the user can force inference of the argument
by replacing it with _
.
-
Error
Cannot infer a term for this placeholder.
¶ Coq was not able to deduce an instantiation of a “_”.
Declaration of implicit arguments¶
Implicit arguments can be declared when a function is declared or
afterwards, using the Arguments
command.
Implicit Argument Binders¶
implicit_binders::=
{ name+ : type? }
|
[ name+ : type? ]
In the context of a function definition, these forms specify that
name
is an implicit argument. The first form, with curly
braces, makes name
a maximally inserted implicit argument. The second
form, with square brackets, makes name
a non-maximally inserted implicit argument.
For example:
- Definition id {A : Type} (x : A) : A := x.
- id is defined
declares the argument A
of id
as a maximally
inserted implicit argument. A
may be omitted
in applications of id
but may be specified 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 goal ============================ forall A : Type, compose id id = id
For non-maximally inserted implicit arguments, use square brackets:
- 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 f t) end.
- map is defined map is recursively defined (guarded on 4th argument)
- Print Implicit map.
- map : forall [A B : Type], (A -> B) -> list A -> list B Arguments A, B are implicit
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.
The syntax is also supported in internal binders. For instance, in the
following kinds of expressions, the type of each declaration present
in binder*
can be bracketed to mark the declaration as
implicit:
* fun (ident:forall binder*, type) => term
,
* forall (ident:forall binder*, type), type
,
* let ident binder* := term in term
,
* fix ident binder* := term in term
and
* cofix ident binder* := term in term
.
Here is an example:
- Axiom Ax : forall (f:forall {A} (a:A), A * A), let g {A} (x y:A) := (x,y) in f 0 = g 0 0.
- Ax is declared
-
Warning
Ignoring implicit binder declaration in unexpected position
¶ This is triggered when setting an argument implicit in an expression which does not correspond to the type of an assumption or to the body of a definition. Here is an example:
- Definition f := forall {y}, y = 0.
- map is defined map is recursively defined (guarded on 4th argument) list is defined list_rect is defined list_ind is defined list_rec is defined list_sind is defined Ax is declared Toplevel input, characters 24-25: > Definition f := forall {y}, y = 0. > ^ Warning: Ignoring implicit binder declaration in unexpected position. [unexpected-implicit-declaration,syntax] f is defined
-
Warning
Making shadowed name of implicit argument accessible by position
¶ This is triggered when two variables of same name are set implicit in the same block of binders, in which case the first occurrence is considered to be unnamed. Here is an example:
- Check let g {x:nat} (H:x=x) {x} (H:x=x) := x in 0.
- Toplevel input, characters 0-50: > Check let g {x:nat} (H:x=x) {x} (H:x=x) := x in 0. > ^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^ Warning: Making shadowed name of implicit argument accessible by position. [shadowed-implicit-name,syntax] let g := fun (x : nat) (H : x = x) (x0 : ?A@{x0:=x}) (_ : x0 = x0) => x0 in 0 : nat where ?A : [x0 : nat H : x0 = x0 x : ?A |- Type] (x cannot be used)
Mode for automatic declaration of implicit arguments¶
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.
Controlling contextual implicit arguments¶
Controlling reversible-pattern implicit arguments¶
Controlling the insertion of implicit arguments not followed by explicit arguments¶
Combining manual declaration and automatic declaration¶
When some arguments are manually specified implicit with binders in a definition and the automatic declaration mode in on, the manual implicit arguments are added to the automatically declared ones.
In that case, and when the flag Maximal Implicit Insertion
is set to off,
some trailing implicit arguments can be inferred to be non-maximally inserted. In
this case, they are converted to maximally inserted ones.
Example
- Set Implicit Arguments.
- Axiom eq0_le0 : forall (n : nat) (x : n = 0), n <= 0.
- eq0_le0 is declared
- Print Implicit eq0_le0.
- eq0_le0 : forall [n : nat], n = 0 -> n <= 0 Argument n is implicit
- Axiom eq0_le0' : forall (n : nat) {x : n = 0}, n <= 0.
- Argument n is a trailing implicit, so it has been declared maximally inserted. eq0_le0' is declared
- Print Implicit eq0_le0'.
- eq0_le0' : forall {n : nat}, n = 0 -> n <= 0 Arguments n, x are implicit and maximally inserted
Explicit applications¶
In presence of non-strict or contextual arguments, or in presence of
partial applications, the synthesis of implicit arguments may fail, so
one may have to explicitly give certain implicit arguments of an
application. Use the (ident := term)
form of arg
to do so,
where ident
is the name of the implicit argument and term
is its corresponding explicit term. Alternatively, one can deactivate
the hiding of implicit arguments for a single function application using the
@qualid_annotated term1+
form of term_application
.
Example: Syntax for explicitly giving implicit arguments (continued)
- Parameter X : Type.
- X is declared
- Definition Relation := X -> X -> Prop.
- eq0_le0 is declared Argument n is a trailing implicit, so it has been declared maximally inserted. eq0_le0' is declared X is declared 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 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 (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
Displaying implicit arguments¶
Displaying implicit arguments when 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 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: Cannot enforce Set+1 <= Set).
but succeeds in
- Check Prop = nat.
- Prop = nat : Prop
Deactivation of implicit arguments for parsing¶
term_explicit::=
@ qualid_annotated
This syntax can be used to disable implicit arguments for a single function.
Example
The function id
has one implicit argument and one explicit
argument.
- Check (id 0).
- id 0 : nat
- Definition id' := @id.
- id' is defined
The function id'
has no implicit argument.
- Check (id' nat 0).
- id' nat 0 : nat
-
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.
Example
We can reproduce the example above using the Parsing
Explicit
flag:
- Set Parsing Explicit.
- Definition id' := id.
- id' is defined
- Unset Parsing Explicit.
- Check (id 1).
- id 1 : nat
- Check (id' nat 1).
- id' nat 1 : nat
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 TypeTypes reserv_list
¶ - reserv_list
::=
( simple_reserv )+|
simple_reservsimple_reserv
::=
ident+ : typeSets the type of bound variables starting with
ident
(eitherident
itself orident
followed by one or more single quotes, underscore or digits) totype
(unless the bound variable is already declared with an explicit type, in which case, that type will be used).
Example
- Require Import List.
- Implicit Types m n : nat.
- Lemma cons_inj_nat : forall m n l, n :: l = m :: l -> n = m.
- 1 goal ============================ forall m n (l : list nat), n :: l = m :: l -> n = m
- Proof. intros m n. Abort.
- 1 goal m, n : nat ============================ forall l : list nat, n :: l = m :: l -> n = m
- Lemma cons_inj_bool : forall (m n:bool) l, n :: l = m :: l -> n = m.
- 1 goal ============================ forall (m n : bool) (l : list bool), n :: l = m :: l -> n = m
- Abort.
Implicit generalization¶
generalizing_binder::=
`( typeclass_constraint+, )
|
`{ typeclass_constraint+, }
|
`[ typeclass_constraint+, ]
typeclass_constraint::=
!? term
|
{ name } : !? term
|
name : !? term
term_generalizing::=
`{ term }
|
`( term )
Implicit generalization is an automatic elaboration of a statement
with free variables into a closed statement where these variables are
quantified explicitly. Use the Generalizable
command to designate
which variables should be generalized.
It is activated within a binder by prefixing it with `, and for terms by surrounding it with `{ }, or `[ ] or `( ).
Terms surrounded by `{ } introduce their free variables as maximally inserted implicit arguments, terms surrounded by `[ ] introduce them as non-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).
- Set Warnings "-deprecated-instance-without-locality".
- 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
-
Command
Generalizable VariableVariables ident+All VariablesNo Variables
¶ Controls the set of generalizable identifiers. By default, no variables are generalizable.
This command supports the
global
attribute.The
VariableVariables ident+
form allows generalization of only the givenident
s. Using this command multiple times adds to the allowed identifiers. The other forms clear the list ofident
s.The
All Variables
form generalizes all free variables in the context that appear 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 variables.The
No Variables
form disables implicit generalization entirely. This is the default behavior (before anyGeneralizable
command has been entered).