aboutsummaryrefslogtreecommitdiffhomepage
path: root/toplevel
diff options
context:
space:
mode:
Diffstat (limited to 'toplevel')
-rw-r--r--toplevel/coqargs.ml15
-rw-r--r--toplevel/coqargs.mli10
-rw-r--r--toplevel/coqinit.ml18
-rw-r--r--toplevel/coqinit.mli12
-rw-r--r--toplevel/coqloop.ml72
-rw-r--r--toplevel/coqloop.mli17
-rw-r--r--toplevel/coqtop.ml95
-rw-r--r--toplevel/coqtop.mli18
-rw-r--r--toplevel/usage.ml10
-rw-r--r--toplevel/usage.mli10
-rw-r--r--toplevel/vernac.ml223
-rw-r--r--toplevel/vernac.mli24
12 files changed, 286 insertions, 238 deletions
diff --git a/toplevel/coqargs.ml b/toplevel/coqargs.ml
index 1fdcd2bd4..a1a07fce8 100644
--- a/toplevel/coqargs.ml
+++ b/toplevel/coqargs.ml
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2018 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * 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) *)
(************************************************************************)
let warning s = Flags.(with_option warn Feedback.msg_warning (Pp.strbrk s))
@@ -146,14 +148,13 @@ let add_vo_require opts d p export =
let add_compat_require opts v =
match v with
- | Flags.V8_5 -> add_vo_require opts "Coq.Compat.Coq85" None (Some false)
| Flags.V8_6 -> add_vo_require opts "Coq.Compat.Coq86" None (Some false)
| Flags.V8_7 -> add_vo_require opts "Coq.Compat.Coq87" None (Some false)
- | Flags.VOld | Flags.Current -> opts
+ | Flags.Current -> opts
let set_batch_mode opts =
Flags.quiet := true;
- System.trust_file_cache := false;
+ System.trust_file_cache := true;
{ opts with batch_mode = true }
let add_compile opts verbose s =
diff --git a/toplevel/coqargs.mli b/toplevel/coqargs.mli
index 8ee1a8f55..de9b6a682 100644
--- a/toplevel/coqargs.mli
+++ b/toplevel/coqargs.mli
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2018 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * 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) *)
(************************************************************************)
type compilation_mode = BuildVo | BuildVio | Vio2Vo
diff --git a/toplevel/coqinit.ml b/toplevel/coqinit.ml
index 8574092b8..96a0bd5ec 100644
--- a/toplevel/coqinit.ml
+++ b/toplevel/coqinit.ml
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * 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) *)
(************************************************************************)
open Util
@@ -21,12 +23,12 @@ let set_debug () =
let rcdefaultname = "coqrc"
-let load_rcfile ~rcfile ~time doc sid =
+let load_rcfile ~rcfile ~time ~state =
try
match rcfile with
| Some rcfile ->
if CUnix.file_readable_p rcfile then
- Vernac.load_vernac ~time ~verbosely:false ~interactive:false ~check:true doc sid rcfile
+ Vernac.load_vernac ~time ~echo:false ~interactive:false ~check:true ~state rcfile
else raise (Sys_error ("Cannot read rcfile: "^ rcfile))
| None ->
try
@@ -37,8 +39,8 @@ let load_rcfile ~rcfile ~time doc sid =
Envars.home ~warn / "."^rcdefaultname^"."^Coq_config.version;
Envars.home ~warn / "."^rcdefaultname
] in
- Vernac.load_vernac ~time ~verbosely:false ~interactive:false ~check:true doc sid inferedrc
- with Not_found -> doc, sid
+ Vernac.load_vernac ~time ~echo:false ~interactive:false ~check:true ~state inferedrc
+ with Not_found -> state
(*
Flags.if_verbose
mSGNL (str ("No coqrc or coqrc."^Coq_config.version^
diff --git a/toplevel/coqinit.mli b/toplevel/coqinit.mli
index 0ceba5b38..71b5523cd 100644
--- a/toplevel/coqinit.mli
+++ b/toplevel/coqinit.mli
@@ -1,16 +1,18 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * 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) *)
(************************************************************************)
(** Initialization. *)
val set_debug : unit -> unit
-val load_rcfile : rcfile:(string option) -> time:bool -> Stm.doc -> Stateid.t -> Stm.doc * Stateid.t
+val load_rcfile : rcfile:(string option) -> time:bool -> state:Vernac.State.t -> Vernac.State.t
val init_ocaml_path : unit -> unit
diff --git a/toplevel/coqloop.ml b/toplevel/coqloop.ml
index aade101a4..a103cfe7f 100644
--- a/toplevel/coqloop.ml
+++ b/toplevel/coqloop.ml
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * 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) *)
(************************************************************************)
open Pp
@@ -258,8 +260,9 @@ let rec discard_to_dot () =
| Stm.End_of_input -> raise Stm.End_of_input
| e when CErrors.noncritical e -> ()
-let read_sentence ~doc sid input =
- try Stm.parse_sentence ~doc sid input
+let read_sentence ~state input =
+ let open Vernac.State in
+ try Stm.parse_sentence ~doc:state.doc state.sid input
with reraise ->
let reraise = CErrors.push reraise in
discard_to_dot ();
@@ -300,19 +303,20 @@ let coqloop_feed (fb : Feedback.feedback) = let open Feedback in
is caught and handled (i.e. not re-raised).
*)
-let do_vernac ~time doc sid =
+let do_vernac ~time ~state =
+ let open Vernac.State in
top_stderr (fnl());
- if !print_emacs then top_stderr (str (top_buffer.prompt doc));
+ if !print_emacs then top_stderr (str (top_buffer.prompt state.doc));
resynch_buffer top_buffer;
try
let input = (top_buffer.tokens, None) in
- Vernac.process_expr ~time doc sid (read_sentence ~doc sid (fst input))
+ Vernac.process_expr ~time ~state (read_sentence ~state (fst input))
with
| Stm.End_of_input | CErrors.Quit ->
top_stderr (fnl ()); raise CErrors.Quit
| CErrors.Drop -> (* Last chance *)
if Mltop.is_ocaml_top() then raise CErrors.Drop
- else (Feedback.msg_warning (str "There is no ML toplevel."); doc, sid)
+ else (Feedback.msg_warning (str "There is no ML toplevel."); state)
(* Exception printing should be done by the feedback listener,
however this is not yet ready so we rely on the exception for
now. *)
@@ -321,7 +325,7 @@ let do_vernac ~time doc sid =
let loc = Loc.get_loc info in
let msg = CErrors.iprint (e, info) in
TopErr.print_error_for_buffer ?loc Feedback.Error msg top_buffer;
- doc, sid
+ state
(** Main coq loop : read vernacular expressions until Drop is entered.
Ctrl-C is handled internally as Sys.Break instead of aborting Coq.
@@ -337,20 +341,50 @@ let loop_flush_all () =
Format.pp_print_flush !Topfmt.std_ft ();
Format.pp_print_flush !Topfmt.err_ft ()
-let rec loop ~time doc =
+let pr_open_cur_subgoals () =
+ try
+ let proof = Proof_global.give_me_the_proof () in
+ Printer.pr_open_subgoals ~proof
+ with Proof_global.NoCurrentProof -> Pp.str ""
+
+(* Goal equality heuristic. *)
+let pequal cmp1 cmp2 (a1,a2) (b1,b2) = cmp1 a1 b1 && cmp2 a2 b2
+let evleq e1 e2 = CList.equal Evar.equal e1 e2
+let cproof p1 p2 =
+ let (a1,a2,a3,a4,_),(b1,b2,b3,b4,_) = Proof.proof p1, Proof.proof p2 in
+ evleq a1 b1 &&
+ CList.equal (pequal evleq evleq) a2 b2 &&
+ CList.equal Evar.equal a3 b3 &&
+ CList.equal Evar.equal a4 b4
+
+let drop_last_doc = ref None
+
+let rec loop ~time ~state =
+ let open Vernac.State in
Sys.catch_break true;
try
- reset_input_buffer doc stdin top_buffer;
+ reset_input_buffer state.doc stdin top_buffer;
(* Be careful to keep this loop tail-recursive *)
- let rec vernac_loop doc sid =
- let ndoc, nsid = do_vernac ~time doc sid in
+ let rec vernac_loop ~state =
+ let nstate = do_vernac ~time ~state in
+ let proof_changed = not (Option.equal cproof nstate.proof state.proof) in
+ let print_goals = not !Flags.quiet &&
+ proof_changed && Proof_global.there_are_pending_proofs () in
+ if print_goals then Feedback.msg_notice (pr_open_cur_subgoals ());
loop_flush_all ();
- vernac_loop ndoc nsid
+ vernac_loop ~state:nstate
(* We recover the current stateid, threading from the caller is
not possible due exceptions. *)
- in vernac_loop doc (Stm.get_current_state ~doc)
+ in vernac_loop ~state
with
- | CErrors.Drop -> doc
+ | CErrors.Drop ->
+ (* Due to using exceptions as a form of control, state here goes
+ out of sync as [do_vernac] will never return. We must thus do
+ this hack until we make `Drop` a toplevel-only command. See
+ bug #6872. *)
+ let state = { state with sid = Stm.get_current_state ~doc:state.doc } in
+ drop_last_doc := Some state;
+ state
| CErrors.Quit -> exit 0
| any ->
top_stderr (str "Anomaly: main loop exited with exception: " ++
@@ -358,4 +392,4 @@ let rec loop ~time doc =
fnl() ++
str"Please report" ++
strbrk" at " ++ str Coq_config.wwwbugtracker ++ str ".");
- loop ~time doc
+ loop ~time ~state
diff --git a/toplevel/coqloop.mli b/toplevel/coqloop.mli
index 09240ec82..bbb9b1383 100644
--- a/toplevel/coqloop.mli
+++ b/toplevel/coqloop.mli
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * 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) *)
(************************************************************************)
(** The Coq toplevel loop. *)
@@ -31,9 +33,10 @@ val set_prompt : (unit -> string) -> unit
val coqloop_feed : Feedback.feedback -> unit
(** Parse and execute one vernac command. *)
-
-val do_vernac : time:bool -> Stm.doc -> Stateid.t -> Stm.doc * Stateid.t
+val do_vernac : time:bool -> state:Vernac.State.t -> Vernac.State.t
(** Main entry point of Coq: read and execute vernac commands. *)
+val loop : time:bool -> state:Vernac.State.t -> Vernac.State.t
-val loop : time:bool -> Stm.doc -> Stm.doc
+(** Last document seen after `Drop` *)
+val drop_last_doc : Vernac.State.t option ref
diff --git a/toplevel/coqtop.ml b/toplevel/coqtop.ml
index 9058f0846..341888d09 100644
--- a/toplevel/coqtop.ml
+++ b/toplevel/coqtop.ml
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * 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) *)
(************************************************************************)
open Pp
@@ -34,19 +36,17 @@ let warning s = Flags.(with_option warn Feedback.msg_warning (strbrk s))
will not be generally be initialized, thus stateid, etc... may be
bogus. For now we just print to the console too *)
let coqtop_init_feed = Coqloop.coqloop_feed
-let drop_last_doc = ref None
(* Default toplevel loop *)
-let console_toploop_run opts doc =
+let console_toploop_run opts ~state =
(* We initialize the console only if we run the toploop_run *)
let tl_feed = Feedback.add_feeder Coqloop.coqloop_feed in
if Dumpglob.dump () then begin
Flags.if_verbose warning "Dumpglob cannot be used in interactive mode.";
Dumpglob.noglob ()
end;
- let doc = Coqloop.loop ~time:opts.time doc in
+ let _ = Coqloop.loop ~time:opts.time ~state in
(* Initialise and launch the Ocaml toplevel *)
- drop_last_doc := Some doc;
Coqinit.init_ocaml_path();
Mltop.ocaml_toploop();
(* We let the feeder in place for users of Drop *)
@@ -90,24 +90,26 @@ let outputstate opts =
(******************************************************************************)
(* Interactive Load File Simulation *)
(******************************************************************************)
-let load_vernacular opts doc sid =
+let load_vernacular opts ~state =
List.fold_left
- (fun (doc,sid) (f_in, verbosely) ->
+ (fun state (f_in, echo) ->
let s = Loadpath.locate_file f_in in
- if !Flags.beautify then
- Flags.with_option Flags.beautify_file (Vernac.load_vernac ~time:opts.time ~verbosely ~interactive:false ~check:true doc sid) f_in
- else
- Vernac.load_vernac ~time:opts.time ~verbosely ~interactive:false ~check:true doc sid s)
- (doc, sid) (List.rev opts.load_vernacular_list)
-
-let load_init_vernaculars opts doc sid =
- let doc, sid = if opts.load_rcfile then
- Coqinit.load_rcfile ~rcfile:opts.rcfile ~time:opts.time doc sid
+ (* Should make the beautify logic clearer *)
+ let load_vernac f = Vernac.load_vernac ~time:opts.time ~echo ~interactive:false ~check:true ~state f in
+ if !Flags.beautify
+ then Flags.with_option Flags.beautify_file load_vernac f_in
+ else load_vernac s
+ ) state (List.rev opts.load_vernacular_list)
+
+let load_init_vernaculars opts ~state =
+ let state = if opts.load_rcfile then
+ Coqinit.load_rcfile ~rcfile:opts.rcfile ~time:opts.time ~state
else begin
Flags.if_verbose Feedback.msg_info (str"Skipping rcfile loading.");
- doc,sid
+ state
end in
- load_vernacular opts doc sid
+
+ load_vernacular opts ~state
(******************************************************************************)
(* Startup LoadPath and Modules *)
@@ -192,7 +194,8 @@ let ensure_exists f =
fatal_error (hov 0 (str "Can't find file" ++ spc () ++ str f))
(* Compile a vernac file *)
-let compile opts ~verbosely ~f_in ~f_out =
+let compile opts ~echo ~f_in ~f_out =
+ let open Vernac.State in
let check_pending_proofs () =
let pfs = Proof_global.get_all_proof_names () in
if not (CList.is_empty pfs) then
@@ -220,16 +223,17 @@ let compile opts ~verbosely ~f_in ~f_out =
iload_path; require_libs; stm_options;
}) in
- let doc, sid = load_init_vernaculars opts doc sid in
- let ldir = Stm.get_ldir ~doc in
+ let state = { doc; sid; proof = None } in
+ let state = load_init_vernaculars opts ~state in
+ let ldir = Stm.get_ldir ~doc:state.doc in
Aux_file.(start_aux_file
~aux_file:(aux_file_name_for long_f_dot_vo)
~v_file:long_f_dot_v);
Dumpglob.start_dump_glob ~vfile:long_f_dot_v ~vofile:long_f_dot_vo;
Dumpglob.dump_string ("F" ^ Names.DirPath.to_string ldir ^ "\n");
let wall_clock1 = Unix.gettimeofday () in
- let doc, _ = Vernac.load_vernac ~time:opts.time ~verbosely ~check:true ~interactive:false doc (Stm.get_current_state ~doc) long_f_dot_v in
- let _doc = Stm.join ~doc in
+ let state = Vernac.load_vernac ~time:opts.time ~echo ~check:true ~interactive:false ~state long_f_dot_v in
+ let _doc = Stm.join ~doc:state.doc in
let wall_clock2 = Unix.gettimeofday () in
check_pending_proofs ();
Library.save_library_to ldir long_f_dot_vo (Global.opaque_tables ());
@@ -266,11 +270,11 @@ let compile opts ~verbosely ~f_in ~f_out =
iload_path; require_libs; stm_options;
}) in
- let doc, sid = load_init_vernaculars opts doc sid in
-
- let ldir = Stm.get_ldir ~doc in
- let doc, _ = Vernac.load_vernac ~time:opts.time ~verbosely ~check:false ~interactive:false doc (Stm.get_current_state ~doc) long_f_dot_v in
- let doc = Stm.finish ~doc in
+ let state = { doc; sid; proof = None } in
+ let state = load_init_vernaculars opts ~state in
+ let ldir = Stm.get_ldir ~doc:state.doc in
+ let state = Vernac.load_vernac ~time:opts.time ~echo ~check:false ~interactive:false ~state long_f_dot_v in
+ let doc = Stm.finish ~doc:state.doc in
check_pending_proofs ();
let _doc = Stm.snapshot_vio ~doc ldir long_f_dot_vio in
Stm.reset_task_queue ()
@@ -284,17 +288,17 @@ let compile opts ~verbosely ~f_in ~f_out =
let univs, proofs = Stm.finish_tasks lfdv univs disch proofs tasks in
Library.save_library_raw lfdv sum lib univs proofs
-let compile opts ~verbosely ~f_in ~f_out =
+let compile opts ~echo ~f_in ~f_out =
ignore(CoqworkmgrApi.get 1);
- compile opts ~verbosely ~f_in ~f_out;
+ compile opts ~echo ~f_in ~f_out;
CoqworkmgrApi.giveback 1
-let compile_file opts (f_in, verbosely) =
+let compile_file opts (f_in, echo) =
if !Flags.beautify then
Flags.with_option Flags.beautify_file
- (fun f_in -> compile opts ~verbosely ~f_in ~f_out:None) f_in
+ (fun f_in -> compile opts ~echo ~f_in ~f_out:None) f_in
else
- compile opts ~verbosely ~f_in ~f_out:None
+ compile opts ~echo ~f_in ~f_out:None
let compile_files opts =
let compile_list = List.rev opts.compile_list in
@@ -465,12 +469,15 @@ let init_toplevel arglist =
let iload_path = build_load_path opts in
let require_libs = require_libs opts in
let stm_options = opts.stm_flags in
- try let doc, sid =
- Stm.(new_doc
- { doc_type = Interactive opts.toplevel_name;
- iload_path; require_libs; stm_options;
- }) in
- Some (load_init_vernaculars opts doc sid), opts
+ try
+ let open Vernac.State in
+ let doc, sid =
+ Stm.(new_doc
+ { doc_type = Interactive opts.toplevel_name;
+ iload_path; require_libs; stm_options;
+ }) in
+ let state = { doc; sid; proof = None } in
+ Some (load_init_vernaculars opts ~state), opts
with any -> flush_all(); fatal_error_exn any
(* Non interactive: we perform a sequence of compilation steps *)
end else begin
@@ -496,8 +503,8 @@ let init_toplevel arglist =
let start () =
match init_toplevel (List.tl (Array.to_list Sys.argv)) with
(* Batch mode *)
- | Some (doc, sid), opts when not opts.batch_mode ->
- !toploop_run opts doc;
+ | Some state, opts when not opts.batch_mode ->
+ !toploop_run opts ~state;
exit 1
| _ , opts ->
flush_all();
diff --git a/toplevel/coqtop.mli b/toplevel/coqtop.mli
index 4d625a03d..056279bbd 100644
--- a/toplevel/coqtop.mli
+++ b/toplevel/coqtop.mli
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * 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) *)
(************************************************************************)
(** The Coq main module. The following function [start] will parse the
@@ -11,14 +13,10 @@
state, load the files given on the command line, load the resource file,
produce the output state if any, and finally will launch [Coqloop.loop]. *)
-val init_toplevel : string list -> (Stm.doc * Stateid.t) option * Coqargs.coq_cmdopts
+val init_toplevel : string list -> Vernac.State.t option * Coqargs.coq_cmdopts
val start : unit -> unit
-(* Last document seen after `Drop` *)
-val drop_last_doc : Stm.doc option ref
-
(* For other toploops *)
val toploop_init : (Coqargs.coq_cmdopts -> string list -> string list) ref
-val toploop_run : (Coqargs.coq_cmdopts -> Stm.doc -> unit) ref
-
+val toploop_run : (Coqargs.coq_cmdopts -> state:Vernac.State.t -> unit) ref
diff --git a/toplevel/usage.ml b/toplevel/usage.ml
index e28637f2e..504ffa521 100644
--- a/toplevel/usage.ml
+++ b/toplevel/usage.ml
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * 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) *)
(************************************************************************)
let version ret =
diff --git a/toplevel/usage.mli b/toplevel/usage.mli
index 48b4792de..fbb0117d4 100644
--- a/toplevel/usage.mli
+++ b/toplevel/usage.mli
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * 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) *)
(************************************************************************)
(** {6 Prints the version number on the standard output and exits (with 0). } *)
diff --git a/toplevel/vernac.ml b/toplevel/vernac.ml
index 6b45a94bc..56bdcc7e5 100644
--- a/toplevel/vernac.ml
+++ b/toplevel/vernac.ml
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * 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) *)
(************************************************************************)
(* Parsing of vernacular. *)
@@ -40,37 +42,6 @@ let vernac_echo ?loc in_chan = let open Loc in
Feedback.msg_notice @@ str @@ really_input_string in_chan len
) loc
-(* vernac parses the given stream, executes interpfun on the syntax tree it
- * parses, and is verbose on "primitives" commands if verbosely is true *)
-
-let beautify_suffix = ".beautified"
-
-let set_formatter_translator ch =
- let out s b e = output_substring ch s b e in
- let ft = Format.make_formatter out (fun () -> flush ch) in
- Format.pp_set_max_boxes ft max_int;
- ft
-
-let pr_new_syntax_in_context ?loc ft_beautify ocom =
- let loc = Option.cata Loc.unloc (0,0) loc in
- let fs = States.freeze ~marshallable:`No in
- (* Side-effect: order matters *)
- let before = comment (CLexer.extract_comments (fst loc)) in
- let com = match ocom with
- | Some com -> Ppvernac.pr_vernac com
- | None -> mt() in
- let after = comment (CLexer.extract_comments (snd loc)) in
- if !Flags.beautify_file then
- (Pp.pp_with ft_beautify (hov 0 (before ++ com ++ after));
- Format.pp_print_flush ft_beautify ())
- else
- Feedback.msg_info (hov 4 (str"New Syntax:" ++ fnl() ++ (hov 0 com)));
- States.unfreeze fs
-
-let pr_new_syntax ?loc po ft_beautify ocom =
- (* Reinstall the context of parsing which includes the bindings of comments to locations *)
- Pcoq.Gram.with_parsable po (pr_new_syntax_in_context ?loc ft_beautify) ocom
-
(* For coqtop -time, we display the position in the file,
and a glimpse of the executed command *)
@@ -99,42 +70,29 @@ let print_cmd_header ?loc com =
Pp.pp_with !Topfmt.std_ft (pp_cmd_header ?loc com);
Format.pp_print_flush !Topfmt.std_ft ()
-let pr_open_cur_subgoals () =
- try
- let proof = Proof_global.give_me_the_proof () in
- Printer.pr_open_subgoals ~proof
- with Proof_global.NoCurrentProof -> Pp.str ""
-
(* Reenable when we get back to feedback printing *)
(* let is_end_of_input any = match any with *)
(* Stm.End_of_input -> true *)
(* | _ -> false *)
-let rec interp_vernac ~time ~check ~interactive doc sid (loc,com) =
- let interp v =
- match under_control v with
- | VernacLoad (verbosely, fname) ->
- let fname = Envars.expand_path_macros ~warn:(fun x -> Feedback.msg_warning (str x)) fname in
- let fname = CUnix.make_suffix fname ".v" in
- let f = Loadpath.locate_file fname in
- load_vernac ~time ~verbosely ~check ~interactive doc sid f
- | _ ->
- (* XXX: We need to run this before add as the classification is
- highly dynamic and depends on the structure of the
- document. Hopefully this is fixed when VtMeta can be removed
- and Undo etc... are just interpreted regularly. *)
-
- (* XXX: The classifier can emit warnings so we need to guard
- against that... *)
- let wflags = CWarnings.get_flags () in
- CWarnings.set_flags "none";
- let is_proof_step = match fst (Vernac_classifier.classify_vernac v) with
- | VtProofStep _ | VtMeta | VtStartProof _ -> true
- | _ -> false
- in
- CWarnings.set_flags wflags;
+module State = struct
+
+ type t = {
+ doc : Stm.doc;
+ sid : Stateid.t;
+ proof : Proof.t option;
+ }
- let doc, nsid, ntip = Stm.add ~doc ~ontop:sid (not !Flags.quiet) (loc,v) in
+end
+
+let interp_vernac ~time ~check ~interactive ~state (loc,com) =
+ let open State in
+ try
+ (* The -time option is only supported from console-based clients
+ due to the way it prints. *)
+ if time then print_cmd_header ?loc com;
+ let com = if time then VernacTime(time,(CAst.make ?loc com)) else com in
+ let doc, nsid, ntip = Stm.add ~doc:state.doc ~ontop:state.sid (not !Flags.quiet) (loc,com) in
(* Main STM interaction *)
if ntip <> `NewTip then
@@ -144,26 +102,12 @@ let rec interp_vernac ~time ~check ~interactive doc sid (loc,com) =
it otherwise reveals bugs *)
(* Stm.observe nsid; *)
let ndoc = if check then Stm.finish ~doc else doc in
-
- (* We could use a more refined criteria that depends on the
- vernac. For now we imitate the old approach and rely on the
- classification. *)
- let print_goals = interactive && not !Flags.quiet &&
- is_proof_step && Proof_global.there_are_pending_proofs () in
-
- if print_goals then Feedback.msg_notice (pr_open_cur_subgoals ());
- ndoc, nsid
- in
- try
- (* The -time option is only supported from console-based
- clients due to the way it prints. *)
- if time then print_cmd_header ?loc com;
- let com = if time then VernacTime(time,(loc,com)) else com in
- interp com
+ let new_proof = Proof_global.give_me_the_proof_opt () in
+ { doc = ndoc; sid = nsid; proof = new_proof }
with reraise ->
(* XXX: In non-interactive mode edit_at seems to do very weird
things, so we better avoid it while we investigate *)
- if interactive then ignore(Stm.edit_at ~doc sid);
+ if interactive then ignore(Stm.edit_at ~doc:state.doc state.sid);
let (reraise, info) = CErrors.push reraise in
let info = begin
match Loc.get_loc info with
@@ -172,25 +116,23 @@ let rec interp_vernac ~time ~check ~interactive doc sid (loc,com) =
end in iraise (reraise, info)
(* Load a vernac file. CErrors are annotated with file and location *)
-and load_vernac ~time ~verbosely ~check ~interactive doc sid file =
- let ft_beautify, close_beautify =
- if !Flags.beautify_file then
- let chan_beautify = open_out (file^beautify_suffix) in
- set_formatter_translator chan_beautify, fun () -> close_out chan_beautify;
- else
- !Topfmt.std_ft, fun () -> ()
- in
+let load_vernac_core ~time ~echo ~check ~interactive ~state file =
+ (* Keep in sync *)
let in_chan = open_utf8_file_in file in
- let in_echo = if verbosely then Some (open_utf8_file_in file) else None in
+ let in_echo = if echo then Some (open_utf8_file_in file) else None in
+ let input_cleanup () = close_in in_chan; Option.iter close_in in_echo in
+
let in_pa = Pcoq.Gram.parsable ~file:(Loc.InFile file) (Stream.of_channel in_chan) in
- let rsid = ref sid in
- let rdoc = ref doc in
+ let rstate = ref state in
+ (* For beautify, list of parsed sids *)
+ let rids = ref [] in
+ let open State in
try
(* we go out of the following infinite loop when a End_of_input is
* raised, which means that we raised the end of the file being loaded *)
while true do
let loc, ast =
- Stm.parse_sentence ~doc:!rdoc !rsid in_pa
+ Stm.parse_sentence ~doc:!rstate.doc !rstate.sid in_pa
(* If an error in parsing occurs, we propagate the exception
so the caller of load_vernac will take care of it. However,
in the future it could be possible that we want to handle
@@ -210,37 +152,78 @@ and load_vernac ~time ~verbosely ~check ~interactive doc sid file =
*)
in
(* Printing of vernacs *)
- if !Flags.beautify then pr_new_syntax ?loc in_pa ft_beautify (Some ast);
Option.iter (vernac_echo ?loc) in_echo;
checknav_simple (loc, ast);
- let ndoc, nsid = Flags.silently (interp_vernac ~time ~check ~interactive !rdoc !rsid) (loc, ast) in
- rsid := nsid;
- rdoc := ndoc
+ let state = Flags.silently (interp_vernac ~time ~check ~interactive ~state:!rstate) (loc, ast) in
+ rids := state.sid :: !rids;
+ rstate := state;
done;
- !rdoc, !rsid
+ input_cleanup ();
+ !rstate, !rids, Pcoq.Gram.comment_state in_pa
with any -> (* whatever the exception *)
let (e, info) = CErrors.push any in
- close_in in_chan;
- Option.iter close_in in_echo;
+ input_cleanup ();
match e with
- | Stm.End_of_input ->
- (* Is this called so comments at EOF are printed? *)
- if !Flags.beautify then
- pr_new_syntax ~loc:(Loc.make_loc (max_int,max_int)) in_pa ft_beautify None;
- if !Flags.beautify_file then close_beautify ();
- !rdoc, !rsid
- | reraise ->
- if !Flags.beautify_file then close_beautify ();
- iraise (disable_drop e, info)
-
-(** [eval_expr : ?preserving:bool -> Loc.t * Vernacexpr.vernac_expr -> unit]
- It executes one vernacular command. By default the command is
- considered as non-state-preserving, in which case we add it to the
- Backtrack stack (triggering a save of a frozen state and the generation
- of a new state label). An example of state-preserving command is one coming
- from the query panel of Coqide. *)
-
-let process_expr ~time doc sid loc_ast =
+ | Stm.End_of_input -> !rstate, !rids, Pcoq.Gram.comment_state in_pa
+ | reraise -> iraise (disable_drop e, info)
+
+let process_expr ~time ~state loc_ast =
checknav_deep loc_ast;
- interp_vernac ~time ~interactive:true ~check:true doc sid loc_ast
+ interp_vernac ~time ~interactive:true ~check:true ~state loc_ast
+
+(******************************************************************************)
+(* Beautify-specific code *)
+(******************************************************************************)
+
+(* vernac parses the given stream, executes interpfun on the syntax tree it
+ * parses, and is verbose on "primitives" commands if verbosely is true *)
+let beautify_suffix = ".beautified"
+
+let set_formatter_translator ch =
+ let out s b e = output_substring ch s b e in
+ let ft = Format.make_formatter out (fun () -> flush ch) in
+ Format.pp_set_max_boxes ft max_int;
+ ft
+
+let pr_new_syntax ?loc ft_beautify ocom =
+ let loc = Option.cata Loc.unloc (0,0) loc in
+ let before = comment (Pputils.extract_comments (fst loc)) in
+ let com = Option.cata Ppvernac.pr_vernac (mt ()) ocom in
+ let after = comment (Pputils.extract_comments (snd loc)) in
+ if !Flags.beautify_file then
+ (Pp.pp_with ft_beautify (hov 0 (before ++ com ++ after));
+ Format.pp_print_flush ft_beautify ())
+ else
+ Feedback.msg_info (hov 4 (str"New Syntax:" ++ fnl() ++ (hov 0 com)))
+
+(* load_vernac with beautify *)
+let beautify_pass ~doc ~comments ~ids ~filename =
+ let ft_beautify, close_beautify =
+ if !Flags.beautify_file then
+ let chan_beautify = open_out (filename^beautify_suffix) in
+ set_formatter_translator chan_beautify, fun () -> close_out chan_beautify;
+ else
+ !Topfmt.std_ft, fun () -> ()
+ in
+ (* The interface to the comment printer is imperative, so we first
+ set the comments, then we call print. This has to be done for
+ each file. *)
+ Pputils.beautify_comments := comments;
+ List.iter (fun id ->
+ Option.iter (fun (loc,ast) ->
+ pr_new_syntax ?loc ft_beautify (Some ast))
+ (Stm.get_ast ~doc id)) ids;
+
+ (* Is this called so comments at EOF are printed? *)
+ pr_new_syntax ~loc:(Loc.make_loc (max_int,max_int)) ft_beautify None;
+ close_beautify ()
+
+(* Main driver for file loading. For now, we only do one beautify
+ pass. *)
+let load_vernac ~time ~echo ~check ~interactive ~state filename =
+ let ostate, ids, comments = load_vernac_core ~time ~echo ~check ~interactive ~state filename in
+ (* Pass for beautify *)
+ if !Flags.beautify then beautify_pass ~doc:ostate.State.doc ~comments ~ids:List.(rev ids) ~filename;
+ (* End pass *)
+ ostate
diff --git a/toplevel/vernac.mli b/toplevel/vernac.mli
index b77b024fa..19bac45c3 100644
--- a/toplevel/vernac.mli
+++ b/toplevel/vernac.mli
@@ -1,20 +1,32 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * 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) *)
(************************************************************************)
(** Parsing of vernacular. *)
+module State : sig
+
+ type t = {
+ doc : Stm.doc;
+ sid : Stateid.t;
+ proof : Proof.t option;
+ }
+
+end
(** [process_expr sid cmd] Executes vernac command [cmd]. Callers are
expected to handle and print errors in form of exceptions, however
care is taken so the state machine is left in a consistent
state. *)
-val process_expr : time:bool -> Stm.doc -> Stateid.t -> Vernacexpr.vernac_control Loc.located -> Stm.doc * Stateid.t
+val process_expr : time:bool -> state:State.t -> Vernacexpr.vernac_control Loc.located -> State.t
(** [load_vernac echo sid file] Loads [file] on top of [sid], will
echo the commands if [echo] is set. Callers are expected to handle
and print errors in form of exceptions. *)
-val load_vernac : time:bool -> verbosely:bool -> check:bool -> interactive:bool -> Stm.doc -> Stateid.t -> string -> Stm.doc * Stateid.t
+val load_vernac : time:bool -> echo:bool -> check:bool -> interactive:bool ->
+ state:State.t -> string -> State.t