Setting properties of a function's arguments¶
- Command Arguments reference arg_specs* , implicits_alt** : args_modifier+,?¶
- arg_specs
::=
argument_spec|
/|
&|
( argument_spec+ ) % scope%_ scope*|
[ argument_spec+ ] % scope%_ scope*|
{ argument_spec+ } % scope%_ scope*argument_spec
::=
!? name % scope%_ scope*implicits_alt
::=
name|
[ name+ ]|
{ name+ }args_modifier
::=
simpl nomatch|
simpl never|
default implicits|
clear implicits|
clear scopes|
clear bidirectionality hint|
rename|
assert|
extra scopes|
clear scopes and implicits|
clear implicits and scopesSpecifies properties of the arguments of a function after the function has already been defined. It gives fine-grained control over the elaboration process (i.e. the translation of Gallina language extensions into the core language used by the kernel). The command's effects include:
Making arguments implicit. Afterward, implicit arguments must be omitted in any expression that applies
reference
.Declaring that some arguments of a given function should be interpreted in a given notation scope.
Affecting when the
simpl
andcbn
tactics unfold the function. See Effects of Arguments on unfolding.Providing bidirectionality hints. See Bidirectionality hints.
This command supports the
local
andglobal
attributes. Default behavior is to limit the effect to the current section but also to extend their effect outside the current module or library file. Applyinglocal
limits the effect of the command to the current module if it's not in a section. Applyingglobal
within a section extends the effect outside the current sections and current module in which the command appears./
the function will be unfolded only if it's applied to at least the arguments appearing before the
/
. See Effects of Arguments on unfolding.- Error The / modifier may only occur once.¶
&
tells the type checking algorithm to first type check the arguments before the
&
and then to propagate information from that typing context to type check the remaining arguments. See Bidirectionality hints.- Error The & modifier may only occur once.¶
( argument_spec+ ) %_ scope*
(name1 name2 ...)%scope*
is shorthand forname1%scope* name2%scope* ...
[ argument_spec+ ] %_ scope*
declares the enclosed names as implicit, non-maximally inserted.
[name1 name2 ... ]%_scope*
is equivalent to[name1]%_scope* [name2]%_scope* ...
{ argument_spec+ } %_ scope*
declares the enclosed names as implicit, maximally inserted.
{name1 name2 ... }%_scope*
is equivalent to{name1}%_scope* {name2}%_scope* ...
!
the function will be unfolded only if all the arguments marked with
!
evaluate to constructors. See Effects of Arguments on unfolding.name %_ scope*
a formal parameter of the function
reference
(i.e. the parameter name used in the function definition). Unlessrename
is specified, the list ofname
s must be a prefix of the formal parameters, including all implicit arguments._
can be used to skip over a formal parameter. This construct declaresname
as non-implicit ifclear implicits
is specified or any othername
in theArguments
command is declared implicit.scope
can be either scope names or their delimiting keys. When multiple scopes are present, notations are interpreted in the leftmost scope containing them. See Binding arguments to scopes.Deprecated since version 8.19: The
% scope
syntax is deprecated in favor of the currently equivalent%_ scope
. It will be reused in future versions with the same semantics as in terms.- Error To rename arguments the 'rename' flag must be specified.¶
clear implicits
makes all implicit arguments into explicit arguments
- Error The 'clear implicits' flag must be omitted if implicit annotations are given.¶
default implicits
automatically determine the implicit arguments of the object. See Automatic declaration of implicit arguments.
- Error The 'default implicits' flag is incompatible with implicit annotations.¶
rename
rename implicit arguments for the object. See the example here.
assert
assert that the object has the expected number of arguments with the expected names. See the example here: Renaming implicit arguments.
clear scopes
clears argument scopes of
reference
extra scopes
defines extra argument scopes, to be used in case of coercion to
Funclass
(see Implicit Coercions) or with a computed type.simpl nomatch
prevents performing a simplification step for
reference
that would expose a match construct in the head position. See Effects of Arguments on unfolding.simpl never
prevents performing a simplification step for
reference
. See Effects of Arguments on unfolding.clear bidirectionality hint
removes the bidirectionality hint, the
&
implicits_alt
use to specify alternative implicit argument declarations for functions that can only be applied to a fixed number of arguments (excluding, for instance, functions whose type is polymorphic). For parsing, the longest list of implicit arguments matching the function application is used to select which implicit arguments are inserted. For printing, the alternative with the most implicit arguments is used; the implict arguments will be omitted if
Printing Implicit
is not set. See the example here.
Use
About
to view the current implicit arguments setting for areference
.Or use the
Print Implicit
command to see the implicit arguments of an object (see Displaying implicit arguments).
Manual declaration of implicit arguments¶
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 (guarded 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 (guarded on 2nd argument)
- Arguments map [A B] f l.
- Arguments length {A} l. (* A has to be maximally inserted *)
- Check (fun l:list (list nat) => map length l).
- fun l : list (list nat) => map length l : list (list nat) -> list nat
Example: Multiple alternatives with implicits_alt
- 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
Automatic declaration of implicit arguments¶
The "
default implicits
"args_modifier
clause tells Coq to automatically determine the implicit arguments of the object.Auto-detection is governed by flags specifying whether strict, contextual, or reversible-pattern implicit arguments must be considered or not (see Controlling strict implicit arguments, Controlling contextual implicit arguments, Controlling reversible-pattern implicit arguments and also Controlling the insertion of implicit arguments not followed by explicit arguments).
Example: Default implicits
- 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
Renaming implicit arguments¶
Example: (continued) Renaming implicit arguments
- 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.
Binding arguments to scopes¶
The following command declares that the first two arguments of
plus_fct
are interpreted in thescope
delimited by the keyF
and the third argument is first interpreted in the scope delimited by the keyR
, then inF
(when a notation has no interpretation inR
).Arguments plus_fct (f1 f2)%_F x%_R%_F.When interpreting a term, if some of the arguments of
reference
are built from a notation, then this notation is interpreted in the scope stack extended by the scopes bound (if any) to this argument. The effect of these scopes is limited to the argument itself. It does not propagate to subterms but the subterms that, after interpretation of the notation, turn to be themselves arguments of a reference are interpreted according to the argument scopes bound to this reference.
Note
In notations, the subterms matching the identifiers of the notations are interpreted in the scope in which the identifiers occurred at the time of the declaration of the notation. Here is an example:
- Parameter g : bool -> bool.
- g is declared
- Declare Scope mybool_scope.
- Notation "@@" := true (only parsing) : bool_scope.
- Notation "@@" := false (only parsing): mybool_scope.
- Bind Scope bool_scope with bool.
- Notation "# x #" := (g x) (at level 40).
- Check # @@ #.
- # true # : bool
- Arguments g _%_mybool_scope.
- Check # @@ #.
- # true # : bool
- Delimit Scope mybool_scope with mybool.
- Check # @@%mybool #.
- # false # : bool
Effects of Arguments
on unfolding¶
simpl never
indicates that a constant should not be unfolded bycbn
orsimpl
when in head position. Note that in the case ofsimpl
, the modifier does not apply to reduction of the main argument of amatch
,fix
, primitive projection, or of an unfoldable constant hiding amatch
,fix
or primitive projection.Example
- Arguments Nat.sub n m : simpl never.
After that command an expression like
(Nat.sub (S x) y)
is left untouched by the tacticscbn
andsimpl
.Otherwise, an expression like
(Nat.sub (S x) 0) + 1
reduces toS (x + 1)
forsimpl
becauseNat.sub
is the main argument of+
in this case.A constant can be marked to be unfolded only if it's applied to at least the arguments appearing before the
/
in aArguments
command.Example
- Definition fcomp A B C f (g : A -> B) (x : A) : C := f (g x).
- fcomp is defined
- Arguments fcomp {A B C} f g x /.
- Notation "f \o g" := (fcomp f g) (at level 50).
After that command the expression
(f \o g)
is left untouched bysimpl
while((f \o g) t)
is reduced to(f (g t))
. The same mechanism can be used to make a constant volatile, i.e. always unfolded.Example
- Definition volatile := fun x : nat => x.
- volatile is defined
- Arguments volatile / x.
A constant can be marked to be unfolded only if an entire set of arguments evaluates to a constructor. The
!
symbol can be used to mark such arguments.Example
- Arguments minus !n !m.
After that command, the expression
(minus (S x) y)
is left untouched bysimpl
, while(minus (S x) (S y))
is reduced to(minus x y)
.simpl nomatch
indicates that a constant should not be unfolded if it would expose amatch
construct in the head position. This affects thecbn
,simpl
andhnf
tactics.Example
- Arguments minus n m : simpl nomatch.
In this case,
(minus (S (S x)) (S y))
is simplified to(minus (S x) y)
even if an extra simplification is possible.In detail: the tactic
simpl
first applies βι-reduction. Then, it expands transparent constants and tries to reduce further using βι-reduction. But, when no ι rule is applied after unfolding then δ-reductions are not applied. For instance trying to usesimpl
on(plus n O) = n
changes nothing.
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.
An Arguments
command containing arg_specs1 & arg_specs2
provides bidirectionality hints.
It tells the typechecking algorithm, when type checking
applications of qualid
, to first type check the arguments in
arg_specs1
and then propagate information from the typing context to
type check the remaining arguments (in arg_specs2
).
Example: Bidirectionality hints
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.