summaryrefslogtreecommitdiff
path: root/toplevel/vernacentries.ml
diff options
context:
space:
mode:
Diffstat (limited to 'toplevel/vernacentries.ml')
-rw-r--r--toplevel/vernacentries.ml861
1 files changed, 472 insertions, 389 deletions
diff --git a/toplevel/vernacentries.ml b/toplevel/vernacentries.ml
index 72dd967b..41ee165f 100644
--- a/toplevel/vernacentries.ml
+++ b/toplevel/vernacentries.ml
@@ -9,7 +9,7 @@
(* Concrete syntax of the mathematical vernacular MV V2.6 *)
open Pp
-open Errors
+open CErrors
open Util
open Flags
open Names
@@ -20,7 +20,6 @@ open Tacmach
open Constrintern
open Prettyp
open Printer
-open Tacinterp
open Command
open Goptions
open Libnames
@@ -32,10 +31,14 @@ open Redexpr
open Lemmas
open Misctypes
open Locality
+open Sigma.Notations
+
+(** TODO: make this function independent of Ltac *)
+let (f_interp_redexp, interp_redexp_hook) = Hook.make ()
let debug = false
-let prerr_endline =
- if debug then prerr_endline else fun _ -> ()
+let prerr_endline x =
+ if debug then prerr_endline (x ()) else ()
(* Misc *)
@@ -45,7 +48,7 @@ let cl_of_qualid = function
| RefClass r -> Class.class_of_global (Smartlocate.smart_global ~head:true r)
let scope_class_of_qualid qid =
- Notation.scope_class_of_reference (Smartlocate.smart_global qid)
+ Notation.scope_class_of_class (cl_of_qualid qid)
(*******************)
(* "Show" commands *)
@@ -54,7 +57,7 @@ let show_proof () =
(* spiwack: this would probably be cooler with a bit of polishing. *)
let p = Proof_global.give_me_the_proof () in
let pprf = Proof.partial_proof p in
- msg_notice (Pp.prlist_with_sep Pp.fnl Printer.pr_constr pprf)
+ Feedback.msg_notice (Pp.prlist_with_sep Pp.fnl Printer.pr_constr pprf)
let show_node () =
(* spiwack: I'm have little clue what this function used to do. I deactivated it,
@@ -62,22 +65,22 @@ let show_node () =
()
let show_thesis () =
- msg_error (anomaly (Pp.str "TODO") )
+ Feedback.msg_error (anomaly (Pp.str "TODO") )
let show_top_evars () =
(* spiwack: new as of Feb. 2010: shows goal evars in addition to non-goal evars. *)
let pfts = get_pftreestate () in
let gls = Proof.V82.subgoals pfts in
let sigma = gls.Evd.sigma in
- msg_notice (pr_evars_int sigma 1 (Evarutil.non_instantiated sigma))
+ Feedback.msg_notice (pr_evars_int sigma 1 (Evarutil.non_instantiated sigma))
let show_universes () =
let pfts = get_pftreestate () in
let gls = Proof.V82.subgoals pfts in
let sigma = gls.Evd.sigma in
let ctx = Evd.universe_context_set (Evd.nf_constraints sigma) in
- msg_notice (Evd.pr_evar_universe_context (Evd.evar_universe_context sigma));
- msg_notice (str"Normalized constraints: " ++ Univ.pr_universe_context_set (Evd.pr_evd_level sigma) ctx)
+ Feedback.msg_notice (Evd.pr_evar_universe_context (Evd.evar_universe_context sigma));
+ Feedback.msg_notice (str"Normalized constraints: " ++ Univ.pr_universe_context_set (Evd.pr_evd_level sigma) ctx)
let show_prooftree () =
(* Spiwack: proof tree is currently not working *)
@@ -88,11 +91,10 @@ let enable_goal_printing = ref true
let print_subgoals () =
if !enable_goal_printing && is_verbose ()
then begin
- msg_notice (pr_open_subgoals ())
+ Feedback.msg_notice (pr_open_subgoals ())
end
let try_print_subgoals () =
- Pp.flush_all();
try print_subgoals () with Proof_global.NoCurrentProof | UserError _ -> ()
@@ -106,10 +108,10 @@ let show_intro all =
let l,_= decompose_prod_assum (strip_outer_cast (pf_concl gl)) in
if all then
let lid = Tactics.find_intro_names l gl in
- msg_notice (hov 0 (prlist_with_sep spc pr_id lid))
+ Feedback.msg_notice (hov 0 (prlist_with_sep spc pr_id lid))
else if not (List.is_empty l) then
let n = List.last l in
- msg_notice (pr_id (List.hd (Tactics.find_intro_names [n] gl)))
+ Feedback.msg_notice (pr_id (List.hd (Tactics.find_intro_names [n] gl)))
end
(** Prepare a "match" template for a given inductive type.
@@ -118,9 +120,7 @@ let show_intro all =
[Not_found] is raised if the given string isn't the qualid of
a known inductive type. *)
-let make_cases s =
- let qualified_name = Libnames.qualid_of_string s in
- let glob_ref = Nametab.locate qualified_name in
+let make_cases_aux glob_ref =
match glob_ref with
| Globnames.IndRef i ->
let {Declarations.mind_nparams = np}
@@ -133,30 +133,35 @@ let make_cases s =
let rec rename avoid = function
| [] -> []
| (n,_)::l ->
- let n' = Namegen.next_name_away_in_cases_pattern ([],mkMeta 0) n avoid in
+ let n' = Namegen.next_name_away_with_default (Id.to_string Namegen.default_dependent_ident) n avoid in
Id.to_string n' :: rename (n'::avoid) l in
let al' = rename [] al in
(Id.to_string consname :: al') :: l)
carr tarr []
| _ -> raise Not_found
+let make_cases s =
+ let qualified_name = Libnames.qualid_of_string s in
+ let glob_ref = Nametab.locate qualified_name in
+ make_cases_aux glob_ref
+
(** Textual display of a generic "match" template *)
let show_match id =
let patterns =
- try make_cases (Id.to_string (snd id))
+ try make_cases_aux (Nametab.global id)
with Not_found -> error "Unknown inductive type."
in
let pr_branch l =
str "| " ++ hov 1 (prlist_with_sep spc str l) ++ str " =>"
in
- msg_notice (v 1 (str "match # with" ++ fnl () ++
+ Feedback.msg_notice (v 1 (str "match # with" ++ fnl () ++
prlist_with_sep fnl pr_branch patterns ++ fnl () ++ str "end" ++ fnl ()))
(* "Print" commands *)
let print_path_entry p =
- let dir = str (DirPath.to_string (Loadpath.logical p)) in
+ let dir = pr_dirpath (Loadpath.logical p) in
let path = str (Loadpath.physical p) in
(dir ++ str " " ++ tbrk (0, 0) ++ path)
@@ -191,23 +196,23 @@ let print_module r =
let globdir = Nametab.locate_dir qid in
match globdir with
DirModule (dirpath,(mp,_)) ->
- msg_notice (Printmod.print_module (Printmod.printable_body dirpath) mp)
+ Feedback.msg_notice (Printmod.print_module (Printmod.printable_body dirpath) mp)
| _ -> raise Not_found
with
- Not_found -> msg_error (str"Unknown Module " ++ pr_qualid qid)
+ Not_found -> Feedback.msg_error (str"Unknown Module " ++ pr_qualid qid)
let print_modtype r =
let (loc,qid) = qualid_of_reference r in
try
let kn = Nametab.locate_modtype qid in
- msg_notice (Printmod.print_modtype kn)
+ Feedback.msg_notice (Printmod.print_modtype kn)
with Not_found ->
(* Is there a module of this name ? If yes we display its type *)
try
let mp = Nametab.locate_module qid in
- msg_notice (Printmod.print_module false mp)
+ Feedback.msg_notice (Printmod.print_module false mp)
with Not_found ->
- msg_error (str"Unknown Module Type or Module " ++ pr_qualid qid)
+ Feedback.msg_error (str"Unknown Module Type or Module " ++ pr_qualid qid)
let print_namespace ns =
let ns = List.rev (Names.DirPath.repr ns) in
@@ -276,7 +281,7 @@ let print_namespace ns =
acc
) constants (str"")
in
- msg_notice ((print_list pr_id ns)++str":"++fnl()++constants_in_namespace)
+ Feedback.msg_notice ((print_list pr_id ns)++str":"++fnl()++constants_in_namespace)
let print_strategy r =
let open Conv_oracle in
@@ -306,7 +311,7 @@ let print_strategy r =
else str "Constant strategies" ++ fnl () ++
hov 0 (prlist_with_sep fnl pr_strategy cst_lvl)
in
- msg_notice (var_msg ++ cst_msg)
+ Feedback.msg_notice (var_msg ++ cst_msg)
| Some r ->
let r = Smartlocate.smart_global r in
let key = match r with
@@ -315,7 +320,7 @@ let print_strategy r =
| IndRef _ | ConstructRef _ -> error "The reference is not unfoldable"
in
let lvl = get_strategy oracle key in
- msg_notice (pr_strategy (r, lvl))
+ Feedback.msg_notice (pr_strategy (r, lvl))
let dump_universes_gen g s =
let output = open_out s in
@@ -333,7 +338,7 @@ let dump_universes_gen g s =
| Univ.Eq ->
Printf.fprintf output " \"%s\" -> \"%s\" [style=dashed];\n" left right
end, begin fun () ->
- if Lazy.lazy_is_val init then Printf.fprintf output "}\n";
+ if Lazy.is_val init then Printf.fprintf output "}\n";
close_out output
end
end else begin
@@ -347,11 +352,11 @@ let dump_universes_gen g s =
end
in
try
- Univ.dump_universes output_constraint g;
+ UGraph.dump_universes output_constraint g;
close ();
- msg_info (str "Universes written to file \"" ++ str s ++ str "\".")
+ Feedback.msg_info (str "Universes written to file \"" ++ str s ++ str "\".")
with reraise ->
- let reraise = Errors.push reraise in
+ let reraise = CErrors.push reraise in
close ();
iraise reraise
@@ -364,11 +369,11 @@ let locate_file f =
let msg_found_library = function
| Library.LibLoaded, fulldir, file ->
- msg_info (hov 0
+ Feedback.msg_info (hov 0
(pr_dirpath fulldir ++ strbrk " has been loaded from file " ++
str file))
| Library.LibInPath, fulldir, file ->
- msg_info (hov 0
+ Feedback.msg_info (hov 0
(pr_dirpath fulldir ++ strbrk " is bound to file " ++ str file))
let err_unmapped_library loc ?from qid =
@@ -409,7 +414,7 @@ let dump_global r =
try
let gr = Smartlocate.smart_global r in
Dumpglob.add_glob (Constrarg.loc_of_or_by_notation loc_of_reference r) gr
- with e when Errors.noncritical e -> ()
+ with e when CErrors.noncritical e -> ()
(**********)
(* Syntax *)
@@ -443,7 +448,27 @@ let vernac_notation locality local =
(***********)
(* Gallina *)
-let start_proof_and_print k l hook = start_proof_com k l hook
+let start_proof_and_print k l hook =
+ let inference_hook =
+ if Flags.is_program_mode () then
+ let hook env sigma ev =
+ let tac = !Obligations.default_tactic in
+ let evi = Evd.find sigma ev in
+ let env = Evd.evar_filtered_env evi in
+ try
+ let concl = Evarutil.nf_evars_universes sigma evi.Evd.evar_concl in
+ if Evarutil.has_undefined_evars sigma concl then raise Exit;
+ let c, _, ctx =
+ Pfedit.build_by_tactic env (Evd.evar_universe_context sigma)
+ concl (Tacticals.New.tclCOMPLETE tac)
+ in Evd.set_universe_context sigma ctx, c
+ with Logic_monad.TacticFailure e when Logic.catchable_exception e ->
+ error "The statement obligations could not be resolved \
+ automatically, write a statement definition first."
+ in Some hook
+ else None
+ in
+ start_proof_com ?inference_hook k l hook
let no_hook = Lemmas.mk_hook (fun _ _ -> ())
@@ -463,14 +488,14 @@ let vernac_definition locality p (local,k) ((loc,id as lid),pl) def =
in
(match def with
| ProveBody (bl,t) -> (* local binders, typ *)
- start_proof_and_print (local,p,DefinitionBody Definition)
- [Some (lid,pl), (bl,t,None)] no_hook
+ start_proof_and_print (local,p,DefinitionBody k)
+ [Some (lid,pl), (bl,t,None)] hook
| DefineBody (bl,red_option,c,typ_opt) ->
let red_option = match red_option with
| None -> None
| Some r ->
let (evc,env)= get_current_context () in
- Some (snd (interp_redexp env evc r)) in
+ Some (snd (Hook.get f_interp_redexp env evc r)) in
do_definition id (local,p,k) pl bl red_option c typ_opt hook)
let vernac_start_proof locality p kind l lettop =
@@ -501,9 +526,9 @@ let vernac_end_proof ?proof = function
let vernac_exact_proof c =
(* spiwack: for simplicity I do not enforce that "Proof proof_term" is
called only at the begining of a proof. *)
- let status = by (Tactics.New.exact_proof c) in
+ let status = by (Tactics.exact_proof c) in
save_proof (Vernacexpr.(Proved(Opaque None,None)));
- if not status then Pp.feedback Feedback.AddedAxiom
+ if not status then Feedback.feedback Feedback.AddedAxiom
let vernac_assumption locality poly (local, kind) l nl =
let local = enforce_locality_exp locality local in
@@ -515,7 +540,7 @@ let vernac_assumption locality poly (local, kind) l nl =
if global then Dumpglob.dump_definition lid false "ax"
else Dumpglob.dump_definition lid true "var") idl) l;
let status = do_assumptions kind nl l in
- if not status then Pp.feedback Feedback.AddedAxiom
+ if not status then Feedback.feedback Feedback.AddedAxiom
let vernac_record k poly finite struc binders sort nameopt cfs =
let const = match nameopt with
@@ -530,6 +555,10 @@ let vernac_record k poly finite struc binders sort nameopt cfs =
| _ -> ()) cfs);
ignore(Record.definition_structure (k,poly,finite,struc,binders,cfs,const,sort))
+(** When [poly] is true the type is declared polymorphic. When [lo] is true,
+ then the type is declared private (as per the [Private] keyword). [finite]
+ indicates whether the type is inductive, co-inductive or
+ neither. *)
let vernac_inductive poly lo finite indl =
if Dumpglob.dump () then
List.iter (fun (((coe,(lid,_)), _, _, _, cstrs), _) ->
@@ -542,29 +571,29 @@ let vernac_inductive poly lo finite indl =
indl;
match indl with
| [ ( _ , _ , _ ,Record, Constructors _ ),_ ] ->
- Errors.error "The Record keyword cannot be used to define a variant type. Use Variant instead."
+ CErrors.error "The Record keyword cannot be used to define a variant type. Use Variant instead."
| [ (_ , _ , _ ,Variant, RecordDecl _),_ ] ->
- Errors.error "The Variant keyword cannot be used to define a record type. Use Record instead."
+ CErrors.error "The Variant keyword cannot be used to define a record type. Use Record instead."
| [ ( id , bl , c , b, RecordDecl (oc,fs) ), [] ] ->
- vernac_record (match b with Class true -> Class false | _ -> b)
+ vernac_record (match b with Class _ -> Class false | _ -> b)
poly finite id bl c oc fs
- | [ ( id , bl , c , Class true, Constructors [l]), _ ] ->
+ | [ ( id , bl , c , Class _, Constructors [l]), [] ] ->
let f =
let (coe, ((loc, id), ce)) = l in
let coe' = if coe then Some true else None in
(((coe', AssumExpr ((loc, Name id), ce)), None), [])
in vernac_record (Class true) poly finite id bl c None [f]
- | [ ( id , bl , c , Class true, _), _ ] ->
- Errors.error "Definitional classes must have a single method"
- | [ ( id , bl , c , Class false, Constructors _), _ ] ->
- Errors.error "Inductive classes not supported"
+ | [ ( _ , _, _, Class _, Constructors _), [] ] ->
+ CErrors.error "Inductive classes not supported"
+ | [ ( id , bl , c , Class _, _), _ :: _ ] ->
+ CErrors.error "where clause not supported for classes"
| [ ( _ , _ , _ , _, RecordDecl _ ) , _ ] ->
- Errors.error "where clause not supported for (co)inductive records"
+ CErrors.error "where clause not supported for (co)inductive records"
| _ -> let unpack = function
| ( (false, id) , bl , c , _ , Constructors l ) , ntn -> ( id , bl , c , l ) , ntn
| ( (true,_),_,_,_,Constructors _),_ ->
- Errors.error "Variant types do not handle the \"> Name\" syntax, which is reserved for records. Use the \":>\" syntax on constructors instead."
- | _ -> Errors.error "Cannot handle mutually (co)inductive records."
+ CErrors.error "Variant types do not handle the \"> Name\" syntax, which is reserved for records. Use the \":>\" syntax on constructors instead."
+ | _ -> CErrors.error "Cannot handle mutually (co)inductive records."
in
let indl = List.map unpack indl in
do_mutual_inductive indl poly lo finite
@@ -633,7 +662,7 @@ let vernac_declare_module export (loc, id) binders_ast mty_ast =
id binders_ast (Enforce mty_ast) []
in
Dumpglob.dump_moddef loc mp "mod";
- if_verbose msg_info (str "Module " ++ pr_id id ++ str " is declared");
+ if_verbose Feedback.msg_info (str "Module " ++ pr_id id ++ str " is declared");
Option.iter (fun export -> vernac_import export [Ident (Loc.ghost,id)]) export
let vernac_define_module export (loc, id) binders_ast mty_ast_o mexpr_ast_l =
@@ -654,7 +683,7 @@ let vernac_define_module export (loc, id) binders_ast mty_ast_o mexpr_ast_l =
export id binders_ast mty_ast_o
in
Dumpglob.dump_moddef loc mp "mod";
- if_verbose msg_info
+ if_verbose Feedback.msg_info
(str "Interactive Module " ++ pr_id id ++ str " started");
List.iter
(fun (export,id) ->
@@ -672,7 +701,7 @@ let vernac_define_module export (loc, id) binders_ast mty_ast_o mexpr_ast_l =
id binders_ast mty_ast_o mexpr_ast_l
in
Dumpglob.dump_moddef loc mp "mod";
- if_verbose msg_info
+ if_verbose Feedback.msg_info
(str "Module " ++ pr_id id ++ str " is defined");
Option.iter (fun export -> vernac_import export [Ident (Loc.ghost,id)])
export
@@ -680,7 +709,7 @@ let vernac_define_module export (loc, id) binders_ast mty_ast_o mexpr_ast_l =
let vernac_end_module export (loc,id as lid) =
let mp = Declaremods.end_module () in
Dumpglob.dump_modref loc mp "mod";
- if_verbose msg_info (str "Module " ++ pr_id id ++ str " is defined");
+ if_verbose Feedback.msg_info (str "Module " ++ pr_id id ++ str " is defined");
Option.iter (fun export -> vernac_import export [Ident lid]) export
let vernac_declare_module_type (loc,id) binders_ast mty_sign mty_ast_l =
@@ -701,7 +730,7 @@ let vernac_declare_module_type (loc,id) binders_ast mty_sign mty_ast_l =
id binders_ast mty_sign
in
Dumpglob.dump_moddef loc mp "modtype";
- if_verbose msg_info
+ if_verbose Feedback.msg_info
(str "Interactive Module Type " ++ pr_id id ++ str " started");
List.iter
(fun (export,id) ->
@@ -720,13 +749,13 @@ let vernac_declare_module_type (loc,id) binders_ast mty_sign mty_ast_l =
id binders_ast mty_sign mty_ast_l
in
Dumpglob.dump_moddef loc mp "modtype";
- if_verbose msg_info
+ if_verbose Feedback.msg_info
(str "Module Type " ++ pr_id id ++ str " is defined")
let vernac_end_modtype (loc,id) =
let mp = Declaremods.end_modtype () in
Dumpglob.dump_modref loc mp "modtype";
- if_verbose msg_info (str "Module Type " ++ pr_id id ++ str " is defined")
+ if_verbose Feedback.msg_info (str "Module Type " ++ pr_id id ++ str " is defined")
let vernac_include l =
Declaremods.declare_include Modintern.interp_module_ast l
@@ -794,7 +823,7 @@ let vernac_coercion locality poly local ref qids qidt =
let source = cl_of_qualid qids in
let ref' = smart_global ref in
Class.try_add_new_coercion_with_target ref' ~local poly ~source ~target;
- if_verbose msg_info (pr_global ref' ++ str " is now a coercion")
+ if_verbose Feedback.msg_info (pr_global ref' ++ str " is now a coercion")
let vernac_identity_coercion locality poly local id qids qidt =
let local = enforce_locality locality local in
@@ -810,11 +839,11 @@ let vernac_instance abst locality poly sup inst props pri =
ignore(Classes.new_instance ~abstract:abst ~global poly sup inst props pri)
let vernac_context poly l =
- if not (Classes.context poly l) then Pp.feedback Feedback.AddedAxiom
+ if not (Classes.context poly l) then Feedback.feedback Feedback.AddedAxiom
-let vernac_declare_instances locality ids pri =
+let vernac_declare_instances locality insts =
let glob = not (make_section_locality locality) in
- List.iter (fun id -> Classes.existing_instance glob id pri) ids
+ List.iter (fun (id, info) -> Classes.existing_instance glob id (Some info)) insts
let vernac_declare_class id =
Record.declare_existing_class (Nametab.global id)
@@ -825,35 +854,6 @@ let vernac_declare_class id =
let command_focus = Proof.new_focus_kind ()
let focus_command_cond = Proof.no_cond command_focus
-
-let print_info_trace = ref None
-
-let _ = let open Goptions in declare_int_option {
- optsync = true;
- optdepr = false;
- optname = "print info trace";
- optkey = ["Info" ; "Level"];
- optread = (fun () -> !print_info_trace);
- optwrite = fun n -> print_info_trace := n;
-}
-
-let vernac_solve n info tcom b =
- if not (refining ()) then
- error "Unknown command of the non proof-editing mode.";
- let status = Proof_global.with_current_proof (fun etac p ->
- let with_end_tac = if b then Some etac else None in
- let global = match n with SelectAll -> true | _ -> false in
- let info = Option.append info !print_info_trace in
- let (p,status) =
- solve n info (Tacinterp.hide_interp global tcom None) ?with_end_tac p
- in
- (* in case a strict subtree was completed,
- go back to the top of the prooftree *)
- let p = Proof.maximal_unfocus command_focus p in
- p,status) in
- if not status then Pp.feedback Feedback.AddedAxiom
-
-
(* A command which should be a tactic. It has been
added by Christine to patch an error in the design of the proof
machine, and enables to instantiate existential variables when
@@ -871,31 +871,36 @@ let vernac_set_end_tac tac =
(* TO DO verifier s'il faut pas mettre exist s | TacId s ici*)
let vernac_set_used_variables e =
+ let open Context.Named.Declaration in
let env = Global.env () in
let tys =
List.map snd (Proof.initial_goals (Proof_global.give_me_the_proof ())) in
let l = Proof_using.process_expr env e tys in
let vars = Environ.named_context env in
List.iter (fun id ->
- if not (List.exists (fun (id',_,_) -> Id.equal id id') vars) then
+ if not (List.exists (Id.equal id % get_id) vars) then
errorlabstrm "vernac_set_used_variables"
(str "Unknown variable: " ++ pr_id id))
l;
let _, to_clear = set_used_variables l in
- vernac_solve
- SelectAll None Tacexpr.(TacAtom (Loc.ghost,TacClear(false,to_clear))) false
-
+ let to_clear = List.map snd to_clear in
+ Proof_global.with_current_proof begin fun _ p ->
+ if List.is_empty to_clear then (p, ())
+ else
+ let tac = Tactics.clear to_clear in
+ fst (solve SelectAll None tac p), ()
+ end
(*****************************)
(* Auxiliary file management *)
let expand filename =
- Envars.expand_path_macros ~warn:(fun x -> msg_warning (str x)) filename
+ Envars.expand_path_macros ~warn:(fun x -> Feedback.msg_warning (str x)) filename
let vernac_add_loadpath implicit pdir ldiropt =
let pdir = expand pdir in
let alias = Option.default Nameops.default_root_prefix ldiropt in
- Mltop.add_rec_path ~unix_path:pdir ~coq_root:alias ~implicit
+ Mltop.add_rec_path Mltop.AddTopML ~unix_path:pdir ~coq_root:alias ~implicit
let vernac_remove_loadpath path =
Loadpath.remove_load_path (expand path)
@@ -910,13 +915,17 @@ let vernac_declare_ml_module locality l =
Mltop.declare_ml_modules local (List.map expand l)
let vernac_chdir = function
- | None -> msg_notice (str (Sys.getcwd()))
+ | None -> Feedback.msg_notice (str (Sys.getcwd()))
| Some path ->
begin
try Sys.chdir (expand path)
- with Sys_error err -> msg_warning (str "Cd failed: " ++ str err)
+ with Sys_error err ->
+ (* Cd is typically used to control the output directory of
+ extraction. A failed Cd could lead to overwriting .ml files
+ so we make it an error. *)
+ CErrors.error ("Cd failed: " ^ err)
end;
- if_verbose msg_info (str (Sys.getcwd()))
+ if_verbose Feedback.msg_info (str (Sys.getcwd()))
(********************)
@@ -935,85 +944,6 @@ let vernac_restore_state file =
(************)
(* Commands *)
-type tacdef_kind =
- | NewTac of Id.t
- | UpdateTac of Nametab.ltac_constant
-
-let is_defined_tac kn =
- try ignore (Tacenv.interp_ltac kn); true with Not_found -> false
-
-let make_absolute_name ident repl =
- let loc = loc_of_reference ident in
- if repl then
- let kn =
- try Nametab.locate_tactic (snd (qualid_of_reference ident))
- with Not_found ->
- Errors.user_err_loc (loc, "",
- str "There is no Ltac named " ++ pr_reference ident ++ str ".")
- in
- UpdateTac kn
- else
- let id = Constrexpr_ops.coerce_reference_to_id ident in
- let kn = Lib.make_kn id in
- let () = if is_defined_tac kn then
- Errors.user_err_loc (loc, "",
- str "There is already an Ltac named " ++ pr_reference ident ++ str".")
- in
- let is_primitive =
- try
- match Pcoq.parse_string Pcoq.Tactic.tactic (Id.to_string id) with
- | Tacexpr.TacArg _ -> false
- | _ -> true (* most probably TacAtom, i.e. a primitive tactic ident *)
- with e when Errors.noncritical e -> true (* prim tactics with args, e.g. "apply" *)
- in
- let () = if is_primitive then
- msg_warning (str "The Ltac name " ++ pr_reference ident ++
- str " may be unusable because of a conflict with a notation.")
- in
- NewTac id
-
-let register_ltac local isrec tacl =
- let map (ident, repl, body) =
- let name = make_absolute_name ident repl in
- (name, body)
- in
- let rfun = List.map map tacl in
- let recvars =
- let fold accu (op, _) = match op with
- | UpdateTac _ -> accu
- | NewTac id -> (Lib.make_path id, Lib.make_kn id) :: accu
- in
- if isrec then List.fold_left fold [] rfun
- else []
- in
- let ist = Tacintern.make_empty_glob_sign () in
- let map (name, body) =
- let body = Flags.with_option Tacintern.strict_check (Tacintern.intern_tactic_or_tacarg ist) body in
- (name, body)
- in
- let defs () =
- (** Register locally the tactic to handle recursivity. This function affects
- the whole environment, so that we transactify it afterwards. *)
- let iter_rec (sp, kn) = Nametab.push_tactic (Nametab.Until 1) sp kn in
- let () = List.iter iter_rec recvars in
- List.map map rfun
- in
- let defs = Future.transactify defs () in
- let iter (def, tac) = match def with
- | NewTac id ->
- Tacenv.register_ltac false local id tac;
- Flags.if_verbose msg_info (Nameops.pr_id id ++ str " is defined")
- | UpdateTac kn ->
- Tacenv.redefine_ltac local kn tac;
- let name = Nametab.shortest_qualid_of_tactic kn in
- Flags.if_verbose msg_info (Libnames.pr_qualid name ++ str " is redefined")
- in
- List.iter iter defs
-
-let vernac_declare_tactic_definition locality (x,def) =
- let local = make_module_locality locality in
- register_ltac local x def
-
let vernac_create_hintdb locality id b =
let local = make_module_locality locality in
Hints.create_hint_db local id full_transparent_state b
@@ -1040,141 +970,265 @@ let vernac_declare_implicits locality r l =
Impargs.declare_manual_implicits local (smart_global r) ~enriching:false
(List.map (List.map (fun (ex,b,f) -> ex, (b,true,f))) imps)
-let vernac_declare_arguments locality r l nargs flags =
- let extra_scope_flag = List.mem `ExtraScopes flags in
- let names = List.map (List.map (fun (id, _,_,_,_) -> id)) l in
- let names, rest = List.hd names, List.tl names in
- let scopes = List.map (List.map (fun (_,_, s, _,_) -> s)) l in
- if List.exists (fun na -> not (List.equal Name.equal na names)) rest then
- error "All arguments lists must declare the same names.";
- if not (List.distinct_f Name.compare (List.filter ((!=) Anonymous) names))
- then error "Arguments names must be distinct.";
- let sr = smart_global r in
+let warn_arguments_assert =
+ CWarnings.create ~name:"arguments-assert" ~category:"vernacular"
+ (fun sr ->
+ strbrk "This command is just asserting the names of arguments of " ++
+ pr_global sr ++ strbrk". If this is what you want add " ++
+ strbrk "': assert' to silence the warning. If you want " ++
+ strbrk "to clear implicit arguments add ': clear implicits'. " ++
+ strbrk "If you want to clear notation scopes add ': clear scopes'")
+
+(* [nargs_for_red] is the number of arguments required to trigger reduction,
+ [args] is the main list of arguments statuses,
+ [more_implicits] is a list of extra lists of implicit statuses *)
+let vernac_arguments locality reference args more_implicits nargs_for_red flags =
+ let assert_flag = List.mem `Assert flags in
+ let rename_flag = List.mem `Rename flags in
+ let clear_scopes_flag = List.mem `ClearScopes flags in
+ let extra_scopes_flag = List.mem `ExtraScopes flags in
+ let clear_implicits_flag = List.mem `ClearImplicits flags in
+ let default_implicits_flag = List.mem `DefaultImplicits flags in
+ let never_unfold_flag = List.mem `ReductionNeverUnfold flags in
+
+ let err_incompat x y =
+ error ("Options \""^x^"\" and \""^y^"\" are incompatible.") in
+
+ if assert_flag && rename_flag then
+ err_incompat "assert" "rename";
+ if Option.has_some nargs_for_red && never_unfold_flag then
+ err_incompat "simpl never" "/";
+ if never_unfold_flag && List.mem `ReductionDontExposeCase flags then
+ err_incompat "simpl never" "simpl nomatch";
+ if clear_scopes_flag && extra_scopes_flag then
+ err_incompat "clear scopes" "extra scopes";
+ if clear_implicits_flag && default_implicits_flag then
+ err_incompat "clear implicits" "default implicits";
+
+ let sr = smart_global reference in
let inf_names =
let ty = Global.type_of_global_unsafe sr in
- Impargs.compute_implicits_names (Global.env ()) ty in
- let rec check li ld ls = match li, ld, ls with
- | [], [], [] -> ()
- | [], Anonymous::ld, (Some _)::ls when extra_scope_flag -> check li ld ls
- | [], _::_, (Some _)::ls when extra_scope_flag ->
- error "Extra notation scopes can be set on anonymous arguments only"
- | [], x::_, _ -> errorlabstrm "vernac_declare_arguments"
- (str "Extra argument " ++ pr_name x ++ str ".")
- | l, [], _ -> errorlabstrm "vernac_declare_arguments"
- (str "The following arguments are not declared: " ++
- prlist_with_sep pr_comma pr_name l ++ str ".")
- | _::li, _::ld, _::ls -> check li ld ls
- | _ -> assert false in
- let () = match l with
- | [[]] -> ()
- | _ ->
- List.iter2 (fun l -> check inf_names l) (names :: rest) scopes
+ Impargs.compute_implicits_names (Global.env ()) ty
in
- (* we take extra scopes apart, and we check they are consistent *)
- let l, scopes =
- let scopes, rest = List.hd scopes, List.tl scopes in
- if List.exists (List.exists ((!=) None)) rest then
- error "Notation scopes can be given only once";
- if not extra_scope_flag then l, scopes else
- let l, _ = List.split (List.map (List.chop (List.length inf_names)) l) in
- l, scopes in
- (* we interpret _ as the inferred names *)
- let l = match l with
- | [[]] -> l
- | _ ->
- let name_anons = function
- | (Anonymous, a,b,c,d), Name id -> Name id, a,b,c,d
- | x, _ -> x in
- List.map (fun ns -> List.map name_anons (List.combine ns inf_names)) l in
- let names_decl = List.map (List.map (fun (id, _,_,_,_) -> id)) l in
- let renamed_arg = ref None in
- let set_renamed a b =
- if Option.is_empty !renamed_arg && not (Id.equal a b) then renamed_arg := Some(b,a) in
- let some_renaming_specified =
- try
- let names = Arguments_renaming.arguments_names sr in
- not (List.equal (List.equal Name.equal) names names_decl)
- with Not_found -> false in
- let some_renaming_specified, implicits =
- match l with
- | [[]] -> false, [[]]
+ let prev_names =
+ try Arguments_renaming.arguments_names sr with Not_found -> inf_names
+ in
+ let num_args = List.length inf_names in
+ assert (Int.equal num_args (List.length prev_names));
+
+ let names_of args = List.map (fun a -> a.name) args in
+
+ (* Checks *)
+
+ let err_extra_args names =
+ errorlabstrm "vernac_declare_arguments"
+ (strbrk "Extra arguments: " ++
+ prlist_with_sep pr_comma pr_name names ++ str ".")
+ in
+ let err_missing_args names =
+ errorlabstrm "vernac_declare_arguments"
+ (strbrk "The following arguments are not declared: " ++
+ prlist_with_sep pr_comma pr_name names ++ str ".")
+ in
+
+ let rec check_extra_args extra_args =
+ match extra_args with
+ | [] -> ()
+ | { notation_scope = None } :: _ -> err_extra_args (names_of extra_args)
+ | { name = Anonymous; notation_scope = Some _ } :: args ->
+ check_extra_args args
| _ ->
- List.fold_map (fun sr il ->
- let sr', impl = List.fold_map (fun b -> function
- | (Anonymous, _,_, true, max), Name id -> assert false
- | (Name x, _,_, true, _), Anonymous ->
- errorlabstrm "vernac_declare_arguments"
- (str "Argument " ++ pr_id x ++ str " cannot be declared implicit.")
- | (Name iid, _,_, true, max), Name id ->
- set_renamed iid id;
- b || not (Id.equal iid id), Some (ExplByName id, max, false)
- | (Name iid, _,_, _, _), Name id ->
- set_renamed iid id;
- b || not (Id.equal iid id), None
- | _ -> b, None)
- sr (List.combine il inf_names) in
- sr || sr', List.map_filter (fun x -> x) impl)
- some_renaming_specified l in
- if some_renaming_specified then
- if not (List.mem `Rename flags) then
- errorlabstrm "vernac_declare_arguments"
- (str "To rename arguments the \"rename\" flag must be specified." ++
- match !renamed_arg with
- | None -> mt ()
- | Some (o,n) ->
- str "\nArgument " ++ pr_id o ++ str " renamed to " ++ pr_id n ++ str ".")
- else
- Arguments_renaming.rename_arguments
- (make_section_locality locality) sr names_decl;
- (* All other infos are in the first item of l *)
- let l = List.hd l in
- let some_implicits_specified = match implicits with
- | [[]] -> false | _ -> true in
- let scopes = List.map (function
- | None -> None
- | Some (o, k) ->
- try ignore (Notation.find_scope k); Some k
- with UserError _ ->
- Some (Notation.find_delimiters_scope o k)) scopes
+ error "Extra notation scopes can be set on anonymous and explicit arguments only."
+ in
+
+ let args, scopes =
+ let scopes = List.map (fun { notation_scope = s } -> s) args in
+ if List.length args > num_args then
+ let args, extra_args = List.chop num_args args in
+ if extra_scopes_flag then
+ (check_extra_args extra_args; (args, scopes))
+ else err_extra_args (names_of extra_args)
+ else args, scopes
+ in
+
+ if Option.cata (fun n -> n > num_args) false nargs_for_red then
+ error "The \"/\" modifier should be put before any extra scope.";
+
+ let scopes_specified = List.exists Option.has_some scopes in
+
+ if scopes_specified && clear_scopes_flag then
+ error "The \"clear scopes\" flag is incompatible with scope annotations.";
+
+ let names = List.map (fun { name } -> name) args in
+ let names = names :: List.map (List.map fst) more_implicits in
+
+ let rename_flag_required = ref false in
+ let example_renaming = ref None in
+ let save_example_renaming renaming =
+ rename_flag_required := !rename_flag_required
+ || not (Name.equal (fst renaming) Anonymous);
+ if Option.is_empty !example_renaming then
+ example_renaming := Some renaming
in
- let some_scopes_specified = List.exists ((!=) None) scopes in
+
+ let rec names_union names1 names2 =
+ match names1, names2 with
+ | [], [] -> []
+ | _ :: _, [] -> names1
+ | [], _ :: _ -> names2
+ | (Name _ as name) :: names1, Anonymous :: names2
+ | Anonymous :: names1, (Name _ as name) :: names2 ->
+ name :: names_union names1 names2
+ | name1 :: names1, name2 :: names2 ->
+ if Name.equal name1 name2 then
+ name1 :: names_union names1 names2
+ else error "Argument lists should agree on the names they provide."
+ in
+
+ let names = List.fold_left names_union [] names in
+
+ let rec rename prev_names names =
+ match prev_names, names with
+ | [], [] -> []
+ | [], _ :: _ -> err_extra_args names
+ | _ :: _, [] when assert_flag ->
+ (* Error messages are expressed in terms of original names, not
+ renamed ones. *)
+ err_missing_args (List.lastn (List.length prev_names) inf_names)
+ | _ :: _, [] -> prev_names
+ | prev :: prev_names, Anonymous :: names ->
+ prev :: rename prev_names names
+ | prev :: prev_names, (Name id as name) :: names ->
+ if not (Name.equal prev name) then save_example_renaming (prev,name);
+ name :: rename prev_names names
+ in
+
+ let names = rename prev_names names in
+ let renaming_specified = Option.has_some !example_renaming in
+
+ if !rename_flag_required && not rename_flag then
+ errorlabstrm "vernac_declare_arguments"
+ (strbrk "To rename arguments the \"rename\" flag must be specified."
+ ++ spc () ++
+ match !example_renaming with
+ | None -> mt ()
+ | Some (o,n) ->
+ str "Argument " ++ pr_name o ++
+ str " renamed to " ++ pr_name n ++ str ".");
+
+ let duplicate_names =
+ List.duplicates Name.equal (List.filter ((!=) Anonymous) names)
+ in
+ if not (List.is_empty duplicate_names) then begin
+ let duplicates = prlist_with_sep pr_comma pr_name duplicate_names in
+ errorlabstrm "_" (strbrk "Some argument names are duplicated: " ++ duplicates)
+ end;
+
+ (* Parts of this code are overly complicated because the implicit arguments
+ API is completely crazy: positions (ExplByPos) are elaborated to
+ names. This is broken by design, since not all arguments have names. So
+ even though we eventually want to map only positions to implicit statuses,
+ we have to check whether the corresponding arguments have names, not to
+ trigger an error in the impargs code. Even better, the names we have to
+ check are not the current ones (after previous renamings), but the original
+ ones (inferred from the type). *)
+
+ let implicits =
+ List.map (fun { name; implicit_status = i } -> (name,i)) args
+ in
+ let implicits = implicits :: more_implicits in
+
+ let open Vernacexpr in
+ let rec build_implicits inf_names implicits =
+ match inf_names, implicits with
+ | _, [] -> []
+ | _ :: inf_names, (_, NotImplicit) :: implicits ->
+ build_implicits inf_names implicits
+
+ (* With the current impargs API, it is impossible to make an originally
+ anonymous argument implicit *)
+ | Anonymous :: _, (name, _) :: _ ->
+ errorlabstrm "vernac_declare_arguments"
+ (strbrk"Argument "++ pr_name name ++
+ strbrk " cannot be declared implicit.")
+
+ | Name id :: inf_names, (name, impl) :: implicits ->
+ let max = impl = MaximallyImplicit in
+ (ExplByName id,max,false) :: build_implicits inf_names implicits
+
+ | _ -> assert false (* already checked in [names_union] *)
+ in
+
+ let implicits = List.map (build_implicits inf_names) implicits in
+ let implicits_specified = match implicits with [[]] -> false | _ -> true in
+
+ if implicits_specified && clear_implicits_flag then
+ error "The \"clear implicits\" flag is incompatible with implicit annotations";
+
+ if implicits_specified && default_implicits_flag then
+ error "The \"default implicits\" flag is incompatible with implicit annotations";
+
let rargs =
Util.List.map_filter (function (n, true) -> Some n | _ -> None)
- (Util.List.map_i (fun i (_, b, _,_,_) -> i, b) 0 l) in
- if some_scopes_specified || List.mem `ClearScopes flags then
- vernac_arguments_scope locality r scopes;
- if not some_implicits_specified && List.mem `DefaultImplicits flags then
- vernac_declare_implicits locality r []
- else if some_implicits_specified || List.mem `ClearImplicits flags then
- vernac_declare_implicits locality r implicits;
- if nargs >= 0 && nargs < List.fold_left max 0 rargs then
- error "The \"/\" option must be placed after the last \"!\".";
- let no_flags = List.is_empty flags in
+ (Util.List.map_i (fun i { recarg_like = b } -> i, b) 0 args)
+ in
+
let rec narrow = function
| #Reductionops.ReductionBehaviour.flag as x :: tl -> x :: narrow tl
- | [] -> [] | _ :: tl -> narrow tl in
- let flags = narrow flags in
- let some_simpl_flags_specified =
- not (List.is_empty rargs) || nargs >= 0 || not (List.is_empty flags) in
- if some_simpl_flags_specified then begin
+ | [] -> [] | _ :: tl -> narrow tl
+ in
+ let red_flags = narrow flags in
+ let red_modifiers_specified =
+ not (List.is_empty rargs) || Option.has_some nargs_for_red
+ || not (List.is_empty red_flags)
+ in
+
+ if not (List.is_empty rargs) && never_unfold_flag then
+ err_incompat "simpl never" "!";
+
+
+ (* Actions *)
+
+ if renaming_specified then begin
+ let local = make_section_locality locality in
+ Arguments_renaming.rename_arguments local sr names
+ end;
+
+ if scopes_specified || clear_scopes_flag then begin
+ let scopes = List.map (Option.map (fun (o,k) ->
+ try ignore (Notation.find_scope k); k
+ with UserError _ ->
+ Notation.find_delimiters_scope o k)) scopes
+ in
+ vernac_arguments_scope locality reference scopes
+ end;
+
+ if implicits_specified || clear_implicits_flag then
+ vernac_declare_implicits locality reference implicits;
+
+ if default_implicits_flag then
+ vernac_declare_implicits locality reference [];
+
+ if red_modifiers_specified then begin
match sr with
| ConstRef _ as c ->
Reductionops.ReductionBehaviour.set
- (make_section_locality locality) c (rargs, nargs, flags)
- | _ -> errorlabstrm "" (strbrk "Modifiers of the behavior of the simpl tactic are relevant for constants only.")
+ (make_section_locality locality) c
+ (rargs, Option.default ~-1 nargs_for_red, red_flags)
+ | _ -> errorlabstrm ""
+ (strbrk "Modifiers of the behavior of the simpl tactic "++
+ strbrk "are relevant for constants only.")
end;
- if not (some_renaming_specified ||
- some_implicits_specified ||
- some_scopes_specified ||
- some_simpl_flags_specified) &&
- no_flags then
- msg_warning (strbrk "This command is just asserting the number and names of arguments of " ++ pr_global sr ++ strbrk". If this is what you want add ': assert' to silence the warning. If you want to clear implicit arguments add ': clear implicits'. If you want to clear notation scopes add ': clear scopes'")
+ if not (renaming_specified ||
+ implicits_specified ||
+ scopes_specified ||
+ red_modifiers_specified) && (List.is_empty flags) then
+ warn_arguments_assert sr
let default_env () = {
Notation_term.ninterp_var_type = Id.Map.empty;
ninterp_rec_vars = Id.Map.empty;
- ninterp_only_parse = false;
}
let vernac_reserve bl =
@@ -1183,7 +1237,7 @@ let vernac_reserve bl =
let sigma = Evd.from_env env in
let t,ctx = Constrintern.interp_type env sigma c in
let t = Detyping.detype false [] env (Evd.from_ctx ctx) t in
- let t = Notation_ops.notation_constr_of_glob_constr (default_env ()) t in
+ let t,_ = Notation_ops.notation_constr_of_glob_constr (default_env ()) t in
Reserve.declare_reserved_type idl t)
in List.iter sb_decl bl
@@ -1370,8 +1424,8 @@ let _ =
optdepr = false;
optname = "kernel term sharing";
optkey = ["Kernel"; "Term"; "Sharing"];
- optread = (fun () -> !Closure.share);
- optwrite = (fun b -> Closure.share := b) }
+ optread = (fun () -> !CClosure.share);
+ optwrite = (fun b -> CClosure.share := b) }
(* No more undo limit in the new proof engine.
The command still exists for compatibility (e.g. with ProofGeneral) *)
@@ -1430,18 +1484,6 @@ let _ =
optread = Flags.get_dump_bytecode;
optwrite = Flags.set_dump_bytecode }
-let vernac_debug b =
- set_debug (if b then Tactic_debug.DebugOn 0 else Tactic_debug.DebugOff)
-
-let _ =
- declare_bool_option
- { optsync = false;
- optdepr = false;
- optname = "Ltac debug";
- optkey = ["Ltac";"Debug"];
- optread = (fun () -> get_debug () != Tactic_debug.DebugOff);
- optwrite = vernac_debug }
-
let _ =
declare_bool_option
{ optsync = true;
@@ -1451,6 +1493,15 @@ let _ =
optread = (fun () -> !Constrintern.parsing_explicit);
optwrite = (fun b -> Constrintern.parsing_explicit := b) }
+let _ =
+ declare_string_option ~preprocess:CWarnings.normalize_flags_string
+ { optsync = true;
+ optdepr = false;
+ optname = "warnings display";
+ optkey = ["Warnings"];
+ optread = CWarnings.get_flags;
+ optwrite = CWarnings.set_flags }
+
let vernac_set_strategy locality l =
let local = make_locality locality in
let glob_ref r =
@@ -1480,6 +1531,9 @@ let vernac_set_option locality key = function
| IntValue n -> set_int_option_value_gen locality key n
| BoolValue b -> set_bool_option_value_gen locality key b
+let vernac_set_append_option locality key s =
+ set_string_option_append_value_gen locality key s
+
let vernac_unset_option locality key =
unset_option_value_gen locality key
@@ -1519,7 +1573,7 @@ let get_current_context_of_args = function
let vernac_check_may_eval redexp glopt rc =
let (sigma, env) = get_current_context_of_args glopt in
let sigma', c = interp_open_constr env sigma rc in
- let sigma' = Evarconv.consider_remaining_unif_problems env sigma' in
+ let sigma' = Evarconv.solve_unif_constraints_with_heuristics env sigma' in
Evarconv.check_problems_are_solved env sigma';
let sigma',nf = Evarutil.nf_evars_and_universes sigma' in
let pl, uctx = Evd.universe_context sigma' in
@@ -1535,18 +1589,22 @@ let vernac_check_may_eval redexp glopt rc =
| None ->
let l = Evar.Set.union (Evd.evars_of_term j.Environ.uj_val) (Evd.evars_of_term j.Environ.uj_type) in
let j = { j with Environ.uj_type = Reductionops.nf_betaiota sigma' j.Environ.uj_type } in
- msg_notice (print_judgment env sigma' j ++
+ Feedback.msg_notice (print_judgment env sigma' j ++
pr_ne_evar_set (fnl () ++ str "where" ++ fnl ()) (mt ()) sigma' l ++
Printer.pr_universe_ctx sigma uctx)
| Some r ->
- Tacintern.dump_glob_red_expr r;
- let (sigma',r_interp) = interp_redexp env sigma' r in
- let redfun env evm c = snd (fst (reduction_of_red_expr env r_interp) env evm c) in
- msg_notice (print_eval redfun env sigma' rc j)
+ let (sigma',r_interp) = Hook.get f_interp_redexp env sigma' r in
+ let redfun env evm c =
+ let (redfun, _) = reduction_of_red_expr env r_interp in
+ let evm = Sigma.Unsafe.of_evar_map evm in
+ let Sigma (c, _, _) = redfun.Reductionops.e_redfun env evm c in
+ c
+ in
+ Feedback.msg_notice (print_eval redfun env sigma' rc j)
let vernac_declare_reduction locality s r =
let local = make_locality locality in
- declare_red_expr local s (snd (interp_redexp (Global.env()) Evd.empty r))
+ declare_red_expr local s (snd (Hook.get f_interp_redexp (Global.env()) Evd.empty r))
(* The same but avoiding the current goal context if any *)
let vernac_global_check c =
@@ -1554,11 +1612,11 @@ let vernac_global_check c =
let sigma = Evd.from_env env in
let c,ctx = interp_constr env sigma c in
let senv = Global.safe_env() in
- let cstrs = snd (Evd.evar_universe_context_set Univ.UContext.empty ctx) in
+ let cstrs = snd (UState.context_set ctx) in
let senv = Safe_typing.add_constraints cstrs senv in
let j = Safe_typing.typing senv c in
let env = Safe_typing.env_of_safe_env senv in
- msg_notice (print_safe_judgment env sigma j)
+ Feedback.msg_notice (print_safe_judgment env sigma j)
let get_nth_goal n =
@@ -1572,6 +1630,7 @@ exception NoHyp
We only print the type and a small statement to this comes from the
goal. Precondition: there must be at least one current goal. *)
let print_about_hyp_globs ref_or_by_not glnumopt =
+ let open Context.Named.Declaration in
try
let gl,id =
match glnumopt,ref_or_by_not with
@@ -1584,17 +1643,17 @@ let print_about_hyp_globs ref_or_by_not glnumopt =
(str "No such goal: " ++ int n ++ str "."))
| _ , _ -> raise NoHyp in
let hyps = pf_hyps gl in
- let (id,bdyopt,typ) = Context.lookup_named id hyps in
- let natureofid = match bdyopt with
- | None -> "Hypothesis"
- | Some bdy ->"Constant (let in)" in
- v 0 (str (Id.to_string id) ++ str":" ++ pr_constr typ ++ fnl() ++ fnl()
+ let decl = Context.Named.lookup id hyps in
+ let natureofid = match decl with
+ | LocalAssum _ -> "Hypothesis"
+ | LocalDef (_,bdy,_) ->"Constant (let in)" in
+ v 0 (pr_id id ++ str":" ++ pr_constr (get_type decl) ++ fnl() ++ fnl()
++ str natureofid ++ str " of the goal context.")
with (* fallback to globals *)
| NoHyp | Not_found -> print_about ref_or_by_not
-let vernac_print = function
+let vernac_print = let open Feedback in function
| PrintTables -> msg_notice (print_tables ())
| PrintFullContext-> msg_notice (print_full_context_typ ())
| PrintSectionContext qid -> msg_notice (print_sec_context_typ qid)
@@ -1613,26 +1672,24 @@ let vernac_print = function
| PrintClasses -> msg_notice (Prettyp.print_classes())
| PrintTypeClasses -> msg_notice (Prettyp.print_typeclasses())
| PrintInstances c -> msg_notice (Prettyp.print_instances (smart_global c))
- | PrintLtac qid -> msg_notice (Tacintern.print_ltac (snd (qualid_of_reference qid)))
| PrintCoercions -> msg_notice (Prettyp.print_coercions())
| PrintCoercionPaths (cls,clt) ->
msg_notice (Prettyp.print_path_between (cl_of_qualid cls) (cl_of_qualid clt))
| PrintCanonicalConversions -> msg_notice (Prettyp.print_canonical_projections ())
| PrintUniverses (b, dst) ->
let univ = Global.universes () in
- let univ = if b then Univ.sort_universes univ else univ in
+ let univ = if b then UGraph.sort_universes univ else univ in
let pr_remaining =
if Global.is_joined_environment () then mt ()
else str"There may remain asynchronous universe constraints"
in
begin match dst with
- | None -> msg_notice (Univ.pr_universes Universes.pr_with_global_universes univ ++ pr_remaining)
+ | None -> msg_notice (UGraph.pr_universes Universes.pr_with_global_universes univ ++ pr_remaining)
| Some s -> dump_universes_gen univ s
end
| PrintHint r -> msg_notice (Hints.pr_hint_ref (smart_global r))
| PrintHintGoal -> msg_notice (Hints.pr_applicable_hint ())
| PrintHintDbName s -> msg_notice (Hints.pr_hint_db_by_name s)
- | PrintRewriteHintDbName s -> msg_notice (Autorewrite.print_rewrite_hintdb s)
| PrintHintDb -> msg_notice (Hints.pr_searchtable ())
| PrintScopes ->
msg_notice (Notation.pr_scopes (Constrextern.without_symbols pr_lglob_constr))
@@ -1684,6 +1741,28 @@ let interp_search_about_item env =
errorlabstrm "interp_search_about_item"
(str "Unable to interp \"" ++ str s ++ str "\" either as a reference or as an identifier component")
+(* 05f22a5d6d5b8e3e80f1a37321708ce401834430 introduced the
+ `search_output_name_only` option to avoid excessive printing when
+ searching.
+
+ The motivation was to make search usable for IDE completion,
+ however, it is still too slow due to the non-indexed nature of the
+ underlying search mechanism.
+
+ In the future we should deprecate the option and provide a fast,
+ indexed name-searching interface.
+*)
+let search_output_name_only = ref false
+
+let _ =
+ declare_bool_option
+ { optsync = true;
+ optdepr = false;
+ optname = "output-name-only search";
+ optkey = ["Search";"Output";"Name";"Only"];
+ optread = (fun () -> !search_output_name_only);
+ optwrite = (:=) search_output_name_only }
+
let vernac_search s gopt r =
let r = interp_search_restriction r in
let env,gopt =
@@ -1695,18 +1774,28 @@ let vernac_search s gopt r =
| Some g -> snd (Pfedit.get_goal_context g) , Some g
in
let get_pattern c = snd (intern_constr_pattern env c) in
+ let pr_search ref env c =
+ let pr = pr_global ref in
+ let pp = if !search_output_name_only
+ then pr
+ else begin
+ let pc = pr_lconstr_env env Evd.empty c in
+ hov 2 (pr ++ str":" ++ spc () ++ pc)
+ end
+ in Feedback.msg_notice pp
+ in
match s with
| SearchPattern c ->
- msg_notice (Search.search_pattern gopt (get_pattern c) r)
+ Search.search_pattern gopt (get_pattern c) r pr_search
| SearchRewrite c ->
- msg_notice (Search.search_rewrite gopt (get_pattern c) r)
+ Search.search_rewrite gopt (get_pattern c) r pr_search
| SearchHead c ->
- msg_notice (Search.search_by_head gopt (get_pattern c) r)
+ Search.search_by_head gopt (get_pattern c) r pr_search
| SearchAbout sl ->
- msg_notice (Search.search_about gopt (List.map (on_snd (interp_search_about_item env)) sl) r)
+ Search.search_about gopt (List.map (on_snd (interp_search_about_item env)) sl) r pr_search
-let vernac_locate = function
- | LocateAny (AN qid) -> msg_notice (print_located_qualid qid)
+let vernac_locate = let open Feedback in function
+ | LocateAny (AN qid) -> msg_notice (print_located_qualid qid)
| LocateTerm (AN qid) -> msg_notice (print_located_term qid)
| LocateAny (ByNotation (_, ntn, sc)) (** TODO : handle Ltac notations *)
| LocateTerm (ByNotation (_, ntn, sc)) ->
@@ -1736,7 +1825,7 @@ let vernac_focus gln =
match gln with
| None -> Proof.focus focus_command_cond () 1 p
| Some 0 ->
- Errors.error "Invalid goal number: 0. Goal numbering starts with 1."
+ CErrors.error "Invalid goal number: 0. Goal numbering starts with 1."
| Some n ->
Proof.focus focus_command_cond () n p)
@@ -1749,7 +1838,7 @@ let vernac_unfocus () =
let vernac_unfocused () =
let p = Proof_global.give_me_the_proof () in
if Proof.unfocused p then
- msg_notice (str"The proof is indeed fully unfocused.")
+ Feedback.msg_notice (str"The proof is indeed fully unfocused.")
else
error "The proof is not fully unfocused."
@@ -1777,7 +1866,7 @@ let vernac_bullet (bullet:Proof_global.Bullet.t) =
Proof_global.simple_with_current_proof (fun _ p ->
Proof_global.Bullet.put p bullet)
-let vernac_show = function
+let vernac_show = let open Feedback in function
| ShowGoal goalref ->
let info = match goalref with
| OpenSubgoals -> pr_open_subgoals ()
@@ -1815,33 +1904,36 @@ let vernac_check_guard () =
with UserError(_,s) ->
(str ("Condition violated: ") ++s)
in
- msg_notice message
+ Feedback.msg_notice message
exception End_of_input
let vernac_load interp fname =
+ let interp x =
+ let proof_mode = Proof_global.get_default_proof_mode_name () in
+ Proof_global.activate_proof_mode proof_mode;
+ interp x in
let parse_sentence = Flags.with_option Flags.we_are_parsing
(fun po ->
match Pcoq.Gram.entry_parse Pcoq.main_entry po with
| Some x -> x
| None -> raise End_of_input) in
let fname =
- Envars.expand_path_macros ~warn:(fun x -> msg_warning (str x)) fname in
+ Envars.expand_path_macros ~warn:(fun x -> Feedback.msg_warning (str x)) fname in
let fname = CUnix.make_suffix fname ".v" in
let input =
let longfname = Loadpath.locate_file fname in
let in_chan = open_utf8_file_in longfname in
- Pcoq.Gram.parsable (Stream.of_channel in_chan) in
+ Pcoq.Gram.parsable ~file:longfname (Stream.of_channel in_chan) in
try while true do interp (snd (parse_sentence input)) done
with End_of_input -> ()
-
(* "locality" is the prefix "Local" attribute, while the "local" component
* is the outdated/deprecated "Local" attribute of some vernacular commands
* still parsed as the obsolete_locality grammar entry for retrocompatibility.
* loc is the Loc.t of the vernacular command being interpreted. *)
let interp ?proof ~loc locality poly c =
- prerr_endline ("interpreting: " ^ Pp.string_of_ppcmds (Ppvernac.pr_vernac c));
+ prerr_endline (fun () -> "interpreting: " ^ Pp.string_of_ppcmds (Ppvernac.pr_vernac c));
match c with
(* Done later in this file *)
| VernacLoad _ -> assert false
@@ -1854,8 +1946,6 @@ let interp ?proof ~loc locality poly c =
| VernacError e -> raise e
(* Syntax *)
- | VernacTacticNotation (n,r,e) ->
- Metasyntax.add_tactic_notation (make_module_locality locality,n,r,e)
| VernacSyntaxExtension (local,sl) ->
vernac_syntax_extension locality local sl
| VernacDelimiters (sc,lr) -> vernac_delimiters sc lr
@@ -1906,14 +1996,13 @@ let interp ?proof ~loc locality poly c =
vernac_identity_coercion locality poly local id s t
(* Type classes *)
- | VernacInstance (abst, sup, inst, props, pri) ->
- vernac_instance abst locality poly sup inst props pri
+ | VernacInstance (abst, sup, inst, props, info) ->
+ vernac_instance abst locality poly sup inst props info
| VernacContext sup -> vernac_context poly sup
- | VernacDeclareInstances (ids, pri) -> vernac_declare_instances locality ids pri
+ | VernacDeclareInstances insts -> vernac_declare_instances locality insts
| VernacDeclareClass id -> vernac_declare_class id
(* Solving *)
- | VernacSolve (n,info,tac,b) -> vernac_solve n info tac b
| VernacSolveExistential (n,c) -> vernac_solve_existential n c
(* Auxiliary file and library management *)
@@ -1934,8 +2023,6 @@ let interp ?proof ~loc locality poly c =
| VernacBackTo _ -> anomaly (str "VernacBackTo not handled by Stm")
(* Commands *)
- | VernacDeclareTacticDefinition def ->
- vernac_declare_tactic_definition locality def
| VernacCreateHintDb (dbname,b) -> vernac_create_hintdb locality dbname b
| VernacRemoveHints (dbnames,ids) -> vernac_remove_hints locality dbnames ids
| VernacHints (local,dbnames,hints) ->
@@ -1944,13 +2031,14 @@ let interp ?proof ~loc locality poly c =
vernac_syntactic_definition locality id c local b
| VernacDeclareImplicits (qid,l) ->
vernac_declare_implicits locality qid l
- | VernacArguments (qid, l, narg, flags) ->
- vernac_declare_arguments locality qid l narg flags
+ | VernacArguments (qid, args, more_implicits, nargs, flags) ->
+ vernac_arguments locality qid args more_implicits nargs flags
| VernacReserve bl -> vernac_reserve bl
| VernacGeneralizable gen -> vernac_generalizable locality gen
| VernacSetOpacity qidl -> vernac_set_opacity locality qidl
| VernacSetStrategy l -> vernac_set_strategy locality l
| VernacSetOption (key,v) -> vernac_set_option locality key v
+ | VernacSetAppendOption (key,v) -> vernac_set_append_option locality key v
| VernacUnsetOption key -> vernac_unset_option locality key
| VernacRemoveOption (key,v) -> vernac_remove_option key v
| VernacAddOption (key,v) -> vernac_add_option key v
@@ -1963,16 +2051,15 @@ let interp ?proof ~loc locality poly c =
| VernacSearch (s,g,r) -> vernac_search s g r
| VernacLocate l -> vernac_locate l
| VernacRegister (id, r) -> vernac_register id r
- | VernacComments l -> if_verbose msg_info (str "Comments ok\n")
- | VernacNop -> ()
+ | VernacComments l -> if_verbose Feedback.msg_info (str "Comments ok\n")
(* The STM should handle that, but LOAD bypasses the STM... *)
- | VernacAbort id -> msg_warning (str "VernacAbort not handled by Stm")
- | VernacAbortAll -> msg_warning (str "VernacAbortAll not handled by Stm")
- | VernacRestart -> msg_warning (str "VernacRestart not handled by Stm")
- | VernacUndo _ -> msg_warning (str "VernacUndo not handled by Stm")
- | VernacUndoTo _ -> msg_warning (str "VernacUndoTo not handled by Stm")
- | VernacBacktrack _ -> msg_warning (str "VernacBacktrack not handled by Stm")
+ | VernacAbort id -> CErrors.errorlabstrm "" (str "Abort cannot be used through the Load command")
+ | VernacAbortAll -> CErrors.errorlabstrm "" (str "AbortAll cannot be used through the Load command")
+ | VernacRestart -> CErrors.errorlabstrm "" (str "Restart cannot be used through the Load command")
+ | VernacUndo _ -> CErrors.errorlabstrm "" (str "Undo cannot be used through the Load command")
+ | VernacUndoTo _ -> CErrors.errorlabstrm "" (str "UndoTo cannot be used through the Load command")
+ | VernacBacktrack _ -> CErrors.errorlabstrm "" (str "Backtrack cannot be used through the Load command")
(* Proof management *)
| VernacGoal t -> vernac_start_proof locality poly Theorem [None,([],t,None)] false
@@ -2012,25 +2099,23 @@ let check_vernac_supports_locality c l =
match l, c with
| None, _ -> ()
| Some _, (
- VernacTacticNotation _
- | VernacOpenCloseScope _
+ VernacOpenCloseScope _
| VernacSyntaxExtension _ | VernacInfix _ | VernacNotation _
| VernacDefinition _ | VernacFixpoint _ | VernacCoFixpoint _
| VernacAssumption _ | VernacStartTheoremProof _
| VernacCoercion _ | VernacIdentityCoercion _
| VernacInstance _ | VernacDeclareInstances _
| VernacDeclareMLModule _
- | VernacDeclareTacticDefinition _
| VernacCreateHintDb _ | VernacRemoveHints _ | VernacHints _
| VernacSyntacticDefinition _
| VernacArgumentsScope _ | VernacDeclareImplicits _ | VernacArguments _
| VernacGeneralizable _
| VernacSetOpacity _ | VernacSetStrategy _
- | VernacSetOption _ | VernacUnsetOption _
+ | VernacSetOption _ | VernacSetAppendOption _ | VernacUnsetOption _
| VernacDeclareReduction _
| VernacExtend _
| VernacInductive _) -> ()
- | Some _, _ -> Errors.error "This command does not support Locality"
+ | Some _, _ -> CErrors.error "This command does not support Locality"
(* Vernaculars that take a polymorphism flag *)
let check_vernac_supports_polymorphism c p =
@@ -2044,13 +2129,13 @@ let check_vernac_supports_polymorphism c p =
| VernacInstance _ | VernacDeclareInstances _
| VernacHints _ | VernacContext _
| VernacExtend _ | VernacUniverse _ | VernacConstraint _) -> ()
- | Some _, _ -> Errors.error "This command does not support Polymorphism"
+ | Some _, _ -> CErrors.error "This command does not support Polymorphism"
let enforce_polymorphism = function
| None -> Flags.is_universe_polymorphism ()
| Some b -> Flags.make_polymorphic_flag b; b
-(** A global default timeout, controled by option "Set Default Timeout n".
+(** A global default timeout, controlled by option "Set Default Timeout n".
Use "Unset Default Timeout" to deactivate it (or set it to 0). *)
let default_timeout = ref None
@@ -2098,17 +2183,17 @@ let with_fail b f =
with
| HasNotFailed as e -> raise e
| e ->
- let e = Errors.push e in
- raise (HasFailed (Errors.iprint
- (Cerrors.process_vernac_interp_error ~allow_uncaught:false ~with_header:false e))))
+ let e = CErrors.push e in
+ raise (HasFailed (CErrors.iprint
+ (ExplainErr.process_vernac_interp_error ~allow_uncaught:false ~with_header:false e))))
()
- with e when Errors.noncritical e ->
- let (e, _) = Errors.push e in
+ with e when CErrors.noncritical e ->
+ let (e, _) = CErrors.push e in
match e with
| HasNotFailed ->
errorlabstrm "Fail" (str "The command has not failed!")
| HasFailed msg ->
- if is_verbose () || !test_mode || !ide_slave then msg_info
+ if is_verbose () || !test_mode || !ide_slave then Feedback.msg_info
(str "The command has indeed failed with message:" ++ fnl () ++ msg)
| _ -> assert false
end
@@ -2117,13 +2202,13 @@ let interp ?(verbosely=true) ?proof (loc,c) =
let orig_program_mode = Flags.is_program_mode () in
let rec aux ?locality ?polymorphism isprogcmd = function
| VernacProgram c when not isprogcmd -> aux ?locality ?polymorphism true c
- | VernacProgram _ -> Errors.error "Program mode specified twice"
+ | VernacProgram _ -> CErrors.error "Program mode specified twice"
| VernacLocal (b, c) when Option.is_empty locality ->
aux ~locality:b ?polymorphism isprogcmd c
| VernacPolymorphic (b, c) when polymorphism = None ->
aux ?locality ~polymorphism:b isprogcmd c
- | VernacPolymorphic (b, c) -> Errors.error "Polymorphism specified twice"
- | VernacLocal _ -> Errors.error "Locality specified twice"
+ | VernacPolymorphic (b, c) -> CErrors.error "Polymorphism specified twice"
+ | VernacLocal _ -> CErrors.error "Locality specified twice"
| VernacStm (Command c) -> aux ?locality ?polymorphism isprogcmd c
| VernacStm (PGLast c) -> aux ?locality ?polymorphism isprogcmd c
| VernacStm _ -> assert false (* Done by Stm *)
@@ -2132,11 +2217,11 @@ let interp ?(verbosely=true) ?proof (loc,c) =
| VernacTimeout (n,v) ->
current_timeout := Some n;
aux ?locality ?polymorphism isprogcmd v
- | VernacRedirect (s, v) ->
- Pp.with_output_to_file s (aux_list ?locality ?polymorphism isprogcmd) v;
- | VernacTime v ->
+ | VernacRedirect (s, (_,v)) ->
+ Feedback.with_output_to_file s (aux false) v
+ | VernacTime (_,v) ->
System.with_time !Flags.time
- (aux_list ?locality ?polymorphism isprogcmd) v;
+ (aux ?locality ?polymorphism isprogcmd) v;
| VernacLoad (_,fname) -> vernac_load (aux false) fname
| c ->
check_vernac_supports_locality c locality;
@@ -2156,16 +2241,14 @@ let interp ?(verbosely=true) ?proof (loc,c) =
| reraise when
(match reraise with
| Timeout -> true
- | e -> Errors.noncritical e)
+ | e -> CErrors.noncritical e)
->
- let e = Errors.push reraise in
+ let e = CErrors.push reraise in
let e = locate_if_not_already loc e in
let () = restore_timeout () in
Flags.program_mode := orig_program_mode;
ignore (Flags.use_polymorphic_flag ());
iraise e
- and aux_list ?locality ?polymorphism isprogcmd l =
- List.iter (aux false) (List.map snd l)
in
if verbosely then Flags.verbosely (aux false) c
else aux false c