\[\begin{split}\newcommand{\as}{\kw{as}} \newcommand{\case}{\kw{case}} \newcommand{\cons}{\textsf{cons}} \newcommand{\consf}{\textsf{consf}} \newcommand{\emptyf}{\textsf{emptyf}} \newcommand{\End}{\kw{End}} \newcommand{\kwend}{\kw{end}} \newcommand{\even}{\textsf{even}} \newcommand{\evenO}{\textsf{even}_\textsf{O}} \newcommand{\evenS}{\textsf{even}_\textsf{S}} \newcommand{\Fix}{\kw{Fix}} \newcommand{\fix}{\kw{fix}} \newcommand{\for}{\textsf{for}} \newcommand{\forest}{\textsf{forest}} \newcommand{\Functor}{\kw{Functor}} \newcommand{\In}{\kw{in}} \newcommand{\ind}[3]{\kw{Ind}~[#1]\left(#2\mathrm{~:=~}#3\right)} \newcommand{\Indp}[4]{\kw{Ind}_{#4}[#1](#2:=#3)} \newcommand{\Indpstr}[5]{\kw{Ind}_{#4}[#1](#2:=#3)/{#5}} \newcommand{\injective}{\kw{injective}} \newcommand{\kw}[1]{\textsf{#1}} \newcommand{\length}{\textsf{length}} \newcommand{\letin}[3]{\kw{let}~#1:=#2~\kw{in}~#3} \newcommand{\List}{\textsf{list}} \newcommand{\lra}{\longrightarrow} \newcommand{\Match}{\kw{match}} \newcommand{\Mod}[3]{{\kw{Mod}}({#1}:{#2}\,\zeroone{:={#3}})} \newcommand{\ModImp}[3]{{\kw{Mod}}({#1}:{#2}:={#3})} \newcommand{\ModA}[2]{{\kw{ModA}}({#1}=={#2})} \newcommand{\ModS}[2]{{\kw{Mod}}({#1}:{#2})} \newcommand{\ModType}[2]{{\kw{ModType}}({#1}:={#2})} \newcommand{\mto}{.\;} \newcommand{\nat}{\textsf{nat}} \newcommand{\Nil}{\textsf{nil}} \newcommand{\nilhl}{\textsf{nil\_hl}} \newcommand{\nO}{\textsf{O}} \newcommand{\node}{\textsf{node}} \newcommand{\nS}{\textsf{S}} \newcommand{\odd}{\textsf{odd}} \newcommand{\oddS}{\textsf{odd}_\textsf{S}} \newcommand{\ovl}[1]{\overline{#1}} \newcommand{\Pair}{\textsf{pair}} \newcommand{\plus}{\mathsf{plus}} \newcommand{\SProp}{\textsf{SProp}} \newcommand{\Prop}{\textsf{Prop}} \newcommand{\return}{\kw{return}} \newcommand{\Set}{\textsf{Set}} \newcommand{\Sort}{\mathcal{S}} \newcommand{\Str}{\textsf{Stream}} \newcommand{\Struct}{\kw{Struct}} \newcommand{\subst}[3]{#1\{#2/#3\}} \newcommand{\tl}{\textsf{tl}} \newcommand{\tree}{\textsf{tree}} \newcommand{\trii}{\triangleright_\iota} \newcommand{\Type}{\textsf{Type}} \newcommand{\WEV}[3]{\mbox{$#1[] \vdash #2 \lra #3$}} \newcommand{\WEVT}[3]{\mbox{$#1[] \vdash #2 \lra$}\\ \mbox{$ #3$}} \newcommand{\WF}[2]{{\mathcal{W\!F}}(#1)[#2]} \newcommand{\WFE}[1]{\WF{E}{#1}} \newcommand{\WFT}[2]{#1[] \vdash {\mathcal{W\!F}}(#2)} \newcommand{\WFTWOLINES}[2]{{\mathcal{W\!F}}\begin{array}{l}(#1)\\\mbox{}[{#2}]\end{array}} \newcommand{\with}{\kw{with}} \newcommand{\WS}[3]{#1[] \vdash #2 <: #3} \newcommand{\WSE}[2]{\WS{E}{#1}{#2}} \newcommand{\WT}[4]{#1[#2] \vdash #3 : #4} \newcommand{\WTE}[3]{\WT{E}{#1}{#2}{#3}} \newcommand{\WTEG}[2]{\WTE{\Gamma}{#1}{#2}} \newcommand{\WTM}[3]{\WT{#1}{}{#2}{#3}} \newcommand{\zeroone}[1]{[{#1}]} \end{split}\]

Recent changes

Unreleased changes

Kernel

Specification language, type inference

Notations

Tactics

Ltac language

Ltac2 language

SSReflect

Commands and options

Command-line tools

CoqIDE

Standard library

Infrastructure and dependencies

Extraction

Miscellaneous

Version 8.20

Summary of changes

Coq version 8.20 adds a new rewrite rule mechanism along with a few new features, a host of improvements to the virtual machine, the notation system, Ltac2 and the standard library.

We highlight some of the most impactful changes here:

Notable breaking changes:

  • Syntactic global references passed through the using clauses of auto-like tactics are now handled as plain references rather than interpreted terms. In particular, their typeclass arguments will not be inferred. In general, the previous behaviour can be emulated by replacing auto using foo with pose proof foo; auto.

  • Argument order for the Ltac2 combinators List.fold_left2 and List.fold_right2 changed to be the same as in OCaml.

  • Importing a module containing a mutable Ltac2 definition does not undo its mutations. Replace Ltac2 mutable foo := some_expr. with Ltac2 mutable foo := some_expr. Ltac2 Set foo := some_expr. to recover the previous behaviour.

  • Some renaming in the standard library. Deprecations are provided for a smooth transition.

See the Changes in 8.20.0 section below for the detailed list of changes, including potentially breaking changes marked with Changed. Coq's reference manual for 8.20, documentation of the 8.20 standard library and developer documentation of the 8.20 ML API are also available.

Théo Zimmermann with help from Ali Caglayan and Jason Gross maintained coqbot used to run Coq's CI and other pull request management tasks.

Jason Gross maintained the bug minimizer and its automatic use through coqbot.

Erik Martin-Dorel maintained the Coq Docker images and the docker-keeper compiler used to build and keep those images up to date (note that the tool is not Coq specific). Cyril Cohen, Vincent Laporte, Pierre Roux and Théo Zimmermann maintained the Nix toolbox used by many Coq projects for continuous integration.

Ali Caglayan, Emilio Jesús Gallego Arias, Rudi Grinberg and Rodolphe Lepigre maintained the Dune build system for OCaml and Coq used to build Coq itself and many Coq projects.

The opam repository for Coq packages has been maintained by Guillaume Claret, Guillaume Melquiond, Karl Palmskog and Enrico Tassi with contributions from many users. A list of packages is available on the Coq website.

Coq 8.20 was made possible thanks to the following reviewers: Frédéric Besson, Lasse Blaauwbroek, Ali Caglayan, Cyril Cohen, Andrej Dudenhefner, Andres Erbsen, Jim Fehrle, Emilio Jesús Gallego Arias, Gaëtan Gilbert, Jason Gross, Hugo Herbelin, Ralf Jung, Jan-Oliver Kaiser, Chantal Keller, Olivier Laurent, Rodolphe Lepigre, Yishuai Li, Ralph Matthes, Guillaume Melquiond, Pierre-Marie Pédrot, Karl Palmskog, Clément Pit-Claudel, Pierre Rousselin, Pierre Roux, Michael Soegtrop, soukouki, Matthieu Sozeau, Nicolas Tabareau, Enrico Tassi, Niels van der Weide, Nickolai Zeldovich and Théo Zimmermann. See the Coq Team face book page for more details on Coq's development team.

The 59 contributors to the 8.20 version are: Timur Aminev, Frédéric Besson, Lasse Blaauwbroek, Björn Brandenburg, Ali Caglayan, Nikolaos Chatzikonstantinou, Sylvain Chiron, chluebi, Cyril Cohen, Anton Danilkin, Louise Dubois de Prisque, Andrej Dudenhefner, Maxime Dénès, Andres Erbsen, Jim Fehrle, Davide Fissore, Andreas Florath, Yannick Forster, Mario Frank, Gaëtan Gilbert, Georges Gonthier, Jason Gross, Stefan Haan, Hugo Herbelin, Lennart Jablonka, Emilio Jesús Gallego Arias, Ralf Jung, Jan-Oliver Kaiser, Evgenii Kosogorov, Rodolphe Lepigre, Yann Leray, David M. Cooke, Erik Martin-Dorel, Guillaume Melquiond, Guillaume Munch-Maccagnoni, Karl Palmskog, Julien Puydt, Pierre-Marie Pédrot, Ramkumar Ramachandra, Pierre Rousselin, Pierre Roux, Kazuhiko Sakaguchi, Bernhard Schommer, Remy Seassau, Matthieu Sozeau, Enrico Tassi, Romain Tetley, Laurent Théry, Alexey Trilis, Oliver Turner, Quentin Vermande, Li-yao Xia and Théo Zimmermann,

The Coq community at large helped improve this new version via the GitHub issue and pull request system, the coq-club@inria.fr mailing list, the Discourse forum and the Coq Zulip chat.

Version 8.20's development spanned 7 months from the release of Coq 8.19.0 (9 months since the branch for 8.19.0). Pierre Roux and Guillaume Melquiond are the release managers of Coq 8.20. This release is the result of 470 merged PRs, closing 113 issues.

Toulouse, September 2024
Pierre Roux and Guillaume Melquiond for the Coq development team

Changes in 8.20.0

Kernel

  • Changed: The guard checker now recognizes uniform parameters of a fixpoint and treats their instances as constant over the recursive call (#17986, grants #16040, by Hugo Herbelin).

  • Added: A mechanism to add user-defined rewrite rules to Coq's reduction mechanisms; see chapter User-defined rewrite rules (#18038, by Yann Leray).

  • Added: Support for primitive strings in terms (#18973, by Rodolphe Lepigre).

Specification language, type inference

  • Changed: warnings future-coercion-class-constructor and future-coercion-class-field about :> in Class as errors by default. This offers a last opportunity to replace :> with :: (available since Coq 8.18) to declare typeclass instances before making :> consistently declare coercions in all records in next version. To adapt huge codebases, you can try this script or the one below. But beware that both are incomplete.

      #!/bin/awk -f
      BEGIN {
        startclass = 0;
        inclass = 0;
        indefclass = 0;  # definitionalclasses (single field, without { ... })
      }
      {
        if ($0 ~ "[ ]*Class") {
          startclass = 1;
        }
        if (startclass == 1 && $0 ~ ":=") {
          inclass = 1;
          indefclass = 1;
        }
        if (startclass == 1 && $0 ~ ":=.*{") {
          indefclass = 0;
        }
        if (inclass == 1) startclass = 0;
    
        if (inclass == 1 && $0 ~ ":>") {
          if ($0 ~ "{ .*:>") {  # first field on a single line
            sub("{ ", "{ #[global] ");
          } else if ($0 ~ ":=.*:>") {  # definitional classes on a single line
            sub(":= ", ":= #[global] ");
          } else if ($0 ~ "^  ") {
            sub("  ", "  #[global] ");
          } else {
            $0 = "#[global] " $0;
          }
          sub(":>", "::")
        }
      print $0;
    
      if ($0 ~ ".*}[.]" || indefclass == 1 && $0 ~ "[.]$") inclass = 0;
    }
    

    (#18590, by Pierre Roux).

  • Changed: Mutually-proved theorems with statements in different coinductive types now supported (#18743, by Hugo Herbelin).

  • Added: CoFixpoint supports attributes bypass_guard, clearbody, deprecated and warn (#18754, by Hugo Herbelin).

  • Added: Program Fixpoint with measure or wf (see Program Fixpoint) now supports the where clause for notations, the local and clearbody attributes, as well as non-atomic conclusions (#18834, by Hugo Herbelin, fixes in particular #13812 and #14841).

  • Fixed: Anomaly on the absence of remaining obligations of some name now an error (#18873, fixes #3889, by Hugo Herbelin).

  • Fixed: Universe polymorphic Program's obligations are now generalized only over the universe variables that effectively occur in the obligation (#18915, fixes #11766 and #11988, by Hugo Herbelin).

  • Fixed: Anomaly assertion failed in pattern-matching compilation, with Program Mode or with let-ins in the arity of an inductive type (#18921, fixes #5777 and #11030 and #11586, by Hugo Herbelin).

  • Fixed: Support for Program-style pattern-matching on more than one argument in an inductive family (#18929, fixes #1956 and #5777, by Hugo Herbelin).

  • Fixed: anomaly with obligations in the binders of a measure- or wf-based Program Fixpoint (#18958, fixes #18920, by Hugo Herbelin).

  • Fixed: Incorrect registration of universe names attached to a primitive polymorphic constant (#19100, fixes #19099, by Hugo Herbelin).

Notations

  • Changed: an only printing interpretation of a notation with a specific format does no longer change the printing rule of other interpretations of the notation; to globally change the default printing rule of all interpretations of a notation, use Reserved Notation instead (#16329, fixes #16262, by Hugo Herbelin).

  • Changed: levels of Reserved Notation now default to levels of previous notations with longest common prefix, if any. This helps to factorize notations with common prefixes (#19149, by Pierre Roux).

  • Added: closed-notation-not-level-0 and postfix-notation-not-level-1 warnings about closed and postfix notations at unusual levels (#18588, by Pierre Roux).

  • Added: notation-incompatible-prefix warning when two notation definitions have incompatible prefixes (#19049, by Pierre Roux).

  • Fixed: Notations for applied constants equipped with multiple signatures of implicit arguments were not correctly inserting as many maximal implicit arguments as they should have (#18445, by Hugo Herbelin).

  • Fixed: Add support for printing notations applied to extra arguments in custom entries, thus eliminating an anomaly (#18447, fixes #18342, by Hugo Herbelin).

Tactics

  • Changed: When using Z.to_euclidean_division_equations, nia can now relate Z.div/Z.modulo to Z.quot/Z.rem a bit better, by virtue of being noticing when there are two equations of the form x = y * q₁ + _ and x = y * q₂ + _ (or minor variations thereof), suggesting that q₁ = q₂. Users can replace Z.to_euclidean_division_equations with let flags := Z.euclidean_division_equations_flags.default_with Z.euclidean_division_equations_flags.find_duplicate_quotients false in Z.to_euclidean_division_equations_with flags or, using Import Z.euclidean_division_equations_flags., with Z.to_euclidean_division_equations_with ltac:(default_with find_duplicate_quotients false) (#17934, by Jason Gross).

  • Changed: The opacity/transparency of primitive projections is now attached to the projections themselves, not the compatibility constants, and compatibility constants are always considered transparent (#18327, fixes #18281, by Jan-Oliver Kaiser and Rodolphe Lepigre).

  • Changed: Tactic intro z on an existential variable goal forces the resolution of the existential variable into a goal forall z:?T, ?P, which becomes ?P in context z:?T after introduction. The existential variable ?P itself is now defined in a context where the variable of type ?T is also named z, as specified by intro instead of x as it was conventionally the case before (#18395, by Hugo Herbelin).

  • Changed: syntactic global references passed through the using clauses of auto-like tactics are now handled as plain references rather than interpreted terms. In particular, their typeclass arguments will not be inferred. In general, the previous behaviour can be emulated by replacing auto using foo with pose proof foo; auto (#18909, by Pierre-Marie Pédrot).

  • Changed: Use Coqlib's Register mechanism for the generalized rewriting tactic and make the (C)RelationClasses/(C)Morphisms independent of the rewrite tactic to ease maintainance. (#19115, by Matthieu Sozeau).

  • Removed: the clear modifier which was deprecated since 8.17 (#18887, by Pierre-Marie Pédrot).

  • Removed: the cutrewrite tactic, which was deprecated since Coq 8.5 (#19027, by Pierre-Marie Pédrot).

  • Deprecated: non-reference hints in using clauses of auto-like tactics (#19006, by Pierre-Marie Pédrot).

  • Deprecated: the gintuition tactic, which used to be undocumented until Coq 8.16 (#19129, by Pierre-Marie Pédrot).

  • Deprecated: destauto, see #11537 (#19179, by Jim Fehrle).

  • Added: When using Z.to_euclidean_division_equations, you can now pose equations of the form x = y * q using Z.divide (#17927, by Evgenii Kosogorov).

  • Added: support for Nat.double and Nat.div2 to zify and lia (#18729, by Andres Erbsen).

  • Added: the replace tactic now accepts -> and <- to specify the direction of the replacement when used with a with clause (#19060, fixes #13480, by Pierre-Marie Pédrot).

  • Fixed: The name of a cofixpoint globally defined with a name is now systematically reused by simpl after reduction, even when the named cofixpoint is mutually defined or defined in a section (#18576, fixes #4056, by Hugo Herbelin).

  • Fixed: The reduction of primitive projections of cofixpoints by simpl is now implemented (#18577, fixes #7982, by Hugo Herbelin).

  • Fixed: Support for refolding reduced global mutual fixpoints/cofixpoints with parameters in cbn (#18601, fixes part of #4056, by Hugo Herbelin).

  • Fixed: cbn was leaving behind unnamable constants when refolding mutual fixpoints/cofixpoints from aliased modules (#18616, fixes #17897, by Hugo Herbelin).

  • Fixed: cbv of primitive projections applied to a tuple now ignores beta like it does for cbn, lazy and simpl (#18618, fixes #9086, by Hugo Herbelin).

Ltac language

Ltac2 language

  • Changed: recursive let and non mutable projections of syntactic values are considered syntactic values (#18411, by Gaëtan Gilbert).

  • Changed: Ltac2 notations are typechecked at declaration time by default. This should produce better errors when a notation argument does not have the expected type (e.g. wrong branch type in match! goal). In the previous behaviour of typechecking, only the expansion result can be recovered using Ltac2 Typed Notations. We believe there are no real use cases for this, please report if you have any (#18432, fixes #17477, by Gaëtan Gilbert).

  • Changed: argument order for the Ltac2 combinators List.fold_left2 and List.fold_right2 changed to be the same as in OCaml (#18706, by Gaëtan Gilbert).

  • Changed: Importing a module containing a mutable Ltac2 definition does not undo its mutations. Replace Ltac2 mutable foo := some_expr. with Ltac2 mutable foo := some_expr. Ltac2 Set foo := some_expr. to recover the previous behaviour (#18713, by Gaëtan Gilbert).

  • Changed: the using clause argument of auto-like tactics in Ltac2 now take a global reference rather than arbitrary constr (#18940, by Pierre-Marie Pédrot).

  • Deprecated: Ltac2.Constr.Pretype.Flags.open_constr_flags whose name is misleading as it runs typeclass inference unlike open_constr:() (#18765, by Gaëtan Gilbert).

  • Added: fst and snd in Ltac2.Init (#18370, by Gaëtan Gilbert).

  • Added: Ltac2.Ltac1.of_preterm and to_preterm (#18551, by Gaëtan Gilbert).

  • Added: of_intro_pattern and to_intro_pattern in Ltac2.Ltac1 (#18558, by Gaëtan Gilbert).

  • Added: basic APIs in Ltac2.Ltac1 to produce slightly more informative errors when failing to convert a Ltac1 value to some Ltac2 type (#18558, by Gaëtan Gilbert).

  • Added: APIs Ltac2.Control.unshelve and Ltac2.Notations.unshelve (#18604, by Gaëtan Gilbert).

  • Added: warning on unused Ltac2 variables (except when starting with _) (#18641, by Gaëtan Gilbert).

  • Added: Ltac2.Control.numgoals (#18690, by Gaëtan Gilbert).

  • Added: intropattern and intropatterns notation scopes support views (foo%bar) (#18757, by Gaëtan Gilbert).

  • Added: open recursion combinators in Ltac2.Constr.Unsafe (#18764, by Gaëtan Gilbert).

  • Added: APIs in Ltac2.Constr.Pretype.Flags to customize pretyping flags. (#18765, by Gaëtan Gilbert).

  • Added: abstract attribute for Ltac2 Type to turn types abstract at the end of the current module (#18766, fixes #18656, by Gaëtan Gilbert).

  • Added: APIs in Ltac2.Message to interact with the boxing system of the pretty printer (#18988, by Gaëtan Gilbert).

  • Added: Automatic Proposition Inductives, Dependent Proposition Eliminators and warning when automatically lowering an inductive declared with Type to Prop (#18989, by Gaëtan Gilbert).

  • Added: String.sub (#19204, by Rodolphe Lepigre).

  • Fixed: Ltac2.Control.new_goal removes the new goal from the shelf and future goals (#19141, fixes #19138, by Gaëtan Gilbert).

SSReflect

  • Changed: ssreflect no longer relies on the recovery mechanism of the parsing engine, this can slightly change the parsing priorities in rare occurences, for instance when combining unshelve and => (#18224, by Pierre Roux).

  • Changed: notations _.1 and _.2 are now defined in the prelude at level 1 rather than in ssrfun at level 2 (#18224, by Pierre Roux).

  • Changed: The have tactic generates a proof term containing an opaque constant, as it did up to PR #15121 included in Coq 8.16.0. See the variant have @H to generate a (transparent) let-in instead (Generating let in context entries with have). (#18449, fixes #18017, by Enrico Tassi).

  • Deprecated: The fun_scope notation scope declared in ssrfun.v is deprecated. Use function_scope instead (#18374, by Kazuhiko Sakaguchi).

  • Fixed: handling of primitive projections in ssrewrite (#19213, fixes #19229, by Pierre Roux, Kazuhiko Sakaguchi, Enrico Tassi and Quentin Vermande).

Commands and options

  • Changed: the default reversibility status of most coercions. The refman states that

    By default coercions are not reversible except for Record fields specified using :>.

    The previous code was making way too many coercion reversible by default. The new behavior should be closer from the spec in the doc (#18705, by Pierre Roux).

  • Changed: focus commands such as 1:{ and goal selection for query commands such as 1: Check do not need Classic (Ltac1) proof mode to function. In particular they function in Ltac2 mode (#18707, fixes #18351, by Gaëtan Gilbert).

  • Changed: inductives declared with : Type or no annotation and automatically put in Prop are not declared template polymorphic (#18867, by Gaëtan Gilbert).

  • Changed: Clarify the warning about use of Let, Variable, Hypothesis and Context outside sections and make it an error by default (#18880, by Pierre Roux).

  • Changed: The "fragile-hint-constr" warning is now an error by default, as the corresponding feature will be removed in a later version (#18895, by Pierre-Marie Pédrot).

  • Changed: Scheme automatically registers the resulting schemes in the Register Scheme database (#19016, fixes #3132, by Gaëtan Gilbert).

  • Changed: Typeclasses Transparent and Typeclasses Opaque default locality outside section is now export (#19069, by Gaëtan Gilbert).

  • Deprecated: The Cd command. Instead use the command line option -output-directory (see Command line options) or, for extraction, Extraction Output Directory (#17403, by Ali Caglayan and Hugo Herbelin).

  • Added: warn attribute generalizing the deprecation machinery to other forms of comments (#18248, by Hugo Herbelin and Pierre Roux).

  • Added: Register Scheme to add entries to the scheme database used by some tactics (#18299, by Gaëtan Gilbert).

  • Added: Print reference now shows the implicit arguments of a reference directly on the type of reference, using {...} and [...] markers for respectively maximally-inserted and non-maximally-inserted implicit arguments, as About does (#18444, by Hugo Herbelin).

  • Added: import_categories supports category options controlling Flags, Options and Tables (#18536, by Gaëtan Gilbert).

  • Added: When a name is a projection, About and Print now indicate it (#18725, by Hugo Herbelin).

  • Added: Hint Projections command that sets the transparency flag for projections for the specified hint databases (#18785, by Jan-Oliver Kaiser and Rodolphe Lepigre).

  • Added: Search now admits the is:Fixpoint and is:CoFixpoint logical kinds to search for constants defined with the Fixpoint and CoFixpoint keywords (#18983, by Pierre Rousselin).

  • Added: The Include command can now include module types with a with clause (with_declaration) to instantiate some parameters (#19144, by Pierre Rousselin).

  • Fixed: Fixes missing implicit arguments coming after a -> in the main type printed by Print and About (#18442, fixes #15020, by Hugo Herbelin).

  • Fixed: Cumulativity Weak Constraints can unify universes to Set when Universe Minimization ToSet is enabled (#18458, by Gaëtan Gilbert).

  • Fixed: Search with modifier is:Scheme restricted the search to inductive types which have schemes instead of the schemes themselves. For instance Search nat is:Scheme with just the prelude loaded would return le i.e. the only inductive type whose type mentions nat (#18537, fixes #18298, by Gaëtan Gilbert).

  • Fixed: Search now searches also in included module types (#18662, fixes #18657, by Hugo Herbelin).

  • Fixed: Eval and Definition with := Eval work without needing to load the Ltac plugin (#18852, fixes #12948, by Gaëtan Gilbert).

  • Fixed: Scheme declares non-recursive schemes for scheme_type Case and Elimination (#19017, fixes #10816, by Gaëtan Gilbert).

  • Fixed: Cumulativity Weak Constraints had its meaning flipped since 8.12 (#19201, by Gaëtan Gilbert).

Command-line tools

  • Changed: signal SIGINT interrupts the process with " "user interrupt" error instead of aborting. This is intended to produce better messages when interrupting Coq (#18716, by Gaëtan Gilbert).

  • Added: Command line option -output-directory dir to set the default output directory for extraction, Redirect and Print Universes (#17392, fixes #8649, by Hugo Herbelin).

  • Fixed: coqdoc links to section variables introduced with Context (#18527, fixes #18516, by Pierre Roux).

CoqIDE

  • Changed: Find/replace UI was improved: margins, icons for found/not found (#18523, fixes #11024, by Sylvain Chiron).

  • Changed: The default key binding modifier for the Navigation menu was changed to Alt on non-macOS systems. The previous default, Ctrl, hid some conventional cursor movement bindings such as Ctrl-Left, Ctrl-Right, Ctrl-Home and Ctrl-End. The new default generally has no effect if you've previously installed Coq on your system. See Shortcuts to change the default.

    The Edit/Undo key binding was changed from Ctrl-U to Ctrl-Z to be more consistent with common conventions. View/Previous Tab and View/Next Tab were changed from Alt-Left/Right to Ctrl-PgUp/PgDn (Cmd-PgUp/PgDn on macOS). To change key bindings on your system (e.g. back to Ctrl-U), see Key bindings (#18717, by Sylvain Chiron).

  • Changed: Changing modifiers for the View menu only applies to toggleable items; View/Show Proof was changed to Shift-F2 (#18717, by Sylvain Chiron).

  • Added: Edit/Select All and Navigation/Fully Check menu items (#18717, fixes #16141, by Sylvain Chiron).

  • Fixed: Opening a file with drag and drop now works correctly (fixed regression) (#18524, fixes #3977, by Sylvain Chiron).

  • Fixed: Incorrect highlight locations and line numbers for errors and warnings, especially in the presence of unicode characters. This updates the XML protocol (#19040, fixes #18682, by Hugo Herbelin).

  • Fixed: Show tooltips for syntax errors (#19153, fixes #19152, by Jim Fehrle).

Standard library

  • Changed: names of "push" lemmas for List.length to follow the same convention as push lemmas for other operations. For example, app_length became length_app. The standard library was migrated using the following script:

    find theories -name '*.v' | xargs sed -i -E '
      s/\<app_length\>/length_app/g;
      s/\<rev_length\>/length_rev/g;
      s/\<map_length\>/length_map/g;
      s/\<fold_left_length\>/fold_left_S_O/g;
      s/\<split_length_l\>/length_fst_split/g;
      s/\<split_length_r\>/length_snd_split/g;
      s/\<combine_length\>/length_combine/g;
      s/\<prod_length\>/length_prod/g;
      s/\<firstn_length\>/length_firstn/g;
      s/\<skipn_length\>/length_skipn/g;
      s/\<seq_length\>/length_seq/g;
      s/\<concat_length\>/length_concat/g;
      s/\<flat_map_length\>/length_flat_map/g;
      s/\<list_power_length\>/length_list_power/g;
    '
    

    (#18564, by Andres Erbsen).

  • Changed: Coq.CRelationClasses.arrow, Coq.CRelationClasses.iffT and Coq.CRelationClasses.flip are now Typeclasses Opaque (#18910, by Pierre-Marie Pédrot).

  • Removed: The library files Coq.NArith.Ndigits, Coq.NArith.Ndist, and Coq.Strings.ByteVector which were deprecated since 8.19 (#18936, by Andres Erbsen).

  • Deprecated: The library files

    • Coq.Numbers.Integer.Binary.ZBinary

    • Coq.Numbers.Integer.NatPairs.ZNatPairs

    • Coq.Numbers.Natural.Binary.NBinary

    have been deprecated. Users should require Coq.Arith.PeanoNat or Coq.Arith.NArith.BinNat if they want implementations of natural numbers and Coq.Arith.ZArith.BinInt if they want an implementation of integers (#18500, by Pierre Rousselin).

  • Deprecated: The library file Coq.Numbers.NatInt.NZProperties is deprecated. Users can require Coq.Numbers.NatInt.NZMulOrder instead and replace the module NZProperties.NZProp with NZMulOrder.NZMulOrderProp (#18501, by Pierre Rousselin).

  • Deprecated: The library file Coq.Arith.Bool_nat has been deprecated (#18538, by Pierre Rousselin).

  • Deprecated: The library file Coq.Numbers.NatInt.NZDomain is deprecated (#18539, by Pierre Rousselin).

  • Deprecated: The library files Coq.Numbers.Integers.Abstract.ZDivEucl and Coq.ZArith.Zeuclid are deprecated (#18544, by Pierre Rousselin).

  • Deprecated: The library files Coq.Numbers.Natural.Abstract.NIso and Coq.Numbers.Natural.Abstract.NDefOps are deprecated (#18668, by Pierre Rousselin).

  • Deprecated: Bool.Bvector. Users are encouraged to consider list bool instead. Please open an issue if you would like to keep using Bvector. (#18947, by Andres Erbsen).

  • Added: A warning on Vector.t to make its new users aware that using this dependently typed representation of fixed-length lists is more technically difficult, compared to bundling lists with a proof of their length. This is not a deprecation and there is no intent to remove it from the standard library. Use option -w -stdlib-vector to silence the warning (#18032, by Pierre Roux, reviewed by Andres Erbsen, Jim Fehrle, Emilio Jesús Gallego Arias, Gaëtan Gilbert, Hugo Herbelin, Olivier Laurent, Yishuai Li, Pierre-Marie Pédrot and Michael Soegtrop).

  • Added: lemmas NoDup_app, NoDup_iff_ForallOrdPairs, NoDup_map_NoDup_ForallPairs and NoDup_concat (#18172, by Stefan Haani and Andrej Dudenhefner).

  • Added: lemmas In_iff_nth_error nth_error_app, nth_error_cons_0, nth_error_cons_succ, nth_error_rev, nth_error_firstn, nth_error_skipn, hd_error_skipn, nth_error_seq (#18563, by Andres Erbsen)

  • Added: to N and Nat lemmas strong_induction_le, binary_induction, strong_induction_le, even_even, odd_even, odd_odd, even_odd, b2n_le_1, testbit_odd_succ', testbit_even_succ', testbit_div2, div2_0, div2_1, div2_le_mono, div2_even, div2_odd', le_div2_diag_l, div2_le_upper_bound, div2_le_lower_bound, lt_div2_diag_l, le_div2, lt_div2, div2_decr, land_even_l, land_even_r, land_odd_l, land_odd_r, land_even_even, land_odd_even, land_even_odd, land_odd_odd, land_le_l, land_le_r, ldiff_even_l, ldiff_odd_l, ldiff_even_r, ldiff_odd_r, ldiff_even_even, ldiff_odd_even, ldiff_even_odd, ldiff_odd_odd, ldiff_le_l, shiftl_lower_bound, shiftr_upper_bound, ones_0, ones_succ, pow_lower_bound (#18628, by Pierre Rousselin).

  • Fixed: Z.euclidean_division_equations_cleanup has been reordered so that zify (and lia, nia, etc) are no longer as slow when the context contains many assumptions of the form 0 <= ... < ... (#18818, fixes #18770, by Jason Gross).

Infrastructure and dependencies

  • Changed: Bump minimal Dune version required to build Coq to 3.6.1 (#18359, by Emilio Jesus Gallego Arias).

  • Removed: Support for .vio files and for .vio2vo transformation has been removed, compilation to .vos is the supported method for quick compilation now (#18424, fixes #4007 and #4013 and #4123 and #5308 and #5223 and #6720 and #8402 and #9637 and #11471 and #18380, by Emilio Jesus Gallego Arias).

  • Added: The coq-doc opam / Dune package will now build and install Coq's documentation (#17808, by Emilio Jesus Gallego Arias).

  • Added: Coq is now compatible with memprof-limits interruption methods. This means that Coq will be recompiled when the library is installed / removed from an OPAM switch. (#18906, fixes #17760, by Emilio Jesus Gallego Arias).

  • Added: ability to exit from Drop. in Coq toplevel by a simple Ctrl + D, without leaving the OCaml toplevel on the stack. Also add a custom OCaml toplevel directory #go which does the same action as go (), but with a more native syntax (#18771, by Anton Danilkin).

Extraction

  • Added: Extension for OCaml extraction: Commands to extract foreign function calls to C (external) and ML function exposition (Callback.register) for calling being able to call them by C functions (#18270, fixes #18212, by Mario Frank).

  • Fixed: Wrongly self-referencing extraction of primitive projections to OCaml in functors (#17321, fixes #16288, by Hugo Herbelin). Note that OCaml wrappers assuming that the applicative syntax of projections is provided may have to use the dot notation instead.

Version 8.19

Summary of changes

Coq version 8.19 extends the kernel universe polymorphism to polymorphism over sorts (e.g. Prop, SProp) along with a few new features, a host of improvements to the notation system, the Ltac2 standard library, and the removal of some standard library files after a long deprecation period.

We highlight some of the most impactful changes here:

  • Sort polymorphism makes it possible to share common constructs over Type Prop and SProp.

  • The notation term%_scope to set a scope only temporarily (in addition to term%scope for opening a scope applying to all subterms).

  • lazy, simpl, cbn and cbv and the associated Eval and eval reductions learned to do head reduction when given flag head.

  • New Ltac2 APIs, improved Ltac2 exact and dynamic building of Ltac2 term patterns.

  • New performance evaluation facilities: Instructions to count CPU instructions used by a command (Linux only) and Profiling system to produce trace files.

  • New command Attributes to assign attributes such as deprecated to a library file.

Notable breaking changes:

  • replace with by tac does not automatically attempt to solve the generated equality subgoal using the hypotheses. Use by first [assumption | symmetry;assumption | tac] if you need the previous behaviour.

  • Removed old deprecated files from the standard library.

See the Changes in 8.19.0 section below for the detailed list of changes, including potentially breaking changes marked with Changed. Coq's reference manual for 8.19, documentation of the 8.19 standard library and developer documentation of the 8.19 ML API are also available.

Maxime Dénès and Thierry Martinez with support from Erik Martin-Dorel and Théo Zimmermann moved the CI away from gitlab.com to use Inria supported runner machines through gitlab.inria.fr.

Théo Zimmermann with help from Ali Caglayan and Jason Gross maintained coqbot used to run Coq's CI and other pull request management tasks.

Jason Gross maintained the bug minimizer and its automatic use through coqbot.

Jaime Arias and Erik Martin-Dorel maintained the Coq Docker images and Cyril Cohen, Vincent Laporte, Pierre Roux and Théo Zimmermann maintained the Nix toolbox used by many Coq projects for continuous integration.

Ali Caglayan, Emilio Jesús Gallego Arias, Rudi Grinberg and Rodolphe Lepigre maintained the Dune build system for OCaml and Coq used to build Coq itself and many Coq projects.

The opam repository for Coq packages has been maintained by Guillaume Claret, Guillaume Melquiond, Karl Palmskog and Enrico Tassi with contributions from many users. A list of packages is available on the Coq website.

Our current maintainers are Yves Bertot, Frédéric Besson, Ana Borges, Ali Caglayan, Tej Chajed, Cyril Cohen, Pierre Corbineau, Pierre Courtieu, Andres Erbsen, Jim Fehrle, Julien Forest, Emilio Jesús Gallego Arias, Gaëtan Gilbert, Georges Gonthier, Benjamin Grégoire, Jason Gross, Hugo Herbelin, Vincent Laporte, Olivier Laurent, Assia Mahboubi, Kenji Maillard, Guillaume Melquiond, Pierre-Marie Pédrot, Clément Pit-Claudel, Pierre Roux, Kazuhiko Sakaguchi, Vincent Semeria, Michael Soegtrop, Arnaud Spiwack, Matthieu Sozeau, Enrico Tassi, Laurent Théry, Anton Trunov, Li-yao Xia and Théo Zimmermann. See the Coq Team face book page for more details.

The 40 contributors to the 8.19 version are: quarkcool, Khalid Abdullah, Tanaka Akira, Isaac van Bakel, Frédéric Besson, Lasse Blaauwbroek, Ana Borges, Ali Caglayan, Nikolaos Chatzikonstantinou, Maxime Dénès, Andrej Dudenhefner, Andres Erbsen, Jim Fehrle, Gaëtan Gilbert, Jason Gross, Stefan Haan, Hugo Herbelin, Emilio Jesús Gallego Arias, Pierre Jouvelot, Ralf Jung, Jan-Oliver Kaiser, Robbert Krebbers, Jean-Christophe Léchenet, Rodolphe Lepigre, Yann Leray, Yishuai Li, Guillaume Melquiond, Guillaume Munch-Maccagnoni, Sotaro Okada, Karl Palmskog, Pierre-Marie Pédrot, Jim Portegies, Pierre Rousselin, Pierre Roux, Michael Soegtrop, David Swasey, Enrico Tassi, Shengyi Wang and Théo Zimmermann.

The Coq community at large helped improve this new version via the GitHub issue and pull request system, the coq-club@inria.fr mailing list, the Discourse forum and the Coq Zulip chat.

Version 8.19's development spanned 4 months from the release of Coq 8.18.0 (6 months since the branch for 8.18.0). Gaëtan Gilbert and Matthieu Sozeau are the release managers of Coq 8.19. This release is the result of 285 merged PRs, closing 70 issues.

Nantes, January 2024
Gaëtan Gilbert for the Coq development team

Changes in 8.19.0

Kernel

  • Added: Sort polymorphism makes it possible to share common constructs over Type Prop and SProp (#17836, #18331, by Gaëtan Gilbert).

  • Fixed: Primitives being incorrectly considered convertible to anything by module subtyping (#18507, fixes #18503, by Gaëtan Gilbert).

Specification language, type inference

  • Changed: term_forall_or_fun, term_let, term_fix, term_cofix and term_if from term at level 200 to term10 at level 10. This is a first step towards getting rid of the recovery mechanism of camlp5/coqpp. The impact will mostly be limited to rare cases of additional parentheses around the above (#18014, by Hugo Herbelin).

  • Changed: Declarations of the form (id := body) in Context outside a section in a Module Type do not any more try to declare a class instance. Assumptions whose type is a class and declared using Context outside a section in a Module Type are now declared as global, instead of local (#18254, by Hugo Herbelin).

  • Fixed: Anomaly in the presence of duplicate variables within a disjunctive pattern (#17857 and #18005, fixes #17854 and #18004, by Hugo Herbelin).

  • Fixed: Printing of constructors and of in clause of match now respects the Printing Implicit and Printing All flags (#18176, fixes #18163, by Hugo Herbelin).

  • Fixed: Wrong shift of argument names when using Arguments in nested sections (#18393, fixes #12755 and #18392, by Hugo Herbelin).

Notations

  • Changed: More informative message when a notation cannot be intepreted as a reference (#18104, addresses #18096, by Hugo Herbelin).

  • Changed: In casts like term : t where t is bound to some scope t_scope, via Bind Scope, the term is now interpreted in scope t_scope. In particular when t is Type the term is interpreted in type_scope and when t is a product the term is interpreted in fun_scope (#6134, fixes #14959, by Hugo Herbelin, reviewed by Maxime Dénès, Jim Fehrle, Emilio Gallego, Gaëtan Gilbert, Jason Gross, Pierre-Marie Pédrot, Pierre Roux, Bas Spitters and Théo Zimmermann).

  • Added: the notation term%_scope to set a scope only temporarily (in addition to term%scope for opening a scope applying to all subterms) (#14928, fixes #11486 and #12157 and #14305, by Hugo Herbelin, reviewed by Pierre Roux).

  • Removed the ability to declare scopes whose name starts with _ (would be ambiguous with the new %_scope notation) (#14928, by Pierre Roux, reviewed by Hugo Herbelin).

  • Deprecated the notation term%scope in Arguments command. In a few version, we'll make it an error and in next version give it the same semantics as in terms (i.e., deep scope opening for all subterms rather than just temporary opening) (#14928, fixes #11486 and #12157 and #14305, by Hugo Herbelin, reviewed by Pierre Roux).

  • Added: Quoted strings can be used as tokens in notations; double quotes can be used in symbols in only printing notations; see Basic notations for details (#17123, by Hugo Herbelin).

  • Added: Parsing support for notations with recursive binders involving not only variables bound by fun or forall but also by let or match (#17856, fixes #17845, by Hugo Herbelin).

  • Added: Declaring more than once the level of a notation variable is now an error (#17988, fixes #17985, by Hugo Herbelin).

  • Fixed: Various bugs and limitations to using custom binders in non-recursive and recursive notations (#17115, fixes parts of #17094, by Hugo Herbelin).

  • Fixed: An invalid case of eta-expansion in notation pretty-printer (#17841, fixes #15221, by Hugo Herbelin).

  • Fixed: Printing Parentheses now works also when an explicit level is set for the right-hand side of a right-open notation (#17844, fixes #15322, by Hugo Herbelin).

  • Fixed: anomaly when a notation variable denoting a binder occurs nested more than once in a recursive pattern (#17861, fixes #17860, by Hugo Herbelin).

  • Fixed: Anomaly when trying to disable a non-existent custom notation (#17891, fixes #17782, by Hugo Herbelin).

  • Fixed: appropriate error instead of anomaly in the presence of notations with constructors applied to too many arguments in pattern-matching (#17892, fixes #17071, by Hugo Herbelin).

  • Fixed: support constructors with parameters in number or string notations for patterns (#17902, fixes #11237, by Hugo Herbelin).

  • Fixed: Chains of entry coercions possibly printed in the wrong order depending on the order in which they were declared (#18230, fixes #18223, by Hugo Herbelin).

Tactics

  • Changed: open_constr in Ltac1 and Ltac2 does not perform evar normalization. Normalization may be recovered using let c := open_constr:(...) in constr:(c) if necessary for performance (#17704, by Gaëtan Gilbert).

  • Changed: abstract now supports existential variables (#17745, by Gaëtan Gilbert).

  • Changed: instances declared with Typeclasses Unique Instances do not allow backtracking even when the goal contains evars (#17789, fixes #6714, by Jan-Oliver Kaiser).

  • Changed: In rewrite_strat, the syntax for the choice strategy has changed slightly. You may need to add parentheses around its arguments (one such case found in our continuous integration tests) (#17832, by Hugo Herbelin, Jim Fehrle and Jason Gross).

  • Changed: replace with by tac does not automatically attempt to solve the generated equality subgoal using the hypotheses. Use by first [assumption | symmetry;assumption | tac] if you need the previous behaviour (#17964, fixes #17959, by Gaëtan Gilbert).

  • Changed: Z.euclidean_division_equations_cleanup now breaks up hypotheses of the form 0 <= _ < _ for better cleanup in zify (#17984, by Jason Gross).

  • Changed: simpl now refolds applied constants unfolding to reducible fixpoints into the original constant even when this constant would become partially applied (#17991, by Hugo Herbelin).

  • Added: Ltac2 tactic Std.resolve_tc to resolve typeclass evars appearing in a given term (#13071, by Gaëtan Gilbert and Maxime Dénès).

  • Added: lazy, simpl, cbn and cbv and the associated Eval and eval reductions learned to do head reduction when given flag head (eg Eval lazy head in (fun x => Some ((fun y => y) x)) 0 produces Some ((fun y => y) 0)) (#17503, by Gaëtan Gilbert; cbv case added in #18190, by Hugo Herbelin).

  • Fixed: ensure that opaque primitive projections are correctly handled by "Evarconv" unification (#17788, fixes #17774, by Rodolphe Lepigre).

  • Fixed: Useless duplications with Hint Cut and Hint Mode (#17887, fixes #17417, by Hugo Herbelin).

  • Fixed: zify / Z.euclidean_division_equations_cleanup now no longer instantiates dependent hypotheses. This will by necessity make Z.to_euclidean_division_equations a bit weaker, but the previous behavior was overly sensitive to hypothesis ordering. See #17935 for a recipe to recapture the power of the previous behavior in a more robust albeit slower way (#17935, fixes #17936, by Jason Gross).

  • Fixed: simpl now working on reducible named mutual fixpoints with parameters (#17993, fixes #12521 and part of #3488, by Hugo Herbelin).

  • Fixed: support for reasoning up to polymorphic universe variables in congruence and f_equal (#18106, fixes #5481 and #9979, by Hugo Herbelin).

  • Fixed: Only run zify saturation on existing hypotheses of the goal (#18152, fixes #18151, by Frédéric Besson and Rodolphe Lepigre).

  • Fixed: A stack overflow due to a non-tail recursive function in lia (#18159, fixes #18158, by Jan-Oliver Kaiser and Rodolphe Lepigre).

  • Fixed: Apply substitution in Case stack node for cbv reify (#18195, fixes #18194, by Yann Leray).

  • Fixed: Anomaly of simpl on partially applied named mutual fixpoints (#18243, fixes #18239, by Hugo Herbelin).

  • Changed: simpl tries to reduce named mutual fixpoints also when they return functions (#18243, by Hugo Herbelin).

Ltac language

  • Fixed: Fix broken "r <num>" and "r <string>" commands in the coqtop Ltac debugger, which also affected the Proof General Ltac debugger (#18068, fixes #18067, by Jim Fehrle).

Ltac2 language

  • Changed: Array.empty, Message.Format.stop and Pattern.empty_context are not thunked (#17534, by Gaëtan Gilbert).

  • Changed: Ltac2 exact and eexact elaborate their argument using the type of the goal as expected type, instead of elaborating with no expected type then unifying the resulting type with the goal (#18157, fixes #12827, by Gaëtan Gilbert).

  • Changed: argument order for the Ltac2 combinators List.fold_left List.fold_right and Array.fold_right changed to be the same as in OCaml (#18197, fixes #16485, by Gaëtan Gilbert).

  • Changed: Ltac2.Std.red_flags added field rStrength to support head-only reduction (#18273, fixes #18209, by Gaëtan Gilbert).

  • Added: Ltac2 supports pattern quotations when building pattern values. This allows building dynamic patterns, eg Ltac2 eq_pattern a b := pattern:($pattern:a = $pattern:b) (#17667, by Gaëtan Gilbert).

  • Added: new standard library modules Ltac2.Unification and Ltac2.TransparentState providing access to "Evarconv" unification, including the configuration of the transparency state (#17777, by Rodolphe Lepigre).

  • Added: Ltac2.Constr.is_float, Ltac2.Constr.is_uint63, Ltac2.Constr.is_array (#17894, by Jason Gross).

  • Added: new Ltac2 standard library modules Ltac2.Ref, Ltac2.Lazy and Ltac2.RedFlags

  • Added: new Ltac2 standard library functions to Ltac2.Control, Ltac2.Array, and Ltac2.List (#18095, fixes #10112, by Rodolphe Lepigre).

  • Added: Support for the setoid_rewrite tactic (#18102, by quarkcool).

  • Added: Ltac2 Globalize and Ltac2 Check useful to investigate the expansion of Ltac2 notations (#18139, by Gaëtan Gilbert).

  • Added: A new flag Ltac2 In Ltac1 Profiling (unset by default) to control whether Ltac2 stack frames are included in Ltac profiles (#18293, by Rodolphe Lepigre).

  • Added: Ltac2.Message.Format.ikfprintf useful to implement conditional printing efficiently (i.e. without building an unused message when not printing) (#18311, fixes #18292, by Gaëtan Gilbert).

  • Fixed: Ltac2 mutable references are not considered values anymore (#18082, by Gaëtan Gilbert).

Commands and options

  • Changed: Let with Qed produces an opaque side definition instead of being treated as a transparent let after the section is closed. The previous behaviour can be recovered using clearbody and Defined (#17576, by Gaëtan Gilbert).

  • Changed: automatic lowering of record types to Prop now matches the behavior for inductives: no lowering when universe polymorphism is on, more lowering with recursive records (#17795, fixes #17801 and #17796 and #17801 and #17805, by Gaëtan Gilbert).

  • Added: Extraction Output Directory option for specifying the directory in which extracted files are written (#16126, fixes #9148, by Ali Caglayan).

  • Added: -profile command line argument and PROFILE variable in coq_makefile to control a new Profiling system (#17702, by Gaëtan Gilbert).

  • Added: new command modifier Instructions that executes the given command and displays the number of CPU instructions it took to execute it. This command is currently only supported on Linux systems, but it does not fail on other systems, where it simply shows an error message instead of the count. (#17744, by Rodolphe Lepigre).

  • Added: support for instruction counts to the -profile option. (#17744, by Rodolphe Lepigre).

  • Added: New command Attributes to assign attributes such as deprecated to a library file (#18193, fixes #8032, by Hugo Herbelin).

  • Fixed: Anomaly with Search in the context of a goal (#17987, fixes #17963, by Hugo Herbelin).

  • Fixed: The printer for Guarded was possibly raising an anomaly in the presence of existential variables (#18008, fixes #18006, by Hugo Herbelin).

Command-line tools

  • Changed: Add a coqdep option -w to adjust warnings and allow turning then into errors like the corresponding coqc option (#17946, fixes #10156, by David Swasey and Rodolphe Lepigre).

  • Fixed: properly delayed variable expansion when coq_makefile uses the combined rule for .vo and .glob targets, i.e. on GNU Make 4.4 and later. (#18077, fixes #18076, by Gaëtan Gilbert).

  • Fixed: Spurious coqdep warnings due to missing path normalization for plugins (#18165, by Rodolphe Lepigre).

  • Fixed: Regression in option --external of coqdoc, whose two arguments were inadvertently swapped (#18448, fixes #18434, by Hugo Herbelin).

Standard library

  • Changed: reimplemented Ncring_tac reification (used by nsatz, cring, but not ring) in Ltac instead of typeclasses (#18325, by Gaëtan Gilbert).

  • Removed: Numbers.Cyclic.ZModulo from the standard library. This file was deprecated in 8.17 and has no known use cases. It is retained in the test suite to ensure consistency of CyclicAxioms (#17258, by Andres Erbsen).

  • Removed: ZArith.Zdigits in favor of Z.testbit (#18025, by Andres Erbsen).

  • Removed: long deprecated files in Arith: Div2.v, Even.v, Gt.v, Le.v, Lt.v, Max.v, Minus.v, Min.v, Mult.v, Plus.v, Arith_prebase.v (#18164, by Pierre Rousselin).

  • Deprecated: NArith.Ndigits and NArith.Ndist due to disuse. For most uses of Ndigits, N.testbit and similar functions seem more desirable. If you would like to continue using these files, please consider volunteering to maintain them, within stdlib or otherwise (#17732, by Andres Erbsen).

  • Deprecated: Strings.ByteVector in favor of Init.Byte (#18022, by Andres Erbsen).

  • Deprecated: Numbers.NaryFunctions due to disuse. If you are interested in continuting to use this module, please consider volunteering to maintain it, in stdlib or otherwise (#18026, by Andres Erbsen).

  • Added: Lemma cardinal_Add_In says that inserting an existing key with a new value doesn't change the size of a map, lemma Add_transpose_neqkey says that unequal keys can be inserted into a map in any order (#12096, by Isaac van Bakel and Jean-Christophe Léchenet).

  • Added: lemmas app_eq_cons, app_inj_pivot and rev_inj (#17787, by Stefan Haan, with help of Olivier Laurent).

  • Added: unfold_nth_error, nth_error_nil, nth_error_cons, nth_error_O, nth_error_S to Coq.Lists.List (#17998, by Jason Gross).

  • Added: Reflexive, Symmetric, Transitive, Antisymmetric, Asymmetric instances for Rle, Rge, Rlt, Rgt (#18059, by Jason Gross).

Extraction

  • Fixed: In the error message about extraction of sort-polymorphic singleton inductive types, do not specifically refer to OCaml as other languages are also concerned (#17889, fixes #17817, by Hugo Herbelin).

Changes in 8.19.1

Kernel

  • Fixed: incorrect abstraction of sort variables for opaque constants leading to an inconsistency (#18596 and #18630, fixes #18594, by Gaëtan Gilbert).

  • Fixed: memory corruption with vm_compute (rare but more likely with OCaml 5.1) (#18599, by Guillaume Melquiond).

Notations

Tactics

  • Fixed: undeclared universe with multiple uses of abstract (#18640, fixes #18636, by Gaëtan Gilbert).

Ltac2 language

  • Fixed: incorrect printing of constructor values with multiple arguments, and over-parenthesizing of constructor printing (#18560, fixes #18556, by Gaëtan Gilbert).

  • Fixed: incorrect declared type for Ltac2.FMap.fold (#18649, fixes #18635, by Gaëtan Gilbert).

Infrastructure and dependencies

  • Fixed: missing conf- dependencies of the opam packages: coq-core depends on conf-linux-libc-dev when compiled on linux, and coq depends on conf-python-3 and conf-time to run the test suite (#18565, by Gaëtan Gilbert).

  • Fixed: avoid comitting symlinks to git which caused build failures on some Windows setups (#18550, fixes #18548, by Gaëtan Gilbert).

Changes in 8.19.2

Specification language, type inference

  • Fixed: Regression from Coq 8.18 in the presence of a defined field in a primitive Record (#19088, fixes #19082, by Hugo Herbelin).

Notations

  • Fixed: Printer sometimes failing to use a prefix or infix custom notation whose right-hand side refers to a different custom entry (#18089, fixes #18914, by Hugo Herbelin).

Tactics

  • Fixed: abstract failing in the presence of admitted goals in the surrounding proof (#18945, fixes #18942, by Gaëtan Gilbert).

Ltac2 language

  • Fixed: anomalies when using Ltac2 in VsCoq due to incorrect state handling of Ltac2 notations (#19096, fixes coq-community/vscoq#772, by Gaëtan Gilbert)

Commands and options

CoqIDE

  • Fixed: Position error/warning tooltips correctly when multibyte UTF-8 characters are present (#19137, fixes #19136, by Jim Fehrle).

Infrastructure and dependencies

  • Fixed: compatibility with OCaml versions where effect is a keyword (#18863, by Remy Seassau)

  • Added: Coq is now compatible with memprof-limits interruption methods. This means that Coq will be recompiled when the library is installed / removed from an OPAM switch. (#18906, fixes #17760, by Emilio Jesus Gallego Arias).

Version 8.18

Summary of changes

Coq version 8.18 integrates two soundness fixes to the Coq kernel along with a host of improvements. We highlight a few impactful changes:

  • the default locality of Hint and Instance commands was switched to export.

  • the universe unification algorithm can now delay the commitment to a sort (the algorithm used to pick Type). Thanks to this feature many Prop and SProp annotations can be now omitted.

  • Ltac2 supports array literals, maps and sets of primitive datatypes such as names (of constants, inductive types, etc) and fine-grained control over profiling.

  • The warning system offers new categories, enabling finer (de)activation of specific warnings. This should be particularly useful to handle deprecations.

  • Many new lemmas useful for teaching analysis with Coq are now part of the standard library about real numbers.

  • The #[deprecated] attribute can now be applied to definitions.

The 41 contributors to the 8.18 version are: Reynald Affeldt, Tanaka Akira, Matthieu Baty, Yves Bertot, Lasse Blaauwbroek, Ana Borges, Kate Deplaix, Ali Caglayan, Cyril Cohen, Maxime Dénès, Andrej Dudenhefner, Andres Erbsen, Jim Fehrle, Yannick Forster, Paolo G. Giarrusso, Gaëtan Gilbert, Jason Gross, Samuel Gruetter, Stefan Haan, Hugo Herbelin, Yoshihiro Imai, Emilio Jesús Gallego Arias, Olivier Laurent, Meven Lennon-Bertrand, Rodolphe Lepigre, Yishuai Li, Guillaume Melquiond, Karl Palmskog, Pierre-Marie Pédrot, Stefan Radziuk, Ramkumar Ramachandra, Pierre Rousselin, Pierre Roux, Julin Shaji, Kazuhiko Sakaguchi, Weng Shiwei, Michael Soegtrop, Matthieu Sozeau, Enrico Tassi, Hao Yang, Théo Zimmermann.

We are very grateful to the Coq community for their help in creating 8.18 in the 6 months since the release of Coq 8.17.0. Maxime Dénès and Enrico Tassi were the release managers.

Sophia-Antipolis, September 2023,
Enrico Tassi for the Coq development team

Changes in 8.18.0

Kernel

  • Changed: the bad-relevance warning is now an error by default (#17172, by Pierre-Marie Pédrot).

  • Fixed: the kernel now checks that case elimination of private inductive types (cf private(matching)) is not used outside their defining module. Previously this was only checked in elaboration and the check could be avoided through some tactics, breaking consistency in the presence of axioms which rely on the elimination restriction to be consistent (#17452, fixes #9608, by Gaëtan Gilbert).

  • Fixed: a bug enabling native_compute to yield arbitrary floating-point values (#17872, fixes #17871, by Guillaume Melquiond and Pierre Roux, bug found by Jason Gross).

Specification language, type inference

  • Changed: enhance the universe unification algorithm, which is now able to delay the definition of a sort. This allows omitting some explicit Prop and SProp annotations when writing terms. Some minor backwards compatibility issues can arise in rare cases, which can be solved with more explicit sort annotations (#16903, by Pierre-Marie Pédrot).

  • Changed: match compilation for primitive record avoids producing an encoding overhead for matches that are equivalent to a primitive projection (#17008, by Gaëtan Gilbert).

  • Added: volatile casts term :> type which do not leave a trace in the elaborated term. They are used by Printing Match All Subterms to display otherwise hidden subterms of match constructs (#16992, fixes #16918, by Gaëtan Gilbert).

  • Added: when printing uninterpreted terms (for instance through Print Ltac on Ltac foo := exact some_term), extensions to the term language (for instance Solving existential variables using tactics) are now printed correctly instead of as holes (_) (#17221, by Gaëtan Gilbert).

  • Added: Support for the local, global and export locality attributes for the single "field" of definitional typeclasses when using the :> and :: syntaxes for coercion and substructures (#17754, fixes #17451, by Pierre Roux).

  • Added: a hook in the coercion mechanism to enable programming coercions in external metalanguages such as Ltac, Ltac2, Elpi or OCaml plugins (#17794, by Pierre Roux).

  • Fixed: canonical instance matching match terms (#17206, fixes #17079, by Gaëtan Gilbert).

  • Fixed: universe constraint inference in module subtyping can trigger constant unfoldings (#17305, fixes #17303, by Gaëtan Gilbert).

Notations

  • Removed: The '[=' keyword. '[=' tokens in notation definitions should be replaced with the pair of tokens '[' '='. If compatibility with Coq < 8.18 is needed, replace [= in uses of the notation with an added space ([ =) (#16788, fixes #16785, by Pierre Roux).

  • Added: Support for Printing Parentheses in custom notations (#17117, by Hugo Herbelin).

  • Added: Improve printing of reverse coercions. When a term x is elaborated to x' through a reverse coercion, return the term reverse_coercion x' x that is convertible to x' but displayed x thanks to the coercion reverse_coercion (#17484, by Pierre Roux).

  • Fixed: Add support to parse a recursive pattern as a sequence of terms in a recursive notation even when this recursive pattern is used in position of binders; it was formerly raising an anomaly (#16937, fixes #12467, by Hugo Herbelin).

  • Fixed: Improved ability to print notations involving anonymous binders (#17050, by Hugo Herbelin).

  • Fixed: anomaly with notations abbreviating a local variable or record field name (#17217, fixes #14975, by Hugo Herbelin).

  • Fixed: Ensure in all cases that a parsing rule is declared when the only parsing flag is given (#17318, fixes #17316, by Hugo Herbelin).

  • Fixed: In Number Notation, "abstract after N" was applied when number >= N. Now it is applied when number > N (#17478, by Jim Fehrle).

Tactics

  • Changed: in the fringe case where the with clause of a call to specialize depends on a variable bound in the type, the tactic will now fail instead of silently producing a shelved evar (#17322, by Pierre-Marie Pédrot).

  • Changed: extensions to the term syntax through generic arguments (typically ltac:(), ltac2:() or ltac2's $) produce errors when used in term patterns (for instance patterns used to filter hints) instead of being treated as holes (_) (#17352, by Gaëtan Gilbert).

  • Changed: the case tactic and its variants always generate a pattern-matching node, regardless of their argument. In particular, they are now guaranteed to generate as many goals as there are constructors in the inductive type. Previously, they used to reduce to the corresponding branch when the argument βι-normalized to a constructor, resulting in a single goal (#17541, by Pierre-Marie Pédrot).

  • Changed: injection continues working using sigma types when Eqdep_dec has not been required even if an equality scheme was found, instead of failing (#17670, by Gaëtan Gilbert).

  • Changed: the unification heuristics for implicit arguments of the case tactic. We unconditionally recommend using destruct instead, and even more so in case of incompatibility (#17564, by Pierre-Marie Pédrot).

  • Removed: the no-argument form of the instantiate tactic, deprecated since 8.16 (#16910, by Pierre-Marie Pédrot).

  • Removed: undocumented tactics hresolve_core and hget_evar (#17035, by Gaëtan Gilbert).

  • Deprecated: the elimtype and casetype tactics (#16904, by Pierre-Marie Pédrot).

  • Deprecated: revert dependent, which is a misleadingly named alias of generalize dependent (#17669, by Gaëtan Gilbert).

  • Fixed: The simpl tactic now respects the simpl never flag even when the subject function is referred to through another definition (#13448, fixes #13428, by Yves Bertot).

  • Fixed: unification is less sensitive to whether a subterm is an indirection through a defined existential variable or a direct term node. This results in less constant unfoldings in rare cases (#16960, by Gaëtan Gilbert).

  • Fixed: untypable proof states generated by setoid_rewrite, which may cause some backwards-incompatibilities (#17304, fixes #17295, by Lasse Blaauwbroek).

  • Fixed: intropatterns destructing a term whose type is a product cannot silently create shelved evars anymore. Instead, it fails with an unsolvable variable. This can be fixed in a backwards compatible way by using the e-variant of the parent tactic (#17564, by Pierre-Marie Pédrot).

  • Fixed: the field_simplify tactic, so that it no longer introduces side-conditions when working on a hypothesis (#17591, by Guillaume Melquiond).

  • Fixed: the tauto tactic and its variants now try to match types up to universe unification. This makes them compatible with universe-polymorphic code (#8905, fixes #4721 and #5351, by Pierre-Marie Pédrot).

Ltac2 language

  • Added: Support for parsing Ltac2 array literals [| ... |] (#16859, fixes #13976, by Samuel Gruetter).

  • Added: Finite set and map APIs for identifier, string, int, constant, inductive and constructor keys (#17347, c.f. #16409, by Gaëtan Gilbert).

  • Added: Ltac2 preterm antiquotation $preterm: (#17359, fixes #13977, by Gaëtan Gilbert).

  • Added: Ltac Profiling also profiles Ltac2 tactics. Ltac2 also provides tactics start_profiling stop_profiling and show_profile for finer grained control (#17371, fixes #10111, by Gaëtan Gilbert).

  • Added: primitives to build and compare values in Ltac2.Init.cast (#17468, by Gaëtan Gilbert).

  • Added: It is possible to define 0-argument externals (#17475, by Gaëtan Gilbert).

  • Added: Ltac2 quotations ltac2val:(ltac2 tactic) in Ltac1 which produce Ltac1 values (as opposed to ltac2:() quotations which are only useful for their side effects) (#17575, by Gaëtan Gilbert).

  • Fixed: nested notations involving Term Antiquotations (#17232, fixes #15864, by Gaëtan Gilbert).

  • Fixed: Parsing level of by clause of Ltac2's assert (#17508, fixes #17491, by Samuel Gruetter).

  • Fixed: multi_match!, multi_match! goal and the underlying Ltac2.Pattern.multi_match0 and Ltac2.Pattern.multi_goal_match0 now preserve exceptions from backtracking after a branch succeeded instead of replacing them with Match_failure (e.g. multi_match! constr:(tt) with tt => () end; Control.zero Not_found now fails with Not_found instead of Match_failure) (#17597, fixes #17594, by Gaëtan Gilbert).

Commands and options

  • Changed: the default locality of Hint and Instance commands was switched to export (#16258, by Pierre-Marie Pédrot).

  • Changed: warning non-primitive-record is now in category records instead of record. This was the only use of record but the plural version is also used by cannot-define-projection future-coercion-class-constructor and future-coercion-class-field. (#16989, by Gaëtan Gilbert).

  • Changed: Eval prints information about existential variables like Check (#17274, by Gaëtan Gilbert).

  • Changed: The names of deprecation warnings now depend on the version in which they were introduced, using their "since" field. This enables deprecation warnings to be selectively enabled, disabled, or treated as an error, according to the version number provided in the deprecated attribute (#17489, fixes #16287, by Pierre Roux, reviewed by Ali Caglayan, Théo Zimmermann and Gaëtan Gilbert).

  • Changed: warnings can now have multiple categories allowing for finer user control on which warning to enable, disable or treat as an error (#17585, by Gaëtan Gilbert).

  • Changed: Template polymorphic inductive types are not implicitly added to the Keep Equalities table anymore when defined. This may change the behavior of equality-related tactics on such types (#17718, by Pierre-Marie Pédrot).

  • Changed: Warnings and warnings now emit a warning when trying to enable an unknown warning (there is still no warning when disabling an unknown warning as this behavior is useful for compatibility, or when enabling an unknown warning through the command line -w as the warning may be in a yet to be loaded plugin) (#17747, by Gaëtan Gilbert).

  • Removed: the flag Apply With Renaming which was deprecated since 8.15 (#16909, by Pierre-Marie Pédrot).

  • Removed: the Typeclasses Filtered Unification flag, deprecated since 8.16 (#16911, by Pierre-Marie Pédrot).

  • Removed: program attribute is not accepted anymore with commands Add Relation, Add Parametric Relation, Add Setoid, Add Parametric Setoid, Add Morphism, Add Parametric Morphism, Declare Morphism. Previously, it was accepted but ignored (#17042, by Théo Zimmermann).

  • Removed: the Elaboration StrictProp Cumulativity and Cumulative SProp flags. These flags became counterproductive after the introduction of sort variables in unification (#17114, fixes #17108, by Pierre-Marie Pédrot).

  • Removed: The Add LoadPath, Add Rec LoadPath, Add ML Path, and Remove LoadPath commands have been removed following deprecation. Users are encouraged to use the existing mechanisms in coq_makefile or dune to configure workspaces of Coq theories (#17394, by Emilio Jesus Gallego Arias).

  • Deprecated: Export modifier for Set. Use attribute export instead (#17333, by Gaëtan Gilbert).

  • Deprecated: the nonuniform attribute, now subsumed by warnings with "-uniform-inheritance" (#17716, by Pierre Roux).

  • Deprecated: Using Qed with Let. End the proof with Defined and use clearbody instead to get the same behavior (#17544, by Gaëtan Gilbert).

  • Added: About now prints information when a constant or inductive is syntactically equal to another through module aliasing (#16796, by Gaëtan Gilbert).

  • Added: Final Obligation command (#16817, by Gaëtan Gilbert).

  • Added: The deprecated attribute is now supported for definition-like constructions (#16890, fixes #12266, by Maxime Dénès and Gaëtan Gilbert).

  • Added: attributes warnings and alias warning to set warnings locally for a command (#16902, fixes #15893, by Gaëtan Gilbert).

  • Added: flag Printing Unfolded Projection As Match (off by default) to be able to distinguish unfolded and folded primitive projections (#16994, by Gaëtan Gilbert).

  • Added: option -time-file, like time but outputting to a file (#17430, by Gaëtan Gilbert).

  • Added: Validate Proof runs the type checker on the current proof, complementary with Guarded which runs the guard checker (#17467, by Gaëtan Gilbert).

  • Added: clearbody for Let to clear the body of a let-in in an interactive proof without kernel enforcement. (This is the behavior that was previously provided by using Qed, which is now deprecated for Lets.) (#17544, by Gaëtan Gilbert).

  • Added: option -time-file, like time but outputting to a file (#17430, by Gaëtan Gilbert).

  • Fixed: universe monomorphic inductives and records do not ignore Universe Minimization ToSet (#17285, fixes #13927, by Gaëtan Gilbert).

Command-line tools

  • Changed: Do not pass the -rectypes flag by default in coq_makefile when compiling OCaml code, since it is no longer required by Coq. To re-enable passing the flag, put CAMLFLAGS+=-rectypes in the local makefile, e.g., CoqMakefile.local (see CoqMakefile.local) (#17038, by Karl Palmskog with help from Gaëtan Gilbert).

  • Changed: disable inclusion of variable binders in coqdoc indexes by default, and provide a new coqdoc option --binder-index for including them (#17045, fixes #13155, by Karl Palmskog).

  • Added: coqdoc handles multiple links to the same source. For example when declaring an inductive type t all occurences of t itself and its elimination principles like t_ind point to its declaration (#17118, by Enrico Tassi).

  • Added: Command line options -require lib (replacing -load-vernac-object lib) and -require-from root lib respectively equivalent to vernacular commands Require lib and From root Require lib (#17364, by Hugo Herbelin).

  • Added: coqtimelog2html command-line tool used to render the timing files produced with -time (which is passed by coq_makefile when environment variable TIMING is defined) (#17411, by Gaëtan Gilbert).

  • Fixed: coq_makefile avoids generating a command containing all files to install in a make rule, which could surpass the maximum single argument size in some developments (#17697, fixes #17721, by Gaëtan Gilbert).

CoqIDE

  • Changed: XML Protocol now sends (and expects) full Coq locations, including line and column information. This makes some IDE operations (such as UTF-8 decoding) more efficient. Clients of the XML protocol can just ignore the new fields if they are not useful for them (#17382, fixes #17023, by Emilio Jesus Gallego Arias).

Standard library

  • Changed: implementation of Vector.nth to follow OCaml and compute strict subterms (#16731, fixes #16738, by Andrej Dudenhefner).

  • Changed: drop the unnecessary second assumption NoDup l' from set_diff_nodup in ListSet.v, with -compat 8.17 providing the old version of set_diff_nodup for compatibility (#16926, by Karl Palmskog with help from Traian Florin Şerbănuţă and Andres Erbsen).

  • Changed: Moved instances from DecidableClass to files that prove the relevant decidability facts: Bool, PeanoNat, and BinInt (#17021, by Andres Erbsen).

  • Changed: Hint Extern btauto.Algebra.bool locality from global to export (#17281, by Andres Erbsen).

  • Changed: xorb to a simpler definition (#17427, by Guillaume Melquiond).

  • Changed lemmas in Reals/RIneq.v

    • completeness_weak renamed as upper_bound_thm,

    • le_epsilon renamed as Rle_epsilon,

    • Rplus_eq_R0 renamed as Rplus_eq_0,

    • Req_EM_T renamed as Req_dec_T,

    • Rinv_r_simpl_m renamed as Rmult_inv_r_id_m,

    • Rinv_r_simpl_l renamed as Rmult_inv_r_id_l,

    • Rinv_r_simpl_r renamed as Rmult_inv_m_id_r,

    • tech_Rgt_minus renamed as Rgt_minus_pos,

    • tech_Rplus renamed as Rplus_le_lt_0_neq_0,

    • IZR_POS_xI modified with 2 instead of 1 + 1,

    • IZR_POS_xO modified with 2 instead of 1 + 1,

    • Rge_refl modified with >= instead of <=

    (#17036, by Pierre Rousselin, reviewer Laurent Théry).

  • Removed: Datatypes.prod_curry, Datatypes.prod_uncurry, Datatypes.prodT_curry, Datatypes.prodT_uncurry, Combinators.prod_curry_uncurry, Combinators.prod_uncurry_curry, Bool.leb, Bool.leb_implb, List.skipn_none, Zdiv.Z_div_mod_eq, Zdiv.div_Zdiv, Zdiv.mod_Zmod, FloatOps.frexp, FloatOps.ldexp, FloatLemmas.frexp_spec, FloatLemmas.ldexp_spec, RList.Rlist, Rlist.cons, Rlist.nil, RList.Rlength, Rtrigo_calc.cos3PI4, Rtrigo_calc.sin3PI4, MSetRBT.filter_app after deprecation for at least two Coq versions (#16920, by Olivier Laurent).

  • Deprecated: List.app_nil_end, List.app_assoc_reverse, List.ass_app, List.app_ass (#16920, by Olivier Laurent).

  • Deprecated: Coq.Lists.List.Forall2_refl (Coq.Lists.List.Forall2_nil has the same type) (#17646, by Gaëtan Gilbert).

  • Deprecated: ZArith.Zdigits in favor of Z.testbit. If you are aware of a use case of this module and would be interested in a drop-in replacement, please comment on the PR with information about the context that would benefit from such functinality (#17733, by Andres Erbsen).

  • Deprecated: Deprecation warnings are now generated for Numbers.Cyclic.Int31.Cyclic31, NNumbers.Cyclic.Int31.Int31, and NNumbers.Cyclic.Int31.Ring31. These modules have been deprecated since Coq 8.10. The modules under Numbers.Cyclic.Int63 remain available (#17734, by Andres Erbsen).

  • Deprecated lemmas in Reals/RIneq.v

    inser_trans_R, IZR_neq, double, double_var, Rinv_mult_simpl, Rle_Rinv, Rlt_Rminus, Rminus_eq_0, Rminus_gt_0_lt, Ropp_div, Ropp_minus_distr', Rplus_sqr_eq_0_l, sum_inequa_Rle_lt_depr, S_O_plus_INR_depr, single_z_r_R1_depr, tech_single_z_r_R1_depr,

    (#17036, by Pierre Rousselin, reviewer Laurent Théry).

  • Added: lemmas L_inj, R_inj, L_R_neq, case_L_R, case_L_R' to Fin.v, and nil_spec, nth_append_L, nth_append_R, In_nth, nth_replace_eq, nth_replace_neq, replace_append_L, replace_append_R, append_const, map_append, map2_ext, append_inj, In_cons_iff, Forall_cons_iff, Forall_map, Forall_append, Forall_nth, Forall2_nth, Forall2_append, map_shiftin, fold_right_shiftin, In_shiftin, Forall_shiftin, rev_nil, rev_cons, rev_shiftin, rev_rev, map_rev, fold_left_rev_right, In_rev, Forall_rev to VectorSpec.v (#16765, closes #6459, by Andrej Dudenhefner).

  • Added: lemmas iter_swap_gen, iter_swap, iter_succ, iter_succ_r, iter_add, iter_ind, iter_rect, iter_invariant for Nat.iter (#17013, by Stefan Haan with help from Jason Gross).

  • Added: module Zbitwise with basic relationships between bitwise and arithmetic operations on integers (#17022, by Andres Erbsen).

  • Added: lemmas forallb_filter, forallb_filter_id, partition_as_filter, filter_length, filter_length_le and filter_length_forallb (#17027, by Stefan Haan with help from Olivier Laurent and Andres Erbsen).

  • Added: lemmas in Reals/RIneq.v:

    eq_IZR_contrapositive, INR_0, INR_1, INR_archimed, INR_unbounded, IPR_2_xH, IPR_2_xI, IPR_2_xO, IPR_eq, IPR_ge_1, IPR_gt_0, IPR_IPR_2, IPR_le, IPR_lt, IPR_not_1, IPR_xH, IPR_xI, IPR_xO, le_IPR, lt_1_IPR, lt_IPR, minus_IPR, mult_IPR, not_1_IPR, not_IPR, plus_IPR, pow_IPR, Rdiv_0_l, Rdiv_0_r, Rdiv_1_l, Rdiv_1_r, Rdiv_def, Rdiv_diag_eq, Rdiv_diag, Rdiv_diag_uniq, Rdiv_eq_compat_l, Rdiv_eq_compat_r, Rdiv_eq_reg_l, Rdiv_eq_reg_r, Rdiv_mult_distr, Rdiv_mult_l_l, Rdiv_mult_l_r, Rdiv_mult_r_l, Rdiv_mult_r_r, Rdiv_neg_neg, Rdiv_neg_pos, Rdiv_opp_l, Rdiv_pos_cases, Rdiv_pos_neg, Rdiv_pos_pos, Rexists_between, Rge_gt_or_eq_dec, Rge_gt_or_eq, Rge_lt_dec, Rge_lt_dec, Rgt_le_dec, Rgt_minus_pos, Rgt_or_le, Rgt_or_not_gt, Rinv_0_lt_contravar, Rinv_eq_compat, Rinv_eq_reg, Rinv_lt_0_contravar, Rinv_neg, Rinv_pos, Rle_gt_dec, Rle_half_plus, Rle_lt_or_eq, Rle_or_gt, Rle_or_not_le, Rlt_0_2, Rlt_0_minus, Rlt_ge_dec, Rlt_half_plus, Rlt_minus_0, Rlt_or_ge, Rlt_or_not_lt, Rminus_def, Rminus_diag, Rminus_eq_compat_l, Rminus_eq_compat_r, Rminus_plus_distr, Rminus_plus_l_l, Rminus_plus_l_r, Rminus_plus_r_l, Rminus_plus_r_r, Rmult_div_assoc, Rmult_div_l, Rmult_div_r, Rmult_div_swap, Rmult_gt_reg_r, Rmult_inv_l, Rmult_inv_m_id_r, Rmult_inv_r, Rmult_inv_r_id_l, Rmult_inv_r_id_m, Rmult_inv_r_uniq, Rmult_neg_cases, Rmult_neg_neg, Rmult_neg_pos, Rmult_pos_cases, Rmult_pos_neg, Rmult_pos_pos, Ropp_div_distr_l, Ropp_eq_reg, Ropp_neg, Ropp_pos, Rplus_0_l_uniq, Rplus_eq_0, Rplus_ge_reg_r, Rplus_gt_reg_r, Rplus_minus_assoc, Rplus_minus_l, Rplus_minus_r, Rplus_minus_swap, Rplus_neg_lt, Rplus_neg_neg, Rplus_neg_npos, Rplus_nneg_ge, Rplus_nneg_nneg, Rplus_nneg_pos, Rplus_npos_le, Rplus_npos_neg, Rplus_npos_npos, Rplus_pos_gt, Rplus_pos_nneg, Rplus_pos_pos, Rsqr_def

    lemmas in Reals/R_Ifp.v: Int_part_spec, Rplus_Int_part_frac_part, Int_part_frac_part_spec

    (#17036, by Pierre Rousselin, reviewer Laurent Théry).

  • Added: lemmas concat_length, flat_map_length, flat_map_constant_length, list_power_length to Lists.List (#17082, by Stefan Haan with help from Olivier Laurent).

Infrastructure and dependencies

  • Changed: Sphinx 4.5.0 or above is now required to build the reference manual, so now / can be used as a quick search shortcut and Esc as a shortcut to remove search highlighting (#17772, fixes #15778, by Ana Borges).

Extraction

  • Fixed: Anomaly when extracting within a module or module type (#17344, fixes #10739, by Hugo Herbelin).

Version 8.17

Summary of changes

Coq version 8.17 integrates a soundness fix to the Coq kernel along with a few new features and a host of improvements to the Ltac2 language and libraries. We highlight some of the most impactful changes here:

  • Fixed a logical inconsistency due to vm_compute in presence of side-effects in the enviroment (e.g. using Back or Fail).

  • It is now possible to dynamically enable or disable notations.

  • Support multiple scopes in Arguments and Bind Scope.

  • The tactics chapter of the manual has many improvements in presentation and wording. The documented grammar is semi-automatically checked for consistency with the implementation.

  • Fixes to the auto and eauto tactics, to respect hint priorities and the documented use of simple apply. This is a potentially breaking change.

  • New Ltac2 APIs, deep pattern-matching with as clauses and handling of literals, support for record types and preterms.

  • Move from :> to :: syntax for declaring typeclass fields as instances, fixing a confusion with declaration of coercions.

  • Standard library improvements.

  • While Coq supports OCaml 5, users are likely to experience slowdowns ranging from +10% to +50% compared to OCaml 4. Moreover, the native_compute machinery is not available when Coq is compiled with OCaml 5. Therefore, OCaml 5 support should still be considered experimental and not production-ready.

See the Changes in 8.17.0 section below for the detailed list of changes, including potentially breaking changes marked with Changed. Coq's reference manual for 8.17, documentation of the 8.17 standard library and developer documentation of the 8.17 ML API are also available.

Ali Caglayan, Emilio Jesús Gallego Arias, Gaëtan Gilbert and Théo Zimmermann worked on maintaining and improving the continuous integration system and package building infrastructure.

Erik Martin-Dorel has maintained the Coq Docker images that are used in many Coq projects for continuous integration.

Maxime Dénès, Paolo G. Giarrusso, Huỳnh Trần Khanh, and Laurent Théry have maintained the VsCoq extension for VS Code.

The opam repository for Coq packages has been maintained by Guillaume Claret, Karl Palmskog, Matthieu Sozeau and Enrico Tassi with contributions from many users. A list of packages is available at https://coq.inria.fr/opam/www/.

The Coq Platform has been maintained by Michael Soegtrop, with help from Karl Palmskog, Pierre Roux, Enrico Tassi and Théo Zimmermann.

Our current maintainers are Yves Bertot, Frédéric Besson, Ana Borges, Ali Caglayan, Tej Chajed, Cyril Cohen, Pierre Corbineau, Pierre Courtieu, Maxime Dénès, Andres Erbsen, Jim Fehrle, Julien Forest, Emilio Jesús Gallego Arias, Gaëtan Gilbert, Georges Gonthier, Benjamin Grégoire, Jason Gross, Hugo Herbelin, Vincent Laporte, Olivier Laurent, Assia Mahboubi, Kenji Maillard, Guillaume Melquiond, Pierre-Marie Pédrot, Clément Pit-Claudel, Pierre Roux, Kazuhiko Sakaguchi, Vincent Semeria, Michael Soegtrop, Arnaud Spiwack, Matthieu Sozeau, Enrico Tassi, Laurent Théry, Anton Trunov, Li-yao Xia and Théo Zimmermann. See the Coq Team face book page for more details.

The 45 contributors to the 8.17 version are: Reynald Affeldt, Tanaka Akira, Lasse Blaauwbroek, Stephan Boyer, Ali Caglayan, Cyril Cohen, Maxime Dénès, Andrej Dudenhefner, Andres Erbsen, František Farka, Jim Fehrle, Paolo G. Giarrusso, Gaëtan Gilbert, Jason Gross, Alban Gruin, Stefan Haan, Hugo Herbelin, Wolf Honore, Bodo Igler, Jerry James, Emilio Jesús Gallego Arias, Ralf Jung, Jan-Oliver Kaiser, Wojciech Karpiel, Chantal Keller, Thomas Klausner, Olivier Laurent, Yishuai Li, Guillaume Melquiond, Karl Palmskog, Sudha Parimala, Pierre-Marie Pédrot, Valentin Robert, Pierre Roux, Julin S, Dmitry Shachnev, Michael Soegtrop, Matthieu Sozeau, Naveen Srinivasan, Sergei Stepanenko, Karolina Surma, Enrico Tassi, Li-yao Xia and Théo Zimmermann.

The Coq community at large helped improve this new version via the GitHub issue and pull request system, the coq-club@inria.fr mailing list, the Discourse forum and the Coq Zulip chat.

Version 8.17's development spanned 5 months from the release of Coq 8.16.0. Théo Zimmermann is the release manager of Coq 8.17. This release is the result of 414 merged PRs, closing 105 issues.

Nantes, February 2023,
Matthieu Sozeau for the Coq development team

Changes in 8.17.0

Kernel

  • Fixed: inconsistency linked to vm_compute. The fix removes a vulnerable cache, thus it may result in slowdowns when vm_compute is used repeatedly, if you encounter such slowdowns please report your use case (#16958, fixes #16957, by Gaëtan Gilbert).

  • Fixed: Unexpected anomaly when checking termination of fixpoints containing match expressions with inaccessible branches (#17116, fixes #17073, by Hugo Herbelin).

Specification language, type inference

  • Changed: Unused variable warning triggers even when catching a single case. This warning used to be triggered only when the unused variable was catching at least two cases (#16135, by Pierre Roux).

  • Fixed: Pattern-matching clauses were possibly lost when matching over a constructor from a singleton inductive type in the presence of implicit coercions (#17138, fixes #17137, by Hugo Herbelin).

  • Fixed: Possible anomaly when using syntax term.(proj) with projections defined in sections (#17174, fixes #17173, by Hugo Herbelin).

Notations

  • Changed: When multiple tokens match the beginning of a sequence of characters, the longest matching token not cutting a subsequence of contiguous letters in the middle is used. Previously, this was only the longest matching token. See lexical conventions for details and examples (#16322, fixes #4712, by Hugo Herbelin).

  • Added: Enable Notation and Disable Notation commands to enable or disable previously defined notations (#12324 and #16945, by Hugo Herbelin and Pierre Roux, extending previous work by Lionel Rieg, review by Jim Fehrle).

  • Added: Support for multiple scopes in the Arguments command (#16472, by Pierre Roux, review by Jim Fehrle, Hugo Herbelin and Enrico Tassi).

  • Added: Attributes add_top and add_bottom to bind multiple scopes through the Bind Scope command (#16472, by Pierre Roux, review by Jim Fehrle, Hugo Herbelin and Enrico Tassi).

Tactics

  • Changed: Documentation in the tactics chapter to give the current correct syntax, consolidate tactic variants for each tactic into a single, unified description for each tactic and many wording improvements. With this change, following similar changes to other chapters in previous releases, the correctness of documented syntax is assured by semi-automated tooling in all chapters except SSReflect (#15015, #16498, and #16659, by Jim Fehrle, reviewed by Théo Zimmermann, with help from many others).

  • Changed: eauto respects priorities of Extern hints (#16289, fixes #5163 and #16282, by Andrej Dudenhefner).

    Warning

    Code that relies on eager evaluation of Extern hints with high assigned cost by eauto will change its performance profile or potentially break. To approximate prior behavior, set to zero the cost of Extern hints, which may solve the goal in one step.

  • Changed: less discrepancies between auto hint evaluation and simple apply, exact tactics (#16293, fixes #16062 and #16323, by Andrej Dudenhefner).

    Warning

    auto may solve more goals. As a result, non-monotone use of auto such as tac1; auto. tac2. may break. For backwards compatibility use explicit goal management.

  • Removed: absurd_hyp tactic, that was marked as obsolete 15 years ago. Use contradict instead (#16670, by Théo Zimmermann).

  • Removed: the undocumented progress_evars tactical (#16843, by Théo Zimmermann).

  • Deprecated: the default intuition_solver (see intuition) now outputs warning intuition-auto-with-star if it solves a goal with auto with * that was not solved with just auto. In a future version it will be changed to just auto. Use intuition tac locally or Ltac Tauto.intuition_solver ::= tac globally to silence the warning in a forward-compatible way with your choice of tactic tac (auto, auto with *, auto with your prefered databases, or any other tactic) (#16026, by Gaëtan Gilbert).

  • Deprecated: > clear modifier that could be used in some tactics like apply and rewrite but was never documented. Open an issue if you actually depend on this feature (#16407, by Théo Zimmermann).

  • Fixed: auto now properly updates local hypotheses after hint application (#16302, fixes #15814 and #6332, by Andrej Dudenhefner).

  • Fixed: Make the behavior of destruct ... using ... more powerful and more similar to destruct ... (#16605, by Lasse Blaauwbroek).

  • Fixed: typeclass inference sometimes caused remaining holes to fail to be detected (#16743, fixes #5239, by Gaëtan Gilbert).

Ltac language

  • Changed: Ltac redefinitions (with ::=) now respect local (#16106, by Gaëtan Gilbert).

  • Changed: In match goal, match goal with hyp := body : typ |- _ is syntax sugar for match goal with hyp := [ body ] : typ |- _ i.e. it matches typ with the type of the hypothesis rather than matching the body as a cast term. This transformation used to be done with any kind of cast (e.g. VM cast <:) and is now done only for default casts : (#16764, by Gaëtan Gilbert).

Ltac2 language

  • Changed: Ltac2.Bool notations are now in a module Ltac2.Bool.BoolNotations (exported by default), so that these notations can be imported separately (#16536, by Jason Gross).

  • Changed: Constr.in_context enforces that the constr passed to it is a type (#16547, fixes #16540, by Gaëtan Gilbert).

  • Changed: goal matching functions from Ltac2.Pattern (matches_goal, lazy_goal_match0, multi_goal_match0 and one_goal_match0) have changed types to support matching hypothesis bodies (#16655, by Gaëtan Gilbert).

  • Added: Deep pattern matching for Ltac2 (#16023, by Gaëtan Gilbert).

  • Added: patterns for Ltac2 matches: as, records and literal integers and strings (#16179, by Gaëtan Gilbert).

  • Added: APIs for working with strings: Message.to_string, String.concat, cat, equal, compare, is_empty (#16217, by Gaëtan Gilbert).

  • Added: Ltac2.Constr.Unsafe.liftn (#16413, by Jason Gross).

  • Added: Ltac2.Constr.Unsafe.closedn, Ltac2.Constr.Unsafe.is_closed, Ltac2.Constr.Unsafe.occur_between, Ltac2.Constr.Unsafe.occurn (#16414, by Jason Gross).

  • Added: Ltac2.List.equal (#16429, by Jason Gross).

  • Added: Print Ltac2, Print Ltac2 Signatures and Locate can now find Ltac2 definitions (#16466, fixes #16418 and #16415, by Gaëtan Gilbert).

  • Added: Ltac2.Array.for_all2 and Ltac2.Array.equal (#16535, by Jason Gross).

  • Added: Ltac2.Constant.equal, Ltac2.Constant.t, Ltac2.Constructor.equal, Ltac2.Constructor.t, Ltac2.Evar.equal, Ltac2.Evar.t, Ltac2.Float.equal, Ltac2.Float.t, Ltac2.Meta.equal, Ltac2.Meta.t, Ltac2.Proj.equal, Ltac2.Proj.t, Ltac2.Uint63.equal, Ltac2.Uint63.t, Ltac2.Char.equal, Ltac2.Char.compare, Ltac2.Constr.Unsafe.Case.equal (#16537, by Jason Gross).

  • Added: Ltac2.Option.equal (#16538, by Jason Gross).

  • Added: syntax for Ltac2 record update { foo with field := bar } (#16552, fixes #10117, by Gaëtan Gilbert).

  • Added: Ltac2 record expressions support punning, i.e. { foo; M.bar } is equivalent to { foo := foo; M.bar := bar } (#16556, by Gaëtan Gilbert).

  • Added: match! goal support for matching hypothesis bodies (#16655, fixes #12803, by Gaëtan Gilbert).

  • Added: quotation and syntax class for preterms (#16740, by Gaëtan Gilbert).

SSReflect

  • Added: port the additions made to ssrfun.v and ssrbool.v in math-comp PR #872 and PR #874, namely definitions olift and pred_oapp as well as lemmas all_sig2_cond, compA, obindEapp, omapEbind, omapEapp, omap_comp, oapp_comp, olift_comp, ocan_comp, eqbLR, eqbRL, can_in_pcan, pcan_in_inj, in_inj_comp, can_in_comp, pcan_in_comp and ocan_in_comp (#16158, by Pierre Roux).

Commands and options

  • Changed: commands which set tactic options (currently Firstorder Solver and Obligation Tactic, as well as any defined by third party plugins) now support export locality. Note that such commands using global without export or using no explicit locality outside sections apply their effects when any module containing it (recursively) is imported. This will change in a future version. (#15274, fixes #15072, by Gaëtan Gilbert).

  • Changed: Hint and Instance commands with no locality attribute are deprecated. Previous versions generated a warning, but this version generates an error by default. This includes all Hint commands described in Creating Hints, Hint Rewrite, and Instance. As mentioned in the error, please add an explicit locality to the hint command. The default was #[global], but we recommend using #[export] where possible (#16004, fixes #13394, by Ali Caglayan).

  • Changed: Transparent obligations generated by Program do not produce an implicit Hint Unfold anymore (#16340, by Pierre-Marie Pédrot).

  • Changed: Print Typeclasses replaces the undocumented Print TypeClasses command which displays the list of typeclasses (#16690, fixes #16686, by Ali Caglayan).

  • Changed: The -async-proofs-tac-j command line option now accepts the argument 0, which makes par block interpreted without spawning any new process (#16837, by Pierre-Marie Pédrot).

  • Removed: the Program Naming flag, which was introduced as an immediately deprecated option in Coq 8.16 (#16519, by Pierre-Marie Pédrot).

  • Removed: undocumented and broken Solve Obligation command (the Solve Obligations command is untouched) (#16842, by Théo Zimmermann).

  • Deprecated :> syntax, to declare fields of Typeclasses as instances, since it is now replaced by :: (see of_type_inst). This will allow, in a future release, making :> declare Implicit Coercions as it does in record definitions (#16230, fixes #16224, by Pierre Roux, reviewed by Ali Caglayan, Jim Fehrle, Gaëtan Gilbert and Pierre-Marie Pédrot).

  • Added: An improved description of Proof using and section variables (#16168, by Jim Fehrle).

  • Added: :: syntax (see of_type_inst) to declare fields of records as typeclass instances (#16230, fixes #16224, by Pierre Roux, reviewed by Ali Caglayan, Jim Fehrle, Gaëtan Gilbert and Pierre-Marie Pédrot).

  • Added: The Print Keywords command, which prints all the currently-defined parser keywords and tokens (#16438, fixes #16375, by Gaëtan Gilbert).

  • Added: Print Grammar can print arbitrary nonterminals or the whole grammar instead of a small adhoc list of nonterminals (#16440, by Gaëtan Gilbert).

  • Fixed: Fast Name Printing flag no longer causes variable name capture when displaying a goal (#16395, fixes #14141, by Wojciech Karpiel).

  • Fixed: vm_compute ignored the bytecode-compiler command line flag (#16931, fixes #16929, by Gaëtan Gilbert).

  • Fixed: The Proof Mode command now gives an error if the specified proof mode doesn't exist. The command was not previously documented (#16981, fixes #16602, by Jim Fehrle).

  • Fixed: Backtracking over grammar modifications from plugins (such as added commands) (#17069, fixes #12575, by Gaëtan Gilbert).

  • Fixed: Anomaly instead of regular error on unsupported applied fix in Function (#17113, fixes #17110, by Hugo Herbelin).

Command-line tools

  • Added: New documentation section Coq configuration basics covering use cases such as setting up Coq with opam, where/how to set up source code for your projects and use of _CoqProject (#15888, by Jim Fehrle).

  • Added: In _CoqProject files, expand paths that are directories to include appropriate files in (sub)directories (#16308, by Jim Fehrle).

  • Fixed: issues when using coq_makefile to build targets requiring both .vo and .glob files (typically documentation targets), where make would run multiple coqc processes on the same source file with racy behaviour (only fixed when using a make supporting "grouped targets" such as GNU Make 4.3) (#16757, by Gaëtan Gilbert).

  • Fixed: Properly process legacy attributes such as Global and Polymorphic in coqdoc to avoid omissions when using the -g (Gallina only) option (#17090, fixes #15933, by Karl Palmskog).

Standard library

  • Changed: Class Saturate in ZifyCLasses.v, PRes now also takes operands (#16355, by František Farka on behalf of BedRock Systems, Inc.).

  • Changed: For uniformity of naming and ease of remembering, R_dist and theorems mentioning R_dist in their name become available with spelling Rdist (#16874, by Hugo Herbelin).

  • Removed: from Nat and N superfluous lemmas rs_rs', rs'_rs'', rbase, A'A_right, ls_ls', ls'_ls'', rs'_rs'', lbase, A'A_left, and also redundant non-negativity assumptions in gcd_unique, gcd_unique_alt, divide_gcd_iff, and gcd_mul_diag_l (#16203, by Andrej Dudenhefner).

  • Deprecated: notation _ ~= _ for JMeq in Coq.Program.Equality (#16436, by Gaëtan Gilbert).

  • Deprecated: lemma Finite_alt in FinFun.v, which is a weaker version of the newly added lemma Finite_dec (#16489, fixes #16479, by Bodo Igler, with help from Olivier Laurent).

  • Deprecated: Zmod, Zdiv_eucl_POS, Zmod_POS_bound, Zmod_pos_bound, and Zmod_neg_bound in ZArith.Zdiv (#16892, by Andres Erbsen).

  • Deprecated: Cyclic.ZModulo.ZModulo because there have been no known use cases for this module and because it does not implement Z/nZ for arbitrary n as one might expect based on the name. The same construction will remain a part of the Coq test suite to ensure consistency of CyclicAxioms (#16914, by Andres Erbsen).

  • Added: lemmas Permutation_incl_cons_inv_r, Permutation_pigeonhole, Permutation_pigeonhole_rel to Permutation.v, and Forall2_cons_iff, Forall2_length, Forall2_impl, Forall2_flip, Forall_Exists_exists_Forall2 to List.v (#15986, by Andrej Dudenhefner, with help from Dominique Larchey-Wendling and Olivier Laurent).

  • Added: modules Nat.Div0 and Nat.Lcm0 in PeanoNat, and N.Div0 and N.Lcm0 in BinNat containing lemmas regarding div and mod, which take into account n div 0 = 0 and n mod 0 = n. Strictly weaker lemmas are deprecated, and will be removed in the future. After the weaker lemmas are removed, the modules Div0 and Lcm0 will be deprecated, and their contents included directly into Nat and N. Locally, you can use Module Nat := Nat.Div0. or Module Nat := Nat.Lcm0. to approximate this inclusion (#16203, fixes #16186, by Andrej Dudenhefner).

  • Added: lemma measure_induction in Nat and N analogous to Wf_nat.induction_ltof1, which is compatible with the using clause for the induction tactic (#16203, by Andrej Dudenhefner).

  • Added: three lemmata related to finiteness and decidability of equality: Listing_decidable_eq, Finite_dec to FinFun.v and lemma NoDup_list_decidable to ListDec.v (#16489, fixes #16479, by Bodo Igler, with help from Olivier Laurent and Andrej Dudenhefner).

  • Added: lemma not_NoDup to ListDec.v and NoDup_app_remove_l, NoDup_app_remove_r to List.v (#16588, by Stefan Haan with a lot of help from Olivier Laurent and Ali Caglayan).

  • Added: the skipn_skipn lemma in Lists/List.v (#16632, by Stephan Boyer).

  • Added: lemmas nth_error_ext, map_repeat, rev_repeat to List.v, and to_list_nil_iff, to_list_inj to VectorSpec.v (#16756, by Stefan Haan).

  • Added: transparent extgcd to replace opaque euclid, euclid_rec, Euclid, and Euclid_intro in Znumtheory. Deprecated compatibility wrappers are provided (#16915, by Andres Erbsen).

Infrastructure and dependencies

  • Changed: Coq is now built entirely using the Dune build system. Packagers and users that build Coq manually must use the new build instructions in the documentation (#15560, by Ali Caglayan, Emilio Jesus Gallego Arias, and Rudi Grinberg).

  • Changed: Coq is not compiled with OCaml's -rectypes option anymore. This means plugins which do not exploit it can also stop passing it to OCaml (#16007, by Gaëtan Gilbert).

  • Changed: Building Coq now requires Dune >= 2.9 (#16118, by Emilio Jesus Gallego Arias).

  • Changed: Coq Makefile targets pretty-timed, make-pretty-timed, make-pretty-timed-before, make-pretty-timed-after, print-pretty-timed, print-pretty-timed-diff, print-pretty-single-time-diff now generate more readable timing tables when absolute paths are used in _CoqProject / the arguments to coq_makefile, by stripping off the absolute prefix (#16268, by Jason Gross).

  • Changed: Coq's configure script now defaults to -native-compiler no. Previously, the default was -native-compiler ondemand, except on Windows. The behavior for users installing through opam does not change, i.e., it is -native-compiler no if the coq-native package is not installed, and -native-compiler yes otherwise (#16997, by Théo Zimmermann).

  • Removed: the -coqide switch to configure in Coq's build infrastructure (it stopped controlling what got compiled in the move to dune) (#16512, by Gaëtan Gilbert).

  • Removed: the -nomacintegration configure flag for CoqIDE. Now CoqIDE will always build with the proper platform-specific integration if available (#16531, by Emilio Jesus Gallego Arias).

  • Added: Coq now supports OCaml 5; note that OCaml 5 is not compatible with Coq's native reduction machine (#15494, #16925, #16947, #16959, #16988, #16991, #16996, #16997, #16999, #17010, and #17015 by Emilio Jesus Gallego Arias, Gaëtan Gilbert, Guillaume Melquiond, Pierre-Marie Pédrot, and others).

  • Added: OCaml 4.14 is now officially supported (#15867, by Gaëtan Gilbert).

Miscellaneous

  • Changed: Module names are now added to the loadpath in alphabetical order for each (sub-)directory. Previously they were added in the order of the directory entries (as shown by "ls -U") (#16725, by Jim Fehrle).

Changes in 8.17.1

A variety of bug fixes and improvements to error messages, including:

  • Fixed: in some cases, coqdep emitted incorrect paths for META files which prevented dune builds for plugins from working correctly (#17270, fixes #16571, by Rodolphe Lepigre).

  • Fixed: Shadowing of record fields in extraction to OCaml (#17324, fixes #12813 and #14843 and #16677, by Hugo Herbelin).

  • Fixed: an impossible to turn off debug message "backtracking and redoing byextend on ..." (#17495, fixes #17488, by Gaëtan Gilbert).

  • Fixed: major memory regression affecting MathComp 2 (#17743, by Enrico Tassi and Pierre Roux).

Version 8.16

Summary of changes

Coq version 8.16 integrates changes to the Coq kernel and performance improvements along with a few new features. We highlight some of the most impactful changes here:

See the Changes in 8.16.0 section below for the detailed list of changes, including potentially breaking changes marked with Changed. Coq's reference manual for 8.16, documentation of the 8.16 standard library and developer documentation of the 8.16 ML API are also available.

Ali Caglayan, Emilio Jesús Gallego Arias, Gaëtan Gilbert and Théo Zimmermann worked on maintaining and improving the continuous integration system and package building infrastructure.

Erik Martin-Dorel has maintained the Coq Docker images that are used in many Coq projects for continuous integration.

The opam repository for Coq packages has been maintained by Guillaume Claret, Karl Palmskog, Matthieu Sozeau and Enrico Tassi with contributions from many users. A list of packages is available at https://coq.inria.fr/opam/www/.

The Coq Platform has been maintained by Michael Soegtrop, with help from Karl Palmskog, Enrico Tassi and Théo Zimmermann.

Our current maintainers are Yves Bertot, Frédéric Besson, Ana Borges, Ali Caglayan, Tej Chajed, Cyril Cohen, Pierre Corbineau, Pierre Courtieu, Maxime Dénès, Jim Fehrle, Julien Forest, Emilio Jesús Gallego Arias, Gaëtan Gilbert, Georges Gonthier, Benjamin Grégoire, Jason Gross, Hugo Herbelin, Vincent Laporte, Olivier Laurent, Assia Mahboubi, Kenji Maillard, Guillaume Melquiond, Pierre-Marie Pédrot, Clément Pit-Claudel, Pierre Roux, Kazuhiko Sakaguchi, Vincent Semeria, Michael Soegtrop, Arnaud Spiwack, Matthieu Sozeau, Enrico Tassi, Laurent Théry, Anton Trunov, Li-yao Xia and Théo Zimmermann. See the Coq Team face book page for more details.

The 57 contributors to the 8.16 versions are Tanaka Akira, Frédéric Besson, Martin Bodin, Ana Borges, Ali Caglayan, Minki Cho, Cyril Cohen, Juan Conejero, "stop-cran", Adrian Dapprich, Maxime Dénès, Stéphane Desarzens, Christian Doczkal, Andrej Dudenhefner, Andres Erbsen, Jim Fehrle, Emilio Jesús Gallego Arias, Attila Gáspár, Paolo G. Giarrusso, Gaëtan Gilbert, Rudi Grinberg, Jason Gross, Hugo Herbelin, Wolf Honore, Jasper Hugunin, Bart Jacobs, Pierre Jouvelot, Ralf Jung, Grant Jurgensen, Jan-Oliver Kaiser, Wojciech Karpiel, Thomas Klausner, Ethan Kuefner, Fabian Kunze, Olivier Laurent, Yishuai Li, Erik Martin-Dorel, Guillaume Melquiond, Jean-Francois Monin, Pierre-Marie Pédrot, Rudy Peterson, Clément Pit-Claudel, Seth Poulsen, Ramkumar Ramachandra, Pierre Roux, Takafumi Saikawa, Kazuhiko Sakaguchi, Gabriel Scherer, Vincent Semeria, Kartik Singhal, Michael Soegtrop, Matthieu Sozeau, Enrico Tassi, Laurent Théry, Anton Trunov, Li-yao Xia and Théo Zimmermann.

The Coq community at large helped improve this new version via the GitHub issue and pull request system, the coq-club@inria.fr mailing list, the Discourse forum and the Coq Zulip chat.

Version 8.16's development spanned 6 months from the release of Coq 8.15.0. Pierre-Marie Pédrot is the release manager of Coq 8.16. This release is the result of 356 merged PRs, closing 99 issues.

Nantes, June 2022,
Matthieu Sozeau for the Coq development team

Changes in 8.16.0

Kernel

  • Changed: Fixpoints are now expected to be guarded even in subterms erasable by reduction, thus getting rid of an artificial obstacle preventing to lift the assumption of weak normalization of Coq to an assumption of strong normalization; for instance (barring implementation bugs) termination of the type-checking algorithm of Coq is now restored (of course, as usual, up to the assumption of the consistency of set theory and type theory, i.e., equivalently, up to the weak normalization of type theory, a "physical" assumption, which has not been contradicted for decades and which specialists commonly believe to be a truth) (#15434, incidentally fixes the complexity issue #5702, by Hugo Herbelin).

  • Changed: Flag Unset Guard Checking nevertheless requires fixpoints to have an argument marked as decreasing in a type which is inductive (#15668, fixes #15621, by Hugo Herbelin).

  • Removed: Template polymorphism is now forbidden for mutual inductive types (#15965, by Gaëtan Gilbert).

  • Fixed: Inlining of non-logical objects (notations, hints, ...) was missing when applying a functor returning one of its arguments as e.g. in Module F (E:T) := E (#15412, fixes #15403, by Hugo Herbelin).

  • Fixed: We introduce a new irrelevant term in the reduction machine. It is used to shortcut computation of terms living in a strict proposition, and behaves as an exception. This restores subject reduction, and also makes conversion of large terms in SProp cheap (#15575, fixes #14015, by Pierre-Marie Pédrot).

  • Fixed: performance blowups while inferring variance information for Cumulative, NonCumulative inductive types (#15662, fixes #11741, by Gaëtan Gilbert).

Specification language, type inference

  • Added: New clause as ident to the Record command to specify the name of the main argument to use by default in the type of projections (#14563, by Hugo Herbelin).

  • Added: Reversible coercions are coercions which cannot be represented by a regular coercion (a Gallina function) but rather a meta procedure, such as type class inference or canonical structure resolution (#15693, by Cyril Cohen, Pierre Roux, Enrico Tassi, reviewed by Ali Caglayan, Jim Fehrle and Gaëtan Gilbert).

  • Added: support for coercions not fulfilling the uniform inheritance condition, allowing more freedom for the parameters that are now inferred using unification, canonical structures or typeclasses (#15789, fixes #2828, #4593, #3115, #5222, #9696 and #8540, by Pierre Roux, reviewed by Ali Caglayan, Enrico Tassi, Kazuhiko Sakaguchi and Jim Fehrle).

  • Fixed: interpretation of {struct} fixpoint annotations when the principal argument comes from an implicit generalization (#15581, fixes #13157, by Gaëtan Gilbert).

Notations

  • Removed: _ in ident entries in notations, which was deprecated in favor of name in 8.13. When you see messages like

    Error: Notation "[ rel _ _ : _ | _ ]" is already defined at level 0
    with arguments name, name, constr, constr while it is now required to be
    at level 0 with arguments ident, ident, constr, constr.
    

    replace ident with name in the Notation command. To ease the change, you can fix the deprecated-ident-entry warnings in Coq 8.15 (or 8.14 or 8.13). The warning can be turned into an error with -arg -w -arg +deprecated-ident-entry in the _CoqProject file (#15754, by Pierre Roux).

  • Added: When defining a recursive notation referring to another recursive notation, expressions of the form x .. y can be used where a sequence of binders is expected (#15291, grants #7911, by Hugo Herbelin).

  • Fixed: Coercions are disabled when typechecking parsers and printers of Number Notation (#15884, fixes #15843, by Pierre Roux).

Tactics

  • Changed: The RewriteRelation type class is now used to declare relations inferable by the setoid_rewrite tactic to construct Proper instances. This can break developments that relied on existing Reflexive instances to infer relations. The fix is to simply add a (backwards compatible) RewriteRelation declaration for the relation. This change allows to set stricter modes on the relation type classes Reflexive, Symmetric, etc. (#13969, fixes #7916, by Matthieu Sozeau).

  • Changed: The setoid_rewrite tactic can now properly recognize homogeneous relations applied to types in different universes (#14138, fixes #13618, by Matthieu Sozeau).

  • Changed: The eauto tactic does not propagate internal Ltac failures with level > 0 anymore. Any failure caused by a hint now behaves as if it were a level 0 error (#15215, fixes #15214, by Pierre-Marie Pédrot).

  • Changed: rewrite when used to rewrite in multiple hypotheses (eg rewrite foo in H,H') requires that the term (foo) does not depend on the hypotheses it rewrites. When using rewrite in *, this means we only rewrite in hypotheses which do not appear in the term (#15426, fixes #3051 and #15448, by Gaëtan Gilbert).

  • Changed: When it fails, assert_succeeds fails with the argument tactic's original error instead of Tactic failure: <tactic closure> fails. (#15728, fixes #10970, by Gaëtan Gilbert).

  • Deprecated: the instantiate tactic without arguments. Since the move to the monadic tactic engine in 8.5, it was behaving as the identity (#15277, by Pierre-Marie Pédrot).

  • Added: generalized rewriting now supports rewriting with (possibly polymorphic) relations valued in Type. Use Classes.CMorphisms instead of Classes.Morphisms to declare Proper instances for rewrite (or setoid_rewrite) to use when rewriting with Type valued relations (#14137, fixes #4632, #5384, #5521, #6278, #7675, #8739, #11011, #12240, and #15279, by Matthieu Sozeau helped by Ali Caglayan).

  • Added: Tactics to obtain a micromega cone expression (aka witness) from an already reified goal. Using those tactics, the user can develop their own micromega tactics for their own types, using their own parsers (#15921, by Pierre Roux, reviewed by Frédéric Besson and Jim Fehrle).

  • Fixed: typeclasses eauto used with multiple hint databases respects priority differences for hints from separate databases (#15289, fixes #5304, by Gaëtan Gilbert).

  • Fixed: cbn has better support for combining simpl nomatch, ! and / specifiers (c.f. Arguments) (#15657, fixes #3989 and #15206, by Gaëtan Gilbert).

Tactic language

  • Changed: Ltac match does not fail when the term to match contains an unfolded primitive projection (#15559, fixes #15554, by Gaëtan Gilbert).

  • Added: Ltac2 understands toplevel_selector and obeys Default Goal Selector. Note that par: is buggy when combined with abstract. Unlike Ltac1 even par: abstract tac is not properly treated (#15378, by Gaëtan Gilbert).

  • Added: Ltac2 Int functions div, mod, asr, lsl, lsr, land, lor , lxor and lnot (#15637, by Michael Soegtrop).

  • Fixed: Ltac2 apply and eapply not unifying with implicit arguments; unification inconsistent with exact and eexact (#15741, by Ramkumar Ramachandra).

SSReflect

Commands and options

  • Changed: Module now only allows parentheses around module arguments. For instance, Module M := (F X). is now a parsing error (#15355, by Gaëtan Gilbert).

  • Changed: Fail no longer catches anomalies, which it has done since Coq version 8.11. Now it only catches user errors (#15366, by Hugo Herbelin).

  • Changed: Program Definition in universe monomorphic mode does not accept non-extensible universe declarations (#15424, fixes #15410, by Gaëtan Gilbert).

  • Changed: The algorithm for name generation of anonymous variables for Program subproofs is now the same as the one used in the general case. This can create incompatibilities in scripts relying on such autogenerated names. The old scheme can be reactivated using the deprecated flag Program Naming (#15442, by Pierre-Marie Pédrot).

  • Removed: Universal Lemma Under Conjunction flag, that was deprecated in 8.15 (#15268, by Théo Zimmermann).

  • Removed: Abort no longer takes an ident as an argument (it has been ignored since 8.5) (#15669, by Gaëtan Gilbert).

  • Removed: Simplex flag, that was deprecated in 8.14. lia and lra will always use the simplex solver (that was already the default behaviour) (#15690, by Frédéric Besson).

  • Deprecated: Add LoadPath and Add Rec LoadPath. If this command is an important feature for you, please open an issue on GitHub <https://github.com/coq/coq/issues> and explain your workflow (#15652, by Gaëtan Gilbert).

  • Deprecated: the Typeclasses Filtered Unification flag. Due to a buggy implementation, it is unlikely this is used in the wild (#15752, by Pierre-Marie Pédrot).

  • Added: Scheme Boolean Equality command to generate the boolean equality for an inductive type whose equality is decidable. It is useful when Coq is able to generate the boolean equality but isn't powerful enough to prove the decidability of equality (unlike Scheme Equality, which tries to prove the decidability of the type) (#15526, by Hugo Herbelin).

  • Added: New more extensive algorithm based on the "parametricity" translation for canonically generating Boolean equalities associated to a decidable inductive type (#15527, by Hugo Herbelin).

  • Added: From Dependency command to declare a dependency of a .v file on an external file. The coqdep tool generates build dependencies accordingly (#15650, fixes #15600, by Enrico Tassi).

  • Added: Print Notation command that prints the level and associativity of a given notation definition string (#15683, fixes #14907 and #4436 and #7730, by Ali Caglayan and Ana Borges, with help from Emilio Jesus Gallego Arias).

  • Added: a warning when trying to deprecate a definition (#15760, by Pierre Roux).

  • Added: A deprecation warning that the Class > syntax, which currently does nothing, will in the future declare coercions as it does when used in Record commands (#15802, by Pierre Roux, reviewed by Gaëtan Gilbert, Ali Caglayan, Jason Gross, Jim Fehrle and Théo Zimmermann).

  • Added: the nonuniform boolean attribute that silences the non-uniform-inheritance warning when user needs to declare such a coercion on purpose (#15853, by Pierre Roux, reviewed by Gaëtan Gilbert and Jim Fehrle).

  • Added: All commands which can import modules (e.g. Module Import M., Module F (Import X : T)., Require Import M., etc) now support import_categories. Require Import and Require Export also support filtered_import (#15945, fixes #14872, by Gaëtan Gilbert).

  • Fixed: Make Require Import M. equivalent to Require M. Import M. (#15347, fixes #3556, by Maxime Dénès).

Command-line tools

  • Added: coq_makefile variable COQPLUGININSTALL to configure the installation of ML plugins (#15788, by Cyril Cohen and Enrico Tassi).

  • Added: Added -bytecode-compiler yesno flag for coqchk enabling vm_compute during checks, which is off by default (#15886, by Ali Caglayan).

  • Fixed: coqdoc confused by the presence of command Load in a file (#15511, fixes #15497, by Hugo Herbelin).

CoqIDE

  • Added: Documentation of editing failed async mode proofs, how to configure key bindings and various previously undocumented details (#16070, by Jim Fehrle).

Standard library

  • Changed: the signature scope of Classes.CMorphisms into signatureT (#15446, by Olivier Laurent).

  • Changed: the locality of typeclass instances Permutation_app' and Permutation_cons from global to export (#15597, fixes #15596, by Gaëtan Gilbert).

  • Removed: Int63, which was deprecated in favor of Uint63 in 8.14 (#15754, by Pierre Roux).

  • Deprecated: some obsolete files from the Arith part of the standard library (Div2, Even, Gt, Le, Lt, Max, Min, Minus, Mult, NPeano, Plus). Import Arith_base instead of these files. References to items in the deprecated files should be replaced with references to PeanoNat.Nat as suggested by the warning messages. Concerning the definitions of parity properties (even and odd), it is recommended to use Nat.Even and Nat.Odd. If an inductive definition of parity is required, the mutually inductive Nat.Even_alt and Nat.Odd_alt can be used. However, induction principles for Nat.Odd and Nat.Even are available as Nat.Even_Odd_ind and Nat.Odd_Even_ind. The equivalence between the non-inductive and mutually inductive definitions of parity can be found in Nat.Even_alt_Even and Nat.Odd_alt_Odd. All Hint declarations in the arith database have been moved to Arith_prebase and Arith_base. To use the results about Peano arithmetic, we recommend importing PeanoNat (or Arith_base to base it on the arith hint database) and using the Nat module. Arith_prebase has been introduced temporarily to ensure compatibility, but it will be removed at the end of the deprecation phase, e.g. in 8.18. Its use is thus discouraged (#14736, #15411, by Olivier Laurent, with help of Karl Palmskog).

  • Deprecated: identity inductive (replaced by the equivalent eq). Init.Logic_Type is removed (the only remaining definition notT is moved to Init.Logic) (#15256, by Olivier Laurent).

  • Deprecated: P_Rmin: use more general Rmin_case instead (#15388, fixes #15382, by Olivier Laurent).

  • Added: lemma count_occ_rev (#15397, by Olivier Laurent).

  • Added: Nat.EvenT and Nat.OddT (almost the same as Nat.Even and Nat.Odd but with output in Type. Decidability of parity (with output Type) is provided EvenT_OddT_dec as well as induction principles Nat.EvenT_OddT_rect and Nat.OddT_EvenT_rect (with output Type) (#15427, by Olivier Laurent).

  • Added: Added a proof of sin x < x for positive x and x < sin x for negative x (#15599, by stop-cran).

  • Added: decidability typeclass instances for Z.le, Z.lt, Z.ge and Z.gt, added lemmas Z.geb_ge and Z.gtb_gt (#15620, by Michael Soegtrop).

  • Added: lemmas Rinv_inv, Rinv_mult, Rinv_opp, Rinv_div, Rdiv_opp_r, Rsqr_div', Rsqr_inv', sqrt_inv, Rabs_inv, pow_inv, powerRZ_inv', powerRZ_neg', powerRZ_mult, cv_infty_cv_0, which are variants of existing lemmas, but without any hypothesis (#15644, by Guillaume Melquiond).

  • Added: a Leibniz equality test for primitive floats (#15719, by Pierre Roux, reviewed by Guillaume Melquiond).

  • Added: support for primitive floats in Scheme Boolean Equality (#15719, by Pierre Roux, reviewed by Hugo Herbelin).

  • Added: lemma le_add_l to NAddOrder.v. Use Nat.le_add_l as replacement for the deprecated Plus.le_plus_r (#16184, by Andrej Dudenhefner).

Infrastructure and dependencies

  • Changed: Bumped lablgtk3 lower bound to 3.1.2 (#15947, by Pierre-Marie Pédrot).

  • Changed: Load plugins using findlib. This requires projects built with coq_makefile to either provide a hand written META file or use the -generate-meta-for-package option when applicable. As a consequence Declare ML Module now uses plugin names according to findlib, e.g. coq-aac-tactics.plugin. coqdep accepts -m META and uses the file to resolve plugin names to actual file names (#15220, fixes #7698, by Enrico Tassi).

  • Changed: Minimum supported zarith version is now 1.11 (#15483 and #16005 and #16030, closes #15496, by Gaëtan Gilbert and Théo Zimmermann and Jason Gross).

  • Changed: Bump the minimum OCaml version to 4.09.0. As a consequence the minimum supported ocamlfind version is now 1.8.1 (#15947 and #16046, fixes #14260 and #16015, by Pierre-Marie Pédrot and Théo Zimmermann).

Extraction

  • Changed: ExtrOCamlInt63 no longer extracts comparison to int in OCaml; the extraction of Uint63.compare and Sint63.compare was also adapted accordingly (#15294, fixes #15280, by Li-yao Xia).

  • Changed: Extraction from nat to OCaml int uses Stdlib instead of Pervasives (#15333, by Rudy Nicolo Peterson).

  • Changed: The empty inductive type is now extracted to OCaml empty type available since OCaml 4.07 (#15967, by Pierre Roux).

  • Added: More extraction definitions for division and comparison of Z and N (#15098, by Li-yao Xia).

  • Fixed: Type int in files Number.v, Decimal.v and Hexadecimal.v have been renamed to signed_int (together with a compatibility alias int) so that they can be used in extraction without conflicting with OCaml's int type (#13460, fixes #7017 and #13288, by Hugo Herbelin).

Changes in 8.16.1

Kernel

  • Fixed: conversion of Prod values in the native compiler (#16651, fixes #16645, by Pierre-Marie Pédrot).

  • Fixed: Coq 8.16.0 missed SProp check for opaque names in conversion (#16768, fixes #16752, by Hugo Herbelin).

  • Fixed: Pass the correct environment to compute η-expansion of cofixpoints in VM and native compilation (#16845, fixes #16831, by Pierre-Marie Pédrot).

  • Fixed: inconsistency with conversion of primitive arrays, and associated incomplete strong normalization of primitive arrays with lazy (#16850, fixes #16829, by Gaëtan Gilbert, reported by Maxime Buyse and Andres Erbsen).

Commands and options

CoqIDE

  • Fixed: "Interrupt computations" now works correctly on Windows—except if you start CoqIDE as a background process, e.g. with coqide & in bash, in which case it won't work at all (#16142, fixes #13550, by Jim Fehrle).

Version 8.15

Summary of changes

Coq version 8.15 integrates many bug fixes, deprecations and cleanups as well as a few new features. We highlight some of the most impactful changes here:

See the Changes in 8.15.0 section below for the detailed list of changes, including potentially breaking changes marked with Changed. Coq's reference manual for 8.15, documentation of the 8.15 standard library and developer documentation of the 8.15 ML API are also available.

Emilio Jesús Gallego Arias, Gaëtan Gilbert, Michael Soegtrop and Théo Zimmermann worked on maintaining and improving the continuous integration system and package building infrastructure.

Erik Martin-Dorel has maintained the Coq Docker images that are used in many Coq projects for continuous integration.

The opam repository for Coq packages has been maintained by Guillaume Claret, Karl Palmskog, Matthieu Sozeau and Enrico Tassi with contributions from many users. A list of packages is available at https://coq.inria.fr/opam/www/.

The Coq Platform has been maintained by Michael Soegtrop and Enrico Tassi.

Our current maintainers are Yves Bertot, Frédéric Besson, Ali Caglayan, Tej Chajed, Cyril Cohen, Pierre Corbineau, Pierre Courtieu, Maxime Dénès, Jim Fehrle, Julien Forest, Emilio Jesús Gallego Arias, Gaëtan Gilbert, Georges Gonthier, Benjamin Grégoire, Jason Gross, Hugo Herbelin, Vincent Laporte, Olivier Laurent, Assia Mahboubi, Kenji Maillard, Guillaume Melquiond, Pierre-Marie Pédrot, Clément Pit-Claudel, Pierre Roux, Kazuhiko Sakaguchi, Vincent Semeria, Michael Soegtrop, Arnaud Spiwack, Matthieu Sozeau, Enrico Tassi, Laurent Théry, Anton Trunov, Li-yao Xia and Théo Zimmermann. See the Coq Team face book page for more details.

The 41 contributors to this version are Tanaka Akira, Frédéric Besson, Juan Conejero, Ali Caglayan, Cyril Cohen, Adrian Dapprich, Maxime Dénès, Stéphane Desarzens, Christian Doczkal, Andrej Dudenhefner, Jim Fehrle, Emilio Jesús Gallego Arias, Attila Gáspár, Gaëtan Gilbert, Jason Gross, Hugo Herbelin, Jasper Hugunin, Bart Jacobs, Ralf Jung, Grant Jurgensen, Jan-Oliver Kaiser, Wojciech Karpiel, Fabian Kunze, Olivier Laurent, Yishuai Li, Erik Martin-Dorel, Guillaume Melquiond, Jean-Francois Monin, Pierre-Marie Pédrot, Rudy Peterson, Clément Pit-Claudel, Seth Poulsen, Pierre Roux, Takafumi Saikawa, Kazuhiko Sakaguchi, Michael Soegtrop, Matthieu Sozeau, Enrico Tassi, Laurent Théry, Anton Trunov and Théo Zimmerman.

The Coq community at large helped improve the design of this new version via the GitHub issue and pull request system, the Coq development mailing list coqdev@inria.fr, the coq-club@inria.fr mailing list, the Discourse forum and the Coq Zulip chat.

Version 8.15's development spanned 3 months from the release of Coq 8.14.0. Gaëtan Gilbert is the release manager of Coq 8.15. This release is the result of 384 merged PRs, closing 143 issues.

Nantes, January 2022,
Matthieu Sozeau for the Coq development team

Changes in 8.15.0

Kernel

  • Fixed: Name clash in a computation of the type of parameters of functorial module types; this computation was provided for the purpose of clients using the algebraic form of module types such as Print Module Type (#15385, fixes #9555, by Hugo Herbelin).

Specification language, type inference

  • Changed: Instance warns about the default locality immediately rather than waiting until the instance is ready to be defined. This changes which command warns when the instance has a separate proof: the Instance command itself warns instead of the proof closing command (such as Defined) (#14705, by Gaëtan Gilbert).

  • Removed: Arguments of section variables may no longer be renamed with Arguments (this was previously applied inconsistently) (#14573, by Gaëtan Gilbert).

  • Added: Non-dependent implicit arguments can be provided explicitly using the syntax (natural := term) where natural is the index of the implicit argument among all non-dependent arguments of the function, starting from 1 (#11099, by Hugo Herbelin).

  • Added: Succeed, a control_command that verifies that the given sentence succeeds without changing the proof state (#14750, by Gaëtan Gilbert).

  • Fixed: The term.(qualid arg*) syntax now takes into account the position of the main argument term when computing the implicit arguments of qualid (#14606, fixes #4167, by Hugo Herbelin).

  • Fixed: Source and target of coercions preserved by module instantiation (#14668, fixes #3527, by Hugo Herbelin).

  • Fixed: Made reference manual consistent with the implementation regarding the role of recursively non-uniform parameters of inductive types in the nested positivity condition (#14967, fixes #14938, by Hugo Herbelin)

Notations

  • Changed: Terms printed in error messages may be more verbose if syntactic sugar would make it appear that the obtained and expected terms only differ in existential variables (#14672, by Gaëtan Gilbert).

  • Removed: the Numeral Notation command that was renamed to Number Notation in 8.13 (#14819, by Pierre Roux).

  • Removed: primitive float notations <, <= and == that were replaced by <?, <=? and =? in 8.13 (#14819, by Pierre Roux).

  • Removed: primitive integer notations \%, <, <= and == that were replaced by mod, <?, <=? and =? in 8.13 (#14819, by Pierre Roux).

  • Added: Include floats in the number notation mechanism (#14525, by Pierre Roux).

  • Added: Coercion entries and ident/global entries in custom notations now respect the only parsing modifier (#15340, fixes #15335, by Hugo Herbelin).

  • Fixed: Reserved Infix now accept further parameters in the infix notation (#14379, fixes #11402, by Hugo Herbelin).

  • Fixed: Useless self reference when printing abbreviations declared in nested modules (#14493, fixes one part of #12777 and #14486, by Hugo Herbelin).

  • Fixed: anomalies with notation applied in match patterns when the notation have a notation variable at head (#14713, fixes #14708, by Hugo Herbelin).

  • Fixed: Regression in parsing error reporting in case of empty custom entry (#15338, fixes #15334, by Hugo Herbelin).

Tactics

  • Changed: apply with does not rename arguments unless using compatibility flag Apply With Renaming (#13837, fixes #13759, by Gaëtan Gilbert).

    Porting hint: if the renaming is because of a goal variable (eg intros x; apply foo with (x0 := bar) where About foo. says the argument is called x) it is probably caused by an interaction with implicit arguments and apply @foo with (x := bar) will usually be a backwards compatible fix.

  • Changed: Hint Unfold in discriminated databases now respects its specification, namely that a constant may be unfolded only when it is the head of the goal. The previous behavior was to perform unfolding on any goal, without any limitation.

    An unexpected side-effect of this was that a database that contained Unfold hints would sometimes trigger silent strong βι-normalization of the goal. Indeed, unfold performs such a normalization regardless of the presence of its argument in the goal. This does introduce a bit of backwards incompatibility, but it occurs in very specific situations and is easily circumvented. Since by default hint bases are not discriminated, it means that incompatibilities are typically observed when adding unfold hints to the typeclass database.

    In order to recover the previous behavior, it is enough to replace instances of Hint Unfold foo. with Hint Extern 4 => progress (unfold foo).. A less compatible but finer-grained change can be achieved by only adding the missing normalization phase with Hint Extern 4 => progress (lazy beta iota). (#14679, fixes #14874, by Pierre-Marie Pédrot).

  • Changed: Correctly consider variables without a body to be rigid for the pattern recognition algorithm of discriminated hints (#14722, by Pierre-Marie Pédrot).

  • Changed: Use discrimination nets for goals containing evars in all auto tactics. It essentially makes the behavior of undiscriminated databases to be the one of discriminated databases where all constants are considered transparent. This may be incompatible with previous behavior in very rare cases (#14848, by Pierre-Marie Pédrot).

  • Changed: The choice strategy for rewrite_strat is now of arbitrary arity (#14989, fixes #6109, by Gaëtan Gilbert).

  • Changed: The exact tactic now takes a uconstr as argument instead of an ad-hoc one. In very rare cases, this can change the order of resolution of dependent evars when used over several goals at once (#15171, by Pierre-Marie Pédrot).

  • Changed: cbn interprets the combination of the ! and / modifiers (from Arguments) to mean "unfold as soon as all arguments before the / are provided and all arguments marked with ! reduce to a constructor". This makes it unfold more often than without the / when all arguments are provided. Previously adding / would only prevent unfolding when insufficient arguments are provided without adding new unfoldings.

    Note that this change only takes effect in default mode (as opposed to when simpl nomatch was used) (#15204, fixes #4555 and #7674, by Gaëtan Gilbert).

  • Removed: the deprecated new auto tactic (#14527, by Pierre-Marie Pédrot).

  • Removed: deprecated syntax for instantiate using capitalized Value or Type (#15193, by Gaëtan Gilbert).

  • Removed: deprecated autoapply ... using syntax for autoapply (#15194, by Gaëtan Gilbert).

  • Deprecated: the bfs eauto tactic. Since its introduction it has behaved exactly like the eauto tactic. Use typeclasses eauto with the bfs flag instead (#15314, fixes #15300, by Pierre-Marie Pédrot).

  • Added: The zify tactic can now recognize Pos.Nsucc_double, Pos.Ndouble, N.succ_double, N.double, N.succ_pos, N.div2, N.pow, N.square, and Z.to_pos. Moreover, importing module ZifyBool lets it recognize Pos.eqb, Pos.leb, Pos.ltb, N.eqb, N.leb, and N.ltb (#10998, by Kazuhiko Sakaguchi).

  • Added: best_effort option to typeclasses eauto, to return a partial solution to its initial proof-search problem. The goals that can remain unsolved are determined according to the modes declared for their head (see Hint Mode). This is used by typeclass resolution during type inference to provide more informative error messages (#13952, fixes #13942 and #14125, by Matthieu Sozeau).

  • Added: A new Keep Equalities table to selectively control the preservation of subterm equalities for the injection tactic. It allows a finer control than the boolean flag Keep Proof Equalities that acts globally (#14439, by Pierre-Marie Pédrot).

  • Added: simple congruence tactic which works like congruence but does not unfold definitions (#14657, fixes #13778 and #5394 and #13189, by Andrej Dudenhefner).

  • Added: Small enhancement of unification in the presence of local definitions (#14673, fixes #4415, by Hugo Herbelin).

  • Added: dfs option in typeclasses eauto to use depth-first search (#14693, fixes #13859, by Ali Caglayan).

  • Fixed: More flexible hypothesis specialization in congruence (#14650, fixes #14651 and #14662, by Andrej Dudenhefner).

  • Fixed: Added caching to congruence initialization to avoid quadratic runtime (#14683, fixes #5548, by Andrej Dudenhefner).

  • Fixed: Correctly handle matching up to η-expansion in discriminated hints (#14732, fixes #14731, by Pierre-Marie Pédrot).

  • Fixed: Old unification understands some inductive cumulativity (#14758, fixes #14734 and #6976, by Gaëtan Gilbert).

  • Fixed: The clear dependent tactic now does not backtrack internally, preventing an exponential blowup (#14984, fixes #11689, by Pierre-Marie Pédrot).

  • Fixed: setoid_rewrite now works when the rewriting lemma has non dependent arguments and rewriting under binders (#14986, fixes #5369, by Gaëtan Gilbert).

  • Fixed: Regression in 8.14.0 and 8.14.1 with action pattern % in as clause of tactic specialize (#15245, fixes #15244, by Hugo Herbelin).

Tactic language

  • Fixed: the parsing level of the Ltac2 tactic now was set to level 6 in order to behave as it did before 8.14 (#15250, fixes #15122, by Pierre-Marie Pédrot).

SSReflect

  • Changed: rewrite generates subgoals in the expected order (side conditions first, by default) also when rewriting with a setoid relation (#14314, fixes #5706, by Enrico Tassi).

  • Removed: The ssrsearch plugin and the ssr Search command (#13760, by Jim Fehrle).

  • Added: port the additions made to ssrbool.v in math-comp PR #757, namely reflect combinators negPP, orPP, andPP and implyPP (#15059, by Christian Doczkal).

  • Fixed: SSR patterns now work with primitive values such as ints, floats or arrays (#14660, fixes #12770, by Juan Conejero).

  • Fixed: A bug where suff would fail due to use of apply under the hood (#14687, fixes #14678, by Ali Caglayan helped by Enrico Tassi).

Commands and options

  • Changed: About and Print now display all known argument names (#14596, grants #13830, by Hugo Herbelin).

  • Changed: Typeclasses Transparent and Typeclasses Opaque support #[local], #[export] and #[global] attributes (#14685, fixes #14513, by Gaëtan Gilbert).

  • Changed: In extraction to OCaml, empty types in Type (such as Empty_set) are now extracted to an abstract type (empty by construction) rather than to the OCaml's unit type (#14802, fixes a remark at #14801, by Hugo Herbelin).

  • Changed: Closed modules now live in a separate namespace from open modules and sections (#15078, fixes #14529, by Gaëtan Gilbert).

  • Removed: boolean attributes monomorphic, noncumulative and notemplate that were replaced by polymorphic=no, cumulative=no and template=no in 8.13 (#14819, by Pierre Roux).

  • Removed: command Grab Existential Variables that was deprecated in 8.13. Use Unshelve that is mostly equivalent, up to the reverse order of the resulting subgoals (#14819, by Pierre Roux).

  • Removed: command Existential that was deprecated in 8.13. Use Unshelve and exact (#14819, by Pierre Roux).

  • Removed: the -outputstate command line argument and the corresponding vernacular commands Write State and Restore State (#14940, by Pierre-Marie Pédrot)

  • Deprecated: ambiguous Proof using and Collection usage (#15056, fixes #13296, by Wojciech Karpiel).

  • Deprecated: Universal Lemma Under Conjunction flag that was introduced for compatibility with Coq versions prior to 8.4 (#15272, by Théo Zimmermann).

  • Deprecated: using Hint Cut, Hint Mode, Hint Transparent, Hint Opaque, Typeclasses Transparent or Typeclasses Opaque without an explicit locality outside sections. (#14697, by Pierre-Marie Pédrot, and #14685, by Gaëtan Gilbert)

  • Added: The Mangle Names Light flag, which changes the behavior of Mangle Names. For example, the name foo becomes _0 with Mangle Names, but with Mangle Names Light set, it will become _foo (#14695, fixes #14548, by Ali Caglayan).

  • Added: The Hint Cut, Hint Mode, Hint Transparent, Hint Opaque, Typeclasses Transparent and Typeclasses Opaque commands now accept the local, export and global locality attributes inside sections. With either attribute, the commands will trigger the non-local-section-hint warning if the arguments refer to local section variables (#14697, by Pierre-Marie Pédrot, and #14685, fixes #14513, by Gaëtan Gilbert).

  • Added: projections(primitive) attribute to make a record use primitive projections (#14699, fixes #13150, by Ali Caglayan).

  • Added: Syntax for import_categories providing selective import of module components (eg Import(notations) M (#14892, by Gaëtan Gilbert).

  • Added: Search understands modifier in as an alias of inside (#15139, fixes #14930, by Gaëtan Gilbert). This is intended to ease transition for ssreflect Search users.

  • Fixed: interaction of Program's obligation state and modules and sections: obligations started in a parent module or section are not available to be solved until the submodules and subsections are closed (#14780, fixes #14446, by Gaëtan Gilbert).

  • Fixed: Eval and Compute now beta-iota-simplify the type of the result, like Check does (#14901, fixes #14899, by Hugo Herbelin)

Command-line tools

  • Changed: Coqdoc options --coqlib and --coqlib_path have been renamed to --coqlib_url and --coqlib to make them more consistent with flags used by other Coq executables (#14059, by Emilio Jesus Gallego Arias).

  • Changed: Syntax of _CoqProject files: -arg is now handled by coq_makefile and not by make. Unquoted # now start line comments (#14558, by Stéphane Desarzens, with help from Jim Fehrle and Enrico Tassi).

  • Changed: Require now selects files whose logical name exactly matches the required name, making it possible to unambiguously select a given file: if several -Q or -R options bind the same logical name to a different file, the option appearing last on the command line takes precedence. Moreover, it is now an error to require a file using a partial logical name which does not resolve to a non-ambiguous path (#14718, by Hugo Herbelin).

  • Changed: coq_makefile now declares variable COQBIN to avoid warnings in make --warn mode (#14787, by Clément Pit-Claudel).

  • Changed: coqchk respects the Kernel Term Sharing flag instead of forcing it on (#14957, by Gaëtan Gilbert)

  • Removed: These options of coq_makefile: -extra, -extra-phony, -custom, -no-install, -install, -no-opt, -byte. Support for subdirectories is also removed (#14558, by Stéphane Desarzens, with help from Jim Fehrle and Enrico Tassi).

  • Added: coq_makefile now takes the -docroot option as alternative to the INSTALLCOQDOCROOT variable (#14558, by Stéphane Desarzens, with help from Jim Fehrle and Enrico Tassi).

  • Fixed: Various coqdep issues with the From clause of Require and a few inconsistencies between coqdep and coqc disambiguation of Require (#14718, fixes #11631 and #14539, by Hugo Herbelin).

  • Fixed: coq_makefile has improved logic when dealing with incorrect _CoqProject files (#13541, fixes #9319, by Fabian Kunze).

  • Fixed: coqdep was confusing periods occurring in comments with periods ending Coq sentences (#14996, fixes #7393, by Hugo Herbelin).

CoqIDE

  • Changed: CoqIDE unicode keys for brackets (e.g. langle) now bind to unicode mathematical symbols rather than unicode CJK brackets (#14452, by Bart Jacobs).

  • Changed: All occurrences of the name CoqIde to CoqIDE. This may cause issues with installing and uninstalling desktop icons, causing apparent duplicates (#14696, fixes #14310, by Ali Caglayan).

  • Added: Initial version of a visual debugger in CoqIDE. Supports setting breakpoints visually and jumping to the stopping point plus continue, step over, step in and step out operations. Displays the call stack and variable values for each stack frame. Currently only for Ltac. See the documentation here (#14644, fixes #13967, by Jim Fehrle)

  • Fixed: It is now possible to deactivate the unicode completion mechanism in CoqIDE (#14863, by Pierre-Marie Pédrot).

Standard library

  • Changed: Permutation-related Proper instances are now at default priority instead of priority 10 (#14574, fixes #14571, by Gaëtan Gilbert).

  • Changed: The new type of epsilon_smallest is (exists n : nat, P n) -> { n : nat | P n /forall k, P k -> n <= k }. Here the minimality of n is expressed by forall k, P k -> n <= k corresponding to the intuitive meaning of minimality "the others are greater", whereas the previous version used the negative equivalent formulation forall k, k < n -> ~P k. Scripts using epsilon_smallest can easily be adapted using lemmas le_not_lt and lt_not_le from the standard library (#14601, by Jean-Francois Monin).

  • Changed: ltb and leb functions for ascii, into comparison-based definition (#14234, by Yishuai Li).

  • Removed: the file Numeral.v that was replaced by Number.v in 8.13 (#14819, by Pierre Roux).

  • Removed: some *_invol functions that were renamed *_involutive for consistency with the remaining of the stdlib in 8.13 (#14819, by Pierre Roux).

  • Deprecated: frexp and ldexp in FloatOps.v, renamed Z.frexp and Z.ldexp (#15085, by Pierre Roux).

  • Added: A proof that incoherent equivalences can be adjusted to adjoint equivalences in Logic.Adjointification (#13408, by Jasper Hugunin).

  • Added: ltb and leb functions for string, and some lemmas about them;

  • Added: simple non dependent product slexprod in Relations/Relation_Operators.v and its proof of well-foundness wf_slexprod in Wellfounded/Lexicographic_Product.v (#14809, by Laurent Thery).

  • Added: The notations (x; y), x.1, x.2 for sigT are now exported and available after Import SigTNotations. (#14813, by Laurent Théry).

  • Added: The function sigT_of_prod turns a pair A * B into {_ : A & B}. Its inverse function is prod_of_sigT. This is shown by theorems sigT_prod_sigT and prod_sigT_prod (#14813, by Laurent Théry).

  • Fixed: split_combine lemma for lists, making it usable (#14458, by Yishuai Li).

Infrastructure and dependencies

  • Changed: Coq's continuous integration now provides a more accessible Windows installer artifact in the "Checks" GitHub tab, both for pull requests and the master branch.

    This facilitates testing Coq's bleeding edge builds on Windows, and should be more reliable than the previous setup (#12425, by Emilio Jesus Gallego Arias).

  • Changed: Coq's ./configure script has gone through a major cleanup. In particular, the following options have been removed:

    • -force-caml-version, -force-findlib-version: Coq won't compile with OCaml or findlib lower than the required versions;

    • -vmbyteflags, -custom, -no-custom: linking options for toplevels are now controlled in topbin/dune;

    • -ocamlfind: Coq will now use the toolchain specified in the Dune configuration; this can be controlled using the workspaces feature;

    • -nodebug: Coq will now follow the standard, which is to always pass -g to OCaml; this can be modified using a custom Dune workspace;

    • -flambda-opts: compilation options are now set in Coq's root dune file, can be updated using a custom Dune workspace;

    • -local, -bindir, -coqdocdir, -annotate, -camldir, -profiling: these flags were deprecated in 8.14, and are now removed.

    Moreover, the -annot and -bin-annot flags only take effect to set coq-makefile's defaults (#14189, by Emilio Jesus Gallego Arias).

  • Changed: Configure will now detect the Dune version, and will correctly pass -etcdir and -docdir to the install procedure if Dune >= 2.9 is available. Note that the -docdir configure option now refers to root path for documentation. If you would like to install Coq documentation in foo/coq, use -docdir foo (#14844, by Emilio Jesus Gallego Arias).

  • Changed: OCaml 4.13 is now officially supported (#14879, by Emilio Jesus Gallego Arias)

  • Changed: Sphinx 3.0.2 or above is now required to build the reference manual (#14963, by Théo Zimmermann)

Extraction

  • Changed: replaced Big module with Big_int_Z functions from zarith.

    OCaml code extracted with the following modules should be linked to the Zarith library.

    • ExtrOcamlNatBigInt

    • ExtrOcamlZBigInt

    Removed ExtrOcamlBigIntConv module.

    (#8252, by Yishuai Li).

  • Fixed: compilation errors in ExtrOcamlString and ExtrOcamlNativeString (#15075, fixes #15076, by Yishuai Li).

Changes in 8.15.1

Kernel

  • Fixed: cases of incompletenesses in the guard condition for fixpoints in the presence of cofixpoints or primitive projections (#15498, fixes #15451, by Hugo Herbelin).

  • Fixed: inconsistency when using module subtyping with squashed inductives (#15839, fixes #15838 (reported by Pierre-Marie Pédrot), by Gaëtan Gilbert).

Notations

  • Fixed: Check for prior declaration of a custom entry was missing for notations in only printing mode (#15628, fixes #15619, by Hugo Herbelin).

Tactics

  • Fixed: rewrite_strat regression in 8.15.0 related to Transitive instances (#15577, fixes #15568, by Gaëtan Gilbert).

  • Fixed: When setoid_rewrite succeeds in rewriting at some occurrence but the resulting equality is the identity, it now tries rewriting in subterms of that occurrence instead of giving up (#15612, fixes #8080, by Gaëtan Gilbert).

  • Fixed: Ill-typed goals created by clearbody in the presence of transitive dependencies in the body of a hypothesis (#15634, fixes #15606, by Hugo Herbelin).

  • Fixed: cbn knows to refold fixpoints when Arguments with / and ! was used (#15653, fixes #15567, by Gaëtan Gilbert).

Command-line tools

  • Fixed: a bug where coqc -vok was not creating an empty '.vok' file (#15745, by Ramkumar Ramachandra).

CoqIDE

  • Fixed: Line numbers shown in the Errors panel were incorrect; they didn't match the error locations in the script panel (#15532, fixes #15531, by Jim Fehrle).

  • Fixed: anomaly when using proof diffs with no focused goal (#15633, fixes #15578, by Jim Fehrle).

  • Fixed: Attempted edits to the processed part of a buffer while Coq is busy processing a request are now ignored to ensure "processed" highlighting is accurate (#15714, fixes #15733 and #15675 and #15725, by Jim Fehrle).

Miscellaneous

  • Fixed: Ensure that the names of arguments of inductive schemes are distinct so that the new Coq 8.15 preservation of argument names in the with clause of tactics in #13837 works as in Coq 8.14 for these schemes (#15537, fixes #15420, by Hugo Herbelin).

Changes in 8.15.2

Tactics

  • Added: intuition and dintuition use Tauto.intuition_solver (defined as auto with *) instead of hardcoding auto with *. This makes it possible to change the default solver with Ltac Tauto.intuition_solver ::= ... (#15866, fixes #7725, by Gaëtan Gilbert).

  • Fixed: uncaught exception UnableToUnify with bidirectionality hints (#16066, fixes #16063, by Gaëtan Gilbert).

CoqIDE

Standard library

  • Fixed: an incorrect implementation of SFClassify, allowing for a proof of False since 8.11.0, due to Axioms present in Float.Axioms (#16101, fixes #16096, by Ali Caglayan).

Version 8.14

Summary of changes

Coq version 8.14 integrates many usability improvements, as well as an important change in the core language. The main changes include:

  • The internal representation of match has changed to a more space-efficient and cleaner structure, allowing the fix of a completeness issue with cumulative inductive types in the type-checker. The internal representation is now closer to the user-level view of match, where the argument context of branches and the inductive binders in and as do not carry type annotations.

  • A new coqnative binary performs separate native compilation of libraries, starting from a .vo file. It is supported by coq_makefile.

  • Improvements to typeclasses and canonical structure resolution, allowing more terms to be considered as classes or keys.

  • More control over notations declarations and support for primitive types in string and number notations.

  • Removal of deprecated tactics, notably omega, which has been replaced by a greatly improved lia, along with many bug fixes.

  • New Ltac2 APIs for interaction with Ltac1, manipulation of inductive types and printing.

  • Many changes and additions to the standard library in the numbers, vectors and lists libraries. A new signed primitive integers library Sint63 is available in addition to the unsigned Uint63 library.

See the Changes in 8.14.0 section below for the detailed list of changes, including potentially breaking changes marked with Changed. Coq's reference manual, documentation of the standard library and developer documentation of the ML API are also available.

Emilio Jesús Gallego Arias, Gaëtan Gilbert, Michael Soegtrop and Théo Zimmermann worked on maintaining and improving the continuous integration system and package building infrastructure.

Erik Martin-Dorel has maintained the Coq Docker images that are used in many Coq projects for continuous integration.

The opam repository for Coq packages has been maintained by Guillaume Claret, Karl Palmskog, Matthieu Sozeau and Enrico Tassi with contributions from many users. A list of packages is available at https://coq.inria.fr/opam/www/.

The Coq Platform has been maintained by Michael Soegtrop and Enrico Tassi.

Our current maintainers are Yves Bertot, Frédéric Besson, Ali Caglayan, Tej Chajed, Cyril Cohen, Pierre Corbineau, Pierre Courtieu, Maxime Dénès, Jim Fehrle, Julien Forest, Emilio Jesús Gallego Arias, Gaëtan Gilbert, Georges Gonthier, Benjamin Grégoire, Jason Gross, Hugo Herbelin, Vincent Laporte, Olivier Laurent, Assia Mahboubi, Kenji Maillard, Guillaume Melquiond, Pierre-Marie Pédrot, Clément Pit-Claudel, Pierre Roux, Kazuhiko Sakaguchi, Vincent Semeria, Michael Soegtrop, Arnaud Spiwack, Matthieu Sozeau, Enrico Tassi, Laurent Théry, Anton Trunov, Li-yao Xia and Théo Zimmermann. See the Coq Team face book page for more details.

The 54 contributors to this version are Reynald Affeldt, Arthur Azevedo de Amorim, Yves Bertot, Frédéric Besson, Lasse Blaauwbroek, Ana Borges, Ali Caglayan, Cyril Cohen, Pierre Courtieu, Maxime Dénès, Stéphane Desarzens, Andrej Dudenhefner, Jim Fehrle, Yannick Forster, Simon Friis Vindum, Gaëtan Gilbert, Jason Gross, Samuel Gruetter, Stefan Haan, Hugo Herbelin, Jasper Hugunin, Emilio Jesús Gallego Arias, Jacques-Henri Jourdan, Ralf Jung, Jan-Oliver Kaiser, Fabian Kunze, Vincent Laporte, Olivier Laurent, Yishuai Li, Barry M. Trager, Kenji Maillard, Erik Martin-Dorel, Guillaume Melquiond, Isaac Oscar Gariano, Pierre-Marie Pédrot, Rudy Peterson, Clément Pit-Claudel, Pierre Roux, Takafumi Saikawa, Kazuhiko Sakaguchi, Gabriel Scherer, Vincent Semeria, shenlebantongying, Avi Shinnar, slrnsc, Michael Soegtrop, Matthieu Sozeau, Enrico Tassi, Hendrik Tews, Anton Trunov, Karolin Varner, Li-yao Xia, Beta Ziliani and Théo Zimmermann.

The Coq community at large helped improve the design of this new version via the GitHub issue and pull request system, the Coq development mailing list coqdev@inria.fr, the coq-club@inria.fr mailing list, the Discourse forum and the Coq Zulip chat.

Version 8.14's development spanned 9 months from the release of Coq 8.13.0. Guillaume Melquiond is the release manager of Coq 8.14. This release is the result of 522 merged PRs, closing ~150 issues.

Nantes, September 2021,
Matthieu Sozeau for the Coq development team

Changes in 8.14.0

Kernel

  • Changed: The term representation of pattern-matchings now uses a compact form that provides a few static guarantees such as eta-expansion of branches and return clauses and is usually more efficient. The most visible user change is that for the time being, the destruct tactic and its variants generate dummy cuts (β redexes) in the branches of the generated proof. This can also generate very uncommon backwards incompatibilities, such as a change of occurrence numbering for subterms, or breakage of unification in complex situations involving pattern-matchings whose underlying inductive type declares let-bindings in parameters, arity or constructor types. For ML plugin developers, an in-depth description of the new representation, as well as porting tips, can be found in dev/doc/case-repr.md (#13563, fixes #3166, by Pierre-Marie Pédrot).

  • Changed: Linking of native-code libraries used by native_compute is now delayed until an actual call to the native_compute machinery is performed. This should make Coq more responsive on some systems (#13853, fixes #13849, by Guillaume Melquiond).

  • Removed: The ability to change typing flags inside sections to prevent exploiting a weakness in Print Assumptions (#14395, fixes #14317, by Gaëtan Gilbert).

Specification language, type inference

  • Changed: The hints mode ! matches a term iff the applicative head is not an existential variable. It now also matches projections applied to any term or a match on any term (#14392, by Matthieu Sozeau).

  • Removed: The little used :> type cast, which was only interpreted in Program-mode (#13911, by Jim Fehrle and Théo Zimmermann).

  • Added: Enable canonical fun _ => _ projections, see Canonical Structures for details (#14041, by Jan-Oliver Kaiser and Pierre Roux, reviewed by Cyril Cohen and Enrico Tassi).

  • Added: Canonical Structure declarations now accept dependent function types forall _, _ as keys (#14386, by Jan-Oliver Kaiser and Kazuhiko Sakaguchi).

  • Added: Ability to declare primitive projections as class, for dependent typeclass resolutions (#9711, fixes #12975, by Matthieu Sozeau).

  • Fixed: Multiple printing of same warning about unused variables catching several cases (#14261, fixes #14207, by Hugo Herbelin).

  • Fixed: Constants id and not were unduly set opaque in some parts of the unification algorithm (#14371, fixes #14374, by Hugo Herbelin).

Notations

  • Changed: Flag Printing Notations no longer controls whether strings and numbers are printed raw (#13840, by Enrico Tassi).

  • Changed: The error Argument X was previously inferred to be in scope XXX_scope but is here used in YYY_scope. is now the warning [inconsistent-scopes,syntax] and can be silenced by specifying the scope of the argument (#13965, by Enrico Tassi).

  • Removed: Decimal-only number notations which were deprecated in 8.12 (#13842, by Pierre Roux).

  • Added: Number Notation and String Notation now support parsing and printing of primitive floats, primitive arrays and type constants of primitive types (#13519, fixes #13484 and #13517, by Fabian Kunze, with help of Jason Gross)

  • Added: Flag Printing Raw Literals to control whether strings and numbers are printed raw (#13840, by Enrico Tassi).

  • Added: Let the user specify a scope for abbreviation arguments, e.g. Notation abbr X := t (X in scope my_scope) (#13965, by Enrico Tassi).

  • Added: Look-ahead of tokens is changed from sequential to tree-based, allowing more automatic rule factorizations in notations (#14070, by Hugo Herbelin).

  • Fixed: Non-local custom entries survive module closing and are declared when a file is Required (#14183, fixes #13654, by Gaëtan Gilbert).

  • Fixed: ident modifier in custom entry notations gave fatal errors at printing time (#14257, fixes #14211, by Hugo Herbelin).

  • Fixed: Anomaly when overriding a notation with different applicability in match patterns (#14377, fixes #13966, by Hugo Herbelin).

Tactics

  • Changed: More systematic checks that occurrences of an at clause are valid in tactics such as rewrite or pattern (#13568, fixes #13566, by Hugo Herbelin).

  • Removed: fail and gfail, which formerly accepted negative values as a parameter, now give syntax errors for negative values (#13469, by Jim Fehrle).

  • Removed: Deprecated flag Bracketing Last Introduction Pattern affecting the behavior of trailing disjunctive introduction patterns is definitively removed (#13509, by Hugo Herbelin).

  • Removed: The omega tactic (deprecated in 8.12) and four * Omega * flags. Use lia instead (#13741, by Jim Fehrle, who addressed the final details, building on much work by Frédéric Besson, who greatly improved lia, Maxime Dénès, Vincent Laporte and with the help of many package maintainers, among others).

  • Removed: convert_concl_no_check. Use change_no_check instead (#13761, by Jim Fehrle).

  • Removed: double induction tactic. Replace double induction ident ident with induction ident; induction ident (or induction ident ; destruct ident depending on the exact needs). Replace double induction natural1 natural2 with induction natural1; induction natural3 where natural3 is the result of natural2 - natural1 (#13762, by Jim Fehrle).

  • Deprecated: In change and change_no_check, the at ... with ... form is deprecated. Use with ... at ... instead. For at ... with ... in H |-, use with ... in H at ... |- (#13696, by Jim Fehrle).

  • Deprecated: The micromega option Simplex, which is currently set by default (#13781, by Frédéric Besson).

  • Deprecated: the undocumented new auto tactic (#14528, by Pierre-Marie Pédrot).

  • Added: lia supports the boolean operator Bool.implb (#13715, by Frédéric Besson).

  • Added: zify (lia/nia) support for div, mod, pow for Nat (via ZifyNat module) and N (via ZifyN module). The signature of Z_div_mod_eq_full has no assumptions (#14037, fixes #11447, by Andrej Dudenhefner, Jason Gross, and Frédéric Besson).

  • Added: Ltac2 now has a unify tactic (#14089, fixes #14083, by Samuel Gruetter).

  • Added: inversion_sigma can now be applied to a specified hypothesis and additionally supports intropatterns, so it can be used much like induction and inversion. Additionally, inversion_sigma now supports the types ex (exists x : A, P x) and ex2 (exists2 x : A, P x & Q x) in cases where the first argument A is a Prop (#14174, by Jason Gross).

  • Added: zify (lia/nia) support for Sint63 (#14408, by Ana Borges, with help from Frédéric Besson).

  • Fixed: Possible collision between a user-level name and an internal name when using the % introduction pattern (#13512, fixes #13413, by Hugo Herbelin).

  • Fixed: simpl and hnf now reduce primitive functions on primitive integers, floats and arrays (#13699, fixes #13579, by Pierre Roux).

  • Fixed: Setoid rewriting now remembers the (invisible) binder names of non-dependent product types. SSReflect's rewrite tactic expects these names to be retained when using rewrite foo in H. This also fixes SSR rewrite foo in H * erroneously reverting H (#13882, fixes #12011, by Gaëtan Gilbert).

  • Fixed: Properly expand projection parameters in hint discrimination nets. (#14033, fixes #9000, #14009, by Pierre-Marie Pédrot).

  • Fixed: anomalies caused by empty strings in Ltac notations are now errors (#14378, fixes #14124, by Hugo Herbelin).

  • Fixed: Print a message instead of a Diff_Failure anomaly when old and new goals can't be matched; show the goal without diff highlights (#14457, fixes #14425, by Jim Fehrle).

  • Fixed: Anomaly of destruct on terms with dependent variables unused in goal (#15099, fixes #11504 and #14090, by Lasse Blaauwbroek and Hugo Herbelin).

  • Fixed: Correct convertibility of multiple terms selected by patterns in tactics such as set when these terms have subterms in SProp (#14610, fixes #14609, by Hugo Herbelin).

Tactic language

  • Changed: Renamed Ltac2 Bool.eq into Bool.equal for uniformity. The old function is now a deprecated alias (#14128, by Pierre-Marie Pédrot).

  • Added: A printf macro to Ltac2. It can be made accessible by importing the Ltac2.Printf module. See the documentation there for more information (#13236, fixes #10108, by Pierre-Marie Pédrot).

  • Added: A function Ltac1.lambda allowing to embed Ltac2 functions into Ltac1 runtime values (#13442, fixes #12871, by Pierre-Marie Pédrot).

  • Added: Ltac2 commands defining terms now accept the deprecated attribute (#13774, fixes #12317, by Pierre-Marie Pédrot).

  • Added: Allow the presence of type casts for function return values, let bindings and global definitions in Ltac2 (#13914, by Pierre-Marie Pédrot).

  • Added: The Ltac2 API Ltac2.Ind for manipulating inductive types (#13920, fixes #10095, by Pierre-Marie Pédrot).

  • Added: Allow scope delimiters in Ltac2 open_constr:(...) quotation (#13939, fixes #12806, by Pierre-Marie Pédrot).

  • Added: A FFI to convert between Ltac1 and Ltac2 identifiers (#13997, fixes #13996, by Pierre-Marie Pédrot).

  • Added: Lazy evaluating boolean operators lazy_and, lazy_or, lazy_impl and infix notations && and || to the Ltac2 Bool.v library l (#14081, fixes #13964, by Michael Soegtrop).

  • Fixed: Ltac2 notations now correctly take into account their assigned level (#14094, fixes #11866, by Pierre-Marie Pédrot).

SSReflect

  • Added: A test that the notations {in _, _} and {pred _} from ssrbool.v are displayed correctly (#13473, by Cyril Cohen).

  • Added: Lemmas about interaction between {in _, _}, {on _, _}, and sig have been backported from Mathematical Components 1.12.0 (#13490, by Kazuhiko Sakaguchi).

Commands and options

  • Changed: Hint Rewrite now supports locality attributes (including export) like other Hint commands (#13725, fixes #13724, by Gaëtan Gilbert).

  • Changed: In Record, alpha-rename the variable associated with the record to avoid alpha-renaming parameters of projections (#13852, fixes #13727, by Li-yao Xia).

  • Changed: Improve the Coercion command to reduce the number of ambiguous paths to report. A pair of multiple inheritance paths that can be reduced to smaller adjoining pairs will not be reported as ambiguous paths anymore (#13909, by Kazuhiko Sakaguchi).

  • Changed: The printing order of Print Classes and Print Graph, due to the changes for the internal tables of coercion classes and coercion paths (#13912, by Kazuhiko Sakaguchi).

  • Removed: The Hide Obligations flag, deprecated in 8.12 (#13758, by Jim Fehrle).

  • Removed: SearchHead command. Use the headconcl: clause of Search instead (#13763, by Jim Fehrle).

  • Removed: Show Zify Spec, Add InjTyp and 11 similar Add * commands. For Show Zify Spec, use Show Zify UnOpSpec or Show Zify BinOpSpec instead. For Add *, Use Add Zify * intead of Add * (#13764, by Jim Fehrle).

  • Deprecated: Like hints, typeclass instances added outside of sections without an explicit locality now generate a deprecation warning. See Hint (#14208, fixes #13562, by Pierre-Marie Pédrot).

  • Deprecated: the Regular Subst Tactic flag (#14336, by Pierre-Marie Pédrot).

  • Added: Debug to control debug messages, functioning similarly to the warning system (#13202, by Maxime Dénès and Gaëtan Gilbert). The following flags have been converted (such that Set Flag becomes Set Debug "flag"):

    • Debug Unification to unification

    • Debug HO Unification to ho-unification

    • Debug Tactic Unification to tactic-unification

    • Congruence Verbose to congruence

    • Debug Cbv to cbv

    • Debug RAKAM to RAKAM

    • Debug Ssreflect to ssreflect

  • Added: The Ltac2 grammar can now be printed using the Print Grammar ltac2 command (#14093, fixes #14092, by Pierre-Marie Pédrot).

  • Added: Instance now accepts the export locality attribute (#14148, by Pierre-Marie Pédrot).

  • Fixed: extraction failure of a parameterized type in Prop exported in an module interface as an assumption in Type (#14102, fixes #14100, by Hugo Herbelin).

  • Fixed: Print Assumptions now treats delayed opaque proofs generated by vos compilation as if they were axioms (#14382, fixes #13589, by Pierre-Marie Pédrot).

  • Fixed: Incorrect de Bruijn index handling in vernac class declaration, preventing users from marking existing instances of existing classes which are primitive projections (#14664, fixes #14652, by Ali Caglayan and Hugo Herbelin).

Command-line tools

  • Changed: coqc now enforces that at most a single .v file can be passed in the command line. Support for multiple .v files in the form of coqc f1.v f2.v didn't properly work in 8.13, tho it was accepted (#13876, by Emilio Jesus Gallego Arias).

  • Changed: coqdep now reports an error if files specified on the command line don't exist or if it encounters unreadable files. Unknown options now generate a warning. Previously these conditions were ignored (#14024, fixes #14023, by Hendrik Tews).

  • Changed: Makefiles produced by coq_makefile now use .DELETE_ON_ERROR (#14238, by Gaëtan Gilbert).

  • Removed: Previously deprecated command line options -sprop-cumulative and -input-state and its alias -is (#13822, by Gaëtan Gilbert).

  • Added: coq_makefile-made Makefiles now support inclusion of a .local-late file at the end, allowing the user to access more variables (#12411, fixes #10912, by Jason Gross).

  • Fixed: Failure of extraction in the presence of inductive types with local definitions in parameters (#13624, fixes #13581, by Hugo Herbelin).

  • Fixed: File name was missing in coqdoc error position reporting (#14285, fixes #14283, by Arthur Charguéraud and Hugo Herbelin).

Native Compilation

  • Changed: coq_makefile now uses the coqnative binary to generate native compilation files. Project files also understand directly the -native-compiler flag without having to wrap it with -arg (#14265, by Pierre-Marie Pédrot).

  • Deprecated: the -native-compiler option for coqc. It is now recommended to use the Split compilation of native computation files binary instead to generate native compilation files ahead of time (#14309, by Pierre-Marie Pédrot).

  • Added: A standalone coqnative binary that performs native compilation out of vo files, allowing to split library compilation from native compilation. See Split compilation of native computation files. The hybrid build system was adapted to perform a split compilation on the stdlib (#13287, by Pierre-Marie Pédrot).

CoqIDE

  • Added: Ltac debugger support in CoqIDE (see Ltac Debug). Debugger output and prompts appear in the Messages panel (#13783, by Jim Fehrle and Emilio J. Gallego Arias).

  • Added: Shift-return in the Find dialog now searches backwards (#13810, by slrnsc).

Standard library

  • Changed: Minor Changes to Rpower: Generalizes exp_ineq1 to hold for all non-zero numbers. Adds exp_ineq1_le, which holds for all reals (but is a <= instead of a <) (#13582, by Avi Shinnar and Barry Trager, with help from Laurent Théry).

  • Changed: set n mod 0 = n uniformly for nat, N, Z, int63, sint63, int31 such that m = (m / n) * n + (m mod n) holds (also for n = 0)

    Warning

    code that relies on n mod 0 = 0 will break; for compatibility with both n mod 0 = n and n mod 0 = 0 you can use n mod 0 = ltac:(match eval hnf in (1 mod 0) with |0 => exact 0 |_ => exact n end)

    (#14086, by Andrej Dudenhefner with help of Guillaume Melquiond, Jason Gross, and Kazuhiko Sakaguchi).

  • Changed: The standard library now contains a more complete theory of equality on types of the form exists x : A, P x and exists2 x : A, P x & Q x when we have A : Prop. To bring this theory more in line with the existing theory about sigma types, eq_ex_uncurried, eq_ex2_uncurried, eq_ex, eq_ex2, eq_ex_hprop, eq_ex2_hprop have been renamed into eq_ex_intro_uncurried, eq_ex_intro2_uncurried, eq_ex_intro, eq_ex_intro2, eq_ex_intro_hprop, eq_ex_intro2_hprop respectively and the implicit status of these lemmas has changed slightly (#14174, by Jason Gross).

  • Changed Moved 39 lemmas and notations about the rationals Q from the constructive reals private file theories/Reals/Cauchy/QExtra.v to appropriate files in theories/QArith. The now public lemmas are mostly about compatibility of multiplication and power with relational operators and simple convenience lemmas e.g. for reduction of Q values. The following moved lemmas have been renamed: Q_factorDenom to Qmult_frac_l, Q_reduce_fl to Qreduce_num_l, Qle_neq to Qlt_leneq, Qmult_lt_le_compat_nonneg to Qmult_le_lt_compat_pos, Qpower_pos_lt to Qpower_0_lt, Qpower_lt_1_increasing to Qpower_1_lt_pos, Qpower_lt_1_increasing' to Qpower_1_lt, Qpower_le_1_increasing to Qpower_1_le_pos, Qpower_le_1_increasing' to Qpower_1_le, Qzero_eq to Qreduce_zero, Qpower_lt_compat to Qpower_lt_compat_l, Qpower_le_compat to Qpower_le_compat_l, Qpower_lt_compat_inv to Qpower_lt_compat_l_inv, Qpower_le_compat_inv to Qpower_le_compat_l_inv, Qpower_decomp' to Qpower_decomp_pos and QarchimedeanExp2_Pos to Qarchimedean_power2_pos. The following lemmas have been renamed and the sides of the equality swapped: Qinv_swap_pos to Qinv_pos, Qinv_swap_neg to Qinv_neg and. The following lemmas have been deleted: Q_factorNum_l and Q_factorNum. The lemma Qopp_lt_compat has been moved from theories/QArith/Qround.v to theories/QArith/QArith_base.v. About 10 additional lemmas have been added for similar cases as the moved lemmas. Compatibility notations are not provided because QExtra is considered internal (excluded from the library documentation) (#14293, by Michael Soegtrop).

  • Changed: Importing ZArith no longer has the side-effect of closing Z_scope (#14343, fixes #13307, by Ralf Jung).

  • Removed: IF_then_else definition and corresponding IF P then Q else R notation (#13871, by Yishuai Li).

  • Removed: from List.v deprecated/unexpected dependencies Setoid, Le, Gt, Minus, Lt (#13986, by Andrej Dudenhefner).

  • Deprecated: Unsigned primitive integers are now named uint63 instead of int63. The Int63 module is replaced by Uint63. The full list of changes is described in the PR (#13895, by Ana Borges).

  • Added: leb and ltb functions for ascii (#13080, by Yishuai Li).

  • Added: Library for signed primitive integers, Sint63. The following operations were added to the kernel: division, remainder, comparison functions, and arithmetic shift right. Everything else works the same for signed and unsigned ints (#13559, fixes #12109, by Ana Borges, Guillaume Melquiond and Pierre Roux).

  • Added: Lemmas about vectors related with to_list: length_to_list, of_list_to_list_opp, to_list_nil, to_list_cons, to_list_hd, to_list_last, to_list_const, to_list_nth_order, to_list_tl, to_list_append, to_list_rev_append_tail, to_list_rev_append, to_list_rev, to_list_map, to_list_fold_left, to_list_fold_right, to_list_Forall, to_list_Exists, to_list_In, to_list_Forall2 (#13671, by Olivier Laurent).

  • Added: Lemmas about count_occ: count_occ_app, count_occ_elt_eq, count_occ_elt_neq, count_occ_bound, count_occ_repeat_eq, count_occ_repeat_neq, count_occ_unique, count_occ_repeat_excl, count_occ_sgt, Permutation_count_occ (#13804, by Olivier Laurent with help of Jean-Christophe Léchenet).

  • Added: Lemmas to List: Exists_map, Exists_concat, Exists_flat_map, Forall_map, Forall_concat, Forall_flat_map, nth_error_map, nth_repeat, nth_error_repeat (#13955, by Andrej Dudenhefner, with help from Olivier Laurent).

  • Added: Cantor.v containing the Cantor pairing function and its inverse. Cantor.to_nat : nat * nat -> nat and Cantor.of_nat : nat -> nat * nat are the respective bijections between nat * nat and nat (#14008, by Andrej Dudenhefner).

  • Added: Lemmas to Q: Qeq_from_parts, Qden_cancel, Qnum_cancel, Qreduce_l, Qreduce_r, Qmult_inject_Z_l, Qmult_inject_Z_r QArith_base Reduction of rationals; establishing equality for Qden/Qnum separately (#14087, by Karolin Varner).

  • Added: Coq.Structures.OrdersEx.String_as_OT and Coq.Structures.OrdersEx.Ascii_as_OT to make strings and ascii ordered types (using lexical order). (#14096, by Jason Gross).

  • Added: Lemmas app_eq_app, Forall_nil_iff, Forall_cons_iff to List.v (#14153, closes #1803, by Andrej Dudenhefner, with help from Olivier Laurent).

  • Added: Z, positive and N constants can now be printed in hexadecimal by opening hex_Z_scope, hex_positive_scope, and hex_N_scope respectively (#14263, by Jason Gross).

  • Added: Absolute value function for Sint63 (#14384, by Ana Borges).

  • Added: Lemmas showing firstn and skipn commute with map (#14406, by Rudy Peterson).

  • Fixed: Haskell extraction is now compatible with GHC versions >= 9.0. Some #if statements have been added to extract unsafeCoerce to its new location in newer versions of GHC. (#14345, fixes #14256, by Jason Gross).

Infrastructure and dependencies

  • Changed: Coq's configure script now requires absolute paths for the -prefix option (#12567, by Emilio Jesus Gallego Arias).

  • Changed: The regular Coq package has been split in two: coq-core, with OCaml-based libraries and tools; and coq-stdlib, which contains the Gallina-based standard library. The package Coq now depends on both for compatiblity (#12567, by Emilio Jesus Gallego Arias, review by Vincent Laporte, Guillaume Melquiond, Enrico Tassi, and Théo Zimmerman).

  • Changed: Coq's OCaml parts and tools [coq-core] are now built using Dune. The main user-facing change is that Dune >= 2.5 is now required to build Coq. This was a large and complex change. If you are packager you may find some minor differences if you were using a lot of custom optimizations. Note that, in particular, the configure option -datadir is not customizable anymore, and -bindir has been removed in favor of $prefix/bin. Moreover, the install procedure will ignore -docdir and -etcdir, unless you patch the makefile and use Dune >= 2.9. We usually recommended using a recent Dune version, if possible. For developers and plugin authors, see the entry in dev/doc/changes.md. For packagers and users, see dev/doc/INSTALL.make.md (#13617, by Emilio Jesús Gallego Arias, Rudi Grinberg, and Théo Zimmerman; review and testing by Gaëtan Gilbert, Guillaume Melquiond, and Enrico Tassi)

  • Changed: Undocumented variables OLDROOT and COQPREFIXINSTALL which added a prefix path to make install have been removed. Now, make install does support the more standard DESTDIR variable, akin to what coq_makefile does (#14258, by Emilio Jesus Gallego Arias).

  • Added: Support OCaml 4.12 (#13885, by Emilio Jesus Gallego Arias, review by Gaëtan Gilbert and Théo Zimmerman).

Miscellaneous

  • Changed: The representation of micromega caches was slightly altered for efficiency purposes. As a consequence all stale caches must be cleaned up (#13405, by Pierre-Marie Pédrot).

  • Fixed: Fix the timeout facility on Unix to allow for nested timeouts. Previous behavior on nested timeouts was that an "inner" timeout would replace an "outer" timeout, so that the outer timeout would no longer fire. With the new behavior, Unix and Windows implementations should be (approximately) equivalent (#13586, by Lasse Blaauwbroek).

Changes in 8.14.1

Kernel

  • Fixed: Fix the implementation of persistent arrays used by the VM and native compute so that it uses a uniform representation. Previously, storing primitive floats inside primitive arrays could cause memory corruption (#15081, closes #15070, by Pierre-Marie Pédrot).

Specification language, type inference

  • Fixed: Missing registration of universe constraints in Module Type elaboration (#14666, fixes #14505, by Hugo Herbelin).

Tactics

  • Fixed: abstract more robust with respect to Ltac constr bindings containing existential variables (#14671, fixes #10796, by Hugo Herbelin).

  • Fixed: correct support of trailing let by tactic specialize (#15046, fixes #15043, by Hugo Herbelin).

Commands and options

Version 8.13

Summary of changes

Coq version 8.13 integrates many usability improvements, as well as extensions of the core language. The main changes include:

  • Introduction of primitive persistent arrays in the core language, implemented using imperative persistent arrays.

  • Introduction of definitional proof irrelevance for the equality type defined in the SProp sort.

  • Cumulative record and inductive type declarations can now specify the variance of their universes.

  • Various bugfixes and uniformization of behavior with respect to the use of implicit arguments and the handling of existential variables in declarations, unification and tactics.

  • New warning for unused variables in catch-all match branches that match multiple distinct patterns.

  • New warning for Hint commands outside sections without a locality attribute, whose goal is to eventually remove the fragile default behavior of importing hints only when using Require. The recommended fix is to declare hints as export, instead of the current default global, meaning that they are imported through Require Import only, not Require. See the following rationale and guidelines for details.

  • General support for boolean attributes.

  • Many improvements to the handling of notations, including number notations, recursive notations and notations with bindings. A new algorithm chooses the most precise notation available to print an expression, which might introduce changes in printing behavior.

  • Tactic improvements in lia and its zify preprocessing step, now supporting reasoning on boolean operators such as Z.leb and supporting primitive integers Int63.

  • Typing flags can now be specified per-constant / inductive.

  • Improvements to the reference manual including updated syntax descriptions that match Coq's grammar in several chapters, and splitting parts of the tactics chapter to independent sections.

See the Changes in 8.13+beta1 section and following sections for the detailed list of changes, including potentially breaking changes marked with Changed.

Coq's documentation is available at https://coq.github.io/doc/v8.13/refman (reference manual), and https://coq.github.io/doc/v8.13/stdlib (documentation of the standard library). Developer documentation of the ML API is available at https://coq.github.io/doc/v8.13/api.

Maxime Dénès, Emilio Jesús Gallego Arias, Gaëtan Gilbert, Michael Soegtrop and Théo Zimmermann worked on maintaining and improving the continuous integration system and package building infrastructure.

Erik Martin-Dorel has maintained the Coq Docker images that are used in many Coq projects for continuous integration.

The opam repository for Coq packages has been maintained by Guillaume Claret, Karl Palmskog, Matthieu Sozeau and Enrico Tassi with contributions from many users. A list of packages is available at https://coq.inria.fr/opam/www/.

Our current 32 maintainers are Yves Bertot, Frédéric Besson, Tej Chajed, Cyril Cohen, Pierre Corbineau, Pierre Courtieu, Maxime Dénès, Jim Fehrle, Julien Forest, Emilio Jesús Gallego Arias, Gaëtan Gilbert, Georges Gonthier, Benjamin Grégoire, Jason Gross, Hugo Herbelin, Vincent Laporte, Olivier Laurent, Assia Mahboubi, Kenji Maillard, Guillaume Melquiond, Pierre-Marie Pédrot, Clément Pit-Claudel, Kazuhiko Sakaguchi, Vincent Semeria, Michael Soegtrop, Arnaud Spiwack, Matthieu Sozeau, Enrico Tassi, Laurent Théry, Anton Trunov, Li-yao Xia and Théo Zimmermann.

The 51 contributors to this version are Reynald Affeldt, Tanaka Akira, Frédéric Besson, Lasse Blaauwbroek, Clément Blaudeau, Martin Bodin, Ali Caglayan, Tej Chajed, Cyril Cohen, Julien Coolen, Matthew Dempsky, Maxime Dénès, Andres Erbsen, Jim Fehrle, Emilio Jesús Gallego Arias, Attila Gáspár, Paolo G. Giarrusso, Gaëtan Gilbert, Jason Gross, Benjamin Grégoire, Hugo Herbelin, Wolf Honore, Jasper Hugunin, Ignat Insarov, Ralf Jung, Fabian Kunze, Vincent Laporte, Olivier Laurent, Larry D. Lee Jr, Thomas Letan, Yishuai Li, James Lottes, Jean-Christophe Léchenet, Kenji Maillard, Erik Martin-Dorel, Yusuke Matsushita, Guillaume Melquiond, Carl Patenaude-Poulin, Clément Pit-Claudel, Pierre-Marie Pédrot, Pierre Roux, Kazuhiko Sakaguchi, Vincent Semeria, Michael Soegtrop, Matthieu Sozeau, Enrico Tassi, Anton Trunov, Edward Wang, Li-yao Xia, Beta Ziliani and Théo Zimmermann.

The Coq community at large helped improve the design of this new version via the GitHub issue and pull request system, the Coq development mailing list coqdev@inria.fr, the coq-club@inria.fr mailing list, the Discourse forum and the Coq Zulip chat.

Version 8.13's development spanned 5 months from the release of Coq 8.12.0. Enrico Tassi and Maxime Dénès are the release managers of Coq 8.13. This release is the result of 400 merged PRs, closing ~100 issues.

Nantes, November 2020,
Matthieu Sozeau for the Coq development team

Changes in 8.13+beta1

Kernel

  • Added: Definitional UIP, only when Definitional UIP is enabled. This models definitional uniqueness of identity proofs for the equality type in SProp. It is deactivated by default as it can lead to non-termination in combination with impredicativity. Use of this flag is also printed by Print Assumptions. See documentation of the flag for details (#10390, by Gaëtan Gilbert).

  • Added: Built-in support for persistent arrays, which expose a functional interface but are implemented using an imperative data structure, for better performance (#11604, by Maxime Dénès and Benjamin Grégoire, with help from Gaëtan Gilbert).

    Primitive arrays are irrelevant in their single polymorphic universe (same as a polymorphic cumulative list inductive would be) (#13356, fixes #13354, by Gaëtan Gilbert).

  • Fixed: A loss of definitional equality for declarations obtained through Include when entering the scope of a Module or Module Type was causing Search not to see the included declarations (#12537, fixes #12525 and #12647, by Hugo Herbelin).

  • Fixed: Fix an incompleteness in the typechecking of match for cumulative inductive types. This could result in breaking subject reduction (#13501, fixes #13495, by Matthieu Sozeau).

Specification language, type inference

  • Changed: Boolean attributes are now specified using key/value pairs, that is to say identattr= yesno?. If the value is missing, the default is yes. The old syntax is still supported, but produces the deprecated-attribute-syntax warning.

    Deprecated attributes are universes(monomorphic), universes(notemplate) and universes(noncumulative), which are respectively replaced by universes(polymorphic=no), universes(template=no) and universes(cumulative=no). Attributes program and canonical are also affected, with the syntax identattr(false) being deprecated in favor of identattr=no (#13312, by Emilio Jesus Gallego Arias).

  • Changed: Heuristics for universe minimization to Set: also use constraints Prop <= i (#10331, by Gaëtan Gilbert with help from Maxime Dénès and Matthieu Sozeau, fixes #12414).

  • Changed: The type given to Instance is no longer automatically generalized over unbound and generalizable variables. Use Instance : `{type} instead of Instance : type to get the old behavior, or enable the compatibility flag Instance Generalized Output (#13188, fixes #6042, by Gaëtan Gilbert).

  • Changed: Tweaked the algorithm giving default names to arguments. Should reduce the frequency that argument names get an unexpected suffix. Also makes Mangle Names not mess up argument names (#12756, fixes #12001 and #6785, by Jasper Hugunin).

  • Removed: Undocumented and experimental forward class hint feature :>>. Use :> (see of_type) instead (#13106, by Pierre-Marie Pédrot).

  • Added: Commands Inductive, Record and synonyms now support syntax Inductive foo@{=i +j *k l} to specify variance information for their universes (in Cumulative mode) (#12653, by Gaëtan Gilbert).

  • Added: Warning on unused variables in pattern-matching branches of match serving as catch-all branches for at least two distinct patterns (#12768, fixes #12762, by Hugo Herbelin).

  • Added: Definition and (Co)Fixpoint now support the using attribute. It has the same effect as Proof using, which is only available in interactive mode (#13183, by Enrico Tassi).

  • Added: Typing flags can now be specified per-constant / inductive, this allows to fine-grain specify them from plugins or attributes. See Controlling Typing Flags for details on attribute syntax (#12586, by Emilio Jesus Gallego Arias).

  • Added: Inference of return predicate of a match by inversion takes sort elimination constraints into account (#13290, grants #13278, by Hugo Herbelin).

  • Fixed: Implicit arguments taken into account in defined fields of a record type declaration (#13166, fixes #13165, by Hugo Herbelin).

  • Fixed: Allow use of typeclass inference for the return predicate of a match (was deactivated in versions 8.10 to 8.12, #13217, fixes #13216, by Hugo Herbelin).

  • Fixed: A case of unification raising an anomaly IllTypedInstance (#13376, fixes #13266, by Hugo Herbelin).

  • Fixed: Using {wf ...} in local fixpoints is an error, not an anomaly (#13383, fixes #11816, by Hugo Herbelin).

  • Fixed: Issue when two expressions involving different projections and one is primitive need to be unified (#13386, fixes #9971, by Hugo Herbelin).

  • Fixed: A bug producing ill-typed instances of existential variables when let-ins interleaved with assumptions (#13387, fixes #12348, by Hugo Herbelin).

Notations

  • Changed: In notations (except in custom entries), the misleading syntax_modifier ident ident (which accepted either an identifier or a _) is deprecated and should be replaced by ident name. If the intent was really to only parse identifiers, this will eventually become possible, but only as of Coq 8.15. In custom entries, the meaning of ident ident is silently changed from parsing identifiers or _ to parsing only identifiers without warning, but this presumably affects only rare, recent and relatively experimental code (#11841, fixes #9514, by Hugo Herbelin).

  • Changed: Improved support for notations/abbreviations with mixed terms and patterns (such as the forcing modality) (#12099, by Hugo Herbelin).

  • Changed Rational and real constants are parsed differently. The exponent is now encoded separately from the fractional part using Z.pow_pos. This way, parsing large exponents can no longer blow up and constants are printed in a form closer to the one in which they were parsed (i.e., 102e-2 is reprinted as such and not 1.02) (#12218, by Pierre Roux).

  • Changed: Scope information is propagated in indirect applications to a reference prefixed with @; this covers for instance the case r.(@p) t where scope information from p is now taken into account for interpreting t (#12685, by Hugo Herbelin).

  • Changed: New model for only parsing and only printing notations with support for at most one parsing-and-printing or only-parsing notation per notation and scope, but an arbitrary number of only-printing notations (#12950, fixes #4738 and #9682 and part 2 of #12908, by Hugo Herbelin).

  • Changed: Redeclaring a notation also reactivates its printing rule; in particular a second Import of the same module reactivates the printing rules declared in this module. In theory, this leads to changes in behavior for printing. However, this is mitigated in general by the adoption in #12986 of a priority given to notations which match a larger part of the term to print (#12984, fixes #7443 and #10824, by Hugo Herbelin).

  • Changed: Use of notations for printing now gives preference to notations which match a larger part of the term to abbreviate (#12986, by Hugo Herbelin).

  • Removed OCaml parser and printer for real constants have been removed. Real constants are now handled with proven Coq code (#12218, by Pierre Roux).

  • Deprecated Numeral.v is deprecated, please use Number.v instead (#12218, by Pierre Roux).

  • Deprecated: Numeral Notation, please use Number Notation instead (#12979, by Pierre Roux).

  • Added: Printing Float flag to print primitive floats as hexadecimal instead of decimal values. This is included in the Printing All flag (#11986, by Pierre Roux).

  • Added: Number Notation and String Notation commands now support parameterized inductive and non-inductive types (#12218, fixes #12035, by Pierre Roux, review by Jason Gross and Jim Fehrle for the reference manual).

  • Added: Added support for encoding notations of the form x y .. z t. This feature is considered experimental (#12765, by Hugo Herbelin).

  • Added: The binder entry of Notation can now be used in notations expecting a single (non-recursive) binder (#13265, by Hugo Herbelin, see section Notations and binders of the reference manual).

  • Fixed: Issues in the presence of notations recursively referring to another applicative notations, such as missing scope propagation, or failure to use a notation for printing (#12960, fixes #9403 and #10803, by Hugo Herbelin).

  • Fixed: Capture the names of global references by binders in the presence of notations for binders (#12965, fixes #9569, by Hugo Herbelin).

  • Fixed: Preventing notations for constructors to involve binders (#13092, fixes #13078, by Hugo Herbelin).

  • Fixed: Notations understand universe names without getting confused by different imported modules between declaration and use locations (#13415, fixes #13303, by Gaëtan Gilbert).

Tactics

  • Changed: In refine, new existential variables unified with existing ones are no longer considered as fresh. The behavior of simple refine no longer depends on the orientation of evar-evar unification problems, and new existential variables are always turned into (unshelved) goals. This can break compatibility in some cases (#7825, by Matthieu Sozeau, with help from Maxime Dénès, review by Pierre-Marie Pédrot and Enrico Tassi, fixes #4095 and #4413).

  • Changed: Giving an empty list of occurrences after in in tactics is no longer permitted. Omitting the in gives the same behavior (#13237, fixes #13235, by Hugo Herbelin).

  • Removed: at occs_nums clauses in tactics such as unfold no longer allow negative values. A "-" before the list (for set complement) is still supported. Ex: "at -1 -2" is no longer supported but "at -1 2" is (#13403, by Jim Fehrle).

  • Removed: A number of tactics that formerly accepted negative numbers as parameters now give syntax errors for negative values. These include {e}constructor, do, timeout, 9 {e}auto tactics and psatz* (#13417, by Jim Fehrle).

  • Removed: The deprecated and undocumented prolog tactic was removed (#12399, by Pierre-Marie Pédrot).

  • Removed: info tactic that was deprecated in 8.5 (#12423, by Jim Fehrle).

  • Deprecated: Undocumented eauto nat_or_var nat_or_var syntax in favor of new bfs eauto. Also deprecated 2-integer syntax for debug eauto and info_eauto (Use bfs eauto with the Info Eauto or Debug Eauto flags instead.) (#13381, by Jim Fehrle).

  • Added: lia is extended to deal with boolean operators e.g. andb or Z.leb (as lia gets more powerful, this may break proof scripts relying on lia failure, #11906, by Frédéric Besson).

  • Added: apply in supports several hypotheses (#12246, by Hugo Herbelin; grants #9816).

  • Added: The zify tactic can now be extended by redefining the zify_pre_hook tactic. (#12552, by Kazuhiko Sakaguchi).

  • Added: The zify tactic provides support for primitive integers (module ZifyInt63) (#12648, by Frédéric Besson).

  • Fixed: Avoid exposing an internal name of the form _tmp when applying the _ introduction pattern which would break a dependency (#13337, fixes #13336, by Hugo Herbelin).

  • Fixed: The case of tactics, such as eapply, producing existential variables under binders with an ill-formed instance (#13373, fixes #13363, by Hugo Herbelin).

Tactic language

  • Added: An if-then-else syntax to Ltac2 (#13232, fixes #10110, by Pierre-Marie Pédrot).

  • Fixed: Printing of the quotation qualifiers when printing Ltac functions (#13028, fixes #9716 and #13004, by Hugo Herbelin).

SSReflect

  • Added: SSReflect intro pattern ltac views /[dup], /[swap] and /[apply] (#13317, by Cyril Cohen).

  • Fixed: Working around a bug of interaction between + and /(ltac:(...)) cf #13458 (#13459, by Cyril Cohen).

Commands and options

  • Changed: Drop prefixes from grammar non-terminal names, e.g. "constr:global" -> "global", "Prim.name" -> "name". Visible in the output of Print Grammar and Print Custom Grammar (#13096, by Jim Fehrle).

  • Changed: When declaring arbitrary terms as hints, unsolved evars are not abstracted implicitly anymore and instead raise an error (#13139, by Pierre-Marie Pédrot).

  • Removed: In the Extraction Language command, remove Ocaml as a valid value. Use OCaml instead. This was deprecated in Coq 8.8, #6261 (#13016, by Jim Fehrle).

  • Deprecated: Hint locality currently defaults to local in a section and global otherwise, but this will change in a future release. Hints added outside of sections without an explicit locality now generate a deprecation warning. We recommend using export where possible (#13384, by Pierre-Marie Pédrot).

  • Deprecated: Grab Existential Variables and Existential commands (#12516, by Maxime Dénès).

  • Added: The export locality can now be used for all Hint commands, including Hint Cut, Hint Mode, Hint Transparent / Opaque and Remove Hints (#13388, by Pierre-Marie Pédrot).

  • Added: Support for automatic insertion of coercions in Search patterns. Additionally, head patterns are now automatically interpreted as types (#13255, fixes #13244, by Hugo Herbelin).

  • Added: The Proof using command can now be used without loading the Ltac plugin (-noinit mode) (#13339, by Théo Zimmermann).

  • Added: Clarify in the documentation that Add ML Path is not exported to compiled files (#13345, fixes #13344, by Hugo Herbelin).

Tools

  • Changed: Option -native-compiler of the configure script now impacts the default value of the -native-compiler option of coqc. The -native-compiler option of the configure script supports a new ondemand value, which becomes the default, thus preserving the previous default behavior. The stdlib is still precompiled when configuring with -native-compiler yes. It is not precompiled otherwise. This an implementation of point 2 of CEP #48 (#13352, by Pierre Roux).

  • Changed: Added the ability for coq_makefile to directly set the installation folders, through the COQLIBINSTALL and COQDOCINSTALL variables. See CoqMakefile.local (#12389, by Martin Bodin, review of Enrico Tassi).

  • Removed: The option -I of coqchk was removed (it was deprecated in Coq 8.8) (#12613, by Gaëtan Gilbert).

  • Fixed: coqchk no longer reports names from inner modules of opaque modules as axioms (#12862, fixes #12845, by Jason Gross).

CoqIDE

Standard library

  • Changed: In the reals theory changed the epsilon in the definition of the modulus of convergence for CReal from 1/n (n in positive) to 2^z (z in Z) so that a precision coarser than one is possible. Also added an upper bound to CReal to enable more efficient computations (#12186, by Michael Soegtrop).

  • Changed: Int63 notations now match up with the rest of the standard library: a \% m, m == n, m < n, m <= n, and m n have been replaced with a mod m, m =? n, m <? n, m <=? n, and m ≤? n. The old notations are still available as deprecated notations. Additionally, there is now a Coq.Numbers.Cyclic.Int63.Int63.Int63Notations module that users can import to get the Int63 notations without unqualifying the various primitives (#12479, fixes #12454, by Jason Gross).

  • Changed: PrimFloat notations now match up with the rest of the standard library: m == n, m < n, and m <= n have been replaced with m =? n, m <? n, and m <=? n. The old notations are still available as deprecated notations. Additionally, there is now a Coq.Floats.PrimFloat.PrimFloatNotations module that users can import to get the PrimFloat notations without unqualifying the various primitives (#12556, fixes #12454, by Jason Gross).

  • Changed: the sort of cyclic numbers from Type to Set. For backward compatibility, a dynamic sort was defined in the 3 packages bignums, coqprime and color. See for example commit 6f62bda in bignums (#12801, by Vincent Semeria).

  • Changed: Require Import Coq.nsatz.NsatzTactic now allows using nsatz with Z and Q without having to supply instances or using Require Import Coq.nsatz.Nsatz, which transitively requires unneeded files declaring axioms used in the reals (#12861, fixes #12860, by Jason Gross).

  • Deprecated: prod_curry and prod_uncurry, in favor of uncurry and curry (#12716, by Yishuai Li).

  • Added: New lemmas about repeat in List and Permutation: repeat_app, repeat_eq_app, repeat_eq_cons, repeat_eq_elt, Forall_eq_repeat, Permutation_repeat (#12799, by Olivier Laurent).

  • Added: Extend some list lemmas to both directions: app_inj_tail_iff, app_inv_head_iff, app_inv_tail_iff (#12094, fixes #12093, by Edward Wang).

  • Added: Decidable instance for negation (#12420, by Yishuai Li).

  • Fixed: Coq.Program.Wf.Fix_F_inv and Coq.Program.Wf.Fix_eq are now axiom-free, and no longer assuming proof irrelevance (#13365, by Li-yao Xia).

Infrastructure and dependencies

  • Changed: When compiled with OCaml >= 4.10.0, Coq will use the new best-fit GC policy, which should provide some performance benefits. Coq's policy is optimized for speed, but could increase memory consumption in some cases. You are welcome to tune it using the OCAMLRUNPARAM variable and report back on good settings so we can improve the defaults (#13040, fixes #11277, by Emilio Jesus Gallego Arias).

  • Changed: Coq now uses the zarith library, based on GNU's gmp instead of num which is deprecated upstream. The custom bigint module is no longer provided (#11742, #13007, by Emilio Jesus Gallego Arias and Vicent Laporte, with help from Frédéric Besson).

Changes in 8.13.0

Commands and options

  • Changed: The warning custom-entry-overriden has been renamed to custom-entry-overridden (with two d's) (#13556, by Simon Friis Vindum).

Changes in 8.13.1

Kernel

  • Fixed: Fix arities of VM opcodes for some floating-point operations that could cause memory corruption (#13867, by Guillaume Melquiond).

CoqIDE

  • Added: Option -v and --version to CoqIDE (#13870, by Guillaume Melquiond).

Changes in 8.13.2

Kernel

  • Fixed: Crash when using vm_compute on an irreducible PArray.set (#14005, fixes #13998, by Guillaume Melquiond).

  • Fixed: Never store persistent arrays as VM / native structured values. This could be used to make vo marshalling crash, and probably breaking some other invariants of the kernel (#14007, fixes #14006, by Pierre-Marie Pédrot).

Tactic language

  • Fixed: Ltac2 Array.init no longer incurs exponential overhead when used recursively (#14012, fixes #14011, by Jason Gross).

Version 8.12

Summary of changes

Coq version 8.12 integrates many usability improvements, in particular with respect to notations, scopes and implicit arguments, along with many bug fixes and major improvements to the reference manual. The main changes include:

  • New binder notation for non-maximal implicit arguments using [ ] allowing to set and see the implicit status of arguments immediately.

  • New notation Inductive I A | x : s := ... to distinguish the uniform from the non-uniform parameters in inductive definitions.

  • More robust and expressive treatment of implicit inductive parameters in inductive declarations.

  • Improvements in the treatment of implicit arguments and partially applied constants in notations, parsing of hexadecimal number notation and better handling of scopes and coercions for printing.

  • A correct and efficient coercion coherence checking algorithm, avoiding spurious or duplicate warnings.

  • An improved Search command which accepts complex queries. Note that this takes precedence over the now deprecated ssreflect search.

  • Many additions and improvements of the standard library.

  • Improvements to the reference manual include a more logical organization of chapters along with updated syntax descriptions that match Coq's grammar in most but not all chapters.

Additionally, the omega tactic is deprecated in this version of Coq, and we recommend users to switch to lia in new proof scripts.

See the Changes in 8.12+beta1 section and following sections for the detailed list of changes, including potentially breaking changes marked with Changed.

Coq's documentation is available at https://coq.github.io/doc/v8.12/refman (reference manual), and https://coq.github.io/doc/v8.12/stdlib (documentation of the standard library). Developer documentation of the ML API is available at https://coq.github.io/doc/v8.12/api.

Maxime Dénès, Emilio Jesús Gallego Arias, Gaëtan Gilbert, Michael Soegtrop and Théo Zimmermann worked on maintaining and improving the continuous integration system and package building infrastructure.

Erik Martin-Dorel has maintained the Coq Docker images that are used in many Coq projects for continuous integration.

The opam repository for Coq packages has been maintained by Guillaume Claret, Karl Palmskog, Matthieu Sozeau and Enrico Tassi with contributions from many users. A list of packages is available at https://coq.inria.fr/opam/www/.

Previously, most components of Coq had a single principal maintainer. This was changed in 8.12 (#11295) so that every component now has a team of maintainers, who are in charge of reviewing and merging incoming pull requests. This gave us a chance to significantly expand the pool of maintainters and provide faster feedback to contributors. Special thanks to all our maintainers!

Our current 31 maintainers are Yves Bertot, Frédéric Besson, Tej Chajed, Cyril Cohen, Pierre Corbineau, Pierre Courtieu, Maxime Dénès, Jim Fehrle, Julien Forest, Emilio Jesús Gallego Arias, Gaëtan Gilbert, Georges Gonthier, Benjamin Grégoire, Jason Gross, Hugo Herbelin, Vincent Laporte, Assia Mahboubi, Kenji Maillard, Guillaume Melquiond, Pierre-Marie Pédrot, Clément Pit-Claudel, Kazuhiko Sakaguchi, Vincent Semeria, Michael Soegtrop, Arnaud Spiwack, Matthieu Sozeau, Enrico Tassi, Laurent Théry, Anton Trunov, Li-yao Xia, Théo Zimmermann

The 59 contributors to this version are Abhishek Anand, Yves Bertot, Frédéric Besson, Lasse Blaauwbroek, Simon Boulier, Quentin Carbonneaux, Tej Chajed, Arthur Charguéraud, Cyril Cohen, Pierre Courtieu, Matthew Dempsky, Maxime Dénès, Andres Erbsen, Erika (@rrika), Nikita Eshkeev, Jim Fehrle, @formalize, Emilio Jesús Gallego Arias, Paolo G. Giarrusso, Gaëtan Gilbert, Jason Gross, Samuel Gruetter, Attila Gáspár, Hugo Herbelin, Jan-Oliver Kaiser, Robbert Krebbers, Vincent Laporte, Olivier Laurent, Xavier Leroy, Thomas Letan, Yishuai Li, Kenji Maillard, Erik Martin-Dorel, Guillaume Melquiond, Ike Mulder, Guillaume Munch-Maccagnoni, Antonio Nikishaev, Karl Palmskog, Pierre-Marie Pédrot, Clément Pit-Claudel, Ramkumar Ramachandra, Lars Rasmusson, Daniel de Rauglaudre, Talia Ringer, Pierre Roux, Kazuhiko Sakaguchi, Vincent Semeria, @scinart, Kartik Singhal, Michael Soegtrop, Matthieu Sozeau, Enrico Tassi, Laurent Théry, Ralf Treinen, Anton Trunov, Bernhard M. Wiedemann, Li-yao Xia, Nickolai Zeldovich and Théo Zimmermann.

Many power users helped to improve the design of this new version via the GitHub issue and pull request system, the Coq development mailing list coqdev@inria.fr, the coq-club@inria.fr mailing list, the Discourse forum and the new Coq Zulip chat (thanks to Cyril Cohen for organizing the move from Gitter).

Version 8.12's development spanned 6 months from the release of Coq 8.11.0. Emilio Jesus Gallego Arias and Théo Zimmermann are the release managers of Coq 8.12. This release is the result of ~500 PRs merged, closing ~100 issues.

Nantes, June 2020,
Matthieu Sozeau for the Coq development team

Changes in 8.12+beta1

Kernel

  • Fixed: Specification of PrimFloat.leb which made (x <= y)%float true for any non-NaN x and y (#12484, fixes #12483, by Pierre Roux).

Specification language, type inference

  • Changed: The deprecation warning raised since Coq 8.10 when a trailing implicit is declared to be non-maximally inserted (with the command Arguments) has been turned into an error (#11368, by SimonBoulier).

  • Changed: Typeclass resolution, accessible through typeclasses eauto, now suspends constraints according to their modes instead of failing. If a typeclass constraint does not match any of the declared modes for its class, the constraint is postponed, and the proof search continues on other goals. Proof search does a fixed point computation to try to solve them at a later stage of resolution. It does not fail if there remain only stuck constraints at the end of resolution. This makes typeclasses with declared modes more robust with respect to the order of resolution (#10858, fixes #9058, by Matthieu Sozeau).

  • Added: Warn when manual implicit arguments are used in unexpected positions of a term (e.g. in Check id (forall {x}, x)) or when an implicit argument name is shadowed (e.g. in Check fun f : forall {x:nat} {x}, nat => f) (#10202, by Hugo Herbelin).

  • Added: Arguments now supports setting implicit an anonymous argument, as e.g. in Arguments id {A} {_} (#11098, by Hugo Herbelin, fixes #4696, #5173, #9098).

  • Added: Syntax for non-maximal implicit arguments in definitions and terms using square brackets. The syntax is [x : A], [x], `[A] to be consistent with the command Arguments (#11235, by Simon Boulier).

  • Added: Implicit Types are now taken into account for printing. To inhibit it, unset the Printing Use Implicit Types flag (#11261, by Hugo Herbelin, granting #10366).

  • Added: New syntax Inductive ident binder* | binder* := ... to specify which parameters of an inductive type are uniform. See Parameterized inductive types (#11600, by Gaëtan Gilbert).

  • Added: Warn when using Fixpoint or CoFixpoint for definitions which are not recursive (#12121, by Hugo Herbelin).

  • Fixed: More robust and expressive treatment of implicit inductive parameters in inductive declarations (#11579, by Maxime Dénès, Gaëtan Gilbert and Jasper Hugunin; fixes #7253 and #11585).

  • Fixed: Anomaly which could be raised when printing binders with implicit types (#12323, by Hugo Herbelin; fixes #12322).

  • Fixed: Case of an anomaly in trying to infer the return clause of an ill-typed match (#12422, fixes #12418, by Hugo Herbelin).

Notations

  • Changed: Notation scopes are now always inherited in notations binding a partially applied constant, including for notations binding an expression of the form @qualid. The latter was not the case beforehand (part of #11120).

  • Changed: The printing algorithm now interleaves search for notations and removal of coercions (#11172, by Hugo Herbelin).

  • Changed: Nicer printing for decimal constants in R and Q. 1.5 is now printed 1.5 rather than 15e-1 (#11848, by Pierre Roux).

  • Removed: deprecated compat modifier of Notation and Infix commands. Use the deprecated attribute instead (#11113, by Théo Zimmermann, with help from Jason Gross).

  • Deprecated: Numeral Notation on Decimal.uint, Decimal.int and Decimal.decimal are replaced respectively by numeral notations on Numeral.uint, Numeral.int and Numeral.numeral (#11948, by Pierre Roux).

  • Added: Notations declared with the where clause in the declaration of inductive types, coinductive types, record fields, fixpoints and cofixpoints now support the only parsing modifier (#11602, by Hugo Herbelin).

  • Added: Printing Parentheses flag to print parentheses even when implied by associativity or precedence (#11650, by Hugo Herbelin and Abhishek Anand).

  • Added: Numeral notations now parse hexadecimal constants such as 0x2a or 0xb.2ap-2. Parsers added for nat, positive, Z, N, Q, R, primitive integers and primitive floats (#11948, by Pierre Roux).

  • Added: Abbreviations support arguments occurring both in term and binder position (#8808, by Hugo Herbelin).

  • Fixed: Different interpretations in different scopes of the same notation string can now be associated with different printing formats (#10832, by Hugo Herbelin, fixes #6092 and #7766).

  • Fixed: Parsing and printing consistently handle inheritance of implicit arguments in notations. With the exception of notations of the form Notation string := @qualid and Notation ident := @qualid which inhibit implicit arguments, all notations binding a partially applied constant, as e.g. in Notation string := (qualid arg+), or Notation string := (@qualid arg+), or Notation ident := (qualid arg+), or Notation ident := (@qualid arg+), inherit the remaining implicit arguments (#11120, by Hugo Herbelin, fixing #4690 and #11091).

  • Fixed: Notations in only printing mode do not uselessly reserve parsing keywords (#11590, by Hugo Herbelin, fixes #9741).

  • Fixed: Numeral Notations now play better with multiple scopes for the same inductive type. Previously, when multiple numeral notations were defined for the same inductive, only the last one was considered for printing. Now, among the notations that are usable for printing and either have a scope delimiter or are open, the selection is made according to the order of open scopes, or according to the last defined notation if no appropriate scope is open (#12163, fixes #12159, by Pierre Roux, review by Hugo Herbelin and Jason Gross).

Tactics

  • Changed: The rapply tactic in Coq.Program.Tactics now handles arbitrary numbers of underscores and takes in a uconstr. In rare cases where users were relying on rapply inserting exactly 15 underscores and no more, due to the lemma having a completely unspecified codomain (and thus allowing for any number of underscores), the tactic will now loop instead (#10760, by Jason Gross).

  • Changed: The auto with zarith tactic and variations (including intuition) may now call lia instead of omega (when the Omega module is loaded); more goals may be automatically solved, fewer section variables will be captured spuriously (#11018, by Vincent Laporte).

  • Changed: The new NativeCompute Timing flag causes calls to native_compute (as well as kernel calls to the native compiler) to emit separate timing information about conversion to native code, compilation, execution, and reification. It replaces the timing information previously emitted when the -debug command-line flag was set, and allows more fine-grained timing of the native compiler (#11025, by Jason Gross). Additionally, the timing information now uses real time rather than user time (fixes #11962, #11963, by Jason Gross)

  • Changed: Improve the efficiency of PreOmega.elim_let using an iterator implemented in OCaml (#11370, by Frédéric Besson).

  • Changed: Improve the efficiency of zify by rewriting the remaining Ltac code in OCaml (#11429, by Frédéric Besson).

  • Changed: Backtrace information for tactics has been improved (#11755, by Emilio Jesus Gallego Arias).

  • Changed: The default tactic used by firstorder is auto with core instead of auto with *; see Solvers for logic and equality for details; old behavior can be reset by using the -compat 8.12 command-line flag; to ease the migration of legacy code, the default solver can be set to debug auto with * with Set Firstorder Solver debug auto with * (#11760, by Vincent Laporte).

  • Changed: autounfold no longer fails when the Opaque command is used on constants in the hint databases (#11883, by Attila Gáspár).

  • Changed: Tactics with qualified name of the form Coq.Init.Notations are now qualified with prefix Coq.Init.Ltac; users of the -noinit option should now import Coq.Init.Ltac if they want to use Ltac (#12023, by Hugo Herbelin; minor source of incompatibilities).

  • Changed: Tactic subst ident now fails over a section variable which is indirectly dependent in the goal; the incompatibility can generally be fixed by first clearing the hypotheses causing an indirect dependency, as reported by the error message, or by using rewrite ... in * instead; similarly, subst has no more effect on such variables (#12146, by Hugo Herbelin; fixes #10812 and #12139).

  • Changed: The check that unfold arguments were indeed unfoldable has been moved to runtime (#12256, by Pierre-Marie Pédrot; fixes #5764, #5159, #4925 and #11727).

  • Changed When the tactic functional induction c1 c2 ... cn is used with no parenthesis around c1 c2 ... cn, c1 c2 ... cn is now read as one single applicative term. In particular implicit arguments should be omitted. Rare source of incompatibility (#12326, by Pierre Courtieu).

  • Changed: When using exists or eexists with multiple arguments, the evaluation of arguments and applications of constructors are now interleaved. This improves unification in some cases (#12366, fixes #12365, by Attila Gáspár).

  • Removed: Undocumented omega with. Using lia is the recommended replacement, although the old semantics of omega with * can also be recovered with zify; omega (#11288, by Emilio Jesus Gallego Arias).

  • Removed: Deprecated syntax _eqn for destruct and remember. Use eqn: syntax instead (#11877, by Hugo Herbelin).

  • Removed: at clauses can no longer be used with autounfold. Since they had no effect, it is safe to remove them (#11883, by Attila Gáspár).

  • Deprecated: The omega tactic is deprecated; use lia from the Micromega plugin instead (#11976, by Vincent Laporte).

  • Added: The zify tactic is now aware of Pos.pred_double, Pos.pred_N, Pos.of_nat, Pos.add_carry, Pos.pow, Pos.square, Z.pow, Z.double, Z.pred_double, Z.succ_double, Z.square, Z.div2, and Z.quot2. Injections for internal definitions in module ZifyBool (isZero and isLeZero) are also added to help users to declare new zify class instances using Micromega tactics (#10998, by Kazuhiko Sakaguchi).

  • Added: Show Lia Profile prints some statistics about lia calls (#11474, by Frédéric Besson).

  • Added: Syntax pose proof (ident:=term) as an alternative to pose proof term as ident, following the model of pose (ident:=term) (#11522, by Hugo Herbelin).

  • Added: New tactical with_strategy which behaves like the command Strategy, with effects local to the given tactic (#12129, by Jason Gross).

  • Added: The zify tactic is now aware of Nat.le, Nat.lt and Nat.eq (#12213, by Frédéric Besson; fixes #12210).

  • Fixed: zify now handles Z.pow_pos by default. In Coq 8.11, this was the case only when loading module ZifyPow because this triggered a regression of lia. The regression is now fixed, and the module kept only for compatibility (#11362, fixes #11191, by Frédéric Besson).

  • Fixed: Efficiency regression of lia (#11474, fixes #11436, by Frédéric Besson).

  • Fixed: The behavior of autounfold no longer depends on the names of terms and modules (#11883, fixes #7812, by Attila Gáspár).

  • Fixed: Wrong type error in tactic functional induction (#12326, by Pierre Courtieu, fixes #11761, reported by Lasse Blaauwbroek).

Tactic language

  • Changed: The "reference" tactic generic argument now accepts arbitrary variables of the goal context (#12254, by Pierre-Marie Pédrot).

  • Added: An array library for Ltac2 (as compatible as possible with OCaml standard library) (#10343, by Michael Soegtrop).

  • Added: The Ltac2 rebinding command Ltac2 Set has been extended with the ability to give a name to the old value so as to be able to reuse it inside the new one (#11503, by Pierre-Marie Pédrot).

  • Added: Ltac2 notations for enough and eenough (#11740, by Michael Soegtrop).

  • Added: New Ltac2 function Fresh.Free.of_goal to return the list of names of declarations of the current goal; new Ltac2 function Fresh.in_goal to return a variable fresh in the current goal (#11882, by Hugo Herbelin).

  • Added: Ltac2 notations for reductions in terms: eval red_expr in term (#11981, by Michael Soegtrop).

  • Fixed: The Ltac Profiling machinery now correctly handles backtracking into multi-success tactics. The call-counts of some tactics are unfortunately inflated by 1, as some tactics are implicitly implemented as tac + fail, which has two entry-points rather than one (fixes #12196, #12197, by Jason Gross).

SSReflect

  • Changed: The Search (ssreflect) command that used to be available when loading the ssreflect plugin has been moved to a separate plugin that needs to be loaded separately: ssrsearch (part of #8855, fixes #12253, by Théo Zimmermann).

  • Deprecated: Search (ssreflect) (available through Require ssrsearch.) in favor of the headconcl: clause of Search (part of #8855, by Théo Zimmermann).

Flags, options and attributes

  • Changed: Legacy attributes can now be passed in any order (#11665, by Théo Zimmermann).

  • Removed: Typeclasses Axioms Are Instances flag, deprecated since 8.10. Use Declare Instance for axioms which should be instances (#11185, by Théo Zimmermann).

  • Removed: Deprecated unsound compatibility Template Check flag that was introduced in 8.10 to help users gradually move their template polymorphic inductive type definitions outside sections (#11546, by Pierre-Marie Pédrot).

  • Removed: Deprecated Shrink Obligations flag (#11828, by Emilio Jesus Gallego Arias).

  • Removed: Unqualified polymorphic, monomorphic, template, notemplate attributes (they were deprecated since Coq 8.10). Use universes(polymorphic), universes(monomorphic), universes(template) and universes(notemplate) instead (#11663, by Théo Zimmermann).

  • Deprecated: Hide Obligations flag (#11828, by Emilio Jesus Gallego Arias).

  • Added: Handle the local attribute in Canonical Structure declarations (#11162, by Enrico Tassi).

  • Added: New attributes supported when defining an inductive type universes(cumulative), universes(noncumulative) and private(matching), which correspond to legacy attributes Cumulative, NonCumulative, and the previously undocumented Private (#11665, by Théo Zimmermann).

  • Added: The Hint commands now accept the export locality as an attribute, allowing to make import-scoped hints (#11812, by Pierre-Marie Pédrot).

  • Added: Cumulative StrictProp to control cumulativity of \(\SProp\) (#12034, by Gaëtan Gilbert).

Commands

  • Changed: The Coercion command has been improved to check the coherence of the inheritance graph. It checks whether a circular inheritance path of C >-> C is convertible with the identity function or not, then report it as an ambiguous path if it is not. The new mechanism does not report ambiguous paths that are redundant with others. For example, checking the ambiguity of [f; g] and [f'; g] is redundant with that of [f] and [f'] thus will not be reported (#11258, by Kazuhiko Sakaguchi).

  • Changed: Several commands (Search, About, ...) now print the implicit arguments in brackets when printing types (#11795, by Simon Boulier).

  • Changed: The warning when using Require inside a section moved from the deprecated category to the fragile category, because there is no plan to remove the functionality at this time (#11972, by Gaëtan Gilbert).

  • Changed: Redirect now obeys the Printing Width and Printing Depth options (#12358, by Emilio Jesus Gallego Arias).

  • Removed: Recursive OCaml loadpaths are not supported anymore; the command Add Rec ML Path has been removed; Add ML Path is now the preferred one. We have also dropped support for the non-qualified version of the Add LoadPath command, that is to say, the Add LoadPath dir version; now, you must always specify a prefix now using Add Loadpath dir as Prefix (#11618, by Emilio Jesus Gallego Arias).

  • Removed: undocumented Chapter command. Use Section instead (#11746, by Théo Zimmermann).

  • Removed: SearchAbout command that was deprecated since 8.5. Use Search instead (#11944, by Jim Fehrle).

  • Deprecated: Declaration of arbitrary terms as hints. Global references are now preferred (#7791, by Pierre-Marie Pédrot).

  • Deprecated: SearchHead in favor of the new headconcl: clause of Search (part of #8855, by Théo Zimmermann).

  • Added: Print Canonical Projections can now take constants as arguments and prints only the unification rules that involve or are synthesized from the given constants (#10747, by Kazuhiko Sakaguchi).

  • Added: A section variable introduced with Let can be declared as a Canonical Structure (#11164, by Enrico Tassi).

  • Added: Support for universe bindings and universe contrainsts in Let definitions (#11534, by Théo Zimmermann).

  • Fixed: A printing bug in the presence of elimination principles with local definitions (#12295, by Hugo Herbelin; fixes #12233).

  • Fixed: Anomalies with Show Proof (#12296, by Hugo Herbelin; fixes #12234).

Tools

  • Changed: Internal options and behavior of coqdep. coqdep no longer works as a replacement for ocamldep, thus .ml files are not supported as input. Also, several deprecated options have been removed: -w, -D, -mldep, -prefix, -slash, and -dumpbox. Passing -boot to coqdep will not load any path by default now, -R/-Q should be used instead (#11523 and #11589, by Emilio Jesus Gallego Arias).

  • Changed: The order in which the require flags -ri, -re, -rfrom, etc. and the option flags -set, -unset are given now matters. In particular, it is now possible to interleave the loading of plugins and the setting of options by choosing the right order for these flags. The load flags -l and -lv are still processed afterward for now (#11851 and #12097, by Lasse Blaauwbroek).

  • Changed: The cleanall target of a makefile generated by coq_makefile now erases .lia.cache and .nia.cache (#12006, by Olivier Laurent).

  • Changed: The output of make TIMED=1 (and therefore the timing targets such as print-pretty-timed and print-pretty-timed-diff) now displays the full name of the output file being built, rather than the stem of the rule (which was usually the filename without the extension, but in general could be anything for user-defined rules involving %) (#12126, by Jason Gross).

  • Changed: When passing TIMED=1 to make with either Coq's own makefile or a coq_makefile-made makefile, timing information is now printed for OCaml files as well (#12211, by Jason Gross).

  • Changed: The pretty-timed scripts and targets now print a newline at the end of their tables, rather than creating text with no trailing newline (#12368, by Jason Gross).

  • Removed: The -load-ml-source and -load-ml-object command-line options have been removed; their use was very limited, you can achieve the same adding additional object files in the linking step or using a plugin (#11409, by Emilio Jesus Gallego Arias).

  • Removed: The confusingly-named -require command-line option, which was deprecated since 8.11. Use the equivalent -require-import / -ri options instead (#12005, by Théo Zimmermann).

  • Deprecated: -cumulative-sprop command-line flag in favor of the new Cumulative StrictProp flag (#12034, by Gaëtan Gilbert).

  • Added: A new documentation environment details to make certain portion of a Coq document foldable. See Hiding / Showing parts of the source (#10592, by Thomas Letan).

  • Added: The make-both-single-timing-files.py script now accepts a --fuzz=N parameter on the command line which determines how many characters two lines may be offset in the "before" and "after" timing logs while still being considered the same line. When invoking this script via the print-pretty-single-time-diff target in a Makefile made by coq_makefile, you can set this argument by passing TIMING_FUZZ=N to make (#11302, by Jason Gross).

  • Added: The make-one-time-file.py and make-both-time-files.py scripts now accept a --real parameter on the command line to print real times rather than user times in the tables. The make-both-single-timing-files.py script accepts a --user parameter to use user times. When invoking these scripts via the print-pretty-timed or print-pretty-timed-diff or print-pretty-single-time-diff targets in a Makefile made by coq_makefile, you can set this argument by passing TIMING_REAL=1 (to pass --real) or TIMING_REAL=0 (to pass --user) to make (#11302, by Jason Gross).

  • Added: Coq's build system now supports both TIMING_FUZZ, TIMING_SORT_BY, and TIMING_REAL just like a Makefile made by coq_makefile (#11302, by Jason Gross).

  • Added: The make-one-time-file.py and make-both-time-files.py scripts now include peak memory usage information in the tables (can be turned off by the --no-include-mem command-line parameter), and a --sort-by-mem parameter to sort the tables by memory rather than time. When invoking these scripts via the print-pretty-timed or print-pretty-timed-diff targets in a Makefile made by coq_makefile, you can set this argument by passing TIMING_INCLUDE_MEM=0 (to pass --no-include-mem) and TIMING_SORT_BY_MEM=1 (to pass --sort-by-mem) to make (#11606, by Jason Gross).

  • Added: Coq's build system now supports both TIMING_INCLUDE_MEM and TIMING_SORT_BY_MEM just like a Makefile made by coq_makefile (#11606, by Jason Gross).

  • Added: New coqc / coqtop option -boot that will not bind the Coq library prefix by default (#11617, by Emilio Jesus Gallego Arias).

  • Added: Definitions in coqdoc link to themselves, giving access in html to their own url (#12026, by Hugo Herbelin; granting #7093).

  • Added: Hyperlinks on bound variables in coqdoc (#12033, by Hugo Herbelin; it incidentally fixes #7697).

  • Added: Highlighting of link targets in coqdoc (#12091, by Hugo Herbelin).

  • Fixed: The various timing targets for Coq's standard library now correctly display and label the "before" and "after" columns, rather than mixing them up (#11302 fixes #11301, by Jason Gross).

  • Fixed: The sorting order of the timing script make-both-time-files.py and the target print-pretty-timed-diff is now deterministic even when the sorting order is absolute or diff; previously the relative ordering of two files with identical times was non-deterministic (#11606, by Jason Gross).

  • Fixed: Fields of a record tuple now link in coqdoc to their definition (#12027, fixes #3415, by Hugo Herbelin).

  • Fixed: coqdoc now reports the location of a mismatched opening [[ instead of throwing an uninformative exception (#12037, fixes #9670, by Xia Li-yao).

  • Fixed: coqchk incorrectly reporting names from opaque modules as axioms (#12076, by Pierre Roux; fixes #5030).

  • Fixed: coq_makefile-generated Makefiles pretty-timed-diff target no longer raises Python exceptions in the rare corner case where the log of times contains no files (#12388, fixes #12387, by Jason Gross).

CoqIDE

  • Removed: "Tactic" menu from CoqIDE which had been unmaintained for a number of years (#11414, by Pierre-Marie Pédrot).

  • Removed: "Revert all buffers" command from CoqIDE which had been broken for a long time (#11415, by Pierre-Marie Pédrot).

Standard library

  • Changed: Notations [|term|] and [||term||] for morphisms from 63-bit integers to Z and zn2z int have been removed in favor of φ(term) and Φ(term) respectively. These notations were breaking Ltac parsing (#11686, by Maxime Dénès).

  • Changed: The names of Sorted_sort and LocallySorted_sort in Coq.Sorting.MergeSort have been swapped to appropriately reflect their meanings (#11885, by Lysxia).

  • Changed: Notations <=? and <? from Coq.Structures.Orders and Coq.Sorting.Mergesort.NatOrder are now at level 70 rather than 35, so as to be compatible with the notations defined everywhere else in the standard library. This may require re-parenthesizing some expressions. These notations were breaking the ability to import modules from the standard library that were otherwise compatible (fixes #11890, #11891, by Jason Gross).

  • Changed: The level of in Coq.Numbers.Cyclic.Int63.Int63 is now 70, no associativity, in line with =. Note that this is a minor incompatibility with developments that declare their own notation and import Int63 (fixes #11905, #11909, by Jason Gross).

  • Changed: No longer re-export ListNotations from Program (Program.Syntax) (#11992, by Antonio Nikishaev).

  • Changed: It is now possible to import the nsatz machinery without transitively depending on the axioms of the real numbers nor of classical logic by loading Coq.nsatz.NsatzTactic rather than Coq.nsatz.Nsatz. Note that some constants have changed kernel names, living in Coq.nsatz.NsatzTactic rather than Coq.nsatz.Nsatz; this might cause minor incompatibilities that can be fixed by actually running Import Nsatz rather than relying on absolute names (#12073, by Jason Gross; fixes #5445).

  • Changed: new lemma NoDup_incl_NoDup in List.v to remove useless hypothesis NoDup l' in Sorting.Permutation.NoDup_Permutation_bis (#12120, by Olivier Laurent).

  • Changed: Fixpoints of the standard library without a recursive call turned into ordinary Definitions (#12121, by Hugo Herbelin; fixes #11903).

  • Deprecated: Bool.leb in favor of Bool.le. The definition of Bool.le is made local to avoid conflicts with Nat.le. As a consequence, previous calls to leb based on importing Bool should now be qualified into Bool.le even if Bool is imported (#12162, by Olivier Laurent).

  • Added: Theorem bezout_comm for natural numbers (#11127, by Daniel de Rauglaudre).

  • Added rew dependent notations for the dependent version of rew in Coq.Init.Logic.EqNotations to improve the display and parsing of match statements on Logic.eq (#11240, by Jason Gross).

  • Added: Lemmas about lists:

    • properties of In: in_elt, in_elt_inv

    • properties of nth: app_nth2_plus, nth_middle, nth_ext

    • properties of last: last_last, removelast_last

    • properties of remove: remove_cons, remove_app, notin_remove, in_remove, in_in_remove, remove_remove_comm, remove_remove_eq, remove_length_le, remove_length_lt

    • properties of concat: in_concat, remove_concat

    • properties of map and flat_map: map_last, map_eq_cons, map_eq_app, flat_map_app, flat_map_ext, nth_nth_nth_map

    • properties of incl: incl_nil_l, incl_l_nil, incl_cons_inv, incl_app_app, incl_app_inv, remove_incl, incl_map, incl_filter, incl_Forall_in_iff

    • properties of NoDup and nodup: NoDup_rev, NoDup_filter, nodup_incl

    • properties of Exists and Forall: Exists_nth, Exists_app, Exists_rev, Exists_fold_right, incl_Exists, Forall_nth, Forall_app, Forall_elt, Forall_rev, Forall_fold_right, incl_Forall, map_ext_Forall, Exists_or, Exists_or_inv, Forall_and, Forall_and_inv, exists_Forall, Forall_image, concat_nil_Forall, in_flat_map_Exists, notin_flat_map_Forall

    • properties of repeat: repeat_cons, repeat_to_concat

    • definitions and properties of list_sum and list_max: list_sum_app, list_max_app, list_max_le, list_max_lt

    • misc: elt_eq_unit, last_length, rev_eq_app, removelast_firstn_len, cons_seq, seq_S

    (#11249, #12237, by Olivier Laurent).

  • Added: Well-founded induction principles for nat: lt_wf_rect1, lt_wf_rect, gt_wf_rect, lt_wf_double_rect (#11335, by Olivier Laurent).

  • Added: remove' and count_occ' over lists, alternatives to remove and count_occ based on filter (#11350, by Yishuai Li).

  • Added: Facts about N.iter and Pos.iter:

    • N.iter_swap_gen, N.iter_swap, N.iter_succ, N.iter_succ_r, N.iter_add, N.iter_ind, N.iter_invariant

    • Pos.iter_succ_r, Pos.iter_ind

    (#11880, by Lysxia).

  • Added: Facts about Permutation:

    • structure: Permutation_refl', Permutation_morph_transp

    • compatibilities: Permutation_app_rot, Permutation_app_swap_app, Permutation_app_middle, Permutation_middle2, Permutation_elt, Permutation_Forall, Permutation_Exists, Permutation_Forall2, Permutation_flat_map, Permutation_list_sum, Permutation_list_max

    • inversions: Permutation_app_inv_m, Permutation_vs_elt_inv, Permutation_vs_cons_inv, Permutation_vs_cons_cons_inv, Permutation_map_inv, Permutation_image, Permutation_elt_map_inv

    • length-preserving definition by means of transpositions Permutation_transp with associated properties: Permutation_transp_sym, Permutation_transp_equiv, Permutation_transp_cons, Permutation_Permutation_transp, Permutation_ind_transp

    (#11946, by Olivier Laurent).

  • Added: Notations for sigma types: { x & P & Q }, { ' pat & P }, { ' pat & P & Q } (#11957, by Olivier Laurent).

  • Added: Order relations lt and compare added in Bool.Bool. Order properties for bool added in Bool.BoolOrder as well as two modules Bool_as_OT and Bool_as_DT in Structures.OrdersEx (#12008, by Olivier Laurent).

  • Added: Properties of some operations on vectors:

    • nth_order: nth_order_hd, nth_order_tl, nth_order_ext

    • replace: nth_order_replace_eq, nth_order_replace_neq, replace_id, replace_replace_eq, replace_replace_neq

    • map: map_id, map_map, map_ext_in, map_ext

    • Forall and Forall2: Forall_impl, Forall_forall, Forall_nth_order, Forall2_nth_order

    (#12014, by Olivier Laurent).

  • Added: Lemmas orb_negb_l, andb_negb_l, implb_true_iff, implb_false_iff, implb_true_r, implb_false_r, implb_true_l, implb_false_l, implb_same, implb_contrapositive, implb_negb, implb_curry, implb_andb_distrib_r, implb_orb_distrib_r, implb_orb_distrib_l in library Bool (#12018, by Hugo Herbelin).

  • Added: Definition and properties of cyclic permutations / circular shifts: CPermutation (#12031, by Olivier Laurent).

  • Added: Structures.OrderedTypeEx.Ascii_as_OT (#12044, by formalize.eth (formalize@protonmail.com)).

  • Fixed: Rewrote Structures.OrderedTypeEx.String_as_OT.compare to avoid huge proof terms (#12044, by formalize.eth (formalize@protonmail.com); fixes #12015).

Reals library

  • Changed: Cleanup of names in the Reals theory: replaced tan_is_inj with tan_inj and replaced atan_right_inv with tan_atan - compatibility notations are provided. Moved various auxiliary lemmas from Ratan.v to more appropriate places (#9803, by Laurent Théry and Michael Soegtrop).

  • Changed: Replace CRzero and CRone by CR_of_Q 0 and CR_of_Q 1 in ConstructiveReals. Use implicit arguments for ConstructiveReals. Move ConstructiveReals into new directory Abstract. Remove imports of implementations inside those Abstract files. Move implementation by means of Cauchy sequences in new directory Cauchy. Split files ConstructiveMinMax and ConstructivePower.

    Warning

    The constructive reals modules are marked as experimental.

    (#11725, #12287 and #12288, by Vincent Semeria).

  • Removed: Type RList has been removed. All uses have been replaced by list R. Functions from RList named In, Rlength, cons_Rlist, app_Rlist have also been removed as they are essentially the same as In, length, app, and map from List, modulo the following changes:

    • RList.In x (RList.cons a l) used to be convertible to (x = a) \/ RList.In x l, but List.In x (a :: l) is convertible to (a = x) \/ List.In l. The equality is reversed.

    • app_Rlist and List.map take arguments in different order.

    (#11404, by Yves Bertot).

  • Added: inverse trigonometric functions asin and acos with lemmas for the derivatives, bounds and special values of these functions; an extensive set of identities between trigonometric functions and their inverse functions; lemmas for the injectivity of sine and cosine; lemmas on the derivative of the inverse of decreasing functions and on the derivative of horizontally mirrored functions; various generic auxiliary lemmas and definitions for Rsqr, sqrt, posreal and others (#9803, by Laurent Théry and Michael Soegtrop).

Extraction

  • Added: Support for better extraction of strings in OCaml and Haskell: ExtOcamlNativeString provides bindings from the Coq String type to the OCaml string type, and string literals can be extracted to literals, both in OCaml and Haskell (#10486, by Xavier Leroy, with help from Maxime Dénès, review by Hugo Herbelin).

  • Fixed: In Haskell extraction with ExtrHaskellString, equality comparisons on strings and characters are now guaranteed to be uniquely well-typed, even in very polymorphic contexts under unsafeCoerce; this is achieved by adding type annotations to the extracted code, and by making ExtrHaskellString export ExtrHaskellBasic (#12263, by Jason Gross, fixes #12257 and #12258).

Reference manual

  • Changed: The reference manual has been restructured to get a more logical organization. In the new version, there are fewer top-level chapters, and, in the HTML format, chapters are split into smaller pages. This is still a work in progress and further restructuring is expected in the next versions of Coq (CEP#43, implemented in #11601, #11871, #11914, #12148, #12172, #12239 and #12330, effort inspired by Matthieu Sozeau, led by Théo Zimmermann, with help and reviews of Jim Fehrle, Clément Pit-Claudel and others).

  • Changed: Most of the grammar is now presented using the notation mechanism that has been used to present commands and tactics since Coq 8.8 and which is documented in Syntax conventions (#11183, #11314, #11423, #11705, #11718, #11720, #11961 and #12103, by Jim Fehrle, reviewed by Théo Zimmermann).

  • Added: A glossary of terms and an index of attributes (#11869, #12150 and #12224, by Jim Fehrle and Théo Zimmermann, reviewed by Clément Pit-Claudel)

  • Added: A selector that allows switching between versions of the reference manual (#12286, by Clément Pit-Claudel).

  • Fixed: Most of the documented syntax has been thoroughly updated to make it accurate and easily understood. This was done using a semi-automated doc_grammar tool introduced for this purpose and through significant revisions to the text (#9884, #10614, #11314, #11423, #11705, #11718, #11720 #11797, #11913, #11958, #11960, #11961 and #12103, by Jim Fehrle, reviewed by Théo Zimmermann and Jason Gross).

Infrastructure and dependencies

  • Changed: Minimal versions of dependencies for building the reference manual: now requires Sphinx >= 2.3.1 & < 3.0.0, sphinx_rtd_theme 0.4.3+ and sphinxcontrib-bibtex 0.4.2+.

    Warning

    The reference manual is known not to build properly with Sphinx 3.

    (#12224, by Jim Fehrle and Théo Zimmermann).

  • Removed: Python 2 is no longer required in any part of the codebase (#11245, by Emilio Jesus Gallego Arias).

Changes in 8.12.0

Notations

  • Added: Simultaneous definition of terms and notations now support custom entries (#12523, fixes #11121 by Maxime Dénès).

  • Fixed: Printing bug with notations for n-ary applications used with applied references (#12683, fixes #12682, by Hugo Herbelin).

Tactics

  • Fixed: typeclasses eauto (and discriminated hint bases) now correctly classify local variables as being unfoldable (#12572, fixes #12571, by Pierre-Marie Pédrot).

Tactic language

  • Fixed: Excluding occurrences was causing an anomaly in tactics (e.g., pattern _ at L where L is -2) (#12541, fixes #12228, by Pierre Roux).

  • Fixed: Parsing of multi-parameters Ltac2 types (#12594, fixes #12595, by Pierre-Marie Pédrot).

SSReflect

  • Fixed: Do not store the full environment inside ssr ast_closure_term (#12708, fixes #12707, by Pierre-Marie Pédrot).

Commands and options

  • Fixed: Properly report the mismatched magic number of vo files (#12677, fixes #12513, by Pierre-Marie Pédrot).

  • Changed: Arbitrary hints have been undeprecated, and their definition now triggers a standard warning instead (#12678, fixes #11970, by Pierre-Marie Pédrot).

CoqIDE

  • Fixed: CoqIDE no longer exits when trying to open a file whose name is not a valid identifier (#12562, fixes #10988, by Vincent Laporte).

Infrastructure and dependencies

  • Fixed: Running make in test-suite/ twice (or more) in a row will no longer rebuild the modules/ tests on subsequent runs, if they have not been modified in the meantime (#12583, fixes #12582, by Jason Gross).

Changes in 8.12.1

Kernel

Notations

  • Fixed: Undetected collision between a lonely notation and a notation in scope at printing time (#12946, fixes the first part of #12908, by Hugo Herbelin).

  • Fixed: Printing of notations in custom entries with variables not mentioning an explicit level (#13026, fixes #12775 and #13018, by Hugo Herbelin).

Tactics

  • Added: replace and inversion support registration of a core.identity-like equality in Type, such as HoTT's path (#12847, partially fixes #12846, by Hugo Herbelin).

  • Fixed: Anomaly with injection involving artificial dependencies disappearing by reduction (#12816, fixes #12787, by Hugo Herbelin).

Tactic language

  • Fixed: Miscellaneous issues with locating tactic errors (#13247, fixes #12773 and #12992, by Hugo Herbelin).

SSReflect

  • Fixed: Regression in error reporting after case. A generic error message "Could not fill dependent hole in apply" was reported for any error following case or elim (#12857, fixes #12837, by Enrico Tassi).

Commands and options

  • Fixed: Failures of Search in the presence of primitive projections (#13301, fixes #13298, by Hugo Herbelin).

  • Fixed: Search supports filtering on parts of identifiers which are not proper identifiers themselves, such as "1" (#13351, fixes #13349, by Hugo Herbelin).

Tools

  • Fixed: Special symbols now escaped in the index produced by coqdoc, avoiding collision with the syntax of the output format (#12754, fixes #12752, by Hugo Herbelin).

  • Fixed: The details environment added in the 8.12 release can now be used as advertised in the reference manual (#12772, by Thomas Letan).

  • Fixed: Targets such as print-pretty-timed in coq_makefile-made Makefiles no longer error in rare cases where --output-sync is not passed to make and the timing output gets interleaved in just the wrong way (#13063, fixes #13062, by Jason Gross).

CoqIDE

  • Fixed: View menu "Display parentheses" (#12794 and #13067, fixes #12793, by Jean-Christophe Léchenet and Hugo Herbelin).

Infrastructure and dependencies

  • Added: Coq is now tested against OCaml 4.11.1 (#12972, by Emilio Jesus Gallego Arias).

  • Fixed: The reference manual can now build with Sphinx 3 (#13011, fixes #12332, by Théo Zimmermann and Jim Fehrle).

Changes in 8.12.2

Notations

  • Fixed: 8.12 regression causing notations mentioning a coercion to be ignored (#13436, fixes #13432, by Hugo Herbelin).

Tactics

  • Fixed: 8.12 regression: incomplete inference of implicit arguments in exists (#13468, fixes #13456, by Hugo Herbelin).

Version 8.11

Summary of changes

The main changes brought by Coq version 8.11 are:

  • Ltac2, a new tactic language for writing more robust larger scale tactics, with built-in support for datatypes and the multi-goal tactic monad.

  • Primitive floats are integrated in terms and follow the binary64 format of the IEEE 754 standard, as specified in the Coq.Float.Floats library.

  • Cleanups of the section mechanism, delayed proofs and further restrictions of template polymorphism to fix soundness issues related to universes.

  • New unsafe flags to disable locally guard, positivity and universe checking. Reliance on these flags is always printed by Print Assumptions.

  • Fixed bugs of Export and Import that can have a significant impact on user developments (common source of incompatibility!).

  • New interactive development method based on vos interface files, allowing to work on a file without recompiling the proof parts of their dependencies.

  • New Arguments annotation for bidirectional type inference configuration for reference (e.g. constants, inductive) applications.

  • New refine attribute for Instance can be used instead of the removed Refine Instance Mode.

  • Generalization of the under and over tactics of SSReflect to arbitrary relations.

  • Revision of the Coq.Reals library, its axiomatisation and instances of the constructive and classical real numbers.

Additionally, while the omega tactic is not yet deprecated in this version of Coq, it should soon be the case and we already recommend users to switch to lia in new proof scripts.

The dev/doc/critical-bugs file documents the known critical bugs of Coq and affected releases. See the Changes in 8.11+beta1 section and following sections for the detailed list of changes, including potentially breaking changes marked with Changed.

Coq's documentation is available at https://coq.github.io/doc/v8.11/api (documentation of the ML API), https://coq.github.io/doc/v8.11/refman (reference manual), and https://coq.github.io/doc/v8.11/stdlib (documentation of the standard library).

Maxime Dénès, Emilio Jesús Gallego Arias, Gaëtan Gilbert, Michael Soegtrop and Théo Zimmermann worked on maintaining and improving the continuous integration system and package building infrastructure.

The opam repository for Coq packages has been maintained by Guillaume Claret, Karl Palmskog, Matthieu Sozeau and Enrico Tassi with contributions from many users. A list of packages is available at https://coq.inria.fr/opam/www/.

The 61 contributors to this version are Michael D. Adams, Guillaume Allais, Helge Bahmann, Langston Barrett, Guillaume Bertholon, Frédéric Besson, Simon Boulier, Michele Caci, Tej Chajed, Arthur Charguéraud, Cyril Cohen, Frédéric Dabrowski, Arthur Azevedo de Amorim, Maxime Dénès, Nikita Eshkeev, Jim Fehrle, Emilio Jesús Gallego Arias, Paolo G. Giarrusso, Gaëtan Gilbert, Georges Gonthier, Jason Gross, Samuel Gruetter, Armaël Guéneau, Hugo Herbelin, Florent Hivert, Jasper Hugunin, Shachar Itzhaky, Jan-Oliver Kaiser, Robbert Krebbers, Vincent Laporte, Olivier Laurent, Samuel Lelièvre, Nicholas Lewycky, Yishuai Li, Jose Fernando Lopez Fernandez, Andreas Lynge, Kenji Maillard, Erik Martin-Dorel, Guillaume Melquiond, Alexandre Moine, Oliver Nash, Wojciech Nawrocki, Antonio Nikishaev, Pierre-Marie Pédrot, Clément Pit-Claudel, Lars Rasmusson, Robert Rand, Talia Ringer, JP Rodi, Pierre Roux, Kazuhiko Sakaguchi, Vincent Semeria, Michael Soegtrop, Matthieu Sozeau, spanjel, Claude Stolze, Enrico Tassi, Laurent Théry, James R. Wilcox, Xia Li-yao, Théo Zimmermann

Many power users helped to improve the design of the new features via the issue and pull request system, the Coq development mailing list, the coq-club@inria.fr mailing list or the Discourse forum. It would be impossible to mention exhaustively the names of everybody who to some extent influenced the development.

Version 8.11 is the sixth release of Coq developed on a time-based development cycle. Its development spanned 3 months from the release of Coq 8.10. Pierre-Marie Pédrot is the release manager and maintainer of this release, assisted by Matthieu Sozeau. This release is the result of 2000+ commits and 300+ PRs merged, closing 75+ issues.

Paris, November 2019,
Matthieu Sozeau for the Coq development team

Changes in 8.11+beta1

Kernel

  • Added: A built-in support of floating-point arithmetic, allowing one to devise efficient reflection tactics involving numerical computation. Primitive floats are added in the language of terms, following the binary64 format of the IEEE 754 standard, and the related operations are implemented for the different reduction engines of Coq by using the corresponding processor operators in rounding-to-nearest-even. The properties of these operators are axiomatized in the theory Coq.Floats.FloatAxioms which is part of the library Coq.Floats.Floats. See Section Primitive Floats (#9867, closes #8276, by Guillaume Bertholon, Erik Martin-Dorel, Pierre Roux).

  • Changed: Internal definitions generated by abstract-like tactics are now inlined inside universe Qed-terminated polymorphic definitions, similarly to what happens for their monomorphic counterparts, (#10439, by Pierre-Marie Pédrot).

  • Fixed: Section data is now part of the kernel. Solves a soundness issue in interactive mode where global monomorphic universe constraints would be dropped when forcing a delayed opaque proof inside a polymorphic section. Also relaxes the nesting criterion for sections, as polymorphic sections can now appear inside a monomorphic one (#10664, by Pierre-Marie Pédrot).

  • Changed: Using SProp is now allowed by default, without needing to pass -allow-sprop or use Allow StrictProp (#10811, by Gaëtan Gilbert).

Specification language, type inference

  • Added: Annotation in Arguments for bidirectionality hints: it is now possible to tell type inference to use type information from the context once the n first arguments of an application are known. The syntax is: Arguments foo x y & z. See Bidirectionality hints (#10049, by Maxime Dénès with help from Enrico Tassi).

  • Added: Record fields can be annotated to prevent them from being used as canonical projections; see Canonical Structures for details (#10076, by Vincent Laporte).

  • Changed: Require parentheses around nested disjunctive patterns, so that pattern and term syntax are consistent; match branch patterns no longer require parentheses for notation at level 100 or more.

    Warning

    Incompatibilities

    • In match p with (_, (0|1)) => ... parentheses may no longer be omitted around 0|1.

    • Notation (p | q) now potentially clashes with core pattern syntax, and should be avoided. -w disj-pattern-notation flags such Notation.

    See Extended pattern matching for details (#10167, by Georges Gonthier).

  • Changed: Function always opens a proof when used with a measure or wf annotation, see Advanced recursive functions for the updated documentation (#10215, by Enrico Tassi).

  • Changed: The legacy command Add Morphism always opens a proof and cannot be used inside a module type. In order to declare a module type parameter that happens to be a morphism, use Declare Morphism. See Deprecated syntax and backward incompatibilities for the updated documentation (#10215, by Enrico Tassi).

  • Changed: The universe polymorphism setting now applies from the opening of a section. In particular, it is not possible anymore to mix polymorphic and monomorphic definitions in a section when there are no variables nor universe constraints defined in this section. This makes the behavior consistent with the documentation. (#10441, by Pierre-Marie Pédrot)

  • Added: The Section command now accepts the "universes" attribute. In addition to setting the section universe polymorphism, it also locally sets the universe polymorphic option inside the section (#10441, by Pierre-Marie Pédrot)

  • Fixed: Program Fixpoint now uses ex and sig to make telescopes involving Prop types (#10758, by Gaëtan Gilbert, fixing #10757 reported by Xavier Leroy).

  • Changed: Output of the Print and About commands. Arguments meta-data is now displayed as the corresponding Arguments command instead of the human-targeted prose used in previous Coq versions. (#10985, by Gaëtan Gilbert).

  • Added: refine attribute for Instance, a more predictable version of the old Refine Instance Mode which unconditionally opens a proof (#10996, by Gaëtan Gilbert).

  • Changed: The unsupported attribute error is now an error-by-default warning, meaning it can be disabled (#10997, by Gaëtan Gilbert).

  • Fixed: Bugs sometimes preventing to define valid (co)fixpoints with implicit arguments in the presence of local definitions, see #3282 (#11132, by Hugo Herbelin).

    Example

    The following features an implicit argument after a local definition. It was wrongly rejected.

    Definition f := fix f (o := true) {n : nat} m {struct m} :=   match m with 0 => 0 | S m' => f (n:=n+1) m' end.
    f is defined

Notations

  • Added: Numeral Notations now support sorts in the input to printing functions (e.g., numeral notations can be defined for terms containing things like @cons Set nat nil). (#9883, by Jason Gross).

  • Added: The Notation and Infix commands now support the deprecated attribute (#10180, by Maxime Dénès).

  • Deprecated: The former compat annotation for notations is deprecated, and its semantics changed. It is now made equivalent to using a deprecated attribute, and is no longer connected with the -compat command-line flag (#10180, by Maxime Dénès).

  • Changed: A simplification of parsing rules could cause a slight change of parsing precedences for the very rare users who defined notations with constr at level strictly between 100 and 200 and used these notations on the right-hand side of a cast operator (:, <:, <<:) (#10963, by Théo Zimmermann, simplification initially noticed by Jim Fehrle).

Tactics

  • Added: Syntax injection term as [= intropattern+ ] as an alternative to injection term as simple_intropattern+ using the standard injection intropattern syntax (#9288, by Hugo Herbelin).

  • Changed: Reimplementation of the zify tactic. The tactic is more efficient and copes with dependent hypotheses. It can also be extended by redefining the tactic zify_post_hook (#9856, fixes #8898, #7886, #9848 and #5155, by Frédéric Besson).

  • Changed: The goal selector tactical only now checks that the goal range it is given is valid instead of ignoring goals out of the focus range (#10318, by Gaëtan Gilbert).

  • Added: Flags Lia Cache, Nia Cache and Nra Cache (#10765, by Frédéric Besson, see #10772 for use case).

  • Added: The zify tactic is now aware of Z.to_N (#10774, grants #9162, by Kazuhiko Sakaguchi).

  • Changed: The assert_succeeds and assert_fails tactics now only run their tactic argument once, even if it has multiple successes. This prevents blow-up and looping from using multisuccess tactics with assert_succeeds. (#10966 fixes #10965, by Jason Gross).

  • Fixed: The assert_succeeds and assert_fails tactics now behave correctly when their tactic fully solves the goal. (#10966 fixes #9114, by Jason Gross).

Tactic language

  • Added: Ltac2, a new version of the tactic language Ltac, that doesn't preserve backward compatibility, has been integrated in the main Coq distribution. It is still experimental, but we already recommend users of advanced Ltac to start using it and report bugs or request enhancements. See its documentation in the dedicated chapter (#10002, plugin authored by Pierre-Marie Pédrot, with contributions by various users, integration by Maxime Dénès, help on integrating / improving the documentation by Théo Zimmermann and Jim Fehrle).

  • Added: Ltac2 tactic notations with “constr” arguments can specify the notation scope for these arguments; see Notations for details (#10289, by Vincent Laporte).

  • Changed: White spaces are forbidden in the &ident syntax for ltac2 references that are described in Built-in quotations (#10324, fixes #10088, authored by Pierre-Marie Pédrot).

SSReflect

  • Added: Generalize tactics under and over for any registered relation. More precisely, assume the given context lemma has type forall f1 f2, .. -> (forall i, R1 (f1 i) (f2 i)) -> R2 f1 f2. The first step performed by under (since Coq 8.10) amounts to calling the tactic rewrite, which itself relies on setoid_rewrite if need be. So this step was already compatible with a double implication or setoid equality for the conclusion head symbol R2. But a further step consists in tagging the generated subgoal R1 (f1 i) (?f2 i) to protect it from unwanted evar instantiation, and get Under_rel _ R1 (f1 i) (?f2 i) that is displayed as 'Under[ f1 i ]. In Coq 8.10, this second (convenience) step was only performed when R1 was Leibniz' eq or iff. Now, it is also performed for any relation R1 which has a RewriteRelation instance (a RelationClasses.Reflexive instance being also needed so over can discharge the 'Under[ _ ] goal by instantiating the hidden evar.) This feature generalizing support for setoid-like relations is enabled as soon as we do both Require Import ssreflect. and Require Setoid. Finally, a rewrite rule UnderE has been added if one wants to "unprotect" the evar, and instantiate it manually with another rule than reflexivity (i.e., without using the over tactic nor the over rewrite rule). See also Section Rewriting under binders (#10022, by Erik Martin-Dorel, with suggestions and review by Enrico Tassi and Cyril Cohen).

  • Added: A void notation for the standard library empty type (Empty_set) (#10932, by Arthur Azevedo de Amorim).

  • Added: Lemma inj_compr to ssr.ssrfun (#11136, by Cyril Cohen).

Commands and options

  • Removed: Deprecated flag Refine Instance Mode (#9530, fixes #3632, #3890 and #4638 by Maxime Dénès, review by Gaëtan Gilbert).

  • Changed: Fail does not catch critical errors (including "stack overflow") anymore (#10173, by Gaëtan Gilbert).

  • Removed: Undocumented Instance : !type syntax (#10185, by Gaëtan Gilbert).

  • Removed: Deprecated Show Script command (#10277, by Gaëtan Gilbert).

  • Added: Unsafe commands to enable/disable guard checking, positivity checking and universes checking (providing a local -type-in-type). See Controlling Typing Flags (#10291 by Simon Boulier).

  • Fixed: Two bugs in Export. This can have an impact on the behavior of the Import command on libraries. Import A when A imports B which exports C was importing C, whereas Import is not transitive. Also, after Import A B, the import of B was sometimes incomplete (#10476, by Maxime Dénès).

    Warning

    This is a common source of incompatibilities in projects migrating to Coq 8.11.

  • Changed: Output generated by Printing Dependent Evars Line flag used by the Prooftree tool in Proof General (#10489, closes #4504, #10399 and #10400, by Jim Fehrle).

  • Added: Optionally highlight the differences between successive proof steps in the Show Proof command. Experimental; only available in coqtop and Proof General for now, may be supported in other IDEs in the future (#10494, by Jim Fehrle).

  • Removed: Legacy commands AddPath, AddRecPath, and DelPath which were undocumented, broken variants of Add LoadPath, Add Rec LoadPath, and Remove LoadPath (#11187, by Maxime Dénès and Théo Zimmermann).

Tools

  • Added: coqc now provides the ability to generate compiled interfaces. Use coqc -vos foo.v to skip all opaque proofs during the compilation of foo.v, and output a file called foo.vos. This feature is experimental. It enables working on a Coq file without the need to first compile the proofs contained in its dependencies (#8642 by Arthur Charguéraud, review by Maxime Dénès and Emilio Gallego).

  • Added: Command-line options -require-import, -require-export, -require-import-from and -require-export-from, as well as their shorthand, -ri, -re, -refrom and -rifrom. Deprecate confusing command line option -require (#10245 by Hugo Herbelin, review by Emilio Gallego).

  • Changed: Renamed VDFILE from .coqdeps.d to .<CoqMakefile>.d in the coq_makefile utility, where <CoqMakefile> is the name of the output file given by the -o option. In this way two generated makefiles can coexist in the same directory (#10947, by Kazuhiko Sakaguchi).

  • Fixed: coq_makefile now supports environment variable COQBIN with no ending / character (#11068, by Gaëtan Gilbert).

Standard library

  • Changed: Moved the auto hints of the OrderedType module into a new ordered_type database (#9772, by Vincent Laporte).

  • Removed: Deprecated modules Coq.ZArith.Zlogarithm and Coq.ZArith.Zsqrt_compat (#9811, by Vincent Laporte).

  • Added: Module Reals.Cauchy.ConstructiveCauchyReals defines constructive real numbers by Cauchy sequences of rational numbers (#10445, by Vincent Semeria, with the help and review of Guillaume Melquiond and Bas Spitters). This module is not meant to be imported directly, please import Reals.Abstract.ConstructiveReals instead.

  • Added: New module Reals.ClassicalDedekindReals defines Dedekind real numbers as boolean-valued functions along with 3 logical axioms: limited principle of omniscience, excluded middle of negations, and functional extensionality. The exposed type R in module Reals.Rdefinitions now corresponds to these Dedekind reals, hidden behind an opaque module, which significantly reduces the number of axioms needed (see Reals.Rdefinitions and Reals.Raxioms), while preserving backward compatibility. Classical Dedekind reals are a quotient of constructive reals, which allows to transport many constructive proofs to the classical case (#10827, by Vincent Semeria, based on discussions with Guillaume Melquiond, Bas Spitters and Hugo Herbelin, code review by Hugo Herbelin).

  • Added: New lemmas on combine, filter, nodup, nth, and nth_error functions on lists (#10651, and #10731, by Oliver Nash).

  • Changed: The lemma filter_app was moved to the List module (#10651, by Oliver Nash).

  • Added: Standard equivalence between weak excluded-middle and the classical instance of De Morgan's law, in module ClassicalFacts (#10895, by Hugo Herbelin).

Infrastructure and dependencies

  • Changed: Coq now officially supports OCaml 4.08. See INSTALL file for details (#10471, by Emilio Jesús Gallego Arias).

Changes in 8.11.0

Kernel

  • Changed: the native compilation (native_compute) now creates a directory to contain temporary files instead of putting them in the root of the system temporary directory (#11081, by Gaëtan Gilbert).

  • Fixed: #11360. Broken section closing when a template polymorphic inductive type depends on a section variable through its parameters (#11361, by Gaëtan Gilbert).

  • Fixed: The type of Set+1 would be computed to be itself, leading to a proof of False (#11422, by Gaëtan Gilbert).

Specification language, type inference

  • Changed: Heuristics for universe minimization to Set: only minimize flexible universes (#10657, by Gaëtan Gilbert with help from Maxime Dénès and Matthieu Sozeau).

  • Fixed: A dependency was missing when looking for default clauses in the algorithm for printing pattern matching clauses (#11233, by Hugo Herbelin, fixing #11231, reported by Barry Jay).

Notations

  • Fixed: Print Visibility was failing in the presence of only-printing notations (#11276, by Hugo Herbelin, fixing #10750).

  • Fixed: Recursive notations with custom entries were incorrectly parsing constr instead of custom grammars (#11311 by Maxime Dénès, fixes #9532, #9490).

Tactics

  • Changed: The tactics eapply, refine and variants no longer allow shelved goals to be solved by typeclass resolution (#10762, by Matthieu Sozeau).

  • Fixed: The optional string argument to time is now properly quoted under Print Ltac (#11203, fixes #10971, by Jason Gross)

  • Fixed: Efficiency regression of lia introduced in 8.10 by PR #9725 (#11263, fixes #11063, and #11242, and #11270, by Frédéric Besson).

  • Deprecated: The undocumented omega with tactic variant has been deprecated. Using lia is the recommended replacement, though the old semantics of omega with * can be recovered with zify; omega (#11337, by Emilio Jesus Gallego Arias).

  • Fixed For compatibility reasons, in 8.11, zify does not support Z.pow_pos by default. It can be enabled by explicitly loading the module ZifyPow (#11430 by Frédéric Besson fixes #11191).

Tactic language

  • Fixed: Syntax of tactic cofix ... with ... was broken since Coq 8.10 (#11241, by Hugo Herbelin).

Commands and options

  • Deprecated: The -load-ml-source and -load-ml-object command line options have been deprecated; their use was very limited, you can achieve the same by adding object files in the linking step or by using a plugin (#11428, by Emilio Jesus Gallego Arias).

Tools

  • Fixed: coqtop --version was broken when called in the middle of an installation process (#11255, by Hugo Herbelin, fixing #11254).

  • Deprecated: The -quick command is renamed to -vio, for consistency with the new -vos and -vok flags. Usage of -quick is now deprecated (#11280, by Arthur Charguéraud).

  • Fixed: coq_makefile does not break when using the CAMLPKGS variable together with an unpacked (mllib) plugin (#11357, by Gaëtan Gilbert).

  • Fixed: coqdoc with option -g (Gallina only) now correctly prints commands with attributes (#11394, fixes #11353, by Karl Palmskog).

CoqIDE

  • Changed: CoqIDE now uses the GtkSourceView native implementation of the autocomplete mechanism (#11400, by Pierre-Marie Pédrot).

Standard library

  • Removed: Export of module RList in Ranalysis and Ranalysis_reg. Module RList is still there but must be imported explicitly where required (#11396, by Michael Soegtrop).

Infrastructure and dependencies

  • Added: Build date can now be overridden by setting the SOURCE_DATE_EPOCH environment variable (#11227, by Bernhard M. Wiedemann).

Changes in 8.11.1

Kernel

  • Fixed: Allow more inductive types in Unset Positivity Checking mode (#11811, by SimonBoulier).

Notations

  • Fixed: Bugs in dealing with precedences of notations in custom entries (#11530, by Hugo Herbelin, fixing in particular #9517, #9519, #9521, #11331).

  • Added: In primitive floats, print a warning when parsing a decimal value that is not exactly a binary64 floating-point number. For instance, parsing 0.1 will print a warning whereas parsing 0.5 won't (#11859, by Pierre Roux).

CoqIDE

  • Fixed: Compiling file paths containing spaces (#10008, by snyke7, fixing #11595).

Infrastructure and dependencies

  • Added: Bump official OCaml support and CI testing to 4.10.0 (#11131, #11123, #11102, by Emilio Jesus Gallego Arias, Jacques-Henri Jourdan, Guillaume Melquiond, and Guillaume Munch-Maccagnoni).

Miscellaneous

Changes in 8.11.2