Writing Coq libraries and plugins¶
This section presents the part of the Coq language that is useful only to library and plugin authors. A tutorial for writing Coq plugins is available in the Coq repository in doc/plugin_tutorial.
Deprecating library objects or tactics¶
You may use the following attribute to deprecate a notation or
tactic. When renaming a definition or theorem, you can introduce a
deprecated compatibility alias using Notation (abbreviation)
(see the example below).
-
Attribute
deprecated ( since = string ,? note = string? )
¶ At least one of
since
ornote
must be present. If both are present, either one may appear first and they must be separated by a comma.This attribute is supported by the following commands:
Ltac
,Tactic Notation
,Notation
,Infix
.It can trigger the following warnings:
-
Warning
Tactic qualid is deprecated since stringsince. stringnote.
¶ -
Warning
Tactic Notation qualid is deprecated since stringsince. stringnote.
¶ -
Warning
Notation string is deprecated since stringsince. stringnote.
¶ qualid
orstring
is the notation,stringsince
is the version number,stringnote
is the note (usually explains the replacement).
-
Warning
Example: Deprecating a tactic.
- #[deprecated(since="0.9", note="Use idtac instead.")] Ltac foo := idtac.
- foo is defined
- Goal True.
- 1 subgoal ============================ True
- Proof.
- now foo.
- Toplevel input, characters 4-7: > now foo. > ^^^ Warning: Tactic foo is deprecated since 0.9. Use idtac instead. [deprecated-tactic,deprecated] No more subgoals.
Example: Introducing a compatibility alias
Let's say your library initially contained:
- Definition foo x := S x.
- foo is defined
and you want to rename foo
into bar
, but you want to avoid breaking
your users' code without advanced notice. To do so, replace the previous
code by the following:
- Definition bar x := S x.
- bar is defined
- #[deprecated(since="1.2", note="Use bar instead.")] Notation foo := bar (only parsing).
Then, the following code still works, but emits a warning:
- Check (foo 0).
- Toplevel input, characters 7-10: > Check (foo 0). > ^^^ Warning: Notation foo is deprecated since 1.2. Use bar instead. [deprecated-syntactic-definition,deprecated] bar 0 : nat