aboutsummaryrefslogtreecommitdiffhomepage
path: root/vernac/vernacinterp.ml
diff options
context:
space:
mode:
authorGravatar Matej Košík <matej.kosik@inria.fr>2017-08-04 18:28:31 +0200
committerGravatar Matej Košík <matej.kosik@inria.fr>2017-11-01 08:05:37 +0100
commitf874a292c94b290a31af8e0f707ffd20ab31001b (patch)
treeeb8c4b67597cbf6adaee337717e9a09eaf9325eb /vernac/vernacinterp.ml
parente5659c8ffe735c530a707a61c692a3af21a79a9a (diff)
provide "loc : Loc.t" binding within "VERNAC COMMAND EXTEND" rules
Originally, it was not possible to define a new vernacular command in the following way: VERNAC COMMAND EXTEND Cmd6 CLASSIFIED AS QUERY [ "SomeCmd" ] -> [ Feedback.msg_notice ?loc (Pp.str "some message") ] END because "loc : Loc.t" was not bound. This commit fixes that, i.e. the location of the custom Vernacular command (within *.v file) is made available as "loc" variable bound on the right side of "->" .
Diffstat (limited to 'vernac/vernacinterp.ml')
-rw-r--r--vernac/vernacinterp.ml14
1 files changed, 7 insertions, 7 deletions
diff --git a/vernac/vernacinterp.ml b/vernac/vernacinterp.ml
index 2d9c0fa36..41fee6bd0 100644
--- a/vernac/vernacinterp.ml
+++ b/vernac/vernacinterp.ml
@@ -11,7 +11,7 @@ open Pp
open CErrors
type deprecation = bool
-type vernac_command = Genarg.raw_generic_argument list -> unit -> unit
+type vernac_command = Genarg.raw_generic_argument list -> Loc.t option -> unit
(* Table of vernac entries *)
let vernac_tab =
@@ -49,8 +49,8 @@ let warn_deprecated_command =
(* Interpretation of a vernac command *)
-let call ?locality (opn,converted_args) =
- let loc = ref "Looking up command" in
+let call ?locality ?loc (opn,converted_args) =
+ let phase = ref "Looking up command" in
try
let depr, callback = vinterp_map opn in
let () = if depr then
@@ -62,16 +62,16 @@ let call ?locality (opn,converted_args) =
let pr = pr_sequence pr_gram rules in
warn_deprecated_command pr;
in
- loc:= "Checking arguments";
+ phase := "Checking arguments";
let hunk = callback converted_args in
- loc:= "Executing command";
+ phase := "Executing command";
Locality.LocalityFixme.set locality;
- hunk();
+ hunk loc;
Locality.LocalityFixme.assert_consumed()
with
| Drop -> raise Drop
| reraise ->
let reraise = CErrors.push reraise in
if !Flags.debug then
- Feedback.msg_debug (str"Vernac Interpreter " ++ str !loc);
+ Feedback.msg_debug (str"Vernac Interpreter " ++ str !phase);
iraise reraise