Chapter 6  Vernacular commands

6.1  Displaying

6.1.1  Print qualid.

This command displays on the screen information about the declared or defined object referred by qualid.


Error messages:

  1. qualid not a defined object


Variants:

  1. Print Term qualid.
    This is a synonym to Print qualid when qualid denotes a global constant.
  2. About qualid.
    This displays various information about the object denoted by qualid: its kind (module, constant, assumption, inductive, constructor, abbreviation, …), long name, type, implicit arguments and argument scopes. It does not print the body of definitions or proofs.

6.1.2  Print All.

This command displays information about the current state of the environment, including sections and modules.


Variants:

  1. Inspect num.
    This command displays the num last objects of the current environment, including sections and modules.
  2. Print Section ident.
    should correspond to a currently open section, this command displays the objects defined since the beginning of this section.

6.2  Flags, Options and Tables

Coq configurability is based on flags (e.g. Set Printing All in Section 2.9), options (e.g. Set Printing Width integer in Section 6.9.3), or tables (e.g. Add Printing Record ident, in Section 2.2.4). The names of flags, options and tables are made of non-empty sequences of identifiers (conventionally with capital initial letter). The general commands handling flags, options and tables are given below.

6.2.1  Set flag.

This command switches flag on. The original state of flag is restored when the current module ends.


Variants:

  1. Local Set flag.
    This command switches flag on. The original state of flag is restored when the current section ends.
  2. Global Set flag.
    This command switches flag on. The original state of flag is not restored at the end of the module. Additionally, if set in a file, flag is switched on when the file is Require-d.

6.2.2  Unset flag.

This command switches flag off. The original state of flag is restored when the current module ends.


Variants:

  1. Local Unset flag.
    This command switches flag off. The original state of flag is restored when the current section ends.
  2. Global Unset flag.
    This command switches flag off. The original state of flag is not restored at the end of the module. Additionally, if set in a file, flag is switched off when the file is Require-d.

6.2.3  Test flag.

This command prints whether flag is on or off.

6.2.4  Set option value.

This command sets option to value. The original value of option is restored when the current module ends.


Variants:

  1. Local Set option value. This command sets option to value. The original value of option is restored at the end of the module.
  2. Global Set option value. This command sets option to value. The original value of option is not restored at the end of the module. Additionally, if set in a file, option is set to value when the file is Require-d.

6.2.5  Unset option.

This command resets option to its default value.


Variants:

  1. Local Unset option.
    This command resets option to its default value. The original state of option is restored when the current section ends.
  2. Global Unset option.
    This command resets option to its default value. The original state of option is not restored at the end of the module. Additionally, if unset in a file, option is reset to its default value when the file is Require-d.

6.2.6  Test option.

This command prints the current value of option.

6.2.7  Tables

The general commands for tables are Add table value, Remove table value, Test table, Test table for value and Print Table table.

6.2.8  Print Options.

This command lists all available flags, options and tables.


Variants:

  1. Print Tables.
    This is a synonymous of Print Options.

6.3  Requests to the environment

6.3.1  Check term.

This command displays the type of term. When called in proof mode, the term is checked in the local context of the current subgoal.


Variants:

  1. selector: Check term.
    specifies on which subgoal to perform typing (see Section 8.1).

6.3.2  Eval convtactic in term.

This command performs the specified reduction on term, and displays the resulting term with its type. The term to be reduced may depend on hypothesis introduced in the first subgoal (if a proof is in progress).


See also: Section 8.7.

6.3.3  Compute term.

This command performs a call-by-value evaluation of term by using the bytecode-based virtual machine. It is a shortcut for Eval vm_compute in term.


See also: Section 8.7.

6.3.4  Extraction term.

This command displays the extracted term from term. The extraction is processed according to the distinction between Set and Prop; that is to say, between logical and computational content (see Section 4.1.1). The extracted term is displayed in Objective Caml syntax, where global identifiers are still displayed as in Coq terms.


Variants:

  1. Recursive Extraction qualid1qualidn.
    Recursively extracts all the material needed for the extraction of global qualid1, …, qualidn.


See also: Chapter 23.

6.3.5  Print Assumptions qualid.

This commands display all the assumptions (axioms, parameters and variables) a theorem or definition depends on. Especially, it informs on the assumptions with respect to which the validity of a theorem relies.


Variants:

  1. Print Opaque Dependencies qualid.
    Displays the set of opaque constants qualid relies on in addition to the assumptions.
  2. Print Transparent Dependencies qualid.
    Displays the set of transparent constants qualid relies on in addition to the assumptions.
  3. Print All Dependencies qualid.
    Displays all assumptions and constants qualid relies on.

6.3.6  Search qualid.

This command displays the name and type of all objects (hypothesis of the current goal, theorems, axioms, etc) of the current context whose statement contains qualid. This command is useful to remind the user of the name of library lemmas.


Error messages:

  1. The reference qualid was not found in the current environment
    There is no constant in the environment named qualid.


Variants:

  1. Search string.

    If string is a valid identifier, this command displays the name and type of all objects (theorems, axioms, etc) of the current context whose name contains string. If string is a notation’s string denoting some reference qualid (referred to by its main symbol as in "+" or by its notation’s string as in "_ + _" or "_ 'U' _", see Section 12.1), the command works like Search qualid.

  2. Search string%key.

    The string string must be a notation or the main symbol of a notation which is then interpreted in the scope bound to the delimiting key key (see Section 12.2.2).

  3. Search term_pattern.

    This searches for all statements or types of definition that contains a subterm that matches the pattern term_pattern (holes of the pattern are either denoted by “_” or by “?ident” when non linear patterns are expected).

  4. Search [-]term_pattern-string [-]term_pattern-string.
    where term_pattern-string is a term_pattern or a string, or a string followed by a scope delimiting key %key.

    This generalization of Search searches for all objects whose statement or type contains a subterm matching term_pattern (or qualid if string is the notation for a reference qualid) and whose name contains all string of the request that correspond to valid identifiers. If a term_pattern or a string is prefixed by “-”, the search excludes the objects that mention that term_pattern or that string.

  5. Search term_pattern-stringterm_pattern-string inside module1modulen.

    This restricts the search to constructions defined in modules module1modulen.

  6. Search term_pattern-string term_pattern-string outside module1...modulen.

    This restricts the search to constructions not defined in modules module1modulen.

  7. selector: Search [-]term_pattern-string [-]term_pattern-string.

    This specifies the goal on which to search hypothesis (see Section 8.1). By default the 1st goal is searched. This variant can be combined with other variants presented here.


Examples:

Coq < Require Import ZArith.

Coq < Search Z.mul Z.add "distr".
fast_Zmult_plus_distr_l:
  forall (n m p : Z) (P : Z -> Prop),
  P (n * p + m * p)%Z -> P ((n + m) * p)%Z
Z.mul_add_distr_r:
  forall n m p : Z, ((n + m) * p)%Z = (n * p + m * p)%Z
Z.mul_add_distr_l:
  forall n m p : Z, (n * (m + p))%Z = (n * m + n * p)%Z

Coq < Search "+"%Z "*"%Z "distr" -positive -Prop.
Z.mul_add_distr_r:
  forall n m p : Z, ((n + m) * p)%Z = (n * p + m * p)%Z
Z.mul_add_distr_l:
  forall n m p : Z, (n * (m + p))%Z = (n * m + n * p)%Z

Coq < Search (?x * _ + ?x * _)%Z outside OmegaLemmas.
Z.mul_add_distr_l:
  forall n m p : Z, (n * (m + p))%Z = (n * m + n * p)%Z


Warning: Up to Coq version 8.4, Search had the behavior of current SearchHead and the behavior of current Search was obtained with command SearchAbout. For compatibility, the deprecated name SearchAbout can still be used as a synonym of Search. For compatibility, the list of objects to search when using SearchAbout may also be enclosed by optional [ ] delimiters.

6.3.7  SearchHead term.

This command displays the name and type of all hypothesis of the current goal (if any) and theorems of the current context whose statement’s conclusion has the form (term t1 .. tn). This command is useful to remind the user of the name of library lemmas.

Coq < SearchHead le.
le_n: forall n : nat, n <= n
le_S: forall n m : nat, n <= m -> n <= S m
le_pred: forall n m : nat, n <= m -> Nat.pred n <= Nat.pred m
le_S_n: forall n m : nat, S n <= S m -> n <= m
le_0_n: forall n : nat, 0 <= n
le_n_S: forall n m : nat, n <= m -> S n <= S m

Coq < SearchHead (@eq bool).
andb_true_intro:
  forall b1 b2 : bool, b1 = true /\ b2 = true -> (b1 && b2)%bool = true


Variants:

  1. SearchHead term inside module1modulen.

    This restricts the search to constructions defined in modules module1modulen.

  2. SearchHead term outside module1modulen.

    This restricts the search to constructions not defined in modules module1modulen.


    Error messages:

    1. Module/section module not found No module module has been required (see Section 6.5.1).
  3. selector: SearchHead term.

    This specifies the goal on which to search hypothesis (see Section 8.1). By default the 1st goal is searched. This variant can be combined with other variants presented here.


Warning: Up to Coq version 8.4, SearchHead was named Search.

6.3.8  SearchPattern term.

This command displays the name and type of all hypothesis of the current goal (if any) and theorems of the current context whose statement’s conclusion or last hypothesis and conclusion matches the expression term where holes in the latter are denoted by “_”. It is a variant of Search term_pattern that does not look for subterms but searches for statements whose conclusion has exactly the expected form, or whose statement finishes by the given series of hypothesis/conclusion.

Coq < Require Import Arith.

Coq < SearchPattern (_ + _ = _ + _).
f_equal2_plus:
  forall x1 y1 x2 y2 : nat, x1 = y1 -> x2 = y2 -> x1 + x2 = y1 + y2
plus_Snm_nSm: forall n m : nat, S n + m = n + S m
plus_assoc_reverse: forall n m p : nat, n + m + p = n + (m + p)
Nat.add_succ_comm: forall n m : nat, S n + m = n + S m
Nat.add_comm: forall n m : nat, n + m = m + n
Nat.add_assoc: forall n m p : nat, n + (m + p) = n + m + p
Nat.add_shuffle0: forall n m p : nat, n + m + p = n + p + m
Nat.add_shuffle1: forall n m p q : nat, n + m + (p + q) = n + p + (m + q)
Nat.add_shuffle2: forall n m p q : nat, n + m + (p + q) = n + q + (m + p)
Nat.add_shuffle3: forall n m p : nat, n + (m + p) = m + (n + p)

Coq < SearchPattern (nat -> bool).
Init.Nat.eqb: nat -> nat -> bool
Init.Nat.leb: nat -> nat -> bool
Init.Nat.ltb: nat -> nat -> bool
Init.Nat.even: nat -> bool
Init.Nat.odd: nat -> bool
Init.Nat.testbit: nat -> nat -> bool
BinNat.N.testbit_nat: BinNums.N -> nat -> bool
BinNatDef.N.testbit_nat: BinNums.N -> nat -> bool
Nat.eqb: nat -> nat -> bool
Nat.leb: nat -> nat -> bool
Nat.ltb: nat -> nat -> bool
Nat.even: nat -> bool
Nat.odd: nat -> bool
Nat.testbit: nat -> nat -> bool
BinPos.Pos.testbit_nat: BinNums.positive -> nat -> bool
BinPosDef.Pos.testbit_nat: BinNums.positive -> nat -> bool

Coq < SearchPattern (forall l : list __ l l).
List.lel_refl: forall (A : Type) (l : list A), List.lel l l
List.incl_refl: forall (A : Type) (l : list A), List.incl l l

Patterns need not be linear: you can express that the same expression must occur in two places by using pattern variables ‘?ident”.

Coq < SearchPattern (?X1 + _ = _ + ?X1).
Nat.add_comm: forall n m : nat, n + m = m + n


Variants:

  1. SearchPattern term inside module1modulen.

    This restricts the search to constructions defined in modules module1modulen.

  2. SearchPattern term outside module1modulen.

    This restricts the search to constructions not defined in modules module1modulen.

  3. selector: SearchPattern term.

    This specifies the goal on which to search hypothesis (see Section 8.1). By default the 1st goal is searched. This variant can be combined with other variants presented here.

6.3.9  SearchRewrite term.

This command displays the name and type of all hypothesis of the current goal (if any) and theorems of the current context whose statement’s conclusion is an equality of which one side matches the expression term. Holes in term are denoted by “_”.

Coq < Require Import Arith.

Coq < SearchRewrite (_ + _ + _).
plus_assoc_reverse: forall n m p : nat, n + m + p = n + (m + p)
Nat.add_assoc: forall n m p : nat, n + (m + p) = n + m + p
Nat.add_shuffle0: forall n m p : nat, n + m + p = n + p + m
Nat.add_shuffle1: forall n m p q : nat, n + m + (p + q) = n + p + (m + q)
Nat.add_shuffle2: forall n m p q : nat, n + m + (p + q) = n + q + (m + p)
Nat.add_carry_div2:
  forall (a b : nat) (c0 : bool),
  (a + b + Nat.b2n c0) / 2 =
  a / 2 + b / 2 +
  Nat.b2n
    (Nat.testbit a 0 && Nat.testbit b 0
     || c0 && (Nat.testbit a 0 || Nat.testbit b 0))


Variants:

  1. SearchRewrite term inside module1modulen.

    This restricts the search to constructions defined in modules module1modulen.

  2. SearchRewrite term outside module1modulen.

    This restricts the search to constructions not defined in modules module1modulen.

  3. selector: SearchRewrite term.

    This specifies the goal on which to search hypothesis (see Section 8.1). By default the 1st goal is searched. This variant can be combined with other variants presented here.

Nota Bene:

For the Search, SearchHead, SearchPattern and SearchRewrite queries, it is possible to globally filter the search results via the command Add Search Blacklist "substring1". A lemma whose fully-qualified name contains any of the declared substrings will be removed from the search results. The default blacklisted substrings are "_subproof" "Private_". The command Remove Search Blacklist ... allows expunging this blacklist.

6.3.10  Locate qualid.

This command displays the full name of objects whose name is a prefix of the qualified identifier qualid, and consequently the Coq module in which they are defined. It searches for objects from the different qualified name spaces of Coq: terms, modules, Ltac, etc.

Coq < Locate nat.
Inductive Coq.Init.Datatypes.nat

Coq < Locate Datatypes.O.
Constructor Coq.Init.Datatypes.O
  (shorter name to refer to it in current context is O)

Coq < Locate Init.Datatypes.O.
Constructor Coq.Init.Datatypes.O
  (shorter name to refer to it in current context is O)

Coq < Locate Coq.Init.Datatypes.O.
Constructor Coq.Init.Datatypes.O
  (shorter name to refer to it in current context is O)

Coq < Locate I.Dont.Exist.
No object of suffix
I.Dont.Exist


Variants:

  1. Locate Term qualid.
    As Locate but restricted to terms.
  2. Locate Module qualid. As Locate but restricted to modules.
  3. Locate Ltac qualid.
    As Locate but restricted to tactics.


See also: Section 12.1.10

6.4  Loading files

Coq offers the possibility of loading different parts of a whole development stored in separate files. Their contents will be loaded as if they were entered from the keyboard. This means that the loaded files are ASCII files containing sequences of commands for Coq’s toplevel. This kind of file is called a script for Coq. The standard (and default) extension of Coq’s script files is .v.

6.4.1  Load ident.

This command loads the file named ident.v, searching successively in each of the directories specified in the loadpath. (see Section 2.6.3)


Variants:

  1. Load string.
    Loads the file denoted by the string string, where string is any complete filename. Then the ~ and .. abbreviations are allowed as well as shell variables. If no extension is specified, Coq will use the default extension .v
  2. Load Verbose ident., Load Verbose string
    Display, while loading, the answers of Coq to each command (including tactics) contained in the loaded file
    See also: Section 6.9.1


Error messages:

  1. Can’t find file ident on loadpath

6.5  Compiled files

This section describes the commands used to load compiled files (see Chapter 14 for documentation on how to compile a file). A compiled file is a particular case of module called library file.

6.5.1  Require qualid.

This command looks in the loadpath for a file containing module qualid and adds the corresponding module to the environment of Coq. As library files have dependencies in other library files, the command Require qualid recursively requires all library files the module qualid depends on and adds the corresponding modules to the environment of Coq too. Coq assumes that the compiled files have been produced by a valid Coq compiler and their contents are then not replayed nor rechecked.

To locate the file in the file system, qualid is decomposed under the form dirpath.ident and the file ident.vo is searched in the physical directory of the file system that is mapped in Coq loadpath to the logical path dirpath (see Section 2.6.3). The mapping between physical directories and logical names at the time of requiring the file must be consistent with the mapping used to compile the file. If several files match, one of them is picked in an unspecified fashion.


Variants:

  1. Require Import qualid.

    This loads and declares the module qualid and its dependencies then imports the contents of qualid as described in Section 2.5.8.

    It does not import the modules on which qualid depends unless these modules were itself required in module qualid using Require Export, as described below, or recursively required through a sequence of Require Export.

    If the module required has already been loaded, Require Import qualid simply imports it, as Import qualid would.

  2. Require Export qualid.

    This command acts as Require Import qualid, but if a further module, say A, contains a command Require Export B, then the command Require Import A also imports the module B.

  3. Require [Import | Export] qualid1qualidn.

    This loads the modules qualid1, …, qualidn and their recursive dependencies. If Import or Export is given, it also imports qualid1, …, qualidn and all the recursive dependencies that were marked or transitively marked as Export.

  4. From dirpath Require qualid.

    This command acts as Require, but picks any library whose absolute name is of the form dirpath.dirpath.qualid for some dirpath’. This is useful to ensure that the qualid library comes from a given package by making explicit its absolute root.


Error messages:

  1. Cannot load qualid: no physical path bound to dirpath
  2. Cannot find library foo in loadpath

    The command did not find the file foo.vo. Either foo.v exists but is not compiled or foo.vo is in a directory which is not in your LoadPath (see Section 2.6.3).

  3. Compiled library ident.vo makes inconsistent assumptions over library qualid

    The command tried to load library file ident.vo that depends on some specific version of library qualid which is not the one already loaded in the current Coq session. Probably ident.v was not properly recompiled with the last version of the file containing module qualid.

  4. Bad magic number

    The file ident.vo was found but either it is not a Coq compiled module, or it was compiled with an older and incompatible version of Coq.

  5. The file ident.vo contains library dirpath and not library dirpath

    The library file dirpath’ is indirectly required by the Require command but it is bound in the current loadpath to the file ident.vo which was bound to a different library name dirpath at the time it was compiled.

  6. Require is not allowed inside a module or a module type

    This command is not allowed inside a module or a module type being defined. It is meant to describe a dependency between compilation units. Note however that the commands Import and Export alone can be used inside modules (see Section 2.5.8).


See also: Chapter 14

6.5.2  Print Libraries.

This command displays the list of library files loaded in the current Coq session. For each of these libraries, it also tells if it is imported.

6.5.3  Declare ML Module string1 .. stringn.

This commands loads the Objective Caml compiled files string1stringn (dynamic link). It is mainly used to load tactics dynamically. The files are searched into the current Objective Caml loadpath (see the command Add ML Path in the Section 2.6.3). Loading of Objective Caml files is only possible under the bytecode version of coqtop (i.e. coqtop called with options -byte, see chapter 14), or when Coq has been compiled with a version of Objective Caml that supports native Dynlink (≥ 3.11).


Variants:

  1. Local Declare ML Module string1 .. stringn.
    This variant is not exported to the modules that import the module where they occur, even if outside a section.


Error messages:

  1. File not found on loadpath : string
  2. Loading of ML object file forbidden in a native Coq

6.5.4  Print ML Modules.

This print the name of all Objective Caml modules loaded with Declare ML Module. To know from where these module were loaded, the user should use the command Locate File (see Section 6.6.10)

6.6  Loadpath

Loadpaths are preferably managed using Coq command line options (see Section 2.6.3) but there remain vernacular commands to manage them for practical purposes. Such commands are only meant to be issued in the toplevel, and using them in source files is discouraged.

6.6.1  Pwd.

This command displays the current working directory.

6.6.2  Cd string.

This command changes the current directory according to string which can be any valid path.


Variants:

  1. Cd.
    Is equivalent to Pwd.

6.6.3  Add LoadPath string as dirpath.

This command is equivalent to the command line option -Q dirpath string. It adds the physical directory string to the current Coq loadpath and maps it to the logical directory dirpath.


Variants:

  1. Add LoadPath string.
    Performs as Add LoadPath string as dirpath but for the empty directory path.

6.6.4  Add Rec LoadPath string as dirpath.

This command is equivalent to the command line option -R dirpath string. It adds the physical directory string and all its subdirectories to the current Coq loadpath.


Variants:

  1. Add Rec LoadPath string.
    Works as Add Rec LoadPath string as dirpath but for the empty logical directory path.

6.6.5  Remove LoadPath string.

This command removes the path string from the current Coq loadpath.

6.6.6  Print LoadPath.

This command displays the current Coq loadpath.


Variants:

  1. Print LoadPath dirpath.
    Works as Print LoadPath but displays only the paths that extend the dirpath prefix.

6.6.7  Add ML Path string.

This command adds the path string to the current Objective Caml loadpath (see the command Declare ML Module in the Section 6.5).

6.6.8  Add Rec ML Path string.

This command adds the directory string and all its subdirectories to the current Objective Caml loadpath (see the command Declare ML Module in the Section 6.5).

6.6.9  Print ML Path string.

This command displays the current Objective Caml loadpath. This command makes sense only under the bytecode version of coqtop, i.e. using option -byte (see the command Declare ML Module in the section 6.5).

6.6.10  Locate File string.

This command displays the location of file string in the current loadpath. Typically, string is a .cmo or .vo or .v file.

6.6.11  Locate Library dirpath.

This command gives the status of the Coq module dirpath. It tells if the module is loaded and if not searches in the load path for a module of logical name dirpath.

6.7  Backtracking

The backtracking commands described in this section can only be used interactively, they cannot be part of a vernacular file loaded via Load or compiled by coqc.

6.7.1  Reset ident.

This command removes all the objects in the environment since ident was introduced, including ident. ident may be the name of a defined or declared object as well as the name of a section. One cannot reset over the name of a module or of an object inside a module.


Error messages:

  1. ident: no such entry


Variants:

  1. Reset Initial.
    Goes back to the initial state, just after the start of the interactive session.

6.7.2  Back.

This commands undoes all the effects of the last vernacular command. Commands read from a vernacular file via a Load are considered as a single command. Proof management commands are also handled by this command (see Chapter 7). For that, Back may have to undo more than one command in order to reach a state where the proof management information is available. For instance, when the last command is a Qed, the management information about the closed proof has been discarded. In this case, Back will then undo all the proof steps up to the statement of this proof.


Variants:

  1. Back n
    Undoes n vernacular commands. As for Back, some extra commands may be undone in order to reach an adequate state. For instance Back n will not re-enter a closed proof, but rather go just before that proof.


Error messages:

  1. Invalid backtrack
    The user wants to undo more commands than available in the history.

6.7.3  BackTo num.

This command brings back the system to the state labeled num, forgetting the effect of all commands executed after this state. The state label is an integer which grows after each successful command. It is displayed in the prompt when in -emacs mode. Just as Back (see above), the BackTo command now handles proof states. For that, it may have to undo some extra commands and end on a state num′ ≤ num if necessary.


Variants:

  1. Backtrack num1 num2 num3.
    Backtrack is a deprecated form of BackTo which allows explicitly manipulating the proof environment. The three numbers num1, num2 and num3 represent the following:
    • num3: Number of Abort to perform, i.e. the number of currently opened nested proofs that must be canceled (see Chapter 7).
    • num2: Proof state number to unbury once aborts have been done. Coq will compute the number of Undo to perform (see Chapter 7).
    • num1: State label to reach, as for BackTo.


Error messages:

  1. Invalid backtrack
    The destination state label is unknown.

6.8  Quitting and debugging

6.8.1  Quit.

This command permits to quit Coq.

6.8.2  Drop.

This is used mostly as a debug facility by Coq’s implementors and does not concern the casual user. This command permits to leave Coq temporarily and enter the Objective Caml toplevel. The Objective Caml command:

#use "include";;

add the right loadpaths and loads some toplevel printers for all abstract types of Coq- section_path, identifiers, terms, judgments, …. You can also use the file base_include instead, that loads only the pretty-printers for section_paths and identifiers. You can return back to Coq with the command:

go();;


Warnings:

  1. It only works with the bytecode version of Coq (i.e. coqtop called with option -byte, see the contents of Section 14.1).
  2. You must have compiled Coq from the source package and set the environment variable COQTOP to the root of your copy of the sources (see Section 14.3.2).

6.8.3  Time command.

This command executes the vernacular command command and display the time needed to execute it.

6.8.4  Redirect "file" command.

This command executes the vernacular command command, redirecting its output to “file.out”.

6.8.5  Timeout int command.

This command executes the vernacular command command. If the command has not terminated after the time specified by the integer (time expressed in seconds), then it is interrupted and an error message is displayed.

6.8.6  Set Default Timeout int.

After using this command, all subsequent commands behave as if they were passed to a Timeout command. Commands already starting by a Timeout are unaffected.

6.8.7  Unset Default Timeout.

This command turns off the use of a default timeout.

6.8.8  Test Default Timeout.

This command displays whether some default timeout has be set or not.

6.9  Controlling display

6.9.1  Set Silent.

This command turns off the normal displaying.

6.9.2  Unset Silent.

This command turns the normal display on.

6.9.3  Set Printing Width integer.

This command sets which left-aligned part of the width of the screen is used for display.

6.9.4  Unset Printing Width.

This command resets the width of the screen used for display to its default value (which is 78 at the time of writing this documentation).

6.9.5  Test Printing Width.

This command displays the current screen width used for display.

6.9.6  Set Printing Depth integer.

This command sets the nesting depth of the formatter used for pretty-printing. Beyond this depth, display of subterms is replaced by dots.

6.9.7  Unset Printing Depth.

This command resets the nesting depth of the formatter used for pretty-printing to its default value (at the time of writing this documentation, the default value is 50).

6.9.8  Test Printing Depth.

This command displays the current nesting depth used for display.

6.10  Controlling the reduction strategies and the conversion algorithm

Coq provides reduction strategies that the tactics can invoke and two different algorithms to check the convertibility of types. The first conversion algorithm lazily compares applicative terms while the other is a brute-force but efficient algorithm that first normalizes the terms before comparing them. The second algorithm is based on a bytecode representation of terms similar to the bytecode representation used in the ZINC virtual machine [98]. It is especially useful for intensive computation of algebraic values, such as numbers, and for reflection-based tactics. The commands to fine-tune the reduction strategies and the lazy conversion algorithm are described first.

6.10.1  Opaque qualid1qualidn.

This command has an effect on unfoldable constants, i.e. on constants defined by Definition or Let (with an explicit body), or by a command assimilated to a definition such as Fixpoint, Program Definition, etc, or by a proof ended by Defined. The command tells not to unfold the constants qualid1qualidn in tactics using δ-conversion (unfolding a constant is replacing it by its definition).

Opaque has also an effect on the conversion algorithm of Coq, telling it to delay the unfolding of a constant as much as possible when Coq has to check the conversion (see Section 4.3) of two distinct applied constants.

The scope of Opaque is limited to the current section, or current file, unless the variant Global Opaque qualid1 qualidn is used.


See also: sections 8.7, 8.16, 7.1


Error messages:

  1. The reference qualid was not found in the current environment
    There is no constant referred by qualid in the environment. Nevertheless, if you asked Opaque foo bar and if bar does not exist, foo is set opaque.

6.10.2  Transparent qualid1qualidn.

This command is the converse of Opaque and it applies on unfoldable constants to restore their unfoldability after an Opaque command.

Note in particular that constants defined by a proof ended by Qed are not unfoldable and Transparent has no effect on them. This is to keep with the usual mathematical practice of proof irrelevance: what matters in a mathematical development is the sequence of lemma statements, not their actual proofs. This distinguishes lemmas from the usual defined constants, whose actual values are of course relevant in general.

The scope of Transparent is limited to the current section, or current file, unless the variant Global Transparent qualid1qualidn is used.


Error messages:

  1. The reference qualid was not found in the current environment
    There is no constant referred by qualid in the environment.


See also: sections 8.7, 8.16, 7.1

6.10.3  Strategy level [ qualid1qualidn ].

This command generalizes the behavior of Opaque and Transparent commands. It is used to fine-tune the strategy for unfolding constants, both at the tactic level and at the kernel level. This command associates a level to qualid1qualidn. Whenever two expressions with two distinct head constants are compared (for instance, this comparison can be triggered by a type cast), the one with lower level is expanded first. In case of a tie, the second one (appearing in the cast type) is expanded.

Levels can be one of the following (higher to lower):

opaque
: level of opaque constants. They cannot be expanded by tactics (behaves like +∞, see next item).
num
: levels indexed by an integer. Level 0 corresponds to the default behavior, which corresponds to transparent constants. This level can also be referred to as transparent. Negative levels correspond to constants to be expanded before normal transparent constants, while positive levels correspond to constants to be expanded after normal transparent constants.
expand
: level of constants that should be expanded first (behaves like −∞)

These directives survive section and module closure, unless the command is prefixed by Local. In the latter case, the behavior regarding sections and modules is the same as for the Transparent and Opaque commands.

6.10.4  Print Strategy qualid.

This command prints the strategy currently associated to qualid. It fails if qualid is not an unfoldable reference, that is, neither a variable nor a constant.


Error messages:

  1. The reference is not unfoldable.


Variants:

  1. Print Strategies
    Print all the currently non-transparent strategies.

6.10.5  Declare Reduction ident := convtactic.

This command allows giving a short name to a reduction expression, for instance lazy beta delta [foo bar]. This short name can then be used in Eval ident in ... or eval directives. This command accepts the Local modifier, for discarding this reduction name at the end of the file or module. For the moment the name cannot be qualified. In particular declaring the same name in several modules or in several functor applications will be refused if these declarations are not local. The name ident cannot be used directly as an Ltac tactic, but nothing prevent the user to also perform a Ltac ident := convtactic.


See also: sections 8.7

6.11  Controlling the locality of commands

6.11.1  Local, Global

Some commands support a Local or Global prefix modifier to control the scope of their effect. There are four kinds of commands: