diff options
| author | Pierre-Marie Pédrot | 2020-06-25 23:35:20 +0200 |
|---|---|---|
| committer | Pierre-Marie Pédrot | 2021-01-22 11:55:21 +0100 |
| commit | 0c1af31e9d338a13b7df2b4468b81e76ef182d32 (patch) | |
| tree | 0fed16037e47c30037f621dad0c63b4f6f17b103 /user-contrib/Ltac2/Printf.v | |
| parent | 7d5618684ef17fbb34246f041b3426d42825b85a (diff) | |
Add a type of format strings to Ltac2.
It provides an abstract type of well-typed format strings, a scope to parse
them and a minimal printf-like API.
Diffstat (limited to 'user-contrib/Ltac2/Printf.v')
| -rw-r--r-- | user-contrib/Ltac2/Printf.v | 56 |
1 files changed, 56 insertions, 0 deletions
diff --git a/user-contrib/Ltac2/Printf.v b/user-contrib/Ltac2/Printf.v new file mode 100644 index 0000000000..e2470ed1c3 --- /dev/null +++ b/user-contrib/Ltac2/Printf.v @@ -0,0 +1,56 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * Copyright INRIA, CNRS and contributors *) +(* <O___,, * (see version control and CREDITS file for authors & dates) *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(* * (see LICENSE file for the text of the license) *) +(************************************************************************) + +Require Import Ltac2.Message. + +(** This file defines a printf notation for easiness of writing messages *) + +(** + + The built-in "format" notation scope can be used to create well-typed variadic + printing commands following a printf-like syntax. The "format" scope parses + quoted strings which contain either raw string data or printing + specifications. Raw strings will be output verbatim as if they were passed + to Ltac2.Message.of_string. + + Printing specifications are of the form + + << '%' type >> + + where the type value defines which kind of arguments will be accepted and + how they will be printed. They can take the following values. + + - << i >>: takes an argument of type int and behaves as Message.of_int + - << I >>: takes an argument of type ident and behaves as Message.of_ident + - << s >>: takes an argument of type string and behaves as Message.of_string + - << t >>: takes an argument of type constr and behaves as Message.of_constr + - << a >>: takes two arguments << f >> of type << (unit -> 'a -> message) >> + and << x >> of type << 'a >> and behaves as << f () x >> + - << % >>: outputs << % >> verbatim + + TODO: add printing modifiers. + +*) + +Ltac2 printf fmt := Format.kfprintf print fmt. +Ltac2 fprintf fmt := Format.kfprintf (fun x => x) fmt. + +(** The two following notations are made available when this module is imported. + + - printf will parse a format and generate a function taking the + corresponding arguments ant printing the resulting message as per + Message.print. In particular when fully applied it has type unit. + - fprintf behaves similarly but return the message as a value instead of + printing it. + +*) + +Ltac2 Notation "printf" fmt(format) := printf fmt. +Ltac2 Notation "fprintf" fmt(format) := fprintf fmt. |
