aboutsummaryrefslogtreecommitdiffhomepage
path: root/vernac/topfmt.ml
diff options
context:
space:
mode:
authorGravatar Emilio Jesus Gallego Arias <e+git@x80.org>2016-09-29 16:30:21 +0200
committerGravatar Emilio Jesus Gallego Arias <e+git@x80.org>2017-03-21 15:51:37 +0100
commit5b8bfee9d80e550cd81e326ec134430b2a4797a5 (patch)
tree779195d25a6c706808ef61b3a78700a65140738b /vernac/topfmt.ml
parentf0341076aa60a84177a6b46db0d8d50df220536b (diff)
[pp] Make feedback the only logging mechanism.
Previously to this patch, Coq featured to distinct logging paths: the console legacy one, based on `Pp.std_ppcmds` and Ocaml's `Format` module, and the `Feedback` one, intended to encapsulate message inside a more general, GUI-based feedback protocol. This patch removes the legacy logging path and makes feedback canonical. Thus, the core of Coq has no dependency on console code anymore. Additionally, this patch resolves the duplication of "document" formats present in the same situation. The original console-based printing code relied on an opaque datatype `std_ppcmds`, (mostly a reification of `Format`'s format strings) that could be then rendered to the console. However, the feedback path couldn't reuse this type due to its opaque nature. The first versions just embedded rending of `std_ppcmds` to a string, however in 8.5 a new "rich printing" type, `Richpp.richpp` was introduced. The idea for this type was to be serializable, however it brought several problems: it didn't have proper document manipulation operations, its format was overly verbose and didn't preserve the full layout, and it still relied on `Format` for generation, making client-side rendering difficult. We thus follow the plan outlined in CEP#9, that is to say, we take a public and refactored version of `std_ppcmds` as the canonical "document type", and move feedback to be over there. The toplevel now is implemented as a feedback listener and has ownership of the console. `richpp` is now IDE-specific, and only used for legacy rendering. It could go away in future versions. `std_ppcmds` carries strictly more information and is friendlier to client-side rendering and display control. Thus, the new panorama is: - `Feedback` has become a very module for event dispatching. - `Pp` contains a target-independent box-based document format. It also contains the `Format`-based renderer. - All console access lives in `toplevel`, with console handlers private to coqtop. _NOTE_: After this patch, many printing parameters such as printing width or depth should be set client-side. This works better IMO, clients don't need to notify Coq about resizing anywmore. Indeed, for box-based capable backends such as HTML or LaTeX, the UI can directly render and let the engine perform the word breaking work. _NOTE_: Many messages could benefit from new features of the output format, however we have chosen not to alter them to preserve output. A Future commits will move console tag handling in `Pp_style` to `toplevel/`, where it logically belongs. The only change with regards to printing is that the "Error:" header was added to console output in several different positions, we have removed some of this duplication, now error messages should be a bit more consistent.
Diffstat (limited to 'vernac/topfmt.ml')
-rw-r--r--vernac/topfmt.ml245
1 files changed, 245 insertions, 0 deletions
diff --git a/vernac/topfmt.ml b/vernac/topfmt.ml
new file mode 100644
index 000000000..85981c386
--- /dev/null
+++ b/vernac/topfmt.ml
@@ -0,0 +1,245 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open Feedback
+open Pp
+
+(** Pp control also belongs here as the terminal is private to the toplevel *)
+
+type pp_global_params = {
+ margin : int;
+ max_indent : int;
+ max_depth : int;
+ ellipsis : string }
+
+(* Default parameters of pretty-printing *)
+
+let dflt_gp = {
+ margin = 78;
+ max_indent = 50;
+ max_depth = 50;
+ ellipsis = "..." }
+
+(* A deeper pretty-printer to print proof scripts *)
+
+let deep_gp = {
+ margin = 78;
+ max_indent = 50;
+ max_depth = 10000;
+ ellipsis = "..." }
+
+(* set_gp : Format.formatter -> pp_global_params -> unit
+ * set the parameters of a formatter *)
+
+let set_gp ft gp =
+ Format.pp_set_margin ft gp.margin ;
+ Format.pp_set_max_indent ft gp.max_indent ;
+ Format.pp_set_max_boxes ft gp.max_depth ;
+ Format.pp_set_ellipsis_text ft gp.ellipsis
+
+let set_dflt_gp ft = set_gp ft dflt_gp
+
+let get_gp ft =
+ { margin = Format.pp_get_margin ft ();
+ max_indent = Format.pp_get_max_indent ft ();
+ max_depth = Format.pp_get_max_boxes ft ();
+ ellipsis = Format.pp_get_ellipsis_text ft () }
+
+(* with_fp : 'a pp_formatter_params -> Format.formatter
+ * returns of formatter for given formatter functions *)
+
+let with_fp chan out_function flush_function =
+ let ft = Format.make_formatter out_function flush_function in
+ Format.pp_set_formatter_out_channel ft chan;
+ ft
+
+(* Output on a channel ch *)
+
+let with_output_to ch =
+ let ft = with_fp ch (output_substring ch) (fun () -> flush ch) in
+ set_gp ft deep_gp;
+ ft
+
+let std_ft = ref Format.std_formatter
+let _ = set_dflt_gp !std_ft
+
+let err_ft = ref Format.err_formatter
+let _ = set_gp !err_ft deep_gp
+
+let deep_ft = ref (with_output_to stdout)
+let _ = set_gp !deep_ft deep_gp
+
+(* For parametrization through vernacular *)
+let default = Format.pp_get_max_boxes !std_ft ()
+let default_margin = Format.pp_get_margin !std_ft ()
+
+let get_depth_boxes () = Some (Format.pp_get_max_boxes !std_ft ())
+let set_depth_boxes v =
+ Format.pp_set_max_boxes !std_ft (match v with None -> default | Some v -> v)
+
+let get_margin () = Some (Format.pp_get_margin !std_ft ())
+let set_margin v =
+ let v = match v with None -> default_margin | Some v -> v in
+ Format.pp_set_margin Format.str_formatter v;
+ Format.pp_set_margin !std_ft v;
+ Format.pp_set_margin !deep_ft v;
+ (* Heuristic, based on usage: the column on the right of max_indent
+ column is 20% of width, capped to 30 characters *)
+ let m = max (64 * v / 100) (v-30) in
+ Format.pp_set_max_indent Format.str_formatter m;
+ Format.pp_set_max_indent !std_ft m;
+ Format.pp_set_max_indent !deep_ft m
+
+(** Console display of feedback *)
+
+type logger = ?loc:Loc.t -> level -> std_ppcmds -> unit
+
+let msgnl_with ?pp_tag fmt strm =
+ pp_with ?pp_tag fmt (strm ++ fnl ());
+ Format.pp_print_flush fmt ()
+
+(* XXX: This is really painful! *)
+module Emacs = struct
+
+ (* Special chars for emacs, to detect warnings inside goal output *)
+ let emacs_quote_start = String.make 1 (Char.chr 254)
+ let emacs_quote_end = String.make 1 (Char.chr 255)
+
+ let emacs_quote_err g =
+ hov 0 (str emacs_quote_start ++ g ++ str emacs_quote_end)
+
+ let emacs_quote_info_start = "<infomsg>"
+ let emacs_quote_info_end = "</infomsg>"
+
+ let emacs_quote_info g =
+ hov 0 (str emacs_quote_info_start++ brk(0,0) ++ g ++ brk(0,0) ++ str emacs_quote_info_end)
+
+end
+
+open Emacs
+
+let dbg_str = tag Ppstyle.debug_tag (str "Debug:") ++ spc ()
+let info_str = mt ()
+let warn_str = tag Ppstyle.warning_tag (str "Warning:") ++ spc ()
+let err_str = tag Ppstyle.error_tag (str "Error:") ++ spc ()
+
+let make_body quoter info ?loc s =
+ let loc = Option.cata Pp.pr_loc (Pp.mt ()) loc in
+ quoter (hov 0 (loc ++ info ++ s))
+
+(* Generic logger *)
+let gen_logger ?pp_tag dbg err ?loc level msg = match level with
+ | Debug -> msgnl_with ?pp_tag !std_ft (make_body dbg dbg_str ?loc msg)
+ | Info -> msgnl_with ?pp_tag !std_ft (make_body dbg info_str ?loc msg)
+ (* XXX: What to do with loc here? *)
+ | Notice -> msgnl_with ?pp_tag !std_ft msg
+ | Warning -> Flags.if_warn (fun () ->
+ msgnl_with ?pp_tag !err_ft (make_body err warn_str ?loc msg)) ()
+ | Error -> msgnl_with ?pp_tag !err_ft (make_body err err_str ?loc msg)
+
+(** Standard loggers *)
+
+(* We provide a generic clear_log_backend callback for backends
+ wanting to do clenaup after the print.
+*)
+let std_logger_tag = ref None
+let std_logger_cleanup = ref (fun () -> ())
+
+let std_logger ?loc level msg =
+ gen_logger (fun x -> x) (fun x -> x) ?pp_tag:!std_logger_tag ?loc level msg;
+ !std_logger_cleanup ()
+
+(** Color logging. Moved from pp_style, it may need some more refactoring *)
+
+(** Not thread-safe. We should put a lock somewhere if we print from
+ different threads. Do we? *)
+let make_style_stack () =
+ (** Default tag is to reset everything *)
+ let empty = Terminal.make () in
+ let default_tag = Terminal.({
+ fg_color = Some `DEFAULT;
+ bg_color = Some `DEFAULT;
+ bold = Some false;
+ italic = Some false;
+ underline = Some false;
+ negative = Some false;
+ })
+ in
+ let style_stack = ref [] in
+ let peek () = match !style_stack with
+ | [] -> default_tag (** Anomalous case, but for robustness *)
+ | st :: _ -> st
+ in
+ let push tag =
+ let style = match Ppstyle.get_style_format tag with
+ | None -> empty
+ | Some st -> st
+ in
+ (** Use the merging of the latest tag and the one being currently pushed.
+ This may be useful if for instance the latest tag changes the background and
+ the current one the foreground, so that the two effects are additioned. *)
+ let style = Terminal.merge (peek ()) style in
+ style_stack := style :: !style_stack;
+ Terminal.eval style
+ in
+ let pop _ = match !style_stack with
+ | [] -> (** Something went wrong, we fallback *)
+ Terminal.eval default_tag
+ | _ :: rem -> style_stack := rem;
+ Terminal.eval (peek ())
+ in
+ let clear () = style_stack := [] in
+ push, pop, clear
+
+let init_color_output () =
+ let push_tag, pop_tag, clear_tag = make_style_stack () in
+ std_logger_cleanup := clear_tag;
+ std_logger_tag := Some Ppstyle.to_format;
+ let tag_handler = {
+ Format.mark_open_tag = push_tag;
+ Format.mark_close_tag = pop_tag;
+ Format.print_open_tag = ignore;
+ Format.print_close_tag = ignore;
+ } in
+ Format.pp_set_mark_tags !std_ft true;
+ Format.pp_set_mark_tags !err_ft true;
+ Format.pp_set_formatter_tag_functions !std_ft tag_handler;
+ Format.pp_set_formatter_tag_functions !err_ft tag_handler
+
+(* Rules for emacs:
+ - Debug/info: emacs_quote_info
+ - Warning/Error: emacs_quote_err
+ - Notice: unquoted
+ *)
+let emacs_logger = gen_logger emacs_quote_info emacs_quote_err
+
+(* Output to file, used only in extraction so a candidate for removal *)
+let ft_logger old_logger ft ?loc level mesg =
+ let id x = x in
+ match level with
+ | Debug -> msgnl_with ft (make_body id dbg_str mesg)
+ | Info -> msgnl_with ft (make_body id info_str mesg)
+ | Notice -> msgnl_with ft mesg
+ | Warning -> old_logger ?loc level mesg
+ | Error -> old_logger ?loc level mesg
+
+let with_output_to_file fname func input =
+ (* XXX FIXME: redirect std_ft *)
+ (* let old_logger = !logger in *)
+ let channel = open_out (String.concat "." [fname; "out"]) in
+ (* logger := ft_logger old_logger (Format.formatter_of_out_channel channel); *)
+ try
+ let output = func input in
+ (* logger := old_logger; *)
+ close_out channel;
+ output
+ with reraise ->
+ let reraise = Backtrace.add_backtrace reraise in
+ (* logger := old_logger; *)
+ close_out channel;
+ Exninfo.iraise reraise