From e8a6467545c2814c9418889201e8be19c0cef201 Mon Sep 17 00:00:00 2001 From: Emilio Jesus Gallego Arias Date: Wed, 18 Jan 2017 15:46:23 +0100 Subject: [location] Make location optional in Loc.located This completes the Loc.ghost removal, the idea is to gear the API towards optional, but uniform, location handling. We don't print anymore in the case there is no location. This is what the test suite expects. The old printing logic for located items was a bit inconsistent as it sometimes printed and other times it printed nothing as the caller checked for `is_ghost` upstream. --- engine/evarutil.mli | 12 +- engine/evd.ml | 6 +- engine/uState.ml | 2 +- ide/coqOps.ml | 23 +-- ide/ide_slave.ml | 10 +- ide/texmacspp.ml | 140 ++++++------- interp/constrexpr_ops.ml | 22 +-- interp/constrexpr_ops.mli | 6 +- interp/constrextern.ml | 66 +++---- interp/constrintern.ml | 338 ++++++++++++++++---------------- interp/dumpglob.ml | 57 +++--- interp/dumpglob.mli | 20 +- interp/implicit_quantifiers.ml | 32 +-- interp/implicit_quantifiers.mli | 6 +- interp/modintern.ml | 12 +- interp/notation_ops.ml | 59 +++--- interp/reserve.ml | 4 +- interp/smartlocate.ml | 12 +- interp/smartlocate.mli | 2 +- interp/stdarg.mli | 2 +- interp/topconstr.ml | 25 +-- interp/topconstr.mli | 4 +- lib/cWarnings.ml | 2 +- lib/cWarnings.mli | 2 +- lib/loc.ml | 25 ++- lib/loc.mli | 14 +- library/declare.ml | 8 +- library/libnames.mli | 2 +- library/library.ml | 6 +- library/nametab.ml | 6 +- parsing/egramcoq.ml | 8 +- parsing/g_constr.ml4 | 66 ++++--- parsing/g_prim.ml4 | 26 +-- parsing/g_vernac.ml4 | 20 +- plugins/extraction/table.ml | 2 +- plugins/funind/g_indfun.ml4 | 2 +- plugins/funind/glob_term_to_relation.ml | 14 +- plugins/funind/glob_termops.ml | 16 +- plugins/funind/indfun.ml | 4 +- plugins/ltac/extratactics.ml4 | 6 +- plugins/ltac/g_ltac.ml4 | 14 +- plugins/ltac/g_tactic.ml4 | 140 ++++++------- plugins/ltac/pptactic.ml | 12 +- plugins/ltac/taccoerce.ml | 4 +- plugins/ltac/taccoerce.mli | 2 +- plugins/ltac/tacentries.ml | 9 +- plugins/ltac/tacintern.ml | 20 +- plugins/ltac/tacinterp.ml | 36 ++-- plugins/ltac/tacsubst.ml | 2 +- plugins/ltac/tactic_debug.ml | 7 +- plugins/ssrmatching/ssrmatching.ml4 | 12 +- plugins/ssrmatching/ssrmatching.mli | 2 +- pretyping/cases.ml | 32 +-- pretyping/detyping.ml | 8 +- pretyping/evarconv.ml | 2 +- pretyping/evardefine.ml | 4 +- pretyping/evardefine.mli | 2 +- pretyping/glob_ops.ml | 8 +- pretyping/glob_ops.mli | 5 +- pretyping/patternops.ml | 16 +- pretyping/pretype_errors.ml | 2 +- pretyping/pretyping.ml | 112 +++++------ printing/ppconstr.ml | 41 ++-- printing/pputils.ml | 5 +- printing/ppvernac.ml | 20 +- printing/prettyp.ml | 16 +- proofs/evar_refiner.ml | 2 +- proofs/pfedit.mli | 2 +- proofs/proof_global.ml | 2 +- proofs/proof_global.mli | 2 +- stm/stm.ml | 22 +-- tactics/autorewrite.ml | 10 +- tactics/autorewrite.mli | 6 +- tactics/hints.ml | 4 +- tactics/inv.ml | 4 +- tactics/tacticals.ml | 6 +- tactics/tactics.ml | 26 +-- toplevel/coqloop.ml | 6 +- toplevel/vernac.ml | 42 ++-- vernac/classes.ml | 6 +- vernac/command.ml | 16 +- vernac/explainErr.ml | 5 +- vernac/lemmas.ml | 4 +- vernac/metasyntax.ml | 2 +- vernac/obligations.ml | 4 +- vernac/record.ml | 8 +- vernac/topfmt.ml | 2 +- vernac/vernacentries.ml | 96 ++++----- 88 files changed, 953 insertions(+), 946 deletions(-) diff --git a/engine/evarutil.mli b/engine/evarutil.mli index ca9591e71..fcc435a2e 100644 --- a/engine/evarutil.mli +++ b/engine/evarutil.mli @@ -22,13 +22,13 @@ val mk_new_meta : unit -> constr (** {6 Creating a fresh evar given their type and context} *) val new_evar : - env -> 'r Sigma.t -> ?src:Loc.t * Evar_kinds.t -> ?filter:Filter.t -> + env -> 'r Sigma.t -> ?src:Evar_kinds.t Loc.located -> ?filter:Filter.t -> ?candidates:constr list -> ?store:Store.t -> ?naming:Misctypes.intro_pattern_naming_expr -> ?principal:bool -> types -> (constr, 'r) Sigma.sigma val new_pure_evar : - named_context_val -> 'r Sigma.t -> ?src:Loc.t * Evar_kinds.t -> ?filter:Filter.t -> + named_context_val -> 'r Sigma.t -> ?src:Evar_kinds.t Loc.located -> ?filter:Filter.t -> ?candidates:constr list -> ?store:Store.t -> ?naming:Misctypes.intro_pattern_naming_expr -> ?principal:bool -> types -> (evar, 'r) Sigma.sigma @@ -37,7 +37,7 @@ val new_pure_evar_full : 'r Sigma.t -> evar_info -> (evar, 'r) Sigma.sigma (** the same with side-effects *) val e_new_evar : - env -> evar_map ref -> ?src:Loc.t * Evar_kinds.t -> ?filter:Filter.t -> + env -> evar_map ref -> ?src:Evar_kinds.t Loc.located -> ?filter:Filter.t -> ?candidates:constr list -> ?store:Store.t -> ?naming:Misctypes.intro_pattern_naming_expr -> ?principal:bool -> types -> constr @@ -45,12 +45,12 @@ val e_new_evar : (** Create a new Type existential variable, as we keep track of them during type-checking and unification. *) val new_type_evar : - env -> 'r Sigma.t -> ?src:Loc.t * Evar_kinds.t -> ?filter:Filter.t -> + env -> 'r Sigma.t -> ?src:Evar_kinds.t Loc.located -> ?filter:Filter.t -> ?naming:Misctypes.intro_pattern_naming_expr -> ?principal:bool -> rigid -> (constr * sorts, 'r) Sigma.sigma val e_new_type_evar : env -> evar_map ref -> - ?src:Loc.t * Evar_kinds.t -> ?filter:Filter.t -> + ?src:Evar_kinds.t Loc.located -> ?filter:Filter.t -> ?naming:Misctypes.intro_pattern_naming_expr -> ?principal:bool -> rigid -> constr * sorts val new_Type : ?rigid:rigid -> env -> 'r Sigma.t -> (constr, 'r) Sigma.sigma @@ -72,7 +72,7 @@ val e_new_global : evar_map ref -> Globnames.global_reference -> constr as a telescope) is [sign] *) val new_evar_instance : named_context_val -> 'r Sigma.t -> types -> - ?src:Loc.t * Evar_kinds.t -> ?filter:Filter.t -> ?candidates:constr list -> + ?src:Evar_kinds.t Loc.located -> ?filter:Filter.t -> ?candidates:constr list -> ?store:Store.t -> ?naming:Misctypes.intro_pattern_naming_expr -> ?principal:bool -> constr list -> (constr, 'r) Sigma.sigma diff --git a/engine/evd.ml b/engine/evd.ml index 9e81ccd36..b0531d581 100644 --- a/engine/evd.ml +++ b/engine/evd.ml @@ -704,10 +704,10 @@ let extract_all_conv_pbs evd = let loc_of_conv_pb evd (pbty,env,t1,t2) = match kind_of_term (fst (decompose_app t1)) with - | Evar (evk1,_) -> Some (fst (evar_source evk1 evd)) + | Evar (evk1,_) -> fst (evar_source evk1 evd) | _ -> match kind_of_term (fst (decompose_app t2)) with - | Evar (evk2,_) -> Some (fst (evar_source evk2 evd)) + | Evar (evk2,_) -> fst (evar_source evk2 evd) | _ -> None (** The following functions return the set of evars immediately @@ -794,7 +794,7 @@ let make_evar_universe_context e l = | Some us -> List.fold_left (fun uctx (loc,id) -> - fst (UState.new_univ_variable ~loc univ_rigid (Some (Id.to_string id)) uctx)) + fst (UState.new_univ_variable ?loc univ_rigid (Some (Id.to_string id)) uctx)) uctx us (****************************************) diff --git a/engine/uState.ml b/engine/uState.ml index c9653b6cd..eb1acb845 100644 --- a/engine/uState.ml +++ b/engine/uState.ml @@ -255,7 +255,7 @@ let universe_context ?names ctx = let l = try UNameMap.find (Id.to_string id) (fst ctx.uctx_names) with Not_found -> - user_err ~loc ~hdr:"universe_context" + user_err ?loc ~hdr:"universe_context" (str"Universe " ++ Nameops.pr_id id ++ str" is not bound anymore.") in (l :: newinst, (id, l) :: map, Univ.LSet.remove l acc)) pl ([], [], levels) diff --git a/ide/coqOps.ml b/ide/coqOps.ml index 7825fb45e..50d0dc491 100644 --- a/ide/coqOps.ml +++ b/ide/coqOps.ml @@ -14,7 +14,7 @@ open Feedback let b2c = byte_offset_to_char_offset -type flag = [ `INCOMPLETE | `UNSAFE | `PROCESSING | `ERROR of Loc.t * string | `WARNING of Loc.t * string ] +type flag = [ `INCOMPLETE | `UNSAFE | `PROCESSING | `ERROR of string Loc.located | `WARNING of string Loc.located ] type mem_flag = [ `INCOMPLETE | `UNSAFE | `PROCESSING | `ERROR | `WARNING ] let mem_flag_of_flag : flag -> mem_flag = function | `ERROR _ -> `ERROR @@ -467,19 +467,15 @@ object(self) | Message(Error, loc, msg), Some (id,sentence) -> log_pp ?id Pp.(str "ErrorMsg" ++ msg); remove_flag sentence `PROCESSING; - let rmsg = Pp.string_of_ppcmds msg in - Option.iter (fun loc -> - add_flag sentence (`ERROR (loc, rmsg)); - ) loc; + let rmsg = Pp.string_of_ppcmds msg in + add_flag sentence (`ERROR (loc, rmsg)); self#mark_as_needed sentence; self#attach_tooltip ?loc sentence rmsg; self#position_tag_at_sentence ?loc Tags.Script.error sentence | Message(Warning, loc, msg), Some (id,sentence) -> log_pp ?id Pp.(str "WarningMsg" ++ msg); let rmsg = Pp.string_of_ppcmds msg in - Option.iter (fun loc -> - add_flag sentence (`WARNING (loc, rmsg)); - ) loc; + add_flag sentence (`WARNING (loc, rmsg)); self#attach_tooltip ?loc sentence rmsg; self#position_tag_at_sentence ?loc Tags.Script.warning sentence; messages#push Warning msg @@ -683,14 +679,13 @@ object(self) let extract_error s = match List.find (function `ERROR _ -> true | _ -> false) s.flags with | `ERROR (loc, msg) -> - let iter = - if Loc.is_ghost loc then - buffer#get_iter_at_mark s.start - else + let iter = begin match loc with + | None -> buffer#get_iter_at_mark s.start + | Some loc -> let (iter, _, phrase) = self#get_sentence s in let (start, _) = Loc.unloc loc in - iter#forward_chars (b2c phrase start) in - iter#line + 1, msg + iter#forward_chars (b2c phrase start) + end in iter#line + 1, msg | _ -> assert false in List.rev (Doc.fold_all document [] (fun acc _ _ s -> diff --git a/ide/ide_slave.ml b/ide/ide_slave.ml index 7f30a4acc..bbc3e4bcb 100644 --- a/ide/ide_slave.ml +++ b/ide/ide_slave.ml @@ -64,8 +64,8 @@ let is_known_option cmd = match cmd with (** Check whether a command is forbidden in the IDE *) let ide_cmd_checks ~id (loc,ast) = - let user_error s = CErrors.user_err ~loc ~hdr:"CoqIde" (str s) in - let warn msg = Feedback.(feedback ~id (Message (Warning, Some loc, strbrk msg))) in + let user_error s = CErrors.user_err ?loc ~hdr:"CoqIde" (str s) in + let warn msg = Feedback.(feedback ~id (Message (Warning, loc, strbrk msg))) in if is_debug ast then user_error "Debug mode not available in the IDE"; if is_known_option ast then @@ -342,8 +342,8 @@ let about () = { let handle_exn (e, info) = let dummy = Stateid.dummy in let loc_of e = match Loc.get_loc e with - | Some loc when not (Loc.is_ghost loc) -> Some (Loc.unloc loc) - | _ -> None in + | Some loc -> Some (Loc.unloc loc) + | _ -> None in let mk_msg () = CErrors.print ~info e in match e with | CErrors.Drop -> dummy, None, Pp.str "Drop is not allowed by coqide!" @@ -391,7 +391,7 @@ let quit = ref false let print_ast id = match Stm.get_ast id with | Some (loc, expr) -> begin - try Texmacspp.tmpp ~loc expr + try Texmacspp.tmpp ?loc expr with e -> Xml_datatype.PCData ("ERROR " ^ Printexc.to_string e) end | None -> Xml_datatype.PCData "ERROR" diff --git a/ide/texmacspp.ml b/ide/texmacspp.ml index 0a07a830b..071861e27 100644 --- a/ide/texmacspp.ml +++ b/ide/texmacspp.ml @@ -108,7 +108,7 @@ let xmlPatvar ?loc id = xmlWithLoc ?loc "patvar" ["id", id] [] let xmlReference ref = let name = Libnames.string_of_reference ref in - let i, j = Loc.unloc (Libnames.loc_of_reference ref) in + let i, j = Option.cata Loc.unloc (0,0) (Libnames.loc_of_reference ref) in let b, e = string_of_int i, string_of_int j in Element("reference",["name", name; "begin", b; "end", e] ,[]) @@ -189,7 +189,7 @@ match sm with | SetOnlyParsing -> ["onlyparsing", ""] | SetCompatVersion v -> ["compat", Flags.pr_version v] | SetFormat (system, (loc, s)) -> - let start, stop = unlock ~loc in + let start, stop = unlock ?loc in ["format-"^system, s; "begin", start; "end", stop] let string_of_assumption_kind l a many = @@ -217,7 +217,7 @@ let rec pp_bindlist bl = let names = (List.map (fun (loc, name) -> - xmlCst ~loc (string_of_name name)) loc_names) in + xmlCst ?loc (string_of_name name)) loc_names) in match e with | _, CHole _ -> names | _ -> names @ [pp_expr e]) @@ -232,12 +232,12 @@ and pp_local_binder lb = (* don't know what it is for now *) | CLocalDef ((loc, nam), ce, ty) -> let attrs = ["name", string_of_name nam] in let value = match ty with - Some t -> Loc.tag ~loc:(Loc.merge (constr_loc ce) (constr_loc t)) @@ CCast (ce, CastConv t) + Some t -> Loc.tag ?loc:(Loc.merge_opt (constr_loc ce) (constr_loc t)) @@ CCast (ce, CastConv t) | None -> ce in pp_expr ~attr:attrs value | CLocalAssum (namll, _, ce) -> let ppl = - List.map (fun (loc, nam) -> (xmlCst ~loc (string_of_name nam))) namll in + List.map (fun (loc, nam) -> (xmlCst ?loc (string_of_name nam))) namll in xmlTyped (ppl @ [pp_expr ce]) | CLocalPattern _ -> assert false @@ -247,7 +247,7 @@ and pp_local_decl_expr lde = (* don't know what it is for now *) | DefExpr (_, ce, _) -> pp_expr ce and pp_inductive_expr ((_, ((l, id),_)), lbl, ceo, _, cl_or_rdexpr) = (* inductive_expr *) - let b,e = Loc.unloc l in + let b,e = Option.cata Loc.unloc (0,0) l in let location = ["begin", string_of_int b; "end", string_of_int e] in [Element ("lident", ["name", Id.to_string id] @ location, [])] @ (* inductive name *) begin match cl_or_rdexpr with @@ -265,7 +265,7 @@ and pp_recursion_order_expr optid roe = (* don't know what it is for now *) match optid with | None -> [] | Some (loc, id) -> - let start, stop = unlock ~loc in + let start, stop = unlock ?loc in ["begin", start; "end", stop ; "name", Id.to_string id] in let kind, expr = match roe with @@ -276,7 +276,7 @@ and pp_recursion_order_expr optid roe = (* don't know what it is for now *) Element ("recursion_order", ["kind", kind] @ attrs, expr) and pp_fixpoint_expr (((loc, id), pl), (optid, roe), lbl, ce, ceo) = (* fixpoint_expr *) - let start, stop = unlock ~loc in + let start, stop = unlock ?loc in let id = Id.to_string id in [Element ("lident", ["begin", start; "end", stop ; "name", id], [])] @ (* fixpoint name *) @@ -290,7 +290,7 @@ and pp_fixpoint_expr (((loc, id), pl), (optid, roe), lbl, ce, ceo) = and pp_cofixpoint_expr (((loc, id), pl), lbl, ce, ceo) = (* cofixpoint_expr *) (* Nota: it is like fixpoint_expr without (optid, roe) * so could be merged if there is no more differences *) - let start, stop = unlock ~loc in + let start, stop = unlock ?loc in let id = Id.to_string id in [Element ("lident", ["begin", start; "end", stop ; "name", id], [])] @ (* cofixpoint name *) @@ -300,23 +300,23 @@ and pp_cofixpoint_expr (((loc, id), pl), lbl, ce, ceo) = (* cofixpoint_expr *) | Some ce -> [pp_expr ce] | None -> [] end -and pp_lident (loc, id) = xmlCst ~loc (Id.to_string id) +and pp_lident (loc, id) = xmlCst ?loc (Id.to_string id) and pp_simple_binder (idl, ce) = List.map pp_lident idl @ [pp_expr ce] and pp_cases_pattern_expr (loc, cpe) = match cpe with | CPatAlias (cpe, id) -> - xmlApply ~loc - (xmlOperator ~loc ~attr:["name", string_of_id id] "alias" :: + xmlApply ?loc + (xmlOperator ?loc ~attr:["name", string_of_id id] "alias" :: [pp_cases_pattern_expr cpe]) | CPatCstr (ref, None, cpel2) -> - xmlApply ~loc - (xmlOperator ~loc "reference" + xmlApply ?loc + (xmlOperator ?loc "reference" ~attr:["name", Libnames.string_of_reference ref] :: [Element ("impargs", [], []); Element ("args", [], (List.map pp_cases_pattern_expr cpel2))]) | CPatCstr (ref, Some cpel1, cpel2) -> - xmlApply ~loc - (xmlOperator ~loc "reference" + xmlApply ?loc + (xmlOperator ?loc "reference" ~attr:["name", Libnames.string_of_reference ref] :: [Element ("impargs", [], (List.map pp_cases_pattern_expr cpel1)); Element ("args", [], (List.map pp_cases_pattern_expr cpel2))]) @@ -324,13 +324,13 @@ and pp_cases_pattern_expr (loc, cpe) = let attrs = match optr with | None -> [] | Some r -> ["name", Libnames.string_of_reference r] in - xmlApply ~loc (xmlOperator ~loc "atom" ~attr:attrs :: []) + xmlApply ?loc (xmlOperator ?loc "atom" ~attr:attrs :: []) | CPatOr cpel -> - xmlApply ~loc (xmlOperator ~loc "or" :: List.map pp_cases_pattern_expr cpel) + xmlApply ?loc (xmlOperator ?loc "or" :: List.map pp_cases_pattern_expr cpel) | CPatNotation (n, (subst_constr, subst_rec), cpel) -> - xmlApply ~loc - (xmlOperator ~loc "notation" :: - [xmlOperator ~loc n; + xmlApply ?loc + (xmlOperator ?loc "notation" :: + [xmlOperator ?loc n; Element ("subst", [], [Element ("subterms", [], List.map pp_cases_pattern_expr subst_constr); @@ -341,29 +341,29 @@ and pp_cases_pattern_expr (loc, cpe) = List.map pp_cases_pattern_expr cpel)) subst_rec)]); Element ("args", [], (List.map pp_cases_pattern_expr cpel))]) - | CPatPrim tok -> pp_token ~loc tok + | CPatPrim tok -> pp_token ?loc tok | CPatRecord rcl -> - xmlApply ~loc - (xmlOperator ~loc "record" :: + xmlApply ?loc + (xmlOperator ?loc "record" :: List.map (fun (r, cpe) -> Element ("field", ["reference", Libnames.string_of_reference r], [pp_cases_pattern_expr cpe])) rcl) | CPatDelimiters (delim, cpe) -> - xmlApply ~loc - (xmlOperator ~loc "delimiter" ~attr:["name", delim] :: + xmlApply ?loc + (xmlOperator ?loc "delimiter" ~attr:["name", delim] :: [pp_cases_pattern_expr cpe]) | CPatCast _ -> assert false and pp_case_expr (e, name, pat) = match name, pat with | None, None -> xmlScrutinee [pp_expr e] | Some (loc, name), None -> - let start, stop= unlock ~loc in + let start, stop= unlock ?loc in xmlScrutinee ~attr:["name", string_of_name name; "begin", start; "end", stop] [pp_expr e] | Some (loc, name), Some p -> - let start, stop= unlock ~loc in + let start, stop= unlock ?loc in xmlScrutinee ~attr:["name", string_of_name name; "begin", start; "end", stop] [Element ("in", [], [pp_cases_pattern_expr p]) ; pp_expr e] @@ -393,77 +393,77 @@ and pp_const_expr_list cel = and pp_expr ?(attr=[]) (loc, e) = match e with | CRef (r, _) -> - xmlCst ~loc:(Libnames.loc_of_reference r) ~attr (Libnames.string_of_reference r) + xmlCst ?loc:(Libnames.loc_of_reference r) ~attr (Libnames.string_of_reference r) | CProdN (bl, e) -> - xmlApply ~loc - (xmlOperator ~loc "forall" :: [pp_bindlist bl] @ [pp_expr e]) + xmlApply ?loc + (xmlOperator ?loc "forall" :: [pp_bindlist bl] @ [pp_expr e]) | CApp ((_, hd), args) -> - xmlApply ~loc ~attr (pp_expr hd :: List.map (fun (e,_) -> pp_expr e) args) + xmlApply ?loc ~attr (pp_expr hd :: List.map (fun (e,_) -> pp_expr e) args) | CAppExpl ((_, r, _), args) -> - xmlApply ~loc ~attr - (xmlCst ~loc:(Libnames.loc_of_reference r) (Libnames.string_of_reference r) + xmlApply ?loc ~attr + (xmlCst ?loc:(Libnames.loc_of_reference r) (Libnames.string_of_reference r) :: List.map pp_expr args) | CNotation (notation, ([],[],[])) -> - xmlOperator ~loc notation + xmlOperator ?loc notation | CNotation (notation, (args, cell, lbll)) -> let fmts = Notation.find_notation_extra_printing_rules notation in - let oper = xmlOperator ~loc notation ~pprules:fmts in + let oper = xmlOperator ?loc notation ~pprules:fmts in let cels = List.map pp_const_expr_list cell in let lbls = List.map pp_local_binder_list lbll in let args = List.map pp_expr args in - xmlApply ~loc (oper :: (List.sort compare_begin_att (args @ cels @ lbls))) + xmlApply ?loc (oper :: (List.sort compare_begin_att (args @ cels @ lbls))) | CSort(s) -> - xmlOperator ~loc (string_of_glob_sort s) + xmlOperator ?loc (string_of_glob_sort s) | CDelimiters (scope, ce) -> - xmlApply ~loc (xmlOperator ~loc "delimiter" ~attr:["name", scope] :: + xmlApply ?loc (xmlOperator ?loc "delimiter" ~attr:["name", scope] :: [pp_expr ce]) - | CPrim tok -> pp_token ~loc tok + | CPrim tok -> pp_token ?loc tok | CGeneralization (kind, _, e) -> let kind= match kind with | Explicit -> "explicit" | Implicit -> "implicit" in - xmlApply ~loc - (xmlOperator ~loc ~attr:["kind", kind] "generalization" :: [pp_expr e]) + xmlApply ?loc + (xmlOperator ?loc ~attr:["kind", kind] "generalization" :: [pp_expr e]) | CCast (e, tc) -> begin match tc with | CastConv t | CastVM t |CastNative t -> - xmlApply ~loc - (xmlOperator ~loc ":" ~attr:["kind", (string_of_cast_sort tc)] :: + xmlApply ?loc + (xmlOperator ?loc ":" ~attr:["kind", (string_of_cast_sort tc)] :: [pp_expr e; pp_expr t]) | CastCoerce -> - xmlApply ~loc - (xmlOperator ~loc ":" ~attr:["kind", "CastCoerce"] :: + xmlApply ?loc + (xmlOperator ?loc ":" ~attr:["kind", "CastCoerce"] :: [pp_expr e]) end | CEvar (ek, cel) -> let ppcel = List.map (fun (id,e) -> xmlAssign id (pp_expr e)) cel in - xmlApply ~loc - (xmlOperator ~loc "evar" ~attr:["id", string_of_id ek] :: + xmlApply ?loc + (xmlOperator ?loc "evar" ~attr:["id", string_of_id ek] :: ppcel) - | CPatVar id -> xmlPatvar ~loc (string_of_id id) - | CHole (_, _, _) -> xmlCst ~loc ~attr "_" + | CPatVar id -> xmlPatvar ?loc (string_of_id id) + | CHole (_, _, _) -> xmlCst ?loc ~attr "_" | CIf (test, (_, ret), th, el) -> let return = match ret with | None -> [] | Some r -> [xmlReturn [pp_expr r]] in - xmlApply ~loc - (xmlOperator ~loc "if" :: + xmlApply ?loc + (xmlOperator ?loc "if" :: return @ [pp_expr th] @ [pp_expr el]) | CLetTuple (names, (_, ret), value, body) -> let return = match ret with | None -> [] | Some r -> [xmlReturn [pp_expr r]] in - xmlApply ~loc - (xmlOperator ~loc "lettuple" :: + xmlApply ?loc + (xmlOperator ?loc "lettuple" :: return @ - (List.map (fun (loc, var) -> xmlCst ~loc (string_of_name var)) names) @ + (List.map (fun (loc, var) -> xmlCst ?loc (string_of_name var)) names) @ [pp_expr value; pp_expr body]) | CCases (sty, ret, cel, bel) -> let return = match ret with | None -> [] | Some r -> [xmlReturn [pp_expr r]] in - xmlApply ~loc - (xmlOperator ~loc ~attr:["style", (string_of_case_style sty)] "match" :: + xmlApply ?loc + (xmlOperator ?loc ~attr:["style", (string_of_case_style sty)] "match" :: (return @ [Element ("scrutinees", [], List.map pp_case_expr cel)] @ [pp_branch_expr_list bel])) @@ -471,18 +471,18 @@ and pp_expr ?(attr=[]) (loc, e) = | CLetIn ((varloc, var), value, typ, body) -> let (loc, value) = match typ with | Some t -> - Loc.tag ~loc:(Loc.merge (constr_loc value) (constr_loc t)) (CCast (value, CastConv t)) + Loc.tag ?loc:(Loc.merge_opt (constr_loc value) (constr_loc t)) (CCast (value, CastConv t)) | None -> value in - xmlApply ~loc - (xmlOperator ~loc "let" :: - [xmlCst ~loc:varloc (string_of_name var) ; pp_expr (Loc.tag ~loc value); pp_expr body]) + xmlApply ?loc + (xmlOperator ?loc "let" :: + [xmlCst ?loc:varloc (string_of_name var) ; pp_expr (Loc.tag ?loc value); pp_expr body]) | CLambdaN (bl, e) -> - xmlApply ~loc - (xmlOperator ~loc "lambda" :: [pp_bindlist bl] @ [pp_expr e]) + xmlApply ?loc + (xmlOperator ?loc "lambda" :: [pp_bindlist bl] @ [pp_expr e]) | CCoFix (_, _) -> assert false | CFix (lid, fel) -> - xmlApply ~loc - (xmlOperator ~loc "fix" :: + xmlApply ?loc + (xmlOperator ?loc "fix" :: List.flatten (List.map (fun (a,b,cl,c,d) -> pp_fixpoint_expr ((a,None),b,cl,c,Some d)) fel)) @@ -499,11 +499,11 @@ let rec tmpp ?loc v = | VernacLoad (verbose,f) -> xmlWithLoc ?loc "load" ["verbose",string_of_bool verbose;"file",f] [] | VernacTime (loc,e) -> - xmlApply ~loc (Element("time",[],[]) :: - [tmpp ~loc e]) + xmlApply ?loc (Element("time",[],[]) :: + [tmpp ?loc e]) | VernacRedirect (s, (loc,e)) -> - xmlApply ~loc (Element("redirect",["path", s],[]) :: - [tmpp ~loc e]) + xmlApply ?loc (Element("redirect",["path", s],[]) :: + [tmpp ?loc e]) | VernacTimeout (s,e) -> xmlApply ?loc (Element("timeout",["val",string_of_int s],[]) :: [tmpp ?loc e]) @@ -586,7 +586,7 @@ let rec tmpp ?loc v = xmlAssumption ?loc kind exprs | VernacInductive (_, _, iednll) -> let kind = - let (_, _, _, k, _),_ = List.hd iednll in + let (_, _, _, k, _), _ = List.hd iednll in begin match k with | Record -> "Record" diff --git a/interp/constrexpr_ops.ml b/interp/constrexpr_ops.ml index 4b61ab494..ce349a63f 100644 --- a/interp/constrexpr_ops.ml +++ b/interp/constrexpr_ops.ml @@ -236,14 +236,14 @@ let raw_cases_pattern_expr_loc (l, _) = l let local_binder_loc = function | CLocalAssum ((loc,_)::_,_,t) - | CLocalDef ((loc,_),t,None) -> Loc.merge loc (constr_loc t) - | CLocalDef ((loc,_),b,Some t) -> Loc.merge loc (Loc.merge (constr_loc b) (constr_loc t)) + | CLocalDef ((loc,_),t,None) -> Loc.merge_opt loc (constr_loc t) + | CLocalDef ((loc,_),b,Some t) -> Loc.merge_opt loc (Loc.merge_opt (constr_loc b) (constr_loc t)) | CLocalAssum ([],_,_) -> assert false | CLocalPattern (loc,_) -> loc let local_binders_loc bll = match bll with | [] -> None - | h :: l -> Some (Loc.merge (local_binder_loc h) (local_binder_loc (List.last bll))) + | h :: l -> Loc.merge_opt (local_binder_loc h) (local_binder_loc (List.last bll)) (** Pseudo-constructors *) @@ -274,27 +274,27 @@ let expand_binders ?loc mkC bl c = | b :: bl -> match b with | CLocalDef ((loc1,_) as n, oty, b) -> - let env, c = loop ~loc:(Loc.opt_merge loc1 loc) bl c in + let env, c = loop ?loc:(Loc.merge_opt loc1 loc) bl c in let env = add_name_in_env env n in (env, Loc.tag ?loc @@ CLetIn (n,oty,b,c)) | CLocalAssum ((loc1,_)::_ as nl, bk, t) -> - let env, c = loop ~loc:(Loc.opt_merge loc1 loc) bl c in + let env, c = loop ?loc:(Loc.merge_opt loc1 loc) bl c in let env = List.fold_left add_name_in_env env nl in (env, mkC ?loc (nl,bk,t) c) | CLocalAssum ([],_,_) -> loop ?loc bl c | CLocalPattern (loc1, (p, ty)) -> - let env, c = loop ~loc:(Loc.opt_merge loc1 loc) bl c in + let env, c = loop ?loc:(Loc.merge_opt loc1 loc) bl c in let ni = Hook.get fresh_var env c in let id = (loc1, Name ni) in let ty = match ty with | Some ty -> ty - | None -> Loc.tag ~loc:loc1 @@ CHole (None, IntroAnonymous, None) + | None -> Loc.tag ?loc:loc1 @@ CHole (None, IntroAnonymous, None) in let e = Loc.tag @@ CRef (Libnames.Ident (loc1, ni), None) in let c = Loc.tag ?loc @@ CCases (LetPatternStyle, None, [(e,None,None)], - [(Loc.tag ~loc:loc1 ([(loc1,[p])], c))]) + [(Loc.tag ?loc:loc1 ([(loc1,[p])], c))]) in (ni :: env, mkC ?loc ([id],Default Explicit,ty) c) in @@ -316,12 +316,12 @@ let prod_constr_expr c bl = mkCProdN ?loc:(local_binders_loc bl) bl c let coerce_reference_to_id = function | Ident (_,id) -> id | Qualid (loc,_) -> - CErrors.user_err ~loc ~hdr:"coerce_reference_to_id" + CErrors.user_err ?loc ~hdr:"coerce_reference_to_id" (str "This expression should be a simple identifier.") let coerce_to_id = function | _loc, CRef (Ident (loc,id),_) -> (loc,id) - | a -> CErrors.user_err ~loc:(constr_loc a) + | a -> CErrors.user_err ?loc:(constr_loc a) ~hdr:"coerce_to_id" (str "This expression should be a simple identifier.") @@ -329,5 +329,5 @@ let coerce_to_name = function | _loc, CRef (Ident (loc,id),_) -> (loc,Name id) | loc, CHole (_,_,_) -> (loc,Anonymous) | a -> CErrors.user_err - ~loc:(constr_loc a) ~hdr:"coerce_to_name" + ?loc:(constr_loc a) ~hdr:"coerce_to_name" (str "This expression should be a name.") diff --git a/interp/constrexpr_ops.mli b/interp/constrexpr_ops.mli index 82e4f54b0..d51576c04 100644 --- a/interp/constrexpr_ops.mli +++ b/interp/constrexpr_ops.mli @@ -34,9 +34,9 @@ val binder_kind_eq : binder_kind -> binder_kind -> bool (** {6 Retrieving locations} *) -val constr_loc : constr_expr -> Loc.t -val cases_pattern_expr_loc : cases_pattern_expr -> Loc.t -val raw_cases_pattern_expr_loc : raw_cases_pattern_expr -> Loc.t +val constr_loc : constr_expr -> Loc.t option +val cases_pattern_expr_loc : cases_pattern_expr -> Loc.t option +val raw_cases_pattern_expr_loc : raw_cases_pattern_expr -> Loc.t option val local_binders_loc : local_binder_expr list -> Loc.t option (** {6 Constructors}*) diff --git a/interp/constrextern.ml b/interp/constrextern.ml index 5960a6baa..30b81ecc4 100644 --- a/interp/constrextern.ml +++ b/interp/constrextern.ml @@ -108,7 +108,7 @@ let is_record indsp = let encode_record r = let indsp = global_inductive r in if not (is_record indsp) then - user_err ~loc:(loc_of_reference r) ~hdr:"encode_record" + user_err ?loc:(loc_of_reference r) ~hdr:"encode_record" (str "This type is not a structure type."); indsp @@ -259,11 +259,11 @@ let make_notation_gen loc ntn mknot mkprim destprim l = let make_notation loc ntn (terms,termlists,binders as subst) = if not (List.is_empty termlists) || not (List.is_empty binders) then - Loc.tag ~loc @@ CNotation (ntn,subst) + Loc.tag ?loc @@ CNotation (ntn,subst) else make_notation_gen loc ntn - (fun (loc,ntn,l) -> Loc.tag ~loc @@ CNotation (ntn,(l,[],[]))) - (fun (loc,p) -> Loc.tag ~loc @@ CPrim p) + (fun (loc,ntn,l) -> Loc.tag ?loc @@ CNotation (ntn,(l,[],[]))) + (fun (loc,p) -> Loc.tag ?loc @@ CPrim p) destPrim terms let make_pat_notation ?loc ntn (terms,termlists as subst) args = @@ -293,9 +293,9 @@ let rec extern_cases_pattern_in_scope (scopes:local_scopes) vars pat = match pat with | loc, PatCstr(cstrsp,args,na) when !Flags.in_debugger||Inductiveops.constructor_has_local_defs cstrsp -> - let c = extern_reference ~loc Id.Set.empty (ConstructRef cstrsp) in + let c = extern_reference ?loc Id.Set.empty (ConstructRef cstrsp) in let args = List.map (extern_cases_pattern_in_scope scopes vars) args in - Loc.tag ~loc @@ CPatCstr (c, Some (add_patt_for_params (fst cstrsp) args), []) + Loc.tag ?loc @@ CPatCstr (c, Some (add_patt_for_params (fst cstrsp) args), []) | _ -> try if !Flags.raw_print || !print_no_symbol then raise No_match; @@ -304,7 +304,7 @@ let rec extern_cases_pattern_in_scope (scopes:local_scopes) vars pat = | None -> raise No_match | Some key -> let loc = cases_pattern_loc pat in - insert_pat_alias ~loc (insert_pat_delimiters ~loc (Loc.tag ~loc @@ CPatPrim p) key) na + insert_pat_alias ?loc (insert_pat_delimiters ?loc (Loc.tag ?loc @@ CPatPrim p) key) na with No_match -> try if !Flags.raw_print || !print_no_symbol then raise No_match; @@ -312,8 +312,8 @@ let rec extern_cases_pattern_in_scope (scopes:local_scopes) vars pat = (uninterp_cases_pattern_notations pat) with No_match -> match pat with - | loc, PatVar (Name id) -> Loc.tag ~loc @@ CPatAtom (Some (Ident (loc,id))) - | loc, PatVar (Anonymous) -> Loc.tag ~loc @@ CPatAtom None + | loc, PatVar (Name id) -> Loc.tag ?loc @@ CPatAtom (Some (Ident (loc,id))) + | loc, PatVar (Anonymous) -> Loc.tag ?loc @@ CPatAtom None | loc, PatCstr(cstrsp,args,na) -> let args = List.map (extern_cases_pattern_in_scope scopes vars) args in let p = @@ -330,22 +330,22 @@ let rec extern_cases_pattern_in_scope (scopes:local_scopes) vars pat = | (_loc, CPatAtom(None)) :: tail -> ip q tail acc (* we don't want to have 'x = _' in our patterns *) | head :: tail -> ip q tail - ((extern_reference ~loc Id.Set.empty (ConstRef c), head) :: acc) + ((extern_reference ?loc Id.Set.empty (ConstRef c), head) :: acc) in - Loc.tag ~loc @@ CPatRecord(List.rev (ip projs args [])) + Loc.tag ?loc @@ CPatRecord(List.rev (ip projs args [])) with Not_found | No_match | Exit -> - let c = extern_reference ~loc Id.Set.empty (ConstructRef cstrsp) in + let c = extern_reference ?loc Id.Set.empty (ConstructRef cstrsp) in if !Topconstr.asymmetric_patterns then if pattern_printable_in_both_syntax cstrsp - then Loc.tag ~loc @@ CPatCstr (c, None, args) - else Loc.tag ~loc @@ CPatCstr (c, Some (add_patt_for_params (fst cstrsp) args), []) + then Loc.tag ?loc @@ CPatCstr (c, None, args) + else Loc.tag ?loc @@ CPatCstr (c, Some (add_patt_for_params (fst cstrsp) args), []) else let full_args = add_patt_for_params (fst cstrsp) args in match drop_implicits_in_patt (ConstructRef cstrsp) 0 full_args with - | Some true_args -> Loc.tag ~loc @@ CPatCstr (c, None, true_args) - | None -> Loc.tag ~loc @@ CPatCstr (c, Some full_args, []) - in insert_pat_alias ~loc p na + | Some true_args -> Loc.tag ?loc @@ CPatCstr (c, None, true_args) + | None -> Loc.tag ?loc @@ CPatCstr (c, Some full_args, []) + in insert_pat_alias ?loc p na and apply_notation_to_pattern ?loc gr ((subst,substlist),(nb_to_drop,more_args)) (tmp_scope, scopes as allscopes) vars = function @@ -398,11 +398,11 @@ and extern_notation_pattern (tmp_scope,scopes as allscopes) vars (loc, t) = func if List.mem keyrule !print_non_active_notations then raise No_match; match t with | PatCstr (cstr,_,na) -> - let p = apply_notation_to_pattern ~loc (ConstructRef cstr) + let p = apply_notation_to_pattern ?loc (ConstructRef cstr) (match_notation_constr_cases_pattern (loc, t) pat) allscopes vars keyrule in - insert_pat_alias ~loc p na - | PatVar Anonymous -> Loc.tag ~loc @@ CPatAtom None - | PatVar (Name id) -> Loc.tag ~loc @@ CPatAtom (Some (Ident (loc,id))) + insert_pat_alias ?loc p na + | PatVar Anonymous -> Loc.tag ?loc @@ CPatAtom None + | PatVar (Name id) -> Loc.tag ?loc @@ CPatAtom (Some (Ident (loc,id))) with No_match -> extern_notation_pattern allscopes vars (loc, t) rules @@ -582,13 +582,13 @@ let rec remove_coercions inctx c = been confused with ordinary application or would have need a surrounding context and the coercion to funclass would have been made explicit to match *) - if List.is_empty l then a' else Loc.tag ~loc @@ GApp (a',l) + if List.is_empty l then a' else Loc.tag ?loc @@ GApp (a',l) | _ -> c with Not_found -> c) | _ -> c let rec flatten_application = function - | (loc, GApp ((_, GApp(a,l')),l)) -> flatten_application (Loc.tag ~loc @@ GApp (a,l'@l)) + | (loc, GApp ((_, GApp(a,l')),l)) -> flatten_application (Loc.tag ?loc @@ GApp (a,l'@l)) | a -> a (**********************************************************************) @@ -600,7 +600,7 @@ let extern_possible_prim_token scopes r = let (sc,n) = uninterp_prim_token r in match availability_of_prim_token n sc scopes with | None -> None - | Some key -> Some (insert_delimiters (Loc.tag ~loc:(loc_of_glob_constr r) @@ CPrim n) key) + | Some key -> Some (insert_delimiters (Loc.tag ?loc:(loc_of_glob_constr r) @@ CPrim n) key) with No_match -> None @@ -644,10 +644,10 @@ let rec extern inctx scopes vars r = let r'' = flatten_application r' in if !Flags.raw_print || !print_no_symbol then raise No_match; extern_notation scopes vars r'' (uninterp_notations r'') - with No_match -> Loc.map_with_loc (fun ~loc -> function + with No_match -> Loc.map_with_loc (fun ?loc -> function | GRef (ref,us) -> extern_global (select_stronger_impargs (implicits_of_global ref)) - (extern_reference ~loc vars ref) (extern_universes us) + (extern_reference ?loc vars ref) (extern_universes us) | GVar id -> CRef (Ident (loc,id),None) @@ -701,7 +701,7 @@ let rec extern inctx scopes vars r = (* we give up since the constructor is not complete *) | (arg, scopes) :: tail -> let head = extern true scopes vars arg in - ip q locs' tail ((extern_reference ~loc Id.Set.empty (ConstRef c), head) :: acc) + ip q locs' tail ((extern_reference ?loc Id.Set.empty (ConstRef c), head) :: acc) in CRecord (List.rev (ip projs locals args [])) with @@ -709,7 +709,7 @@ let rec extern inctx scopes vars r = let args = extern_args (extern true) vars args in extern_app inctx (select_stronger_impargs (implicits_of_global ref)) - (Some ref,extern_reference ~loc:rloc vars ref) (extern_universes us) args + (Some ref,extern_reference ?loc:rloc vars ref) (extern_universes us) args end | _ -> @@ -781,7 +781,7 @@ let rec extern inctx scopes vars r = let listdecl = Array.mapi (fun i fi -> let (bl,ty,def) = blv.(i), tyv.(i), bv.(i) in - let bl = List.map (extended_glob_local_binder_of_decl ~loc) bl in + let bl = List.map (extended_glob_local_binder_of_decl ?loc) bl in let (assums,ids,bl) = extern_local_binder scopes vars bl in let vars0 = List.fold_right (name_fold Id.Set.add) ids vars in let vars1 = List.fold_right (name_fold Id.Set.add) ids vars' in @@ -798,7 +798,7 @@ let rec extern inctx scopes vars r = | GCoFix n -> let listdecl = Array.mapi (fun i fi -> - let bl = List.map (extended_glob_local_binder_of_decl ~loc) blv.(i) in + let bl = List.map (extended_glob_local_binder_of_decl ?loc) blv.(i) in let (_,ids,bl) = extern_local_binder scopes vars bl in let vars0 = List.fold_right (name_fold Id.Set.add) ids vars in let vars1 = List.fold_right (name_fold Id.Set.add) ids vars' in @@ -871,7 +871,7 @@ and extern_local_binder scopes vars = function (assums,ids, CLocalPattern(Loc.tag @@ (p,ty)) :: l) and extern_eqn inctx scopes vars (loc,(ids,pl,c)) = - Loc.tag ~loc ([loc,List.map (extern_cases_pattern_in_scope scopes vars) pl], + Loc.tag ?loc ([loc,List.map (extern_cases_pattern_in_scope scopes vars) pl], extern inctx scopes vars c) and extern_notation (tmp_scope,scopes as allscopes) vars t = function @@ -943,12 +943,12 @@ and extern_notation (tmp_scope,scopes as allscopes) vars t = function extern true (scopt,scl@scopes) vars c, None) terms in let a = CRef (Qualid (loc, shortest_qualid_of_syndef vars kn),None) in - Loc.tag ~loc @@ if List.is_empty l then a else CApp ((None, Loc.tag a),l) in + Loc.tag ?loc @@ if List.is_empty l then a else CApp ((None, Loc.tag a),l) in if List.is_empty args then e else let args = fill_arg_scopes args argsscopes scopes in let args = extern_args (extern true) vars args in - Loc.tag ~loc @@ explicitize false argsimpls (None,e) args + Loc.tag ?loc @@ explicitize false argsimpls (None,e) args with No_match -> extern_notation allscopes vars t rules diff --git a/interp/constrintern.ml b/interp/constrintern.ml index 585f03808..a672771b1 100644 --- a/interp/constrintern.ml +++ b/interp/constrintern.ml @@ -118,7 +118,7 @@ type internalization_error = | NonLinearPattern of Id.t | BadPatternsNumber of int * int -exception InternalizationError of Loc.t * internalization_error +exception InternalizationError of internalization_error Loc.located let explain_variable_capture id id' = pr_id id ++ str " is dependent in the type of " ++ pr_id id' ++ @@ -271,7 +271,7 @@ let error_expect_binder_notation_type ?loc id = (pr_id id ++ str " is expected to occur in binding position in the right-hand side.") -let set_var_scope loc id istermvar env ntnvars = +let set_var_scope ?loc id istermvar env ntnvars = try let isonlybinding,idscopes,typ = Id.Map.find id ntnvars in if istermvar then isonlybinding := false; @@ -282,12 +282,12 @@ let set_var_scope loc id istermvar env ntnvars = | Some (tmp, scope) -> let s1 = make_current_scope tmp scope in let s2 = make_current_scope env.tmp_scope env.scopes in - if not (List.equal String.equal s1 s2) then error_inconsistent_scope ~loc id s1 s2 + if not (List.equal String.equal s1 s2) then error_inconsistent_scope ?loc id s1 s2 end in match typ with | NtnInternTypeBinder -> - if istermvar then error_expect_binder_notation_type ~loc id + if istermvar then error_expect_binder_notation_type ?loc id | NtnInternTypeConstr -> (* We need sometimes to parse idents at a constr level for factorization and we cannot enforce this constraint: @@ -302,14 +302,14 @@ let set_type_scope env = {env with tmp_scope = Notation.current_type_scope_name let reset_tmp_scope env = {env with tmp_scope = None} -let rec it_mkGProd loc2 env body = +let rec it_mkGProd ?loc env body = match env with - (loc1, (na, bk, t)) :: tl -> it_mkGProd loc2 tl (Loc.tag ~loc:(Loc.merge loc1 loc2) @@ GProd (na, bk, t, body)) + (loc2, (na, bk, t)) :: tl -> it_mkGProd ?loc:loc2 tl (Loc.tag ?loc:(Loc.merge_opt loc loc2) @@ GProd (na, bk, t, body)) | [] -> body -let rec it_mkGLambda loc2 env body = +let rec it_mkGLambda ?loc env body = match env with - (loc1, (na, bk, t)) :: tl -> it_mkGLambda loc2 tl (Loc.tag ~loc:(Loc.merge loc1 loc2) @@ GLambda (na, bk, t, body)) + (loc2, (na, bk, t)) :: tl -> it_mkGLambda ?loc:loc2 tl (Loc.tag ?loc:(Loc.merge_opt loc loc2) @@ GLambda (na, bk, t, body)) | [] -> body (**********************************************************************) @@ -371,15 +371,15 @@ let push_name_env ?(global_level=false) ntnvars implargs env = function | loc,Anonymous -> if global_level then - user_err ~loc (str "Anonymous variables not allowed"); + user_err ?loc (str "Anonymous variables not allowed"); env | loc,Name id -> check_hidden_implicit_parameters id env.impls ; if Id.Map.is_empty ntnvars && Id.equal id ldots_var - then error_ldots_var ~loc; - set_var_scope loc id false env ntnvars; + then error_ldots_var ?loc; + set_var_scope ?loc id false env ntnvars; if global_level then Dumpglob.dump_definition (loc,id) true "var" - else Dumpglob.dump_binding loc id; + else Dumpglob.dump_binding ?loc id; {env with ids = Id.Set.add id env.ids; impls = Id.Map.add id implargs env.impls} let intern_generalized_binder ?(global_level=false) intern_type lvar @@ -393,11 +393,11 @@ let intern_generalized_binder ?(global_level=false) intern_type lvar let ty' = intern_type {env with ids = ids; unb = true} ty in let fvs = Implicit_quantifiers.generalizable_vars_of_glob_constr ~bound:ids ~allowed:ids' ty' in let env' = List.fold_left - (fun env (x, l) -> push_name_env ~global_level lvar (Variable,[],[],[])(*?*) env (l, Name x)) + (fun env (l, x) -> push_name_env ~global_level lvar (Variable,[],[],[])(*?*) env (l, Name x)) env fvs in let bl = List.map - (fun (id, loc) -> - (loc, (Name id, b, Loc.tag ~loc @@ GHole (Evar_kinds.BinderType (Name id), Misctypes.IntroAnonymous, None)))) + (fun (loc, id) -> + (loc, (Name id, b, Loc.tag ?loc @@ GHole (Evar_kinds.BinderType (Name id), Misctypes.IntroAnonymous, None)))) fvs in let na = match na with @@ -424,7 +424,7 @@ let intern_assumption intern lvar env nal bk ty = List.fold_left (fun (env, bl) (loc, na as locna) -> (push_name_env lvar impls env locna, - (loc,(na,k,locate_if_hole ~loc na ty))::bl)) + (Loc.tag ?loc (na,k,locate_if_hole ?loc na ty))::bl)) (env, []) nal | Generalized (b,b',t) -> let env, b = intern_generalized_binder intern_type lvar env (List.hd nal) b b' t ty in @@ -448,20 +448,20 @@ let intern_local_pattern intern lvar env p = List.fold_left (fun env (loc, i) -> let bk = Default Implicit in - let ty = Loc.tag ~loc @@ CHole (None, Misctypes.IntroAnonymous, None) in + let ty = Loc.tag ?loc @@ CHole (None, Misctypes.IntroAnonymous, None) in let n = Name i in let env, _ = intern_assumption intern lvar env [(loc, n)] bk ty in env) env (free_vars_of_pat [] p) -let glob_local_binder_of_extended = Loc.with_loc (fun ~loc -> function +let glob_local_binder_of_extended = Loc.with_loc (fun ?loc -> function | GLocalAssum (na,bk,t) -> (na,bk,None,t) | GLocalDef (na,bk,c,Some t) -> (na,bk,Some c,t) | GLocalDef (na,bk,c,None) -> - let t = Loc.tag ~loc @@ GHole(Evar_kinds.BinderType na,Misctypes.IntroAnonymous,None) in + let t = Loc.tag ?loc @@ GHole(Evar_kinds.BinderType na,Misctypes.IntroAnonymous,None) in (na,bk,Some c,t) | GLocalPattern (_,_,_,_) -> - Loc.raise ~loc (Stream.Error "pattern with quote not allowed here.") + Loc.raise ?loc (Stream.Error "pattern with quote not allowed here.") ) let intern_cases_pattern_fwd = ref (fun _ -> failwith "intern_cases_pattern_fwd") @@ -469,18 +469,18 @@ let intern_cases_pattern_fwd = ref (fun _ -> failwith "intern_cases_pattern_fwd" let intern_local_binder_aux ?(global_level=false) intern lvar (env,bl) = function | CLocalAssum(nal,bk,ty) -> let env, bl' = intern_assumption intern lvar env nal bk ty in - let bl' = List.map (fun (loc,(na,c,t)) -> Loc.tag ~loc @@ GLocalAssum (na,c,t)) bl' in + let bl' = List.map (fun (loc,(na,c,t)) -> Loc.tag ?loc @@ GLocalAssum (na,c,t)) bl' in env, bl' @ bl | CLocalDef((loc,na as locna),def,ty) -> let term = intern env def in let ty = Option.map (intern env) ty in (push_name_env lvar (impls_term_list term) env locna, - (Loc.tag ~loc @@ GLocalDef (na,Explicit,term,ty)) :: bl) + (Loc.tag ?loc @@ GLocalDef (na,Explicit,term,ty)) :: bl) | CLocalPattern (loc,(p,ty)) -> let tyc = match ty with | Some ty -> ty - | None -> Loc.tag ~loc @@ CHole(None,Misctypes.IntroAnonymous,None) + | None -> Loc.tag ?loc @@ CHole(None,Misctypes.IntroAnonymous,None) in let env = intern_local_pattern intern lvar env p in let il = List.map snd (free_vars_of_pat [] p) in @@ -495,7 +495,7 @@ let intern_local_binder_aux ?(global_level=false) intern lvar (env,bl) = functio let bk = Default Explicit in let _, bl' = intern_assumption intern lvar env [na] bk tyc in let _,(_,bk,t) = List.hd bl' in - (env, (Loc.tag ~loc @@ GLocalPattern((cp,il),id,bk,t)) :: bl) + (env, (Loc.tag ?loc @@ GLocalPattern((cp,il),id,bk,t)) :: bl) let intern_generalization intern env lvar loc bk ak c = let c = intern {env with unb = true} c in @@ -517,15 +517,15 @@ let intern_generalization intern env lvar loc bk ak c = | None -> false in if pi then - (fun (id, loc') acc -> - Loc.tag ~loc:(Loc.merge loc' loc) @@ - GProd (Name id, bk, Loc.tag ~loc:loc' @@ GHole (Evar_kinds.BinderType (Name id), Misctypes.IntroAnonymous, None), acc)) + (fun (loc', id) acc -> + Loc.tag ?loc:(Loc.merge_opt loc' loc) @@ + GProd (Name id, bk, Loc.tag ?loc:loc' @@ GHole (Evar_kinds.BinderType (Name id), Misctypes.IntroAnonymous, None), acc)) else - (fun (id, loc') acc -> - Loc.tag ~loc:(Loc.merge loc' loc) @@ - GLambda (Name id, bk, Loc.tag ~loc:loc' @@ GHole (Evar_kinds.BinderType (Name id), Misctypes.IntroAnonymous, None), acc)) + (fun (loc', id) acc -> + Loc.tag ?loc:(Loc.merge_opt loc' loc) @@ + GLambda (Name id, bk, Loc.tag ?loc:loc' @@ GHole (Evar_kinds.BinderType (Name id), Misctypes.IntroAnonymous, None), acc)) in - List.fold_right (fun (id, loc as lid) (env, acc) -> + List.fold_right (fun (loc, id as lid) (env, acc) -> let env' = push_name_env lvar (Variable,[],[],[]) env (loc, Name id) in (env', abs lid acc)) fvs (env,c) in c' @@ -566,44 +566,46 @@ let traverse_binder (terms,_,_ as subst) avoid (renaming,env) = function in (renaming',env), Name id' -type letin_param = - | LPLetIn of Loc.t * (Name.t * glob_constr * glob_constr option) - | LPCases of Loc.t * (cases_pattern * Id.t list) * Id.t +type letin_param_r = + | LPLetIn of Name.t * glob_constr * glob_constr option + | LPCases of (cases_pattern * Id.t list) * Id.t +and letin_param = letin_param_r Loc.located let make_letins = List.fold_right (fun a c -> match a with - | LPLetIn (loc,(na,b,t)) -> - Loc.tag ~loc @@ GLetIn(na,b,t,c) - | LPCases (loc,(cp,il),id) -> - let tt = (Loc.tag ~loc @@ GVar id, (Name id,None)) in - Loc.tag ~loc @@ GCases(Misctypes.LetPatternStyle,None,[tt],[(loc,(il,[cp],c))])) + | loc, LPLetIn (na,b,t) -> + Loc.tag ?loc @@ GLetIn(na,b,t,c) + | loc, LPCases ((cp,il),id) -> + let tt = (Loc.tag ?loc @@ GVar id, (Name id,None)) in + Loc.tag ?loc @@ GCases(Misctypes.LetPatternStyle,None,[tt],[(loc,(il,[cp],c))])) let rec subordinate_letins letins = function (* binders come in reverse order; the non-let are returned in reverse order together *) (* with the subordinated let-in in writing order *) | (loc, GLocalDef (na,_,b,t))::l -> - subordinate_letins (LPLetIn (loc,(na,b,t))::letins) l + subordinate_letins ((Loc.tag ?loc @@ LPLetIn (na,b,t))::letins) l | (loc, GLocalAssum (na,bk,t))::l -> let letins',rest = subordinate_letins [] l in letins',((loc,(na,bk,t)),letins)::rest | (loc, GLocalPattern (u,id,bk,t)) :: l -> - subordinate_letins (LPCases (loc,u,id)::letins) ([Loc.tag ~loc @@ GLocalAssum (Name id,bk,t)] @ l) + subordinate_letins ((Loc.tag ?loc @@ LPCases (u,id))::letins) + ([Loc.tag ?loc @@ GLocalAssum (Name id,bk,t)] @ l) | [] -> letins,[] let terms_of_binders bl = - let rec term_of_pat pt = Loc.map_with_loc (fun ~loc -> function + let rec term_of_pat pt = Loc.map_with_loc (fun ?loc -> function | PatVar (Name id) -> CRef (Ident (loc,id), None) | PatVar (Anonymous) -> error "Cannot turn \"_\" into a term." | PatCstr (c,l,_) -> let r = Qualid (loc,qualid_of_path (path_of_global (ConstructRef c))) in - let hole = Loc.tag ~loc @@ CHole (None,Misctypes.IntroAnonymous,None) in + let hole = Loc.tag ?loc @@ CHole (None,Misctypes.IntroAnonymous,None) in let params = List.make (Inductiveops.inductive_nparams (fst c)) hole in CAppExpl ((None,r,None),params @ List.map term_of_pat l)) pt in let rec extract_variables = function - | (loc, GLocalAssum (Name id,_,_))::l -> (Loc.tag ~loc @@ CRef (Ident (loc,id), None)) :: extract_variables l + | (loc, GLocalAssum (Name id,_,_))::l -> (Loc.tag ?loc @@ CRef (Ident (loc,id), None)) :: extract_variables l | (loc, GLocalDef (Name id,_,_,_))::l -> extract_variables l | (loc, GLocalDef (Anonymous,_,_,_))::l | (loc, GLocalAssum (Anonymous,_,_))::l -> error "Cannot turn \"_\" into a term." @@ -663,7 +665,7 @@ let instantiate_notation_constr loc intern ntnvars subst infos c = let bindings = Id.Map.map mk_env terms in Some (Genintern.generic_substitute_notation bindings arg) in - Loc.tag ~loc @@ GHole (knd, naming, arg) + Loc.tag ?loc @@ GHole (knd, naming, arg) | NBinderList (x,y,iter,terminator) -> (try (* All elements of the list are in scopes (scopt,subscopes) *) @@ -681,24 +683,24 @@ let instantiate_notation_constr loc intern ntnvars subst infos c = let a,letins = snd (Option.get binderopt) in let e = make_letins letins (aux subst' infos c') in let (loc,(na,bk,t)) = a in - Loc.tag ~loc @@ GProd (na,bk,t,e) + Loc.tag ?loc @@ GProd (na,bk,t,e) | NLambda (Name id,NHole _,c') when option_mem_assoc id binderopt -> let a,letins = snd (Option.get binderopt) in let (loc,(na,bk,t)) = a in - Loc.tag ~loc @@ GLambda (na,bk,t,make_letins letins (aux subst' infos c')) + Loc.tag ?loc @@ GLambda (na,bk,t,make_letins letins (aux subst' infos c')) (* Two special cases to keep binder name synchronous with BinderType *) | NProd (na,NHole(Evar_kinds.BinderType na',naming,arg),c') when Name.equal na na' -> let subinfos,na = traverse_binder subst avoid subinfos na in - let ty = Loc.tag ~loc @@ GHole (Evar_kinds.BinderType na,naming,arg) in - Loc.tag ~loc @@ GProd (na,Explicit,ty,aux subst' subinfos c') + let ty = Loc.tag ?loc @@ GHole (Evar_kinds.BinderType na,naming,arg) in + Loc.tag ?loc @@ GProd (na,Explicit,ty,aux subst' subinfos c') | NLambda (na,NHole(Evar_kinds.BinderType na',naming,arg),c') when Name.equal na na' -> let subinfos,na = traverse_binder subst avoid subinfos na in - let ty = Loc.tag ~loc @@ GHole (Evar_kinds.BinderType na,naming,arg) in - Loc.tag ~loc @@ GLambda (na,Explicit,ty,aux subst' subinfos c') + let ty = Loc.tag ?loc @@ GHole (Evar_kinds.BinderType na,naming,arg) in + Loc.tag ?loc @@ GLambda (na,Explicit,ty,aux subst' subinfos c') | t -> - glob_constr_of_notation_constr_with_binders ~loc + glob_constr_of_notation_constr_with_binders ?loc (traverse_binder subst avoid) (aux subst') subinfos t and subst_var (terms, _binderopt, _terminopt) (renaming, env) id = (* subst remembers the delimiters stack in the interpretation *) @@ -708,7 +710,7 @@ let instantiate_notation_constr loc intern ntnvars subst infos c = intern {env with tmp_scope = scopt; scopes = subscopes @ env.scopes} a with Not_found -> - Loc.tag ~loc ( + Loc.tag ?loc ( try GVar (Id.Map.find id renaming) with Not_found -> @@ -729,8 +731,8 @@ let make_subst ids l = let intern_notation intern env lvar loc ntn fullargs = let ntn,(args,argslist,bll as fullargs) = contract_notation ntn fullargs in - let ((ids,c),df) = interp_notation ~loc ntn (env.tmp_scope,env.scopes) in - Dumpglob.dump_notation_location (ntn_loc loc fullargs ntn) ntn df; + let ((ids,c),df) = interp_notation ?loc ntn (env.tmp_scope,env.scopes) in + Dumpglob.dump_notation_location (ntn_loc ?loc fullargs ntn) ntn df; let ids,idsl,idsbl = split_by_type ids in let terms = make_subst ids args in let termlists = make_subst idsl argslist in @@ -748,9 +750,9 @@ let string_of_ty = function | Variable -> "var" let gvar (loc, id) us = match us with -| None -> Loc.tag ~loc @@ GVar id +| None -> Loc.tag ?loc @@ GVar id | Some _ -> - user_err ~loc (str "Variable " ++ pr_id id ++ + user_err ?loc (str "Variable " ++ pr_id id ++ str " cannot have a universe instance") let intern_var genv (ltacvars,ntnvars) namedctx loc id us = @@ -758,9 +760,9 @@ let intern_var genv (ltacvars,ntnvars) namedctx loc id us = try let ty,expl_impls,impls,argsc = Id.Map.find id genv.impls in let expl_impls = List.map - (fun id -> Loc.tag ~loc @@ CRef (Ident (loc,id),None), Some (loc,ExplByName id)) expl_impls in + (fun id -> Loc.tag ?loc @@ CRef (Ident (loc,id),None), Some (loc,ExplByName id)) expl_impls in let tys = string_of_ty ty in - Dumpglob.dump_reference loc "<>" (Id.to_string id) tys; + Dumpglob.dump_reference ?loc "<>" (Id.to_string id) tys; gvar (loc,id) us, make_implicits_list impls, argsc, expl_impls with Not_found -> (* Is [id] bound in current term or is an ltac var bound to constr *) @@ -770,15 +772,15 @@ let intern_var genv (ltacvars,ntnvars) namedctx loc id us = (* Is [id] a notation variable *) else if Id.Map.mem id ntnvars then - (set_var_scope loc id true genv ntnvars; gvar (loc,id) us, [], [], []) + (set_var_scope ?loc id true genv ntnvars; gvar (loc,id) us, [], [], []) (* Is [id] the special variable for recursive notations *) else if Id.equal id ldots_var then if Id.Map.is_empty ntnvars - then error_ldots_var ~loc + then error_ldots_var ?loc else gvar (loc,id) us, [], [], [] else if Id.Set.mem id ltacvars.ltac_bound then (* Is [id] bound to a free name in ltac (this is an ltac error message) *) - user_err ~loc ~hdr:"intern_var" + user_err ?loc ~hdr:"intern_var" (str "variable " ++ pr_id id ++ str " should be bound to a term.") else (* Is [id] a goal or section variable *) @@ -789,8 +791,8 @@ let intern_var genv (ltacvars,ntnvars) namedctx loc id us = let ref = VarRef id in let impls = implicits_of_global ref in let scopes = find_arguments_scope ref in - Dumpglob.dump_reference loc "<>" (string_of_qualid (Decls.variable_secpath id)) "var"; - Loc.tag ~loc @@ GRef (ref, us), impls, scopes, [] + Dumpglob.dump_reference ?loc "<>" (string_of_qualid (Decls.variable_secpath id)) "var"; + Loc.tag ?loc @@ GRef (ref, us), impls, scopes, [] with e when CErrors.noncritical e -> (* [id] a goal variable *) gvar (loc,id) us, [], [], [] @@ -820,11 +822,11 @@ let check_no_explicitation l = | [] -> () | (_, None) :: _ -> assert false | (_, Some (loc, _)) :: _ -> - user_err ~loc (str"Unexpected explicitation of the argument of an abbreviation.") + user_err ?loc (str"Unexpected explicitation of the argument of an abbreviation.") let dump_extended_global loc = function - | TrueGlobal ref -> (*feedback_global loc ref;*) Dumpglob.add_glob loc ref - | SynDef sp -> Dumpglob.add_glob_kn loc sp + | TrueGlobal ref -> (*feedback_global loc ref;*) Dumpglob.add_glob ?loc ref + | SynDef sp -> Dumpglob.add_glob_kn ?loc sp let intern_extended_global_of_qualid (loc,qid) = let r = Nametab.locate_extended qid in dump_extended_global loc r; r @@ -833,18 +835,18 @@ let intern_reference ref = let qid = qualid_of_reference ref in let r = try intern_extended_global_of_qualid qid - with Not_found -> error_global_not_found ~loc:(fst qid) (snd qid) + with Not_found -> error_global_not_found ?loc:(fst qid) (snd qid) in Smartlocate.global_of_extended_global r (* Is it a global reference or a syntactic definition? *) let intern_qualid loc qid intern env lvar us args = match intern_extended_global_of_qualid (loc,qid) with - | TrueGlobal ref -> (Loc.tag ~loc @@ GRef (ref, us)), true, args + | TrueGlobal ref -> (Loc.tag ?loc @@ GRef (ref, us)), true, args | SynDef sp -> let (ids,c) = Syntax_def.search_syntactic_definition sp in let nids = List.length ids in - if List.length args < nids then error_not_enough_arguments ~loc; + if List.length args < nids then error_not_enough_arguments ?loc; let args1,args2 = List.chop nids args in check_no_explicitation args1; let terms = make_subst ids (List.map fst args1) in @@ -854,11 +856,11 @@ let intern_qualid loc qid intern env lvar us args = let c = instantiate_notation_constr loc intern lvar subst infos c in let c = match us, c with | None, _ -> c - | Some _, (loc, GRef (ref, None)) -> Loc.tag ~loc @@ GRef (ref, us) + | Some _, (loc, GRef (ref, None)) -> Loc.tag ?loc @@ GRef (ref, us) | Some _, (loc, GApp ((loc', GRef (ref, None)), arg)) -> - Loc.tag ~loc @@ GApp (Loc.tag ~loc:loc' @@ GRef (ref, us), arg) + Loc.tag ?loc @@ GApp (Loc.tag ?loc:loc' @@ GRef (ref, us), arg) | Some _, _ -> - user_err ~loc (str "Notation " ++ pr_qualid qid ++ + user_err ?loc (str "Notation " ++ pr_qualid qid ++ str " cannot have a universe instance, its expanded head does not start with a reference") in @@ -874,7 +876,7 @@ let intern_applied_reference intern env namedctx (_, ntnvars as lvar) us args = | Qualid (loc, qid) -> let r,projapp,args2 = try intern_qualid loc qid intern env ntnvars us args - with Not_found -> error_global_not_found ~loc qid + with Not_found -> error_global_not_found ?loc qid in let x, imp, scopes, l = find_appl_head_data r in (x,imp,scopes,l), args2 @@ -890,7 +892,7 @@ let intern_applied_reference intern env namedctx (_, ntnvars as lvar) us args = (* Extra allowance for non globalizing functions *) if !interning_grammar || env.unb then (gvar (loc,id) us, [], [], []), args - else error_global_not_found ~loc qid + else error_global_not_found ?loc qid let interp_reference vars r = let (r,_,_,_),_ = @@ -952,7 +954,7 @@ let rec has_duplicate = function | x::l -> if Id.List.mem x l then (Some x) else has_duplicate l let loc_of_lhs lhs = - Loc.merge (fst (List.hd lhs)) (fst (List.last lhs)) + Loc.merge_opt (fst (List.hd lhs)) (fst (List.last lhs)) let check_linearity lhs ids = match has_duplicate ids with @@ -968,7 +970,7 @@ let check_number_of_pattern loc n l = let check_or_pat_variables loc ids idsl = if List.exists (fun ids' -> not (List.eq_set Id.equal ids ids')) idsl then - user_err ~loc (str + user_err ?loc (str "The components of this disjunctive pattern must bind the same variables.") (** Use only when params were NOT asked to the user. @@ -977,7 +979,7 @@ let check_constructor_length env loc cstr len_pl pl0 = let n = len_pl + List.length pl0 in if Int.equal n (Inductiveops.constructor_nallargs cstr) then false else (Int.equal n (Inductiveops.constructor_nalldecls cstr) || - (error_wrong_numarg_constructor ~loc env cstr + (error_wrong_numarg_constructor ?loc env cstr (Inductiveops.constructor_nrealargs cstr))) let add_implicits_check_length fail nargs nargs_with_letin impls_st len_pl1 pl2 = @@ -1002,14 +1004,14 @@ let add_implicits_check_constructor_length env loc c len_pl1 pl2 = let nargs = Inductiveops.constructor_nallargs c in let nargs' = Inductiveops.constructor_nalldecls c in let impls_st = implicits_of_global (ConstructRef c) in - add_implicits_check_length (error_wrong_numarg_constructor ~loc env c) + add_implicits_check_length (error_wrong_numarg_constructor ?loc env c) nargs nargs' impls_st len_pl1 pl2 let add_implicits_check_ind_length env loc c len_pl1 pl2 = let nallargs = inductive_nallargs_env env c in let nalldecls = inductive_nalldecls_env env c in let impls_st = implicits_of_global (IndRef c) in - add_implicits_check_length (error_wrong_numarg_inductive ~loc env c) + add_implicits_check_length (error_wrong_numarg_inductive ?loc env c) nallargs nalldecls impls_st len_pl1 pl2 (** Do not raise NotEnoughArguments thanks to preconditions*) @@ -1020,7 +1022,7 @@ let chop_params_pattern loc ind args with_letin = assert (nparams <= List.length args); let params,args = List.chop nparams args in List.iter (function _, PatVar Anonymous -> () - | loc', PatVar _ | loc', PatCstr(_,_,_) -> error_parameter_not_implicit ~loc:loc') params; + | loc', PatVar _ | loc', PatCstr(_,_,_) -> error_parameter_not_implicit ?loc:loc') params; args let find_constructor loc add_params ref = @@ -1028,10 +1030,10 @@ let find_constructor loc add_params ref = | ConstructRef cstr -> cstr | IndRef _ -> let error = str "There is an inductive name deep in a \"in\" clause." in - user_err ~loc ~hdr:"find_constructor" error + user_err ?loc ~hdr:"find_constructor" error | ConstRef _ | VarRef _ -> let error = str "This reference is not a constructor." in - user_err ~loc ~hdr:"find_constructor" error + user_err ?loc ~hdr:"find_constructor" error in cstr, match add_params with | Some nb_args -> @@ -1053,7 +1055,7 @@ let check_duplicate loc fields = match dups with | [] -> () | (r, _) :: _ -> - user_err ~loc (str "This record defines several times the field " ++ + user_err ?loc (str "This record defines several times the field " ++ pr_reference r ++ str ".") (** [sort_fields ~complete loc fields completer] expects a list @@ -1078,7 +1080,7 @@ let sort_fields ~complete loc fields completer = let gr = global_reference_of_reference first_field_ref in (gr, Recordops.find_projection gr) with Not_found -> - user_err ~loc:(loc_of_reference first_field_ref) ~hdr:"intern" + user_err ?loc:(loc_of_reference first_field_ref) ~hdr:"intern" (pr_reference first_field_ref ++ str": Not a projection") in (* the number of parameters *) @@ -1109,7 +1111,7 @@ let sort_fields ~complete loc fields completer = by a let-in in the record declaration (its value is fixed from other fields). *) if first_field && not regular && complete then - user_err ~loc (str "No local fields allowed in a record construction.") + user_err ?loc (str "No local fields allowed in a record construction.") else if first_field then build_proj_list projs proj_kinds (idx+1) ~acc_first_idx:idx acc else if not regular && complete then @@ -1122,7 +1124,7 @@ let sort_fields ~complete loc fields completer = | None :: projs -> if complete then (* we don't want anonymous fields *) - user_err ~loc (str "This record contains anonymous fields.") + user_err ?loc (str "This record contains anonymous fields.") else (* anonymous arguments don't appear in proj_kinds *) build_proj_list projs proj_kinds (idx+1) ~acc_first_idx acc @@ -1136,13 +1138,13 @@ let sort_fields ~complete loc fields completer = | (field_ref, field_value) :: fields -> let field_glob_ref = try global_reference_of_reference field_ref with Not_found -> - user_err ~loc:(loc_of_reference field_ref) ~hdr:"intern" + user_err ?loc:(loc_of_reference field_ref) ~hdr:"intern" (str "The field \"" ++ pr_reference field_ref ++ str "\" does not exist.") in let remaining_projs, (field_index, _) = let the_proj (idx, glob_ref) = eq_gr field_glob_ref glob_ref in try CList.extract_first the_proj remaining_projs with Not_found -> - user_err ~loc + user_err ?loc (str "This record contains fields of different records.") in index_fields fields remaining_projs ((field_index, field_value) :: acc) @@ -1199,12 +1201,12 @@ let alias_of als = match als.alias_ids with let rec subst_pat_iterator y t (loc, p) = match p with | RCPatAtom id -> - begin match id with Some x when Id.equal x y -> t | _ -> Loc.tag ~loc p end + begin match id with Some x when Id.equal x y -> t | _ -> Loc.tag ?loc p end | RCPatCstr (id,l1,l2) -> - Loc.tag ~loc @@ RCPatCstr (id, List.map (subst_pat_iterator y t) l1, + Loc.tag ?loc @@ RCPatCstr (id, List.map (subst_pat_iterator y t) l1, List.map (subst_pat_iterator y t) l2) - | RCPatAlias (p,a) -> Loc.tag ~loc @@ RCPatAlias (subst_pat_iterator y t p,a) - | RCPatOr pl -> Loc.tag ~loc @@ RCPatOr (List.map (subst_pat_iterator y t) pl) + | RCPatAlias (p,a) -> Loc.tag ?loc @@ RCPatAlias (subst_pat_iterator y t p,a) + | RCPatOr pl -> Loc.tag ?loc @@ RCPatOr (List.map (subst_pat_iterator y t) pl) let drop_notations_pattern looked_for = (* At toplevel, Constructors and Inductives are accepted, in recursive calls @@ -1214,7 +1216,7 @@ let drop_notations_pattern looked_for = if top then looked_for g else match g with ConstructRef _ -> () | _ -> raise Not_found with Not_found -> - error_invalid_pattern_notation ~loc () + error_invalid_pattern_notation ?loc () in let test_kind top = if top then looked_for else function ConstructRef _ -> () | _ -> raise Not_found @@ -1240,7 +1242,7 @@ let drop_notations_pattern looked_for = (* Convention: do not deactivate implicit arguments and scopes for further arguments *) test_kind top g; let nvars = List.length vars in - if List.length pats < nvars then error_not_enough_arguments ~loc; + if List.length pats < nvars then error_not_enough_arguments ?loc; let pats1,pats2 = List.chop nvars pats in let subst = make_subst vars pats1 in let idspl1 = List.map (in_not false loc scopes (subst, Id.Map.empty) []) args in @@ -1249,17 +1251,17 @@ let drop_notations_pattern looked_for = | _ -> raise Not_found) | TrueGlobal g -> test_kind top g; - Dumpglob.add_glob loc g; + Dumpglob.add_glob ?loc g; let (_,argscs) = find_remaining_scopes [] pats g in Some (g,[],List.map2 (fun x -> in_pat false (x,snd scopes)) argscs pats) with Not_found -> None and in_pat top scopes (loc, pt) = match pt with - | CPatAlias (p, id) -> Loc.tag ~loc @@ RCPatAlias (in_pat top scopes p, id) + | CPatAlias (p, id) -> Loc.tag ?loc @@ RCPatAlias (in_pat top scopes p, id) | CPatRecord l -> let sorted_fields = sort_fields ~complete:false loc l (fun _idx -> (loc, CPatAtom None)) in begin match sorted_fields with - | None -> Loc.tag ~loc @@ RCPatAtom None + | None -> Loc.tag ?loc @@ RCPatAtom None | Some (n, head, pl) -> let pl = if !asymmetric_patterns then pl else @@ -1272,7 +1274,7 @@ let drop_notations_pattern looked_for = | CPatCstr (head, None, pl) -> begin match drop_syndef top scopes head pl with - | Some (a,b,c) -> Loc.tag ~loc @@ RCPatCstr(a, b, c) + | Some (a,b,c) -> Loc.tag ?loc @@ RCPatCstr(a, b, c) | None -> raise (InternalizationError (loc,NotAConstructor head)) end | CPatCstr (r, Some expl_pl, pl) -> @@ -1281,36 +1283,36 @@ let drop_notations_pattern looked_for = raise (InternalizationError (loc,NotAConstructor r)) in if expl_pl == [] then (* Convention: (@r) deactivates all further implicit arguments and scopes *) - Loc.tag ~loc @@ RCPatCstr (g, List.map (in_pat false scopes) pl, []) + Loc.tag ?loc @@ RCPatCstr (g, List.map (in_pat false scopes) pl, []) else (* Convention: (@r expl_pl) deactivates implicit arguments in expl_pl and in pl *) (* but not scopes in expl_pl *) let (argscs1,_) = find_remaining_scopes expl_pl pl g in - Loc.tag ~loc @@ RCPatCstr (g, List.map2 (in_pat_sc scopes) argscs1 expl_pl @ List.map (in_pat false scopes) pl, []) + Loc.tag ?loc @@ RCPatCstr (g, List.map2 (in_pat_sc scopes) argscs1 expl_pl @ List.map (in_pat false scopes) pl, []) | CPatNotation ("- _",([_loc,CPatPrim(Numeral p)],[]),[]) when Bigint.is_strictly_pos p -> - fst (Notation.interp_prim_token_cases_pattern_expr ~loc (ensure_kind false loc) (Numeral (Bigint.neg p)) scopes) + fst (Notation.interp_prim_token_cases_pattern_expr ?loc (ensure_kind false loc) (Numeral (Bigint.neg p)) scopes) | CPatNotation ("( _ )",([a],[]),[]) -> in_pat top scopes a | CPatNotation (ntn, fullargs,extrargs) -> let ntn,(args,argsl as fullargs) = contract_pat_notation ntn fullargs in - let ((ids',c),df) = Notation.interp_notation ~loc ntn scopes in + let ((ids',c),df) = Notation.interp_notation ?loc ntn scopes in let (ids',idsl',_) = split_by_type ids' in - Dumpglob.dump_notation_location (patntn_loc loc fullargs ntn) ntn df; + Dumpglob.dump_notation_location (patntn_loc ?loc fullargs ntn) ntn df; let substlist = make_subst idsl' argsl in let subst = make_subst ids' args in in_not top loc scopes (subst,substlist) extrargs c | CPatDelimiters (key, e) -> - in_pat top (None,find_delimiters_scope ~loc key::snd scopes) e - | CPatPrim p -> fst (Notation.interp_prim_token_cases_pattern_expr ~loc (test_kind false) p scopes) + in_pat top (None,find_delimiters_scope ?loc key::snd scopes) e + | CPatPrim p -> fst (Notation.interp_prim_token_cases_pattern_expr ?loc (test_kind false) p scopes) | CPatAtom Some id -> begin match drop_syndef top scopes id [] with - | Some (a,b,c) -> Loc.tag ~loc @@ RCPatCstr (a, b, c) - | None -> Loc.tag ~loc @@ RCPatAtom (Some (find_pattern_variable id)) + | Some (a,b,c) -> Loc.tag ?loc @@ RCPatCstr (a, b, c) + | None -> Loc.tag ?loc @@ RCPatAtom (Some (find_pattern_variable id)) end - | CPatAtom None -> Loc.tag ~loc @@ RCPatAtom None - | CPatOr pl -> Loc.tag ~loc @@ RCPatOr (List.map (in_pat top scopes) pl) + | CPatAtom None -> Loc.tag ?loc @@ RCPatAtom None + | CPatOr pl -> Loc.tag ?loc @@ RCPatOr (List.map (in_pat top scopes) pl) | CPatCast _ -> assert false and in_pat_sc scopes x = in_pat false (x,snd scopes) @@ -1324,21 +1326,21 @@ let drop_notations_pattern looked_for = let (a,(scopt,subscopes)) = Id.Map.find id subst in in_pat top (scopt,subscopes@snd scopes) a with Not_found -> - if Id.equal id ldots_var then Loc.tag ~loc @@ RCPatAtom (Some id) else + if Id.equal id ldots_var then Loc.tag ?loc @@ RCPatAtom (Some id) else anomaly (str "Unbound pattern notation variable: " ++ Id.print id) end | NRef g -> ensure_kind top loc g; let (_,argscs) = find_remaining_scopes [] args g in - Loc.tag ~loc @@ RCPatCstr (g, [], List.map2 (in_pat_sc scopes) argscs args) + Loc.tag ?loc @@ RCPatCstr (g, [], List.map2 (in_pat_sc scopes) argscs args) | NApp (NRef g,pl) -> ensure_kind top loc g; let (argscs1,argscs2) = find_remaining_scopes pl args g in - Loc.tag ~loc @@ RCPatCstr (g, + Loc.tag ?loc @@ RCPatCstr (g, List.map2 (fun x -> in_not false loc (x,snd scopes) fullsubst []) argscs1 pl @ List.map (in_pat false scopes) args, []) | NList (x,y,iter,terminator,lassoc) -> - if not (List.is_empty args) then user_err ~loc + if not (List.is_empty args) then user_err ?loc (strbrk "Application of arguments to a recursive notation not supported in patterns."); (try (* All elements of the list are in scopes (scopt,subscopes) *) @@ -1353,8 +1355,8 @@ let drop_notations_pattern looked_for = anomaly (Pp.str "Inconsistent substitution of recursive notation")) | NHole _ -> let () = assert (List.is_empty args) in - Loc.tag ~loc @@ RCPatAtom None - | t -> error_invalid_pattern_notation ~loc () + Loc.tag ?loc @@ RCPatAtom None + | t -> error_invalid_pattern_notation ?loc () in in_pat true let rec intern_pat genv aliases pat = @@ -1362,7 +1364,7 @@ let rec intern_pat genv aliases pat = let idslpl2 = List.map (intern_pat genv empty_alias) pl2 in let (ids',pll) = product_of_cases_patterns aliases.alias_ids (idslpl1@idslpl2) in let pl' = List.map (fun (asubst,pl) -> - (asubst, Loc.tag ~loc @@ PatCstr (c,chop_params_pattern loc (fst c) pl with_letin,alias_of aliases))) pll in + (asubst, Loc.tag ?loc @@ PatCstr (c,chop_params_pattern loc (fst c) pl with_letin,alias_of aliases))) pll in ids',pl' in match pat with | loc, RCPatAlias (p, id) -> @@ -1382,10 +1384,10 @@ let rec intern_pat genv aliases pat = intern_cstr_with_all_args loc c with_letin idslpl1 (expl_pl@pl2) | loc, RCPatAtom (Some id) -> let aliases = merge_aliases aliases id in - (aliases.alias_ids,[aliases.alias_map, Loc.tag ~loc @@ PatVar (alias_of aliases)]) + (aliases.alias_ids,[aliases.alias_map, Loc.tag ?loc @@ PatVar (alias_of aliases)]) | loc, RCPatAtom (None) -> let { alias_ids = ids; alias_map = asubst; } = aliases in - (ids, [asubst, Loc.tag ~loc @@ PatVar (alias_of aliases)]) + (ids, [asubst, Loc.tag ?loc @@ PatVar (alias_of aliases)]) | loc, RCPatOr pl -> assert (not (List.is_empty pl)); let pl' = List.map (intern_pat genv aliases) pl in @@ -1406,7 +1408,7 @@ let rec intern_pat genv aliases pat = [pattern] rule. *) let rec check_no_patcast (loc, pt) = match pt with | CPatCast (_,_) -> - CErrors.user_err ~loc ~hdr:"check_no_patcast" + CErrors.user_err ?loc ~hdr:"check_no_patcast" (Pp.strbrk "Casts are not supported here.") | CPatDelimiters(_,p) | CPatAlias(p,_) -> check_no_patcast p @@ -1440,11 +1442,11 @@ let intern_ind_pattern genv scopes pat = let no_not = try drop_notations_pattern (function (IndRef _ | ConstructRef _) -> () | _ -> raise Not_found) scopes pat - with InternalizationError(loc,NotAConstructor _) -> error_bad_inductive_type ~loc + with InternalizationError(loc,NotAConstructor _) -> error_bad_inductive_type ?loc in match no_not with | loc, RCPatCstr (head, expl_pl, pl) -> - let c = (function IndRef ind -> ind | _ -> error_bad_inductive_type ~loc) head in + let c = (function IndRef ind -> ind | _ -> error_bad_inductive_type ?loc) head in let with_letin, pl2 = add_implicits_check_ind_length genv loc c (List.length expl_pl) pl in let idslpl1 = List.rev_map (intern_pat genv empty_alias) expl_pl in @@ -1452,8 +1454,8 @@ let intern_ind_pattern genv scopes pat = (with_letin, match product_of_cases_patterns [] (List.rev_append idslpl1 idslpl2) with | _,[_,pl] -> (c,chop_params_pattern loc c pl with_letin) - | _ -> error_bad_inductive_type ~loc) - | x -> error_bad_inductive_type ~loc:(raw_cases_pattern_expr_loc x) + | _ -> error_bad_inductive_type ?loc) + | x -> error_bad_inductive_type ?loc:(raw_cases_pattern_expr_loc x) (**********************************************************************) (* Utilities for application *) @@ -1474,8 +1476,8 @@ let get_implicit_name n imps = Some (Impargs.name_of_implicit (List.nth imps (n-1))) let set_hole_implicit i b = function - | (loc, GRef (r,_)) | (_, GApp ((loc, (GRef (r,_))),_)) -> Loc.tag ~loc (Evar_kinds.ImplicitArg (r,i,b),Misctypes.IntroAnonymous,None) - | (loc, GVar id) -> Loc.tag ~loc (Evar_kinds.ImplicitArg (VarRef id,i,b),Misctypes.IntroAnonymous,None) + | (loc, GRef (r,_)) | (_, GApp ((loc, (GRef (r,_))),_)) -> Loc.tag ?loc (Evar_kinds.ImplicitArg (r,i,b),Misctypes.IntroAnonymous,None) + | (loc, GVar id) -> Loc.tag ?loc (Evar_kinds.ImplicitArg (VarRef id,i,b),Misctypes.IntroAnonymous,None) | _ -> anomaly (Pp.str "Only refs have implicits") let exists_implicit_name id = @@ -1492,10 +1494,10 @@ let extract_explicit_arg imps args = let id = match pos with | ExplByName id -> if not (exists_implicit_name id imps) then - user_err ~loc + user_err ?loc (str "Wrong argument name: " ++ pr_id id ++ str "."); if Id.Map.mem id eargs then - user_err ~loc (str "Argument name " ++ pr_id id + user_err ?loc (str "Argument name " ++ pr_id id ++ str " occurs more than once."); id | ExplByPos (p,_id) -> @@ -1505,11 +1507,11 @@ let extract_explicit_arg imps args = if not (is_status_implicit imp) then failwith "imp"; name_of_implicit imp with Failure _ (* "nth" | "imp" *) -> - user_err ~loc + user_err ?loc (str"Wrong argument position: " ++ int p ++ str ".") in if Id.Map.mem id eargs then - user_err ~loc (str"Argument at position " ++ int p ++ + user_err ?loc (str"Argument at position " ++ int p ++ str " is mentioned more than once."); id in (Id.Map.add id (loc, a) eargs, rargs) @@ -1519,7 +1521,7 @@ let extract_explicit_arg imps args = (* Main loop *) let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = - let rec intern env = Loc.with_loc (fun ~loc -> function + let rec intern env = Loc.with_loc (fun ?loc -> function | CRef (ref,us) -> let (c,imp,subscopes,l),_ = intern_applied_reference intern env (Environ.named_context globalenv) @@ -1564,7 +1566,7 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = push_name_env ntnvars (impls_type_list ~args:fix_args tyi) en (Loc.tag @@ Name name)) 0 env' lf in (a,b,c,intern {env'' with tmp_scope = None} bd)) dl idl_temp in - Loc.tag ~loc @@ + Loc.tag ?loc @@ GRec (GFix (Array.map (fun (ro,_,_,_) -> ro) idl,n), Array.of_list lf, @@ -1591,7 +1593,7 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = push_name_env ntnvars (impls_type_list ~args:cofix_args tyi) en (Loc.tag @@ Name name)) 0 env' lf in (b,c,intern {env'' with tmp_scope = None} bd)) dl idl_tmp in - Loc.tag ~loc @@ + Loc.tag ?loc @@ GRec (GCoFix n, Array.of_list lf, Array.map (fun (bl,_,_) -> bl) idl, @@ -1600,30 +1602,30 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = | CProdN ([],c2) -> intern_type env c2 | CProdN ((nal,bk,ty)::bll,c2) -> - iterate_prod loc env bk ty (Loc.tag ~loc @@ CProdN (bll, c2)) nal + iterate_prod ?loc env bk ty (Loc.tag ?loc @@ CProdN (bll, c2)) nal | CLambdaN ([],c2) -> intern env c2 | CLambdaN ((nal,bk,ty)::bll,c2) -> - iterate_lam loc (reset_tmp_scope env) bk ty (Loc.tag ~loc @@ CLambdaN (bll, c2)) nal + iterate_lam loc (reset_tmp_scope env) bk ty (Loc.tag ?loc @@ CLambdaN (bll, c2)) nal | CLetIn (na,c1,t,c2) -> let inc1 = intern (reset_tmp_scope env) c1 in let int = Option.map (intern_type env) t in - Loc.tag ~loc @@ + Loc.tag ?loc @@ GLetIn (snd na, inc1, int, intern (push_name_env ntnvars (impls_term_list inc1) env na) c2) | CNotation ("- _",([_, CPrim (Numeral p)],[],[])) when Bigint.is_strictly_pos p -> - intern env (Loc.tag ~loc @@ CPrim (Numeral (Bigint.neg p))) + intern env (Loc.tag ?loc @@ CPrim (Numeral (Bigint.neg p))) | CNotation ("( _ )",([a],[],[])) -> intern env a | CNotation (ntn,args) -> intern_notation intern env ntnvars loc ntn args | CGeneralization (b,a,c) -> intern_generalization intern env ntnvars loc b a c | CPrim p -> - fst (Notation.interp_prim_token ~loc p (env.tmp_scope,env.scopes)) + fst (Notation.interp_prim_token ?loc p (env.tmp_scope,env.scopes)) | CDelimiters (key, e) -> intern {env with tmp_scope = None; - scopes = find_delimiters_scope ~loc key :: env.scopes} e + scopes = find_delimiters_scope ?loc key :: env.scopes} e | CAppExpl ((isproj,ref,us), args) -> let (f,_,args_scopes,_),args = let args = List.map (fun a -> (a,None)) args in @@ -1631,7 +1633,7 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = lvar us args ref in (* Rem: GApp(_,f,[]) stands for @f *) - Loc.tag ~loc @@ + Loc.tag ?loc @@ GApp (f, intern_args env args_scopes (List.map fst args)) | CApp ((isproj,f), args) -> @@ -1658,15 +1660,15 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = let st = Evar_kinds.Define (not (Program.get_proofs_transparency ())) in let fields = sort_fields ~complete:true loc fs - (fun _idx -> Loc.tag ~loc @@ CHole (Some (Evar_kinds.QuestionMark st), + (fun _idx -> Loc.tag ?loc @@ CHole (Some (Evar_kinds.QuestionMark st), Misctypes.IntroAnonymous, None)) in begin match fields with - | None -> user_err ~loc ~hdr:"intern" (str"No constructor inference.") + | None -> user_err ?loc ~hdr:"intern" (str"No constructor inference.") | Some (n, constrname, args) -> - let pars = List.make n (Loc.tag ~loc @@ CHole (None, Misctypes.IntroAnonymous, None)) in - let app = Loc.tag ~loc @@ CAppExpl ((None, constrname,None), List.rev_append pars args) in + let pars = List.make n (Loc.tag ?loc @@ CHole (None, Misctypes.IntroAnonymous, None)) in + let app = Loc.tag ?loc @@ CAppExpl ((None, constrname,None), List.rev_append pars args) in intern env app end | CCases (sty, rtnpo, tms, eqns) -> @@ -1701,7 +1703,7 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = let main_sub_eqn = Loc.tag @@ ([],thepats, (* "|p1,..,pn" *) Option.cata (intern_type env') - (Loc.tag ~loc @@ GHole(Evar_kinds.CasesType false,Misctypes.IntroAnonymous,None)) + (Loc.tag ?loc @@ GHole(Evar_kinds.CasesType false,Misctypes.IntroAnonymous,None)) rtnpo) (* "=> P" if there were a return predicate P, and "=> _" otherwise *) in let catch_all_sub_eqn = if List.for_all (irrefutable globalenv) thepats then [] else @@ -1710,7 +1712,7 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = Some (Loc.tag @@ GCases(Term.RegularStyle,sub_rtn,sub_tms,main_sub_eqn::catch_all_sub_eqn)) in let eqns' = List.map (intern_eqn (List.length tms) env) eqns in - Loc.tag ~loc @@ + Loc.tag ?loc @@ GCases (sty, rtnpo, tms, List.flatten eqns') | CLetTuple (nal, (na,po), b, c) -> let env' = reset_tmp_scope env in @@ -1720,7 +1722,7 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = let env'' = push_name_env ntnvars (Variable,[],[],[]) (reset_hidden_inductive_implicit_test env') (Loc.tag na') in intern_type env'' u) po in - Loc.tag ~loc @@ + Loc.tag ?loc @@ GLetTuple (List.map snd nal, (na', p'), b', intern (List.fold_left (push_name_env ntnvars (Variable,[],[],[])) (reset_hidden_inductive_implicit_test env) nal) c) | CIf (c, (na,po), b1, b2) -> @@ -1730,7 +1732,7 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = let env'' = push_name_env ntnvars (Variable,[],[],[]) (reset_hidden_inductive_implicit_test env) (Loc.tag na') in intern_type env'' p) po in - Loc.tag ~loc @@ + Loc.tag ?loc @@ GIf (c', (na', p'), intern env b1, intern env b2) | CHole (k, naming, solve) -> let k = match k with @@ -1756,28 +1758,28 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = let (_, glb) = Genintern.generic_intern ist gen in Some glb in - Loc.tag ~loc @@ + Loc.tag ?loc @@ GHole (k, naming, solve) (* Parsing pattern variables *) | CPatVar n when allow_patvar -> - Loc.tag ~loc @@ + Loc.tag ?loc @@ GPatVar (true,n) | CEvar (n, []) when allow_patvar -> - Loc.tag ~loc @@ + Loc.tag ?loc @@ GPatVar (false,n) (* end *) (* Parsing existential variables *) | CEvar (n, l) -> - Loc.tag ~loc @@ + Loc.tag ?loc @@ GEvar (n, List.map (on_snd (intern env)) l) | CPatVar _ -> raise (InternalizationError (loc,IllegalMetavariable)) (* end *) | CSort s -> - Loc.tag ~loc @@ + Loc.tag ?loc @@ GSort s | CCast (c1, c2) -> - Loc.tag ~loc @@ + Loc.tag ?loc @@ GCast (intern env c1, Miscops.map_cast_type (intern_type env) c2) ) and intern_type env = intern (set_type_scope env) @@ -1836,7 +1838,7 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = let rec canonize_args case_rel_ctxt arg_pats forbidden_names match_acc var_acc = let add_name l = function | _,Anonymous -> l - | loc,(Name y as x) -> (y, Loc.tag ~loc @@ PatVar x) :: l in + | loc,(Name y as x) -> (y, Loc.tag ?loc @@ PatVar x) :: l in match case_rel_ctxt,arg_pats with (* LetIn in the rel_context *) | LocalDef _ :: t, l when not with_letin -> @@ -1860,13 +1862,13 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = [], None in (tm',(snd na,typ)), extra_id, match_td - and iterate_prod loc2 env bk ty body nal = + and iterate_prod ?loc env bk ty body nal = let env, bl = intern_assumption intern ntnvars env nal bk ty in - it_mkGProd loc2 bl (intern_type env body) + it_mkGProd ?loc bl (intern_type env body) - and iterate_lam loc2 env bk ty body nal = + and iterate_lam loc env bk ty body nal = let env, bl = intern_assumption intern ntnvars env nal bk ty in - it_mkGLambda loc2 bl (intern env body) + it_mkGLambda ?loc bl (intern env body) and intern_impargs c env l subscopes args = let eargs, rargs = extract_explicit_arg l args in @@ -1898,7 +1900,7 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = | (imp::impl', []) -> if not (Id.Map.is_empty eargs) then (let (id,(loc,_)) = Id.Map.choose eargs in - user_err ~loc (str "Not enough non implicit \ + user_err ?loc (str "Not enough non implicit \ arguments to accept the argument bound to " ++ pr_id id ++ str".")); [] @@ -1915,8 +1917,8 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = and smart_gapp f loc = function | [] -> f | l -> match f with - | (loc', GApp (g, args)) -> Loc.tag ~loc:(Loc.merge loc' loc) @@ GApp (g, args@l) - | _ -> Loc.tag ~loc:(Loc.merge (loc_of_glob_constr f) loc) @@ GApp (f, l) + | (loc', GApp (g, args)) -> Loc.tag ?loc:(Loc.merge_opt loc' loc) @@ GApp (g, args@l) + | _ -> Loc.tag ?loc:(Loc.merge_opt (loc_of_glob_constr f) loc) @@ GApp (f, l) and intern_args env subscopes = function | [] -> [] @@ -1929,7 +1931,7 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = intern env c with InternalizationError (loc,e) -> - user_err ~loc ~hdr:"internalize" + user_err ?loc ~hdr:"internalize" (explain_internalization_error e) (**************************************************************************) @@ -1969,7 +1971,7 @@ let intern_pattern globalenv patt = intern_cases_pattern globalenv (None,[]) empty_alias patt with InternalizationError (loc,e) -> - user_err ~loc ~hdr:"internalize" (explain_internalization_error e) + user_err ?loc ~hdr:"internalize" (explain_internalization_error e) (*********************************************************************) @@ -2055,12 +2057,12 @@ let interp_notation_constr ?(impls=empty_internalization_env) nenv a = let interp_binder env sigma na t = let t = intern_gen IsType env t in - let t' = locate_if_hole ~loc:(loc_of_glob_constr t) na t in + let t' = locate_if_hole ?loc:(loc_of_glob_constr t) na t in understand ~expected_type:IsType env sigma t' let interp_binder_evars env evdref na t = let t = intern_gen IsType env t in - let t' = locate_if_hole ~loc:(loc_of_glob_constr t) na t in + let t' = locate_if_hole ?loc:(loc_of_glob_constr t) na t in understand_tcc_evars env evdref ~expected_type:IsType t' open Environ @@ -2079,7 +2081,7 @@ let intern_context global_level env impl_env binders = tmp_scope = None; scopes = []; impls = impl_env}, []) binders in (lenv.impls, List.map glob_local_binder_of_extended bl) with InternalizationError (loc,e) -> - user_err ~loc ~hdr:"internalize" (explain_internalization_error e) + user_err ?loc ~hdr:"internalize" (explain_internalization_error e) let interp_rawcontext_evars env evdref k bl = let open EConstr in @@ -2087,7 +2089,7 @@ let interp_rawcontext_evars env evdref k bl = List.fold_left (fun (env,params,n,impls) (na, k, b, t) -> let t' = - if Option.is_empty b then locate_if_hole ~loc:(loc_of_glob_constr t) na t + if Option.is_empty b then locate_if_hole ?loc:(loc_of_glob_constr t) na t else t in let t = understand_tcc_evars env evdref ~expected_type:IsType t' in diff --git a/interp/dumpglob.ml b/interp/dumpglob.ml index 9f549b0c0..10621f14d 100644 --- a/interp/dumpglob.ml +++ b/interp/dumpglob.ml @@ -139,30 +139,32 @@ let interval loc = let loc1,loc2 = Loc.unloc loc in loc1, loc2-1 -let dump_ref loc filepath modpath ident ty = +let dump_ref ?loc filepath modpath ident ty = match !glob_output with | Feedback -> - Feedback.feedback (Feedback.GlobRef (loc, filepath, modpath, ident, ty)) + Option.iter (fun loc -> + Feedback.feedback (Feedback.GlobRef (loc, filepath, modpath, ident, ty)) + ) loc | NoGlob -> () - | _ when not (Loc.is_ghost loc) -> + | _ -> Option.iter (fun loc -> let bl,el = interval loc in dump_string (Printf.sprintf "R%d:%d %s %s %s %s\n" bl el filepath modpath ident ty) - | _ -> () + ) loc -let dump_reference loc modpath ident ty = +let dump_reference ?loc modpath ident ty = let filepath = Names.DirPath.to_string (Lib.library_dp ()) in - dump_ref loc filepath modpath ident ty + dump_ref ?loc filepath modpath ident ty -let dump_modref loc mp ty = +let dump_modref ?loc mp ty = let (dp, l) = Lib.split_modpath mp in let filepath = Names.DirPath.to_string dp in let modpath = Names.DirPath.to_string (Names.DirPath.make l) in let ident = "<>" in - dump_ref loc filepath modpath ident ty + dump_ref ?loc filepath modpath ident ty -let dump_libref loc dp ty = - dump_ref loc (Names.DirPath.to_string dp) "<>" "<>" ty +let dump_libref ?loc dp ty = + dump_ref ?loc (Names.DirPath.to_string dp) "<>" "<>" ty let cook_notation df sc = (* We encode notations so that they are space-free and still human-readable *) @@ -208,10 +210,10 @@ let dump_notation_location posl df (((path,secpath),_),sc) = let secpath = Names.DirPath.to_string secpath in let df = cook_notation df sc in List.iter (fun l -> - dump_ref (Loc.make_loc l) path secpath df "not") + dump_ref ~loc:(Loc.make_loc l) path secpath df "not") posl -let add_glob_gen loc sp lib_dp ty = +let add_glob_gen ?loc sp lib_dp ty = if dump () then let mod_dp,id = Libnames.repr_path sp in let mod_dp = remove_sections mod_dp in @@ -219,50 +221,51 @@ let add_glob_gen loc sp lib_dp ty = let filepath = Names.DirPath.to_string lib_dp in let modpath = Names.DirPath.to_string mod_dp_trunc in let ident = Names.Id.to_string id in - dump_ref loc filepath modpath ident ty + dump_ref ?loc filepath modpath ident ty -let add_glob loc ref = - if dump () && not (Loc.is_ghost loc) then +let add_glob ?loc ref = + if dump () then let sp = Nametab.path_of_global ref in let lib_dp = Lib.library_part ref in let ty = type_of_global_ref ref in - add_glob_gen loc sp lib_dp ty + add_glob_gen ?loc sp lib_dp ty let mp_of_kn kn = let mp,sec,l = Names.repr_kn kn in Names.MPdot (mp,l) -let add_glob_kn loc kn = - if dump () && not (Loc.is_ghost loc) then +let add_glob_kn ?loc kn = + if dump () then let sp = Nametab.path_of_syndef kn in let lib_dp = Lib.dp_of_mp (mp_of_kn kn) in - add_glob_gen loc sp lib_dp "syndef" + add_glob_gen ?loc sp lib_dp "syndef" -let dump_binding loc id = () +let dump_binding ?loc id = () -let dump_def ty loc secpath id = +let dump_def ?loc ty secpath id = Option.iter (fun loc -> if !glob_output = Feedback then Feedback.feedback (Feedback.GlobDef (loc, id, secpath, ty)) else let bl,el = interval loc in dump_string (Printf.sprintf "%s %d:%d %s %s\n" ty bl el secpath id) + ) loc let dump_definition (loc, id) sec s = - dump_def s loc (Names.DirPath.to_string (Lib.current_dirpath sec)) (Names.Id.to_string id) + dump_def ?loc s (Names.DirPath.to_string (Lib.current_dirpath sec)) (Names.Id.to_string id) let dump_constraint (((loc, n),_), _, _) sec ty = match n with | Names.Name id -> dump_definition (loc, id) sec ty | Names.Anonymous -> () -let dump_moddef loc mp ty = +let dump_moddef ?loc mp ty = let (dp, l) = Lib.split_modpath mp in let mp = Names.DirPath.to_string (Names.DirPath.make l) in - dump_def ty loc "<>" mp + dump_def ?loc ty "<>" mp -let dump_notation (loc,(df,_)) sc sec = +let dump_notation (loc,(df,_)) sc sec = Option.iter (fun loc -> (* We dump the location of the opening '"' *) let i = fst (Loc.unloc loc) in let location = (Loc.make_loc (i, i+1)) in - dump_def "not" location (Names.DirPath.to_string (Lib.current_dirpath sec)) (cook_notation df sc) - + dump_def ~loc:location "not" (Names.DirPath.to_string (Lib.current_dirpath sec)) (cook_notation df sc) + ) loc diff --git a/interp/dumpglob.mli b/interp/dumpglob.mli index e84a64052..f42055af7 100644 --- a/interp/dumpglob.mli +++ b/interp/dumpglob.mli @@ -22,19 +22,19 @@ val feedback_glob : unit -> unit val pause : unit -> unit val continue : unit -> unit -val add_glob : Loc.t -> Globnames.global_reference -> unit -val add_glob_kn : Loc.t -> Names.kernel_name -> unit - -val dump_definition : Loc.t * Names.Id.t -> bool -> string -> unit -val dump_moddef : Loc.t -> Names.module_path -> string -> unit -val dump_modref : Loc.t -> Names.module_path -> string -> unit -val dump_reference : Loc.t -> string -> string -> string -> unit -val dump_libref : Loc.t -> Names.DirPath.t -> string -> unit +val add_glob : ?loc:Loc.t -> Globnames.global_reference -> unit +val add_glob_kn : ?loc:Loc.t -> Names.kernel_name -> unit + +val dump_definition : Names.Id.t Loc.located -> bool -> string -> unit +val dump_moddef : ?loc:Loc.t -> Names.module_path -> string -> unit +val dump_modref : ?loc:Loc.t -> Names.module_path -> string -> unit +val dump_reference : ?loc:Loc.t -> string -> string -> string -> unit +val dump_libref : ?loc:Loc.t -> Names.DirPath.t -> string -> unit val dump_notation_location : (int * int) list -> Constrexpr.notation -> (Notation.notation_location * Notation_term.scope_name option) -> unit -val dump_binding : Loc.t -> Names.Id.Set.elt -> unit +val dump_binding : ?loc:Loc.t -> Names.Id.Set.elt -> unit val dump_notation : - Loc.t * (Constrexpr.notation * Notation.notation_location) -> + (Constrexpr.notation * Notation.notation_location) Loc.located -> Notation_term.scope_name option -> bool -> unit val dump_constraint : Constrexpr.typeclass_constraint -> bool -> string -> unit diff --git a/interp/implicit_quantifiers.ml b/interp/implicit_quantifiers.ml index fa7712bdc..dd04e2030 100644 --- a/interp/implicit_quantifiers.ml +++ b/interp/implicit_quantifiers.ml @@ -29,11 +29,11 @@ let generalizable_table = Summary.ref Id.Pred.empty ~name:"generalizable-ident" let declare_generalizable_ident table (loc,id) = if not (Id.equal id (root_of_id id)) then - user_err ~loc ~hdr:"declare_generalizable_ident" + user_err ?loc ~hdr:"declare_generalizable_ident" ((pr_id id ++ str " is not declarable as generalizable identifier: it must have no trailing digits, quote, or _")); if Id.Pred.mem id table then - user_err ~loc ~hdr:"declare_generalizable_ident" + user_err ?loc ~hdr:"declare_generalizable_ident" ((pr_id id++str" is already declared as a generalizable identifier")) else Id.Pred.add id table @@ -80,7 +80,7 @@ let is_freevar ids env x = (* Auxiliary functions for the inference of implicitly quantified variables. *) let ungeneralizable loc id = - user_err ~loc ~hdr:"Generalization" + user_err ?loc ~hdr:"Generalization" (str "Unbound and ungeneralizable variable " ++ pr_id id) let free_vars_of_constr_expr c ?(bound=Id.Set.empty) l = @@ -128,8 +128,8 @@ let generalizable_vars_of_glob_constr ?(bound=Id.Set.empty) ?(allowed=Id.Set.emp let rec vars bound vs (loc, t) = match t with | GVar id -> if is_freevar bound (Global.env ()) id then - if Id.List.mem_assoc id vs then vs - else (id, loc) :: vs + if Id.List.mem_assoc_sym id vs then vs + else (Loc.tag ?loc id) :: vs else vs | GApp (f,args) -> List.fold_left (vars bound) vs (f::args) @@ -189,7 +189,7 @@ let generalizable_vars_of_glob_constr ?(bound=Id.Set.empty) ?(allowed=Id.Set.emp vars_option bound' vs tyopt in fun rt -> let vars = List.rev (vars bound [] rt) in - List.iter (fun (id, loc) -> + List.iter (fun (loc, id) -> if not (Id.Set.mem id allowed || find_generalizable_ident id) then ungeneralizable loc id) vars; vars @@ -212,7 +212,7 @@ let combine_params avoid fn applied needed = | Anonymous -> false in if not (List.exists is_id needed) then - user_err ~loc (str "Wrong argument name: " ++ Nameops.pr_id id); + user_err ?loc (str "Wrong argument name: " ++ Nameops.pr_id id); true | _ -> false) applied in @@ -246,7 +246,7 @@ let combine_params avoid fn applied needed = aux (t' :: ids) avoid' app need | (x,_) :: _, [] -> - user_err ~loc:(Constrexpr_ops.constr_loc x) (str "Typeclass does not expect more arguments") + user_err ?loc:(Constrexpr_ops.constr_loc x) (str "Typeclass does not expect more arguments") in aux [] avoid applied needed let combine_params_freevar = @@ -256,21 +256,21 @@ let combine_params_freevar = let destClassApp (loc, cl) = match cl with - | CApp ((None, (_loc, CRef (ref, inst))), l) -> loc, ref, List.map fst l, inst - | CAppExpl ((None, ref, inst), l) -> loc, ref, l, inst - | CRef (ref, inst) -> loc_of_reference ref, ref, [], inst + | CApp ((None, (_loc, CRef (ref, inst))), l) -> Loc.tag ?loc (ref, List.map fst l, inst) + | CAppExpl ((None, ref, inst), l) -> Loc.tag ?loc (ref, l, inst) + | CRef (ref, inst) -> Loc.tag ?loc:(loc_of_reference ref) (ref, [], inst) | _ -> raise Not_found let destClassAppExpl (loc, cl) = match cl with - | CApp ((None, (_loc, CRef (ref, inst))), l) -> loc, ref, l, inst - | CRef (ref, inst) -> loc_of_reference ref, ref, [], inst + | CApp ((None, (_loc, CRef (ref, inst))), l) -> Loc.tag ?loc (ref, l, inst) + | CRef (ref, inst) -> Loc.tag ?loc:(loc_of_reference ref) (ref, [], inst) | _ -> raise Not_found let implicit_application env ?(allow_partial=true) f ty = let is_class = try - let (_, r, _, _ as clapp) = destClassAppExpl ty in + let (_, (r, _, _) as clapp) = destClassAppExpl ty in let (loc, qid) = qualid_of_reference r in let gr = Nametab.locate qid in if Typeclasses.is_class gr then Some (clapp, gr) else None @@ -278,7 +278,7 @@ let implicit_application env ?(allow_partial=true) f ty = in match is_class with | None -> ty, env - | Some ((loc, id, par, inst), gr) -> + | Some ((loc, (id, par, inst)), gr) -> let avoid = Id.Set.union env (ids_of_list (free_vars_of_constr_expr ty ~bound:env [])) in let c, avoid = let c = class_info gr in @@ -296,7 +296,7 @@ let implicit_application env ?(allow_partial=true) f ty = end; let pars = List.rev (List.combine ci rd) in let args, avoid = combine_params avoid f par pars in - Loc.tag ~loc @@ CAppExpl ((None, id, inst), args), avoid + Loc.tag ?loc @@ CAppExpl ((None, id, inst), args), avoid in c, avoid let implicits_of_glob_constr ?(with_products=true) l = diff --git a/interp/implicit_quantifiers.mli b/interp/implicit_quantifiers.mli index 71009ec3c..945bed2aa 100644 --- a/interp/implicit_quantifiers.mli +++ b/interp/implicit_quantifiers.mli @@ -16,8 +16,8 @@ open Globnames val declare_generalizable : Vernacexpr.locality_flag -> (Id.t located) list option -> unit val ids_of_list : Id.t list -> Id.Set.t -val destClassApp : constr_expr -> Loc.t * reference * constr_expr list * instance_expr option -val destClassAppExpl : constr_expr -> Loc.t * reference * (constr_expr * explicitation located option) list * instance_expr option +val destClassApp : constr_expr -> (reference * constr_expr list * instance_expr option) located +val destClassAppExpl : constr_expr -> (reference * (constr_expr * explicitation located option) list * instance_expr option) located (** Fragile, should be used only for construction a set of identifiers to avoid *) @@ -31,7 +31,7 @@ val free_vars_of_binders : order with the location of their first occurrence *) val generalizable_vars_of_glob_constr : ?bound:Id.Set.t -> ?allowed:Id.Set.t -> - glob_constr -> (Id.t * Loc.t) list + glob_constr -> Id.t located list val make_fresh : Id.Set.t -> Environ.env -> Id.t -> Id.t diff --git a/interp/modintern.ml b/interp/modintern.ml index 166711659..45e6cd06c 100644 --- a/interp/modintern.ml +++ b/interp/modintern.ml @@ -26,16 +26,16 @@ let error_not_a_module_loc kind loc qid = | ModType -> Modops.ModuleTypingError (Modops.NotAModuleType s) | ModAny -> ModuleInternalizationError (NotAModuleNorModtype s) in - Loc.raise ~loc e + Loc.raise ?loc e let error_application_to_not_path loc me = - Loc.raise ~loc (Modops.ModuleTypingError (Modops.ApplicationToNotPath me)) + Loc.raise ?loc (Modops.ModuleTypingError (Modops.ApplicationToNotPath me)) let error_incorrect_with_in_module loc = - Loc.raise ~loc (ModuleInternalizationError IncorrectWithInModule) + Loc.raise ?loc (ModuleInternalizationError IncorrectWithInModule) let error_application_to_module_type loc = - Loc.raise ~loc (ModuleInternalizationError IncorrectModuleApplication) + Loc.raise ?loc (ModuleInternalizationError IncorrectModuleApplication) (** Searching for a module name in the Nametab. @@ -47,12 +47,12 @@ let lookup_module_or_modtype kind (loc,qid) = try if kind == ModType then raise Not_found; let mp = Nametab.locate_module qid in - Dumpglob.dump_modref loc mp "modtype"; (mp,Module) + Dumpglob.dump_modref ?loc mp "modtype"; (mp,Module) with Not_found -> try if kind == Module then raise Not_found; let mp = Nametab.locate_modtype qid in - Dumpglob.dump_modref loc mp "mod"; (mp,ModType) + Dumpglob.dump_modref ?loc mp "mod"; (mp,ModType) with Not_found -> error_not_a_module_loc kind loc qid let lookup_module lqid = fst (lookup_module_or_modtype Module lqid) diff --git a/interp/notation_ops.ml b/interp/notation_ops.ml index 32c564156..328fdd519 100644 --- a/interp/notation_ops.ml +++ b/interp/notation_ops.ml @@ -227,8 +227,8 @@ let split_at_recursive_part c = | None -> let () = sub := Some c in begin match l with - | [] -> Loc.tag ~loc @@ GVar ldots_var - | _ :: _ -> Loc.tag ~loc:loc0 @@ GApp (Loc.tag ~loc @@ GVar ldots_var, l) + | [] -> Loc.tag ?loc @@ GVar ldots_var + | _ :: _ -> Loc.tag ?loc:loc0 @@ GApp (Loc.tag ?loc @@ GVar ldots_var, l) end | Some _ -> (* Not narrowed enough to find only one recursive part *) @@ -243,10 +243,13 @@ let split_at_recursive_part c = | GVar v when Id.equal v ldots_var -> (* Not enough context *) raise Not_found | _ -> outer_iterator, c -let subtract_loc loc1 loc2 = Loc.make_loc (fst (Loc.unloc loc1),fst (Loc.unloc loc2)-1) +let subtract_loc loc1 loc2 = + let l1 = fst (Option.cata Loc.unloc (0,0) loc1) in + let l2 = fst (Option.cata Loc.unloc (0,0) loc2) in + Some (Loc.make_loc (l1,l2-1)) let check_is_hole id = function _, GHole _ -> () | t -> - user_err ~loc:(loc_of_glob_constr t) + user_err ?loc:(loc_of_glob_constr t) (strbrk "In recursive notation with binders, " ++ pr_id id ++ strbrk " is expected to come without type.") @@ -298,7 +301,7 @@ let compare_recursive_parts found f f' (iterator,subc) = let loc1 = loc_of_glob_constr iterator in let loc2 = loc_of_glob_constr (Option.get !terminator) in (* Here, we would need a loc made of several parts ... *) - user_err ~loc:(subtract_loc loc1 loc2) + user_err ?loc:(subtract_loc loc1 loc2) (str "Both ends of the recursive pattern are the same.") | Some (x,y,RecursiveTerms lassoc) -> let newfound,x,y,lassoc = @@ -342,7 +345,7 @@ let notation_constr_and_vars_of_glob_constr a = | GApp ((loc, GVar f),[c]) when Id.equal f ldots_var -> (* Fall on the second part of the recursive pattern w/o having found the first part *) - user_err ~loc + user_err ?loc (str "Cannot find where the recursive pattern starts.") | _c -> aux' c @@ -459,7 +462,7 @@ let rec subst_pat subst (loc, pat) = | PatCstr (((kn,i),j),cpl,n) -> let kn' = subst_mind subst kn and cpl' = List.smartmap (subst_pat subst) cpl in - Loc.tag ~loc @@ + Loc.tag ?loc @@ if kn' == kn && cpl' == cpl then pat else PatCstr (((kn',i),j),cpl',n) @@ -749,11 +752,11 @@ let rec map_cases_pattern_name_left f = Loc.map (function ) let rec fold_cases_pattern_eq f x p p' = match p, p' with - | (loc, PatVar na), (_, PatVar na') -> let x,na = f x na na' in x, Loc.tag ~loc @@ PatVar na + | (loc, PatVar na), (_, PatVar na') -> let x,na = f x na na' in x, Loc.tag ?loc @@ PatVar na | (loc, PatCstr (c,l,na)), (_, PatCstr (c',l',na')) when eq_constructor c c' -> let x,l = fold_cases_pattern_list_eq f x l l' in let x,na = f x na na' in - x, Loc.tag ~loc @@ PatCstr (c,l,na) + x, Loc.tag ?loc @@ PatCstr (c,l,na) | _ -> failwith "Not equal" and fold_cases_pattern_list_eq f x pl pl' = match pl, pl' with @@ -799,13 +802,13 @@ let bind_bindinglist_env alp (terms,onlybinders,termlists,binderlists as sigma) match b, b' with | GLocalAssum (na,bk,t), GLocalAssum (na',bk',t') -> let alp, na = unify_name alp na na' in - alp, Loc.tag ~loc @@ GLocalAssum (na, unify_binding_kind bk bk', unify_term alp t t') + alp, Loc.tag ?loc @@ GLocalAssum (na, unify_binding_kind bk bk', unify_term alp t t') | GLocalDef (na,bk,c,t), GLocalDef (na',bk',c',t') -> let alp, na = unify_name alp na na' in - alp, Loc.tag ~loc @@ GLocalDef (na, unify_binding_kind bk bk', unify_term alp c c', unify_opt_term alp t t') + alp, Loc.tag ?loc @@ GLocalDef (na, unify_binding_kind bk bk', unify_term alp c c', unify_opt_term alp t t') | GLocalPattern ((p,ids),id,bk,t), GLocalPattern ((p',_),_,bk',t') -> let alp, p = unify_pat alp p p' in - alp, Loc.tag ~loc @@ GLocalPattern ((p,ids), id, unify_binding_kind bk bk', unify_term alp t t') + alp, Loc.tag ?loc @@ GLocalPattern ((p,ids), id, unify_binding_kind bk bk', unify_term alp t t') | _ -> raise No_match in let rec unify alp bl bl' = match bl, bl' with @@ -832,7 +835,7 @@ let bind_bindinglist_as_term_env alp (terms,onlybinders,termlists,binderlists) v let unify_pat p p' = if cases_pattern_eq (map_cases_pattern_name_left (name_app (rename_var (snd alp))) p) p' then p' else raise No_match in - let unify_term_binder c (loc, b') = Loc.tag ~loc @@ + let unify_term_binder c (loc, b') = Loc.tag ?loc @@ match c, b' with | (_, GVar id), GLocalAssum (na', bk', t') -> GLocalAssum (unify_id id na', bk', t') @@ -895,21 +898,21 @@ let rec match_cases_pattern_binders metas acc (_, pat1) (_, pat2) = let glue_letin_with_decls = true -let rec match_iterated_binders islambda decls bi = Loc.with_loc (fun ~loc -> function +let rec match_iterated_binders islambda decls bi = Loc.with_loc (fun ?loc -> function | GLambda (Name p,bk,t,(_, GCases (LetPatternStyle,None,[((_, GVar e),_)],[(_,(ids,[cp],b))]))) when islambda && Id.equal p e -> - match_iterated_binders islambda ((Loc.tag ~loc @@ GLocalPattern((cp,ids),p,bk,t))::decls) b + match_iterated_binders islambda ((Loc.tag ?loc @@ GLocalPattern((cp,ids),p,bk,t))::decls) b | GLambda (na,bk,t,b) when islambda -> - match_iterated_binders islambda ((Loc.tag ~loc @@ GLocalAssum(na,bk,t))::decls) b + match_iterated_binders islambda ((Loc.tag ?loc @@ GLocalAssum(na,bk,t))::decls) b | GProd (Name p,bk,t,(_, GCases (LetPatternStyle,None,[((_, GVar e),_)],[(_,(ids,[cp],b))]))) when not islambda && Id.equal p e -> - match_iterated_binders islambda ((Loc.tag ~loc @@ GLocalPattern((cp,ids),p,bk,t))::decls) b + match_iterated_binders islambda ((Loc.tag ?loc @@ GLocalPattern((cp,ids),p,bk,t))::decls) b | GProd ((Name _ as na),bk,t,b) when not islambda -> - match_iterated_binders islambda ((Loc.tag ~loc @@ GLocalAssum(na,bk,t))::decls) b + match_iterated_binders islambda ((Loc.tag ?loc @@ GLocalAssum(na,bk,t))::decls) b | GLetIn (na,c,t,b) when glue_letin_with_decls -> match_iterated_binders islambda - ((Loc.tag ~loc @@ GLocalDef (na,Explicit (*?*), c,t))::decls) b - | b -> (decls, Loc.tag ~loc b) + ((Loc.tag ?loc @@ GLocalDef (na,Explicit (*?*), c,t))::decls) b + | b -> (decls, Loc.tag ?loc b) ) bi let remove_sigma x (terms,onlybinders,termlists,binderlists) = @@ -989,13 +992,13 @@ let rec match_ inner u alp metas sigma a1 a2 = (* "λ p, let 'cp = p in t" -> "λ 'cp, t" *) | GLambda (Name p,bk,t1,(_, GCases (LetPatternStyle,None,[((_, GVar e),_)],[(_,(ids,[cp],b1))]))), NBinderList (x,_,NLambda (Name _id2,_,b2),termin) when Id.equal p e -> - let (decls,b) = match_iterated_binders true [Loc.tag ~loc @@ GLocalPattern((cp,ids),p,bk,t1)] b1 in + let (decls,b) = match_iterated_binders true [Loc.tag ?loc @@ GLocalPattern((cp,ids),p,bk,t1)] b1 in let alp,sigma = bind_bindinglist_env alp sigma x decls in match_in u alp metas sigma b termin (* Matching recursive notations for binders: ad hoc cases supporting let-in *) | GLambda (na1,bk,t1,b1), NBinderList (x,_,NLambda (Name _id2,_,b2),termin)-> - let (decls,b) = match_iterated_binders true [Loc.tag ~loc @@ GLocalAssum (na1,bk,t1)] b1 in + let (decls,b) = match_iterated_binders true [Loc.tag ?loc @@ GLocalAssum (na1,bk,t1)] b1 in (* TODO: address the possibility that termin is a Lambda itself *) let alp,sigma = bind_bindinglist_env alp sigma x decls in match_in u alp metas sigma b termin @@ -1003,13 +1006,13 @@ let rec match_ inner u alp metas sigma a1 a2 = (* "∀ p, let 'cp = p in t" -> "∀ 'cp, t" *) | GProd (Name p,bk,t1,(_, GCases (LetPatternStyle,None,[((_, GVar e),_)],[(_,(ids,[cp],b1))]))), NBinderList (x,_,NProd (Name _id2,_,b2),(NVar v as termin)) when Id.equal p e -> - let (decls,b) = match_iterated_binders true [Loc.tag ~loc @@ GLocalPattern ((cp,ids),p,bk,t1)] b1 in + let (decls,b) = match_iterated_binders true [Loc.tag ?loc @@ GLocalPattern ((cp,ids),p,bk,t1)] b1 in let alp,sigma = bind_bindinglist_env alp sigma x decls in match_in u alp metas sigma b termin | GProd (na1,bk,t1,b1), NBinderList (x,_,NProd (Name _id2,_,b2),termin) when na1 != Anonymous -> - let (decls,b) = match_iterated_binders false [Loc.tag ~loc @@ GLocalAssum (na1,bk,t1)] b1 in + let (decls,b) = match_iterated_binders false [Loc.tag ?loc @@ GLocalAssum (na1,bk,t1)] b1 in (* TODO: address the possibility that termin is a Prod itself *) let alp,sigma = bind_bindinglist_env alp sigma x decls in match_in u alp metas sigma b termin @@ -1021,15 +1024,15 @@ let rec match_ inner u alp metas sigma a1 a2 = | GLambda (Name p,bk,t,(_, GCases (LetPatternStyle,None,[((_, GVar e),_)],[(_,(ids,[cp],b1))]))), NLambda (Name id,_,b2) when is_bindinglist_meta id metas -> - let alp,sigma = bind_bindinglist_env alp sigma id [Loc.tag ~loc @@ GLocalPattern ((cp,ids),p,bk,t)] in + let alp,sigma = bind_bindinglist_env alp sigma id [Loc.tag ?loc @@ GLocalPattern ((cp,ids),p,bk,t)] in match_in u alp metas sigma b1 b2 | GLambda (na,bk,t,b1), NLambda (Name id,_,b2) when is_bindinglist_meta id metas -> - let alp,sigma = bind_bindinglist_env alp sigma id [Loc.tag ~loc @@ GLocalAssum (na,bk,t)] in + let alp,sigma = bind_bindinglist_env alp sigma id [Loc.tag ?loc @@ GLocalAssum (na,bk,t)] in match_in u alp metas sigma b1 b2 | GProd (na,bk,t,b1), NProd (Name id,_,b2) when is_bindinglist_meta id metas && na != Anonymous -> - let alp,sigma = bind_bindinglist_env alp sigma id [Loc.tag ~loc @@ GLocalAssum (na,bk,t)] in + let alp,sigma = bind_bindinglist_env alp sigma id [Loc.tag ?loc @@ GLocalAssum (na,bk,t)] in match_in u alp metas sigma b1 b2 (* Matching compositionally *) @@ -1041,7 +1044,7 @@ let rec match_ inner u alp metas sigma a1 a2 = if n1 < n2 then let l21,l22 = List.chop (n2-n1) l2 in f1,l1, NApp (f2,l21), l22 else if n1 > n2 then - let l11,l12 = List.chop (n1-n2) l1 in Loc.tag ~loc @@ GApp (f1,l11),l12, f2,l2 + let l11,l12 = List.chop (n1-n2) l1 in Loc.tag ?loc @@ GApp (f1,l11),l12, f2,l2 else f1,l1, f2, l2 in let may_use_eta = does_not_come_from_already_eta_expanded_var f1 in List.fold_left2 (match_ may_use_eta u alp metas) diff --git a/interp/reserve.ml b/interp/reserve.ml index 1565ba4a9..20fdd6caa 100644 --- a/interp/reserve.ml +++ b/interp/reserve.ml @@ -86,12 +86,12 @@ let in_reserved : Id.t * notation_constr -> obj = let declare_reserved_type_binding (loc,id) t = if not (Id.equal id (root_of_id id)) then - user_err ~loc ~hdr:"declare_reserved_type" + user_err ?loc ~hdr:"declare_reserved_type" ((pr_id id ++ str " is not reservable: it must have no trailing digits, quote, or _")); begin try let _ = Id.Map.find id !reserve_table in - user_err ~loc ~hdr:"declare_reserved_type" + user_err ?loc ~hdr:"declare_reserved_type" ((pr_id id++str" is already bound to a type")) with Not_found -> () end; add_anonymous_leaf (in_reserved (id,t)) diff --git a/interp/smartlocate.ml b/interp/smartlocate.ml index fd9599ec0..a9d94669a 100644 --- a/interp/smartlocate.ml +++ b/interp/smartlocate.ml @@ -46,7 +46,7 @@ let locate_global_with_alias ?(head=false) (loc,qid) = if head then global_of_extended_global_head ref else global_of_extended_global ref with Not_found -> - user_err ~loc (pr_qualid qid ++ + user_err ?loc (pr_qualid qid ++ str " is bound to a notation that does not denote a reference.") let global_inductive_with_alias r = @@ -54,27 +54,27 @@ let global_inductive_with_alias r = try match locate_global_with_alias lqid with | IndRef ind -> ind | ref -> - user_err ~loc:(loc_of_reference r) ~hdr:"global_inductive" + user_err ?loc:(loc_of_reference r) ~hdr:"global_inductive" (pr_reference r ++ spc () ++ str "is not an inductive type.") - with Not_found -> Nametab.error_global_not_found ~loc qid + with Not_found -> Nametab.error_global_not_found ?loc qid let global_with_alias ?head r = let (loc,qid as lqid) = qualid_of_reference r in try locate_global_with_alias ?head lqid - with Not_found -> Nametab.error_global_not_found ~loc qid + with Not_found -> Nametab.error_global_not_found ?loc qid let smart_global ?head = function | AN r -> global_with_alias ?head r | ByNotation (loc,(ntn,sc)) -> - Notation.interp_notation_as_global_reference ~loc (fun _ -> true) ntn sc + Notation.interp_notation_as_global_reference ?loc (fun _ -> true) ntn sc let smart_global_inductive = function | AN r -> global_inductive_with_alias r | ByNotation (loc,(ntn,sc)) -> destIndRef - (Notation.interp_notation_as_global_reference ~loc isIndRef ntn sc) + (Notation.interp_notation_as_global_reference ?loc isIndRef ntn sc) let loc_of_smart_reference = function | AN r -> loc_of_reference r diff --git a/interp/smartlocate.mli b/interp/smartlocate.mli index 0749ca576..acae1a391 100644 --- a/interp/smartlocate.mli +++ b/interp/smartlocate.mli @@ -38,4 +38,4 @@ val smart_global : ?head:bool -> reference or_by_notation -> global_reference val smart_global_inductive : reference or_by_notation -> inductive (** Return the loc of a smart reference *) -val loc_of_smart_reference : reference or_by_notation -> Loc.t +val loc_of_smart_reference : reference or_by_notation -> Loc.t option diff --git a/interp/stdarg.mli b/interp/stdarg.mli index 113fe40ba..44a176d94 100644 --- a/interp/stdarg.mli +++ b/interp/stdarg.mli @@ -22,7 +22,7 @@ open Tactypes open Genarg (** FIXME: nothing to do there. *) -val loc_of_or_by_notation : ('a -> Loc.t) -> 'a or_by_notation -> Loc.t +val loc_of_or_by_notation : ('a -> Loc.t option) -> 'a or_by_notation -> Loc.t option val wit_unit : unit uniform_genarg_type diff --git a/interp/topconstr.ml b/interp/topconstr.ml index 2ffeb1f83..a74e64172 100644 --- a/interp/topconstr.ml +++ b/interp/topconstr.ml @@ -59,7 +59,7 @@ let rec cases_pattern_fold_names f a pt = match snd pt with | CPatAtom (Some (Ident (_,id))) when not (is_constructor id) -> f id a | CPatPrim _ | CPatAtom _ -> a | CPatCast ((loc,_),_) -> - CErrors.user_err ~loc ~hdr:"cases_pattern_fold_names" + CErrors.user_err ?loc ~hdr:"cases_pattern_fold_names" (Pp.strbrk "Casts are not supported here.") let ids_of_pattern = @@ -103,7 +103,7 @@ let rec fold_local_binders g f n acc b = function | [] -> f n acc b -let fold_constr_expr_with_binders g f n acc = Loc.with_loc (fun ~loc -> function +let fold_constr_expr_with_binders g f n acc = Loc.with_loc (fun ?loc -> function | CAppExpl ((_,_,_),l) -> List.fold_left (f n) acc l | CApp ((_,t),l) -> List.fold_left (f n) (f n acc t) (List.map fst l) | CProdN (l,b) | CLambdaN (l,b) -> fold_constr_expr_binders g f n acc b l @@ -181,9 +181,9 @@ let split_at_annot bl na = end | CLocalDef _ as x :: rest -> aux (x :: acc) rest | CLocalPattern (loc,_) :: rest -> - Loc.raise ~loc (Stream.Error "pattern with quote not allowed after fix") + Loc.raise ?loc (Stream.Error "pattern with quote not allowed after fix") | [] -> - user_err ~loc + user_err ?loc (str "No parameter named " ++ Nameops.pr_id id ++ str".") in aux [] bl @@ -271,19 +271,20 @@ let rec replace_vars_constr_expr l = function (* Returns the ranges of locs of the notation that are not occupied by args *) (* and which are then occupied by proper symbols of the notation (or spaces) *) -let locs_of_notation loc locs ntn = - let (bl, el) = Loc.unloc loc in - let locs = List.map Loc.unloc locs in +let locs_of_notation ?loc locs ntn = + let unloc loc = Option.cata Loc.unloc (0,0) loc in + let (bl, el) = unloc loc in + let locs = List.map unloc locs in let rec aux pos = function | [] -> if Int.equal pos el then [] else [(pos,el)] | (ba,ea)::l -> if Int.equal pos ba then aux ea l else (pos,ba)::aux ea l in aux bl (List.sort (fun l1 l2 -> fst l1 - fst l2) locs) -let ntn_loc loc (args,argslist,binderslist) = - locs_of_notation loc +let ntn_loc ?loc (args,argslist,binderslist) = + locs_of_notation ?loc (List.map constr_loc (args@List.flatten argslist)@ - List.map_filter local_binders_loc binderslist) + List.map local_binders_loc binderslist) -let patntn_loc loc (args,argslist) = - locs_of_notation loc +let patntn_loc ?loc (args,argslist) = + locs_of_notation ?loc (List.map cases_pattern_expr_loc (args@List.flatten argslist)) diff --git a/interp/topconstr.mli b/interp/topconstr.mli index b6ac40041..fabb1cb93 100644 --- a/interp/topconstr.mli +++ b/interp/topconstr.mli @@ -40,9 +40,9 @@ val map_constr_expr_with_binders : 'a -> constr_expr -> constr_expr val ntn_loc : - Loc.t -> constr_notation_substitution -> string -> (int * int) list + ?loc:Loc.t -> constr_notation_substitution -> string -> (int * int) list val patntn_loc : - Loc.t -> cases_pattern_notation_substitution -> string -> (int * int) list + ?loc:Loc.t -> cases_pattern_notation_substitution -> string -> (int * int) list (** For cases pattern parsing errors *) diff --git a/lib/cWarnings.ml b/lib/cWarnings.ml index 4e692af36..ececc6308 100644 --- a/lib/cWarnings.ml +++ b/lib/cWarnings.ml @@ -23,7 +23,7 @@ let categories : (string, string list) Hashtbl.t = Hashtbl.create 97 let current_loc = ref None let flags = ref "" -let set_current_loc loc = current_loc := Some loc +let set_current_loc loc = current_loc := loc let get_flags () = !flags diff --git a/lib/cWarnings.mli b/lib/cWarnings.mli index 3f6cee31b..c1fb5d604 100644 --- a/lib/cWarnings.mli +++ b/lib/cWarnings.mli @@ -8,7 +8,7 @@ type status = Disabled | Enabled | AsError -val set_current_loc : Loc.t -> unit +val set_current_loc : Loc.t option -> unit val create : name:string -> category:string -> ?default:status -> ('a -> Pp.std_ppcmds) -> ?loc:Loc.t -> 'a -> unit diff --git a/lib/loc.ml b/lib/loc.ml index 3051ca7b9..e02fe108d 100644 --- a/lib/loc.ml +++ b/lib/loc.ml @@ -45,33 +45,32 @@ let merge loc1 loc2 = bp = loc2.bp; ep = loc1.ep; } else loc2 -let merge_opt l1 l2 = Option.cata (fun l1 -> merge l1 l2) l2 l1 -let opt_merge l1 l2 = Option.cata (fun l2 -> merge l1 l2) l1 l2 +let merge_opt l1 l2 = match l1, l2 with + | None, None -> None + | Some l , None -> Some l + | None, Some l -> Some l + | Some l1, Some l2 -> Some (merge l1 l2) let unloc loc = (loc.bp, loc.ep) let join_loc = merge -(** Located type *) - -type 'a located = t * 'a - -let is_ghost loc = loc.ep = 0 - -let ghost = { +let internal_ghost = { fname = ""; line_nb = -1; bol_pos = 0; line_nb_last = -1; bol_pos_last = 0; bp = 0; ep = 0; } -let internal_ghost = ghost +(** Located type *) +type 'a located = t option * 'a + let to_pair x = x -let tag ?loc x = Option.default ghost loc, x +let tag ?loc x = loc, x let obj (_,x) = x -let with_loc f (loc, x) = f ~loc x +let with_loc f (loc, x) = f ?loc x let with_unloc f (_,x) = f x let map f (l,x) = (l, f x) -let map_with_loc f (loc, x) = (loc, f ~loc x) +let map_with_loc f (loc, x) = (loc, f ?loc x) let located_fold_left f x (_,a) = f x a let located_iter2 f (_,a) (_,b) = f a b diff --git a/lib/loc.mli b/lib/loc.mli index 82792613c..6de6c584d 100644 --- a/lib/loc.mli +++ b/lib/loc.mli @@ -33,13 +33,11 @@ val make_loc : int * int -> t (** Make a location out of its start and end position *) val internal_ghost : t -val is_ghost : t -> bool + (** Test whether the location is meaningful *) val merge : t -> t -> t - -val merge_opt : t option -> t -> t -val opt_merge : t -> t option -> t +val merge_opt : t option -> t option -> t option (** {5 Located exceptions} *) @@ -54,20 +52,20 @@ val raise : ?loc:t -> exn -> 'a (** {5 Objects with location information } *) -type 'a located = t * 'a +type 'a located = t option * 'a (** Embed a location in a type *) (** Warning, this API is experimental *) -val to_pair : 'a located -> t * 'a +val to_pair : 'a located -> t option * 'a val tag : ?loc:t -> 'a -> 'a located val obj : 'a located -> 'a -val with_loc : (loc:t -> 'a -> 'b) -> 'a located -> 'b +val with_loc : (?loc:t -> 'a -> 'b) -> 'a located -> 'b val with_unloc : ('a -> 'b) -> 'a located -> 'b val map : ('a -> 'b) -> 'a located -> 'b located -val map_with_loc : (loc:t -> 'a -> 'b) -> 'a located -> 'b located +val map_with_loc : (?loc:t -> 'a -> 'b) -> 'a located -> 'b located val located_fold_left : ('a -> 'b -> 'a) -> 'a -> 'b located -> 'a val down_located : ('a -> 'b) -> 'a located -> 'b diff --git a/library/declare.ml b/library/declare.ml index 6b505ac09..3adb957fa 100644 --- a/library/declare.ml +++ b/library/declare.ml @@ -522,7 +522,7 @@ let do_constraint poly l = let names, _ = Global.global_universe_names () in try loc, Idmap.find id names with Not_found -> - user_err ~loc ~hdr:"Constraint" (str "Undeclared universe " ++ pr_id id) + user_err ?loc ~hdr:"Constraint" (str "Undeclared universe " ++ pr_id id) in let in_section = Lib.sections_are_opened () in let () = @@ -530,18 +530,18 @@ let do_constraint poly l = user_err ~hdr:"Constraint" (str"Cannot declare polymorphic constraints outside sections") in - let check_poly loc p loc' p' = + let check_poly ?loc p loc' p' = if poly then () else if p || p' then let loc = if p then loc else loc' in - user_err ~loc ~hdr:"Constraint" + user_err ?loc ~hdr:"Constraint" (str "Cannot declare a global constraint on " ++ str "a polymorphic universe, use " ++ str "Polymorphic Constraint instead") in let constraints = List.fold_left (fun acc (l, d, r) -> let ploc, (p, lu) = u_of_id l and rloc, (p', ru) = u_of_id r in - check_poly ploc p rloc p'; + check_poly ?loc:ploc p rloc p'; Univ.Constraint.add (lu, d, ru) acc) Univ.Constraint.empty l in diff --git a/library/libnames.mli b/library/libnames.mli index 58d1da9d6..57013ef82 100644 --- a/library/libnames.mli +++ b/library/libnames.mli @@ -125,7 +125,7 @@ val eq_reference : reference -> reference -> bool val qualid_of_reference : reference -> qualid located val string_of_reference : reference -> string val pr_reference : reference -> std_ppcmds -val loc_of_reference : reference -> Loc.t +val loc_of_reference : reference -> Loc.t option val join_reference : reference -> reference -> reference (** Deprecated synonyms *) diff --git a/library/library.ml b/library/library.ml index 3086e3d18..2b3381fa7 100644 --- a/library/library.ml +++ b/library/library.ml @@ -582,7 +582,7 @@ let require_library_from_dirpath modrefl export = let safe_locate_module (loc,qid) = try Nametab.locate_module qid with Not_found -> - user_err ~loc ~hdr:"import_library" + user_err ?loc ~hdr:"import_library" (pr_qualid qid ++ str " is not a module") let import_module export modl = @@ -597,7 +597,7 @@ let import_module export modl = | [] -> () | modl -> add_anonymous_leaf (in_import_library (List.rev modl, export)) in let rec aux acc = function - | (loc,dir as m) :: l -> + | (loc, dir as m) :: l -> let m,acc = try Nametab.locate_module dir, acc with Not_found-> flush acc; safe_locate_module m, [] in @@ -607,7 +607,7 @@ let import_module export modl = flush acc; try Declaremods.import_module export mp; aux [] l with Not_found -> - user_err ~loc ~hdr:"import_library" + user_err ?loc ~hdr:"import_library" (pr_qualid dir ++ str " is not a module")) | [] -> flush acc in aux [] modl diff --git a/library/nametab.ml b/library/nametab.ml index b76048e89..1027e291c 100644 --- a/library/nametab.ml +++ b/library/nametab.ml @@ -453,11 +453,11 @@ let global r = try match locate_extended qid with | TrueGlobal ref -> ref | SynDef _ -> - user_err ~loc ~hdr:"global" + user_err ?loc ~hdr:"global" (str "Unexpected reference to a notation: " ++ pr_qualid qid) with Not_found -> - error_global_not_found ~loc qid + error_global_not_found ?loc qid (* Exists functions ********************************************************) @@ -532,7 +532,7 @@ let global_inductive r = match global r with | IndRef ind -> ind | ref -> - user_err ~loc:(loc_of_reference r) ~hdr:"global_inductive" + user_err ?loc:(loc_of_reference r) ~hdr:"global_inductive" (pr_reference r ++ spc () ++ str "is not an inductive type") diff --git a/parsing/egramcoq.ml b/parsing/egramcoq.ml index c8fe96f74..49bf267db 100644 --- a/parsing/egramcoq.ml +++ b/parsing/egramcoq.ml @@ -313,13 +313,13 @@ let interp_entry forpat e = match e with | ETBinderList (true, _) -> assert false | ETBinderList (false, tkl) -> TTAny (TTBinderListF tkl) -let constr_expr_of_name (loc,na) = Loc.tag ~loc @@ match na with +let constr_expr_of_name (loc,na) = Loc.tag ?loc @@ match na with | Anonymous -> CHole (None,Misctypes.IntroAnonymous,None) - | Name id -> CRef (Ident (loc,id), None) + | Name id -> CRef (Ident (Loc.tag ?loc id), None) -let cases_pattern_expr_of_name (loc,na) = Loc.tag ~loc @@ match na with +let cases_pattern_expr_of_name (loc,na) = Loc.tag ?loc @@ match na with | Anonymous -> CPatAtom None - | Name id -> CPatAtom (Some (Ident (loc,id))) + | Name id -> CPatAtom (Some (Ident (Loc.tag ?loc id))) type 'r env = { constrs : 'r list; diff --git a/parsing/g_constr.ml4 b/parsing/g_constr.ml4 index e6c82c894..45585d9ce 100644 --- a/parsing/g_constr.ml4 +++ b/parsing/g_constr.ml4 @@ -33,12 +33,12 @@ let _ = List.iter CLexer.add_keyword constr_kw let mk_cast = function (c,(_,None)) -> c | (c,(_,Some ty)) -> - let loc = Loc.merge (constr_loc c) (constr_loc ty) - in Loc.tag ~loc @@ CCast(c, CastConv ty) + let loc = Loc.merge_opt (constr_loc c) (constr_loc ty) + in Loc.tag ?loc @@ CCast(c, CastConv ty) let binder_of_name expl (loc,na) = CLocalAssum ([loc, na], Default expl, - Loc.tag ~loc @@ CHole (Some (Evar_kinds.BinderType na), IntroAnonymous, None)) + Loc.tag ?loc @@ CHole (Some (Evar_kinds.BinderType na), IntroAnonymous, None)) let binders_of_names l = List.map (binder_of_name Explicit) l @@ -51,7 +51,7 @@ let mk_fixb (id,bl,ann,body,(loc,tyc)) = let mk_cofixb (id,bl,ann,body,(loc,tyc)) = let _ = Option.map (fun (aloc,_) -> - CErrors.user_err ~loc:aloc + CErrors.user_err ?loc:aloc ~hdr:"Constr:mk_cofixb" (Pp.str"Annotation forbidden in cofix expression.")) (fst ann) in let ty = match tyc with @@ -131,7 +131,7 @@ GEXTEND Gram [ [ id = Prim.ident -> id ] ] ; Prim.name: - [ [ "_" -> (!@loc, Anonymous) ] ] + [ [ "_" -> Loc.tag ~loc:!@loc Anonymous ] ] ; global: [ [ r = Prim.reference -> r ] ] @@ -183,13 +183,13 @@ GEXTEND Gram | "90" RIGHTA [ ] | "10" LEFTA [ f=operconstr; args=LIST1 appl_arg -> Loc.tag ~loc:(!@loc) @@ CApp((None,f),args) - | "@"; f=global; i = instance; args=LIST0 NEXT -> Loc.tag ~loc:(!@loc) @@ CAppExpl((None,f,i),args) + | "@"; f=global; i = instance; args=LIST0 NEXT -> Loc.tag ~loc:!@loc @@ CAppExpl((None,f,i),args) | "@"; (locid,id) = pattern_identref; args=LIST1 identref -> let args = List.map (fun x -> Loc.tag @@ CRef (Ident x,None), None) args in - Loc.tag ~loc:(!@loc) @@ CApp((None, Loc.tag ~loc:locid @@ CPatVar id),args) ] + Loc.tag ~loc:(!@loc) @@ CApp((None, Loc.tag ?loc:locid @@ CPatVar id),args) ] | "9" [ ".."; c = operconstr LEVEL "0"; ".." -> - Loc.tag ~loc:(!@loc) @@ CAppExpl ((None, Ident (!@loc,ldots_var),None),[c]) ] + Loc.tag ~loc:!@loc @@ CAppExpl ((None, Ident (Loc.tag ~loc:!@loc ldots_var),None),[c]) ] | "8" [ ] | "1" LEFTA [ c=operconstr; ".("; f=global; args=LIST0 appl_arg; ")" -> @@ -236,10 +236,10 @@ GEXTEND Gram | "let"; id=name; bl = binders; ty = type_cstr; ":="; c1 = operconstr LEVEL "200"; "in"; c2 = operconstr LEVEL "200" -> let ty,c1 = match ty, c1 with - | (_,None), (_, CCast(c, CastConv t)) -> (Loc.tag ~loc:(constr_loc t) @@ Some t), c (* Tolerance, see G_vernac.def_body *) + | (_,None), (_, CCast(c, CastConv t)) -> (Loc.tag ?loc:(constr_loc t) @@ Some t), c (* Tolerance, see G_vernac.def_body *) | _, _ -> ty, c1 in - Loc.tag ~loc:!@loc @@ CLetIn(id,mkCLambdaN ~loc:(constr_loc c1) bl c1, - Option.map (mkCProdN ~loc:(fst ty) bl) (snd ty), c2) + Loc.tag ~loc:!@loc @@ CLetIn(id,mkCLambdaN ?loc:(constr_loc c1) bl c1, + Option.map (mkCProdN ?loc:(fst ty) bl) (snd ty), c2) | "let"; fx = single_fix; "in"; c = operconstr LEVEL "200" -> let fixp = mk_single_fix fx in let (li,id) = match snd fixp with @@ -254,14 +254,18 @@ GEXTEND Gram Loc.tag ~loc:!@loc @@ CLetTuple (lb,po,c1,c2) | "let"; "'"; p=pattern; ":="; c1 = operconstr LEVEL "200"; "in"; c2 = operconstr LEVEL "200" -> - Loc.tag ~loc:!@loc @@ CCases (LetPatternStyle, None, [c1, None, None], [Loc.tag ~loc:!@loc ([(!@loc, [p])], c2)]) + Loc.tag ~loc:!@loc @@ + CCases (LetPatternStyle, None, [c1, None, None], [Loc.tag ~loc:!@loc ([(Loc.tag ~loc:!@loc [p])], c2)]) | "let"; "'"; p=pattern; ":="; c1 = operconstr LEVEL "200"; rt = case_type; "in"; c2 = operconstr LEVEL "200" -> - Loc.tag ~loc:!@loc @@ CCases (LetPatternStyle, Some rt, [c1, aliasvar p, None], [Loc.tag ~loc:!@loc ([(!@loc, [p])], c2)]) + Loc.tag ~loc:!@loc @@ + CCases (LetPatternStyle, Some rt, [c1, aliasvar p, None], [Loc.tag ~loc:!@loc ([(Loc.tag ~loc:!@loc [p])], c2)]) + | "let"; "'"; p=pattern; "in"; t = pattern LEVEL "200"; ":="; c1 = operconstr LEVEL "200"; rt = case_type; "in"; c2 = operconstr LEVEL "200" -> - Loc.tag ~loc:!@loc @@ CCases (LetPatternStyle, Some rt, [c1, aliasvar p, Some t], [Loc.tag ~loc:!@loc ([(!@loc, [p])], c2)]) + Loc.tag ~loc:!@loc @@ + CCases (LetPatternStyle, Some rt, [c1, aliasvar p, Some t], [Loc.tag ~loc:!@loc ([(Loc.tag ~loc:!@loc [p])], c2)]) | "if"; c=operconstr LEVEL "200"; po = return_type; "then"; b1=operconstr LEVEL "200"; "else"; b2=operconstr LEVEL "200" -> @@ -270,7 +274,7 @@ GEXTEND Gram ; appl_arg: [ [ id = lpar_id_coloneq; c=lconstr; ")" -> - (c,Some (!@loc,ExplByName id)) + (c,Some (Loc.tag ~loc:!@loc @@ ExplByName id)) | c=operconstr LEVEL "9" -> (c,None) ] ] ; atomic_constr: @@ -345,7 +349,7 @@ GEXTEND Gram [ [ OPT"|"; br=LIST0 eqn SEP "|" -> br ] ] ; mult_pattern: - [ [ pl = LIST1 pattern LEVEL "99" SEP "," -> (!@loc,pl) ] ] + [ [ pl = LIST1 pattern LEVEL "99" SEP "," -> (Loc.tag ~loc:!@loc pl) ] ] ; eqn: [ [ pll = LIST1 mult_pattern SEP "|"; @@ -364,7 +368,7 @@ GEXTEND Gram pattern: [ "200" RIGHTA [ ] | "100" RIGHTA - [ p = pattern; "|"; pl = LIST1 pattern SEP "|" -> !@loc, CPatOr (p::pl) ] + [ p = pattern; "|"; pl = LIST1 pattern SEP "|" -> Loc.tag ~loc:!@loc @@ CPatOr (p::pl) ] | "99" RIGHTA [ ] | "11" LEFTA [ p = pattern; "as"; id = ident -> @@ -374,21 +378,21 @@ GEXTEND Gram (match p with | _, CPatAtom (Some r) -> Loc.tag ~loc:!@loc @@ CPatCstr (r, None, lp) | loc, CPatCstr (r, None, l2) -> - CErrors.user_err ~loc ~hdr:"compound_pattern" + CErrors.user_err ?loc ~hdr:"compound_pattern" (Pp.str "Nested applications not supported.") | _, CPatCstr (r, l1, l2) -> Loc.tag ~loc:!@loc @@ CPatCstr (r, l1 , l2@lp) | _, CPatNotation (n, s, l) -> Loc.tag ~loc:!@loc @@ CPatNotation (n , s, l@lp) | _ -> CErrors.user_err - ~loc:(cases_pattern_expr_loc p) ~hdr:"compound_pattern" - (Pp.str "Such pattern cannot have arguments.")) + ?loc:(cases_pattern_expr_loc p) ~hdr:"compound_pattern" + (Pp.str "Such pattern cannot have arguments.")) |"@"; r = Prim.reference; lp = LIST0 NEXT -> - !@loc, CPatCstr (r, Some lp, []) ] + Loc.tag ~loc:!@loc @@ CPatCstr (r, Some lp, []) ] | "1" LEFTA - [ c = pattern; "%"; key=IDENT -> !@loc, CPatDelimiters (key,c) ] + [ c = pattern; "%"; key=IDENT -> Loc.tag ~loc:!@loc @@ CPatDelimiters (key,c) ] | "0" [ r = Prim.reference -> Loc.tag ~loc:!@loc @@ CPatAtom (Some r) | "{|"; pat = record_patterns; "|}" -> Loc.tag ~loc:!@loc @@ CPatRecord pat - | "_" -> !@loc, CPatAtom None + | "_" -> Loc.tag ~loc:!@loc @@ CPatAtom None | "("; p = pattern LEVEL "200"; ")" -> (match p with | _, CPatPrim (Numeral z) when Bigint.is_pos_or_zero z -> @@ -401,7 +405,7 @@ GEXTEND Gram Loc.tag ~loc:!@loc @@ CPatNotation("( _ )",([p],[]),[]) | _ -> p in - !@loc, CPatCast (p, ty) + Loc.tag ~loc:!@loc @@ CPatCast (p, ty) | n = INT -> Loc.tag ~loc:!@loc @@ CPatPrim (Numeral (Bigint.of_string n)) | s = string -> Loc.tag ~loc:!@loc @@ CPatPrim (String s) ] ] ; @@ -411,7 +415,7 @@ GEXTEND Gram (fun na -> CLocalAssum (na::nal,Default Implicit,c)) | nal=LIST1 name; "}" -> (fun na -> CLocalAssum (na::nal,Default Implicit, - Loc.tag ~loc:(Loc.merge (fst na) !@loc) @@ CHole (Some (Evar_kinds.BinderType (snd na)), IntroAnonymous, None))) + Loc.tag ?loc:(Loc.merge_opt (fst na) (Some !@loc)) @@ CHole (Some (Evar_kinds.BinderType (snd na)), IntroAnonymous, None))) | ":"; c=lconstr; "}" -> (fun na -> CLocalAssum ([na],Default Implicit,c)) ] ] @@ -424,7 +428,7 @@ GEXTEND Gram ] ] ; impl_name_head: - [ [ id = impl_ident_head -> (!@loc,Name id) ] ] + [ [ id = impl_ident_head -> (Loc.tag ~loc:!@loc @@ Name id) ] ] ; binders_fixannot: [ [ na = impl_name_head; assum = impl_ident_tail; bl = binders_fixannot -> @@ -444,7 +448,7 @@ GEXTEND Gram | id = name; idl = LIST0 name; bl = binders -> binders_of_names (id::idl) @ bl | id1 = name; ".."; id2 = name -> - [CLocalAssum ([id1;(!@loc,Name ldots_var);id2], + [CLocalAssum ([id1;(Loc.tag ~loc:!@loc (Name ldots_var));id2], Default Explicit, Loc.tag ~loc:!@loc @@ CHole (None, IntroAnonymous, None))] | bl = closed_binder; bl' = binders -> bl@bl' @@ -490,17 +494,17 @@ GEXTEND Gram ] ] ; typeclass_constraint: - [ [ "!" ; c = operconstr LEVEL "200" -> (!@loc, Anonymous), true, c + [ [ "!" ; c = operconstr LEVEL "200" -> (Loc.tag ~loc:!@loc Anonymous), true, c | "{"; id = name; "}"; ":" ; expl = [ "!" -> true | -> false ] ; c = operconstr LEVEL "200" -> id, expl, c | iid=name_colon ; expl = [ "!" -> true | -> false ] ; c = operconstr LEVEL "200" -> - (!@loc, iid), expl, c + (Loc.tag ~loc:!@loc iid), expl, c | c = operconstr LEVEL "200" -> - (!@loc, Anonymous), false, c + (Loc.tag ~loc:!@loc Anonymous), false, c ] ] ; type_cstr: - [ [ c=OPT [":"; c=lconstr -> c] -> (!@loc,c) ] ] + [ [ c=OPT [":"; c=lconstr -> c] -> Loc.tag ~loc:!@loc c ] ] ; END;; diff --git a/parsing/g_prim.ml4 b/parsing/g_prim.ml4 index ed6a8df4e..d94c5e963 100644 --- a/parsing/g_prim.ml4 +++ b/parsing/g_prim.ml4 @@ -44,13 +44,13 @@ GEXTEND Gram [ [ LEFTQMARK; id = ident -> id ] ] ; pattern_identref: - [ [ id = pattern_ident -> (!@loc, id) ] ] + [ [ id = pattern_ident -> Loc.tag ~loc:!@loc id ] ] ; var: (* as identref, but interpret as a term identifier in ltac *) - [ [ id = ident -> (!@loc, id) ] ] + [ [ id = ident -> Loc.tag ~loc:!@loc id ] ] ; identref: - [ [ id = ident -> (!@loc, id) ] ] + [ [ id = ident -> Loc.tag ~loc:!@loc id ] ] ; field: [ [ s = FIELD -> Id.of_string s ] ] @@ -61,8 +61,8 @@ GEXTEND Gram ] ] ; fullyqualid: - [ [ id = ident; (l,id')=fields -> !@loc,id::List.rev (id'::l) - | id = ident -> !@loc,[id] + [ [ id = ident; (l,id')=fields -> Loc.tag ~loc:!@loc @@ id::List.rev (id'::l) + | id = ident -> Loc.tag ~loc:!@loc [id] ] ] ; basequalid: @@ -71,13 +71,13 @@ GEXTEND Gram ] ] ; name: - [ [ IDENT "_" -> (!@loc, Anonymous) - | id = ident -> (!@loc, Name id) ] ] + [ [ IDENT "_" -> Loc.tag ~loc:!@loc Anonymous + | id = ident -> Loc.tag ~loc:!@loc @@ Name id ] ] ; reference: [ [ id = ident; (l,id') = fields -> - Qualid (!@loc, local_make_qualid (l@[id]) id') - | id = ident -> Ident (!@loc,id) + Qualid (Loc.tag ~loc:!@loc @@ local_make_qualid (l@[id]) id') + | id = ident -> Ident (Loc.tag ~loc:!@loc id) ] ] ; by_notation: @@ -88,15 +88,15 @@ GEXTEND Gram | ntn = by_notation -> Misctypes.ByNotation ntn ] ] ; qualid: - [ [ qid = basequalid -> !@loc, qid ] ] + [ [ qid = basequalid -> Loc.tag ~loc:!@loc qid ] ] ; ne_string: [ [ s = STRING -> - if s="" then CErrors.user_err ~loc:(!@loc) (Pp.str"Empty string."); s + if s="" then CErrors.user_err ~loc:!@loc (Pp.str"Empty string."); s ] ] ; ne_lstring: - [ [ s = ne_string -> (!@loc, s) ] ] + [ [ s = ne_string -> Loc.tag ~loc:!@loc s ] ] ; dirpath: [ [ id = ident; l = LIST0 field -> @@ -106,7 +106,7 @@ GEXTEND Gram [ [ s = STRING -> s ] ] ; lstring: - [ [ s = string -> (!@loc, s) ] ] + [ [ s = string -> (Loc.tag ~loc:!@loc s) ] ] ; integer: [ [ i = INT -> my_int_of_string (!@loc) i diff --git a/parsing/g_vernac.ml4 b/parsing/g_vernac.ml4 index 7a3a2ace1..c8101875d 100644 --- a/parsing/g_vernac.ml4 +++ b/parsing/g_vernac.ml4 @@ -102,7 +102,7 @@ GEXTEND Gram ; located_vernac: - [ [ v = vernac -> !@loc, v ] ] + [ [ v = vernac -> Loc.tag ~loc:!@loc v ] ] ; END @@ -542,8 +542,8 @@ GEXTEND Gram starredidentref: [ [ i = identref -> SsSingl i | i = identref; "*" -> SsFwdClose(SsSingl i) - | "Type" -> SsSingl (!@loc, Id.of_string "Type") - | "Type"; "*" -> SsFwdClose (SsSingl (!@loc, Id.of_string "Type")) ]] + | "Type" -> SsSingl (Loc.tag ~loc:!@loc @@ Id.of_string "Type") + | "Type"; "*" -> SsFwdClose (SsSingl (Loc.tag ~loc:!@loc @@ Id.of_string "Type")) ]] ; ssexpr: [ "35" @@ -720,7 +720,7 @@ GEXTEND Gram ; argument_spec: [ [ b = OPT "!"; id = name ; s = OPT scope -> - snd id, not (Option.is_empty b), Option.map (fun x -> !@loc, x) s + snd id, not (Option.is_empty b), Option.map (fun x -> Loc.tag ~loc:!@loc x) s ] ]; (* List of arguments implicit status, scope, modifiers *) @@ -733,7 +733,7 @@ GEXTEND Gram | "/" -> [`Slash] | "("; items = LIST1 argument_spec; ")"; sc = OPT scope -> let f x = match sc, x with - | None, x -> x | x, None -> Option.map (fun y -> !@loc, y) x + | None, x -> x | x, None -> Option.map (fun y -> Loc.tag ~loc:!@loc y) x | Some _, Some _ -> error "scope declared twice" in List.map (fun (name,recarg_like,notation_scope) -> `Id { name=name; recarg_like=recarg_like; @@ -741,7 +741,7 @@ GEXTEND Gram implicit_status = NotImplicit}) items | "["; items = LIST1 argument_spec; "]"; sc = OPT scope -> let f x = match sc, x with - | None, x -> x | x, None -> Option.map (fun y -> !@loc, y) x + | None, x -> x | x, None -> Option.map (fun y -> Loc.tag ~loc:!@loc y) x | Some _, Some _ -> error "scope declared twice" in List.map (fun (name,recarg_like,notation_scope) -> `Id { name=name; recarg_like=recarg_like; @@ -749,7 +749,7 @@ GEXTEND Gram implicit_status = Implicit}) items | "{"; items = LIST1 argument_spec; "}"; sc = OPT scope -> let f x = match sc, x with - | None, x -> x | x, None -> Option.map (fun y -> !@loc, y) x + | None, x -> x | x, None -> Option.map (fun y -> Loc.tag ~loc:!@loc y) x | Some _, Some _ -> error "scope declared twice" in List.map (fun (name,recarg_like,notation_scope) -> `Id { name=name; recarg_like=recarg_like; @@ -782,7 +782,7 @@ GEXTEND Gram [ [ name = pidentref; sup = OPT binders -> (let ((loc,id),l) = name in ((loc, Name id),l)), (Option.default [] sup) - | -> ((!@loc, Anonymous), None), [] ] ] + | -> ((Loc.tag ~loc:!@loc Anonymous), None), [] ] ] ; hint_info: [ [ "|"; i = OPT natural; pat = OPT constr_pattern -> @@ -1149,8 +1149,8 @@ GEXTEND Gram | IDENT "only"; IDENT "parsing" -> SetOnlyParsing | IDENT "compat"; s = STRING -> SetCompatVersion (Coqinit.get_compat_version s) - | IDENT "format"; s1 = [s = STRING -> (!@loc,s)]; - s2 = OPT [s = STRING -> (!@loc,s)] -> + | IDENT "format"; s1 = [s = STRING -> Loc.tag ~loc:!@loc s]; + s2 = OPT [s = STRING -> Loc.tag ~loc:!@loc s] -> begin match s1, s2 with | (_,k), Some s -> SetFormat(k,s) | s, None -> SetFormat ("text",s) end diff --git a/plugins/extraction/table.ml b/plugins/extraction/table.ml index d6a334c5f..0cdb10a37 100644 --- a/plugins/extraction/table.ml +++ b/plugins/extraction/table.ml @@ -892,7 +892,7 @@ let extract_constant_inline inline r ids s = let extract_inductive r s l optstr = check_inside_section (); let g = Smartlocate.global_with_alias r in - Dumpglob.add_glob (loc_of_reference r) g; + Dumpglob.add_glob ?loc:(loc_of_reference r) g; match g with | IndRef ((kn,i) as ip) -> let mib = Global.lookup_mind kn in diff --git a/plugins/funind/g_indfun.ml4 b/plugins/funind/g_indfun.ml4 index 129c8dc16..ee82d95d0 100644 --- a/plugins/funind/g_indfun.ml4 +++ b/plugins/funind/g_indfun.ml4 @@ -159,7 +159,7 @@ GEXTEND Gram GLOBAL: function_rec_definition_loc ; function_rec_definition_loc: - [ [ g = Vernac.rec_definition -> !@loc, g ]] + [ [ g = Vernac.rec_definition -> Loc.tag ~loc:!@loc g ]] ; END diff --git a/plugins/funind/glob_term_to_relation.ml b/plugins/funind/glob_term_to_relation.ml index 946ee55d4..48ab3dd57 100644 --- a/plugins/funind/glob_term_to_relation.ml +++ b/plugins/funind/glob_term_to_relation.ml @@ -610,7 +610,7 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return = then the one corresponding to the value [t] and combine the two result *) - let v = match typ with None -> v | Some t -> Loc.tag ~loc @@ GCast (v,CastConv t) in + let v = match typ with None -> v | Some t -> Loc.tag ?loc @@ GCast (v,CastConv t) in let v_res = build_entry_lc env funnames avoid v in let v_as_constr,ctx = Pretyping.understand env (Evd.from_env env) v in let v_type = Typing.unsafe_type_of env (Evd.from_env env) (EConstr.of_constr v_as_constr) in @@ -973,7 +973,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt = (mkGHole ())))) in let eq' = - Loc.tag ~loc:loc1 @@ GApp(Loc.tag ~loc:loc2 @@GRef(jmeq,None),[ty;Loc.tag ~loc:loc3 @@ GVar id;rt_typ;rt]) + Loc.tag ?loc:loc1 @@ GApp(Loc.tag ?loc:loc2 @@GRef(jmeq,None),[ty;Loc.tag ?loc:loc3 @@ GVar id;rt_typ;rt]) in observe (str "computing new type for jmeq : " ++ pr_glob_constr eq'); let eq'_as_constr,ctx = Pretyping.understand env (Evd.from_env env) eq' in @@ -1119,7 +1119,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt = end | loc, GLetIn(n,v,t,b) -> begin - let t = match t with None -> v | Some t -> Loc.tag ~loc @@ GCast (v,CastConv t) in + let t = match t with None -> v | Some t -> Loc.tag ?loc @@ GCast (v,CastConv t) in let not_free_in_t id = not (is_free_in id t) in let evd = (Evd.from_env env) in let t',ctx = Pretyping.understand env evd t in @@ -1249,11 +1249,11 @@ let compute_params_name relnames (args : (Name.t * Glob_term.glob_constr * glob_ let rec rebuild_return_type (loc, rt) = match rt with | Constrexpr.CProdN(n,t') -> - Loc.tag ~loc @@ Constrexpr.CProdN(n,rebuild_return_type t') + Loc.tag ?loc @@ Constrexpr.CProdN(n,rebuild_return_type t') | Constrexpr.CLetIn(na,v,t,t') -> - Loc.tag ~loc @@ Constrexpr.CLetIn(na,v,t,rebuild_return_type t') - | _ -> Loc.tag ~loc @@ Constrexpr.CProdN([[Loc.tag Anonymous], - Constrexpr.Default Decl_kinds.Explicit,Loc.tag ~loc rt], + Loc.tag ?loc @@ Constrexpr.CLetIn(na,v,t,rebuild_return_type t') + | _ -> Loc.tag ?loc @@ Constrexpr.CProdN([[Loc.tag Anonymous], + Constrexpr.Default Decl_kinds.Explicit,Loc.tag ?loc rt], Loc.tag @@ Constrexpr.CSort(GType [])) let do_build_inductive diff --git a/plugins/funind/glob_termops.ml b/plugins/funind/glob_termops.ml index 01e607412..66b9897d0 100644 --- a/plugins/funind/glob_termops.ml +++ b/plugins/funind/glob_termops.ml @@ -193,14 +193,14 @@ let rec alpha_pat excluded (loc, pat) = match pat with | PatVar Anonymous -> let new_id = Indfun_common.fresh_id excluded "_x" in - (Loc.tag ~loc @@ PatVar(Name new_id)),(new_id::excluded),Id.Map.empty + (Loc.tag ?loc @@ PatVar(Name new_id)),(new_id::excluded),Id.Map.empty | PatVar(Name id) -> if Id.List.mem id excluded then let new_id = Namegen.next_ident_away id excluded in - (Loc.tag ~loc @@ PatVar(Name new_id)),(new_id::excluded), + (Loc.tag ?loc @@ PatVar(Name new_id)),(new_id::excluded), (Id.Map.add id new_id Id.Map.empty) - else (Loc.tag ~loc pat),excluded,Id.Map.empty + else (Loc.tag ?loc pat),excluded,Id.Map.empty | PatCstr(constr,patl,na) -> let new_na,new_excluded,map = match na with @@ -218,7 +218,7 @@ let rec alpha_pat excluded (loc, pat) = ([],new_excluded,map) patl in - (Loc.tag ~loc @@ PatCstr(constr,List.rev new_patl,new_na)),new_excluded,new_map + (Loc.tag ?loc @@ PatCstr(constr,List.rev new_patl,new_na)),new_excluded,new_map let alpha_patl excluded patl = let patl,new_excluded,map = @@ -255,7 +255,7 @@ let raw_get_pattern_id pat acc = let get_pattern_id pat = raw_get_pattern_id pat [] let rec alpha_rt excluded (loc, rt) = - let new_rt = Loc.tag ~loc @@ + let new_rt = Loc.tag ?loc @@ match rt with | GRef _ | GVar _ | GEvar _ | GPatVar _ -> rt | GLambda(Anonymous,k,t,b) -> @@ -445,7 +445,7 @@ let rec pattern_to_term pt = Loc.with_unloc (function let replace_var_by_term x_id term = - let rec replace_var_by_pattern (loc, rt) = Loc.tag ~loc @@ + let rec replace_var_by_pattern (loc, rt) = Loc.tag ?loc @@ match rt with | GRef _ -> rt | GVar id when Id.compare id x_id == 0 -> Loc.obj term @@ -605,7 +605,7 @@ let ids_of_glob_constr c = let zeta_normalize = - let rec zeta_normalize_term (loc, rt) = Loc.tag ~loc @@ + let rec zeta_normalize_term (loc, rt) = Loc.tag ?loc @@ match rt with | GRef _ -> rt | GVar _ -> rt @@ -673,7 +673,7 @@ let expand_as = | PatCstr(_,patl,_) -> List.fold_left add_as map patl in let rec expand_as map (loc, rt) = - Loc.tag ~loc @@ + Loc.tag ?loc @@ match rt with | GRef _ | GEvar _ | GPatVar _ | GSort _ | GHole _ -> rt | GVar id -> diff --git a/plugins/funind/indfun.ml b/plugins/funind/indfun.ml index 1da85c467..3dc16626c 100644 --- a/plugins/funind/indfun.ml +++ b/plugins/funind/indfun.ml @@ -756,7 +756,7 @@ let rec add_args id new_args = Loc.map (function List.map (fun (b,na,b_option) -> add_args id new_args b, na, b_option) cel, - List.map (fun (loc,(cpl,e)) -> Loc.tag ~loc @@ (cpl,add_args id new_args e)) cal + List.map (fun (loc,(cpl,e)) -> Loc.tag ?loc @@ (cpl,add_args id new_args e)) cal ) | CLetTuple(nal,(na,b_option),b1,b2) -> CLetTuple(nal,(na,Option.map (add_args id new_args) b_option), @@ -882,7 +882,7 @@ let make_graph (f_ref:global_reference) = | Constrexpr.CLocalDef (na,_,_)-> [] | Constrexpr.CLocalAssum (nal,_,_) -> List.map - (fun (loc,n) -> Loc.tag ~loc @@ + (fun (loc,n) -> Loc.tag ?loc @@ CRef(Libnames.Ident(loc, Nameops.out_name n),None)) nal | Constrexpr.CLocalPattern _ -> assert false diff --git a/plugins/ltac/extratactics.ml4 b/plugins/ltac/extratactics.ml4 index 5123d6c40..dc43930e4 100644 --- a/plugins/ltac/extratactics.ml4 +++ b/plugins/ltac/extratactics.ml4 @@ -267,7 +267,7 @@ let add_rewrite_hint bases ort t lcsr = (Declare.declare_universe_context false ctx; Univ.ContextSet.empty) in - Constrexpr_ops.constr_loc ce, (c, ctx), ort, Option.map (in_gen (rawwit wit_ltac)) t in + Loc.tag ?loc:(Constrexpr_ops.constr_loc ce) ((c, ctx), ort, Option.map (in_gen (rawwit wit_ltac)) t) in let eqs = List.map f lcsr in let add_hints base = add_rew_rules base eqs in List.iter add_hints bases @@ -679,8 +679,8 @@ let hResolve id c occ t = with | Pretype_errors.PretypeError (_,_,Pretype_errors.UnsolvableImplicit _) as e -> let (e, info) = CErrors.push e in - let loc = match Loc.get_loc info with None -> 0,0 | Some loc -> Loc.unloc loc in - resolve_hole (subst_hole_with_term (fst loc) c_raw t_hole) + let loc_begin = Option.cata (fun l -> fst (Loc.unloc l)) 0 (Loc.get_loc info) in + resolve_hole (subst_hole_with_term loc_begin c_raw t_hole) in let t_constr,ctx = resolve_hole (subst_var_with_hole occ id t_raw) in let t_constr = EConstr.of_constr t_constr in diff --git a/plugins/ltac/g_ltac.ml4 b/plugins/ltac/g_ltac.ml4 index 23643c5d3..e20beae96 100644 --- a/plugins/ltac/g_ltac.ml4 +++ b/plugins/ltac/g_ltac.ml4 @@ -41,7 +41,7 @@ let in_tac tac = in_gen (rawwit Tacarg.wit_ltac) tac let reference_to_id = function | Libnames.Ident (loc, id) -> (loc, id) | Libnames.Qualid (loc,_) -> - CErrors.user_err ~loc + CErrors.user_err ?loc (str "This expression should be a simple identifier.") let tactic_mode = Gram.entry_create "vernac:tactic_command" @@ -159,9 +159,9 @@ GEXTEND Gram | g=failkw; n = [ n = int_or_var -> n | -> fail_default_value ]; l = LIST0 message_token -> TacFail (g,n,l) | st = simple_tactic -> st - | a = tactic_arg -> TacArg(!@loc,a) + | a = tactic_arg -> TacArg(Loc.tag ~loc:!@loc a) | r = reference; la = LIST0 tactic_arg_compat -> - TacArg(!@loc, TacCall (!@loc,(r,la))) ] + TacArg(Loc.tag ~loc:!@loc @@ TacCall (Loc.tag ~loc:!@loc (r,la))) ] | "0" [ "("; a = tactic_expr; ")" -> a | "["; ">"; (tf,tail) = tactic_then_gen; "]" -> @@ -169,7 +169,7 @@ GEXTEND Gram | Some (t,tl) -> TacExtendTac(Array.of_list tf,t,tl) | None -> TacDispatch tf end - | a = tactic_atom -> TacArg (!@loc,a) ] ] + | a = tactic_atom -> TacArg (Loc.tag ~loc:!@loc a) ] ] ; failkw: [ [ IDENT "fail" -> TacLocal | IDENT "gfail" -> TacGlobal ] ] @@ -203,7 +203,7 @@ GEXTEND Gram verbose most of the time. *) fresh_id: [ [ s = STRING -> ArgArg s (*| id = ident -> ArgVar (!@loc,id)*) - | qid = qualid -> let (_pth,id) = Libnames.repr_qualid (snd qid) in ArgVar (!@loc,id) ] ] + | qid = qualid -> let (_pth,id) = Libnames.repr_qualid (snd qid) in ArgVar (Loc.tag ~loc:!@loc id) ] ] ; constr_eval: [ [ IDENT "eval"; rtc = red_expr; "in"; c = Constr.constr -> @@ -219,7 +219,7 @@ GEXTEND Gram ; tactic_atom: [ [ n = integer -> TacGeneric (genarg_of_int n) - | r = reference -> TacCall (!@loc,(r,[])) + | r = reference -> TacCall (Loc.tag ~loc:!@loc (r,[])) | "()" -> TacGeneric (genarg_of_unit ()) ] ] ; match_key: @@ -470,7 +470,7 @@ let pr_ltac_production_item = function VERNAC ARGUMENT EXTEND ltac_production_item PRINTED BY pr_ltac_production_item | [ string(s) ] -> [ Tacentries.TacTerm s ] | [ ident(nt) "(" ident(p) ltac_production_sep_opt(sep) ")" ] -> - [ Tacentries.TacNonTerm (loc, ((Names.Id.to_string nt, sep), p)) ] + [ Tacentries.TacNonTerm (Loc.tag ~loc ((Names.Id.to_string nt, sep), p)) ] END VERNAC COMMAND EXTEND VernacTacticNotation diff --git a/plugins/ltac/g_tactic.ml4 b/plugins/ltac/g_tactic.ml4 index d1e5c810f..8aaad0566 100644 --- a/plugins/ltac/g_tactic.ml4 +++ b/plugins/ltac/g_tactic.ml4 @@ -166,20 +166,20 @@ let mkTacCase with_evar = function error "Use of numbers as direct arguments of 'case' is not supported."; TacInductionDestruct (false,with_evar,ic) -let rec mkCLambdaN_simple_loc loc bll c = +let rec mkCLambdaN_simple_loc ?loc bll c = match bll with | ((loc1,_)::_ as idl,bk,t) :: bll -> - Loc.tag ~loc @@ CLambdaN ([idl,bk,t],mkCLambdaN_simple_loc (Loc.merge loc1 loc) bll c) - | ([],_,_) :: bll -> mkCLambdaN_simple_loc loc bll c + Loc.tag ?loc @@ CLambdaN ([idl,bk,t],mkCLambdaN_simple_loc ?loc:(Loc.merge_opt loc1 loc) bll c) + | ([],_,_) :: bll -> mkCLambdaN_simple_loc ?loc bll c | [] -> c let mkCLambdaN_simple bl c = match bl with | [] -> c | h :: _ -> - let loc = Loc.merge (fst (List.hd (pi1 h))) (Constrexpr_ops.constr_loc c) in - mkCLambdaN_simple_loc loc bl c + let loc = Loc.merge_opt (fst (List.hd (pi1 h))) (Constrexpr_ops.constr_loc c) in + mkCLambdaN_simple_loc ?loc bl c -let loc_of_ne_list l = Loc.merge (fst (List.hd l)) (fst (List.last l)) +let loc_of_ne_list l = Loc.merge_opt (fst (List.hd l)) (fst (List.last l)) let map_int_or_var f = function | ArgArg x -> ArgArg (f x) @@ -301,7 +301,7 @@ GEXTEND Gram (* (A & B & C) is translated into (A,(B,C)) *) let rec pairify = function | ([]|[_]|[_;_]) as l -> l - | t::q -> [t;(loc_of_ne_list q,IntroAction (IntroOrAndPattern (IntroAndPattern (pairify q))))] + | t::q -> [t; Loc.tag ?loc:(loc_of_ne_list q) (IntroAction (IntroOrAndPattern (IntroAndPattern (pairify q))))] in IntroAndPattern (pairify (si::tc)) ] ] ; equality_intropattern: @@ -316,8 +316,8 @@ GEXTEND Gram ; nonsimple_intropattern: [ [ l = simple_intropattern -> l - | "*" -> !@loc, IntroForthcoming true - | "**" -> !@loc, IntroForthcoming false ]] + | "*" -> Loc.tag ~loc:!@loc @@ IntroForthcoming true + | "**" -> Loc.tag ~loc:!@loc @@ IntroForthcoming false ]] ; simple_intropattern: [ [ pat = simple_intropattern_closed; @@ -325,15 +325,15 @@ GEXTEND Gram let loc0,pat = pat in let f c pat = let loc1 = Constrexpr_ops.constr_loc c in - let loc = Loc.merge loc0 loc1 in + let loc = Loc.merge_opt loc0 loc1 in IntroAction (IntroApplyOn ((loc1,c),(loc,pat))) in - !@loc, List.fold_right f l pat ] ] + Loc.tag ~loc:!@loc @@ List.fold_right f l pat ] ] ; simple_intropattern_closed: - [ [ pat = or_and_intropattern -> !@loc, IntroAction (IntroOrAndPattern pat) - | pat = equality_intropattern -> !@loc, IntroAction pat - | "_" -> !@loc, IntroAction IntroWildcard - | pat = naming_intropattern -> !@loc, IntroNaming pat ] ] + [ [ pat = or_and_intropattern -> Loc.tag ~loc:!@loc @@ IntroAction (IntroOrAndPattern pat) + | pat = equality_intropattern -> Loc.tag ~loc:!@loc @@ IntroAction pat + | "_" -> Loc.tag ~loc:!@loc @@ IntroAction IntroWildcard + | pat = naming_intropattern -> Loc.tag ~loc:!@loc @@ IntroNaming pat ] ] ; simple_binding: [ [ "("; id = ident; ":="; c = lconstr; ")" -> Loc.tag ~loc:!@loc (NamedHyp id, c) @@ -468,7 +468,7 @@ GEXTEND Gram | -> None ] ] ; or_and_intropattern_loc: - [ [ ipat = or_and_intropattern -> ArgArg (!@loc,ipat) + [ [ ipat = or_and_intropattern -> ArgArg (Loc.tag ~loc:!@loc ipat) | locid = identref -> ArgVar locid ] ] ; as_or_and_ipat: @@ -476,13 +476,13 @@ GEXTEND Gram | -> None ] ] ; eqn_ipat: - [ [ IDENT "eqn"; ":"; pat = naming_intropattern -> Some (!@loc, pat) + [ [ IDENT "eqn"; ":"; pat = naming_intropattern -> Some (Loc.tag ~loc:!@loc pat) | IDENT "_eqn"; ":"; pat = naming_intropattern -> let loc = !@loc in - warn_deprecated_eqn_syntax ~loc "H"; Some (loc, pat) + warn_deprecated_eqn_syntax ~loc "H"; Some (Loc.tag ~loc pat) | IDENT "_eqn" -> let loc = !@loc in - warn_deprecated_eqn_syntax ~loc "?"; Some (loc, IntroAnonymous) + warn_deprecated_eqn_syntax ~loc "?"; Some (Loc.tag ~loc IntroAnonymous) | -> None ] ] ; as_name: @@ -521,145 +521,145 @@ GEXTEND Gram [ [ (* Basic tactics *) IDENT "intros"; pl = ne_intropatterns -> - TacAtom (!@loc, TacIntroPattern (false,pl)) + TacAtom (Loc.tag ~loc:!@loc @@ TacIntroPattern (false,pl)) | IDENT "intros" -> - TacAtom (!@loc, TacIntroPattern (false,[!@loc,IntroForthcoming false])) + TacAtom (Loc.tag ~loc:!@loc @@ TacIntroPattern (false,[Loc.tag ~loc:!@loc @@IntroForthcoming false])) | IDENT "eintros"; pl = ne_intropatterns -> - TacAtom (!@loc, TacIntroPattern (true,pl)) + TacAtom (Loc.tag ~loc:!@loc @@ TacIntroPattern (true,pl)) | IDENT "apply"; cl = LIST1 constr_with_bindings_arg SEP ","; - inhyp = in_hyp_as -> TacAtom (!@loc, TacApply (true,false,cl,inhyp)) + inhyp = in_hyp_as -> TacAtom (Loc.tag ~loc:!@loc @@ TacApply (true,false,cl,inhyp)) | IDENT "eapply"; cl = LIST1 constr_with_bindings_arg SEP ","; - inhyp = in_hyp_as -> TacAtom (!@loc, TacApply (true,true,cl,inhyp)) + inhyp = in_hyp_as -> TacAtom (Loc.tag ~loc:!@loc @@ TacApply (true,true,cl,inhyp)) | IDENT "simple"; IDENT "apply"; cl = LIST1 constr_with_bindings_arg SEP ","; - inhyp = in_hyp_as -> TacAtom (!@loc, TacApply (false,false,cl,inhyp)) + inhyp = in_hyp_as -> TacAtom (Loc.tag ~loc:!@loc @@ TacApply (false,false,cl,inhyp)) | IDENT "simple"; IDENT "eapply"; cl = LIST1 constr_with_bindings_arg SEP","; - inhyp = in_hyp_as -> TacAtom (!@loc, TacApply (false,true,cl,inhyp)) + inhyp = in_hyp_as -> TacAtom (Loc.tag ~loc:!@loc @@ TacApply (false,true,cl,inhyp)) | IDENT "elim"; cl = constr_with_bindings_arg; el = OPT eliminator -> - TacAtom (!@loc, TacElim (false,cl,el)) + TacAtom (Loc.tag ~loc:!@loc @@ TacElim (false,cl,el)) | IDENT "eelim"; cl = constr_with_bindings_arg; el = OPT eliminator -> - TacAtom (!@loc, TacElim (true,cl,el)) - | IDENT "case"; icl = induction_clause_list -> TacAtom (!@loc, mkTacCase false icl) - | IDENT "ecase"; icl = induction_clause_list -> TacAtom (!@loc, mkTacCase true icl) + TacAtom (Loc.tag ~loc:!@loc @@ TacElim (true,cl,el)) + | IDENT "case"; icl = induction_clause_list -> TacAtom (Loc.tag ~loc:!@loc @@ mkTacCase false icl) + | IDENT "ecase"; icl = induction_clause_list -> TacAtom (Loc.tag ~loc:!@loc @@ mkTacCase true icl) | "fix"; id = ident; n = natural; "with"; fd = LIST1 fixdecl -> - TacAtom (!@loc, TacMutualFix (id,n,List.map mk_fix_tac fd)) + TacAtom (Loc.tag ~loc:!@loc @@ TacMutualFix (id,n,List.map mk_fix_tac fd)) | "cofix"; id = ident; "with"; fd = LIST1 cofixdecl -> - TacAtom (!@loc, TacMutualCofix (id,List.map mk_cofix_tac fd)) + TacAtom (Loc.tag ~loc:!@loc @@ TacMutualCofix (id,List.map mk_cofix_tac fd)) | IDENT "pose"; (id,b) = bindings_with_parameters -> - TacAtom (!@loc, TacLetTac (Names.Name id,b,Locusops.nowhere,true,None)) + TacAtom (Loc.tag ~loc:!@loc @@ TacLetTac (Names.Name id,b,Locusops.nowhere,true,None)) | IDENT "pose"; b = constr; na = as_name -> - TacAtom (!@loc, TacLetTac (na,b,Locusops.nowhere,true,None)) + TacAtom (Loc.tag ~loc:!@loc @@ TacLetTac (na,b,Locusops.nowhere,true,None)) | IDENT "set"; (id,c) = bindings_with_parameters; p = clause_dft_concl -> - TacAtom (!@loc, TacLetTac (Names.Name id,c,p,true,None)) + TacAtom (Loc.tag ~loc:!@loc @@ TacLetTac (Names.Name id,c,p,true,None)) | IDENT "set"; c = constr; na = as_name; p = clause_dft_concl -> - TacAtom (!@loc, TacLetTac (na,c,p,true,None)) + TacAtom (Loc.tag ~loc:!@loc @@ TacLetTac (na,c,p,true,None)) | IDENT "remember"; c = constr; na = as_name; e = eqn_ipat; p = clause_dft_all -> - TacAtom (!@loc, TacLetTac (na,c,p,false,e)) + TacAtom (Loc.tag ~loc:!@loc @@ TacLetTac (na,c,p,false,e)) (* Alternative syntax for "pose proof c as id" *) | IDENT "assert"; test_lpar_id_coloneq; "("; (loc,id) = identref; ":="; c = lconstr; ")" -> - TacAtom (!@loc, TacAssert (true,None,Some (!@loc,IntroNaming (IntroIdentifier id)),c)) + TacAtom (Loc.tag ~loc:!@loc @@ TacAssert (true,None,Some (Loc.tag ~loc:!@loc @@IntroNaming (IntroIdentifier id)),c)) (* Alternative syntax for "assert c as id by tac" *) | IDENT "assert"; test_lpar_id_colon; "("; (loc,id) = identref; ":"; c = lconstr; ")"; tac=by_tactic -> - TacAtom (!@loc, TacAssert (true,Some tac,Some (!@loc,IntroNaming (IntroIdentifier id)),c)) + TacAtom (Loc.tag ~loc:!@loc @@ TacAssert (true,Some tac,Some (Loc.tag ~loc:!@loc @@IntroNaming (IntroIdentifier id)),c)) (* Alternative syntax for "enough c as id by tac" *) | IDENT "enough"; test_lpar_id_colon; "("; (loc,id) = identref; ":"; c = lconstr; ")"; tac=by_tactic -> - TacAtom (!@loc, TacAssert (false,Some tac,Some (!@loc,IntroNaming (IntroIdentifier id)),c)) + TacAtom (Loc.tag ~loc:!@loc @@ TacAssert (false,Some tac,Some (Loc.tag ~loc:!@loc @@IntroNaming (IntroIdentifier id)),c)) | IDENT "assert"; c = constr; ipat = as_ipat; tac = by_tactic -> - TacAtom (!@loc, TacAssert (true,Some tac,ipat,c)) + TacAtom (Loc.tag ~loc:!@loc @@ TacAssert (true,Some tac,ipat,c)) | IDENT "pose"; IDENT "proof"; c = lconstr; ipat = as_ipat -> - TacAtom (!@loc, TacAssert (true,None,ipat,c)) + TacAtom (Loc.tag ~loc:!@loc @@ TacAssert (true,None,ipat,c)) | IDENT "enough"; c = constr; ipat = as_ipat; tac = by_tactic -> - TacAtom (!@loc, TacAssert (false,Some tac,ipat,c)) + TacAtom (Loc.tag ~loc:!@loc @@ TacAssert (false,Some tac,ipat,c)) | IDENT "generalize"; c = constr -> - TacAtom (!@loc, TacGeneralize [((AllOccurrences,c),Names.Anonymous)]) + TacAtom (Loc.tag ~loc:!@loc @@ TacGeneralize [((AllOccurrences,c),Names.Anonymous)]) | IDENT "generalize"; c = constr; l = LIST1 constr -> let gen_everywhere c = ((AllOccurrences,c),Names.Anonymous) in - TacAtom (!@loc, TacGeneralize (List.map gen_everywhere (c::l))) + TacAtom (Loc.tag ~loc:!@loc @@ TacGeneralize (List.map gen_everywhere (c::l))) | IDENT "generalize"; c = constr; lookup_at_as_comma; nl = occs; na = as_name; l = LIST0 [","; c = pattern_occ; na = as_name -> (c,na)] -> - TacAtom (!@loc, TacGeneralize (((nl,c),na)::l)) + TacAtom (Loc.tag ~loc:!@loc @@ TacGeneralize (((nl,c),na)::l)) (* Derived basic tactics *) | IDENT "induction"; ic = induction_clause_list -> - TacAtom (!@loc, TacInductionDestruct (true,false,ic)) + TacAtom (Loc.tag ~loc:!@loc @@ TacInductionDestruct (true,false,ic)) | IDENT "einduction"; ic = induction_clause_list -> - TacAtom (!@loc, TacInductionDestruct(true,true,ic)) + TacAtom (Loc.tag ~loc:!@loc @@ TacInductionDestruct(true,true,ic)) | IDENT "destruct"; icl = induction_clause_list -> - TacAtom (!@loc, TacInductionDestruct(false,false,icl)) + TacAtom (Loc.tag ~loc:!@loc @@ TacInductionDestruct(false,false,icl)) | IDENT "edestruct"; icl = induction_clause_list -> - TacAtom (!@loc, TacInductionDestruct(false,true,icl)) + TacAtom (Loc.tag ~loc:!@loc @@ TacInductionDestruct(false,true,icl)) (* Equality and inversion *) | IDENT "rewrite"; l = LIST1 oriented_rewriter SEP ","; - cl = clause_dft_concl; t=by_tactic -> TacAtom (!@loc, TacRewrite (false,l,cl,t)) + cl = clause_dft_concl; t=by_tactic -> TacAtom (Loc.tag ~loc:!@loc @@ TacRewrite (false,l,cl,t)) | IDENT "erewrite"; l = LIST1 oriented_rewriter SEP ","; - cl = clause_dft_concl; t=by_tactic -> TacAtom (!@loc, TacRewrite (true,l,cl,t)) + cl = clause_dft_concl; t=by_tactic -> TacAtom (Loc.tag ~loc:!@loc @@ TacRewrite (true,l,cl,t)) | IDENT "dependent"; k = [ IDENT "simple"; IDENT "inversion" -> SimpleInversion | IDENT "inversion" -> FullInversion | IDENT "inversion_clear" -> FullInversionClear ]; hyp = quantified_hypothesis; ids = as_or_and_ipat; co = OPT ["with"; c = constr -> c] -> - TacAtom (!@loc, TacInversion (DepInversion (k,co,ids),hyp)) + TacAtom (Loc.tag ~loc:!@loc @@ TacInversion (DepInversion (k,co,ids),hyp)) | IDENT "simple"; IDENT "inversion"; hyp = quantified_hypothesis; ids = as_or_and_ipat; cl = in_hyp_list -> - TacAtom (!@loc, TacInversion (NonDepInversion (SimpleInversion, cl, ids), hyp)) + TacAtom (Loc.tag ~loc:!@loc @@ TacInversion (NonDepInversion (SimpleInversion, cl, ids), hyp)) | IDENT "inversion"; hyp = quantified_hypothesis; ids = as_or_and_ipat; cl = in_hyp_list -> - TacAtom (!@loc, TacInversion (NonDepInversion (FullInversion, cl, ids), hyp)) + TacAtom (Loc.tag ~loc:!@loc @@ TacInversion (NonDepInversion (FullInversion, cl, ids), hyp)) | IDENT "inversion_clear"; hyp = quantified_hypothesis; ids = as_or_and_ipat; cl = in_hyp_list -> - TacAtom (!@loc, TacInversion (NonDepInversion (FullInversionClear, cl, ids), hyp)) + TacAtom (Loc.tag ~loc:!@loc @@ TacInversion (NonDepInversion (FullInversionClear, cl, ids), hyp)) | IDENT "inversion"; hyp = quantified_hypothesis; "using"; c = constr; cl = in_hyp_list -> - TacAtom (!@loc, TacInversion (InversionUsing (c,cl), hyp)) + TacAtom (Loc.tag ~loc:!@loc @@ TacInversion (InversionUsing (c,cl), hyp)) (* Conversion *) | IDENT "red"; cl = clause_dft_concl -> - TacAtom (!@loc, TacReduce (Red false, cl)) + TacAtom (Loc.tag ~loc:!@loc @@ TacReduce (Red false, cl)) | IDENT "hnf"; cl = clause_dft_concl -> - TacAtom (!@loc, TacReduce (Hnf, cl)) + TacAtom (Loc.tag ~loc:!@loc @@ TacReduce (Hnf, cl)) | IDENT "simpl"; d = delta_flag; po = OPT ref_or_pattern_occ; cl = clause_dft_concl -> - TacAtom (!@loc, TacReduce (Simpl (all_with d, po), cl)) + TacAtom (Loc.tag ~loc:!@loc @@ TacReduce (Simpl (all_with d, po), cl)) | IDENT "cbv"; s = strategy_flag; cl = clause_dft_concl -> - TacAtom (!@loc, TacReduce (Cbv s, cl)) + TacAtom (Loc.tag ~loc:!@loc @@ TacReduce (Cbv s, cl)) | IDENT "cbn"; s = strategy_flag; cl = clause_dft_concl -> - TacAtom (!@loc, TacReduce (Cbn s, cl)) + TacAtom (Loc.tag ~loc:!@loc @@ TacReduce (Cbn s, cl)) | IDENT "lazy"; s = strategy_flag; cl = clause_dft_concl -> - TacAtom (!@loc, TacReduce (Lazy s, cl)) + TacAtom (Loc.tag ~loc:!@loc @@ TacReduce (Lazy s, cl)) | IDENT "compute"; delta = delta_flag; cl = clause_dft_concl -> - TacAtom (!@loc, TacReduce (Cbv (all_with delta), cl)) + TacAtom (Loc.tag ~loc:!@loc @@ TacReduce (Cbv (all_with delta), cl)) | IDENT "vm_compute"; po = OPT ref_or_pattern_occ; cl = clause_dft_concl -> - TacAtom (!@loc, TacReduce (CbvVm po, cl)) + TacAtom (Loc.tag ~loc:!@loc @@ TacReduce (CbvVm po, cl)) | IDENT "native_compute"; po = OPT ref_or_pattern_occ; cl = clause_dft_concl -> - TacAtom (!@loc, TacReduce (CbvNative po, cl)) + TacAtom (Loc.tag ~loc:!@loc @@ TacReduce (CbvNative po, cl)) | IDENT "unfold"; ul = LIST1 unfold_occ SEP ","; cl = clause_dft_concl -> - TacAtom (!@loc, TacReduce (Unfold ul, cl)) + TacAtom (Loc.tag ~loc:!@loc @@ TacReduce (Unfold ul, cl)) | IDENT "fold"; l = LIST1 constr; cl = clause_dft_concl -> - TacAtom (!@loc, TacReduce (Fold l, cl)) + TacAtom (Loc.tag ~loc:!@loc @@ TacReduce (Fold l, cl)) | IDENT "pattern"; pl = LIST1 pattern_occ SEP","; cl = clause_dft_concl -> - TacAtom (!@loc, TacReduce (Pattern pl, cl)) + TacAtom (Loc.tag ~loc:!@loc @@ TacReduce (Pattern pl, cl)) (* Change ne doit pas s'appliquer dans un Definition t := Eval ... *) | IDENT "change"; (oc,c) = conversion; cl = clause_dft_concl -> let p,cl = merge_occurrences (!@loc) cl oc in - TacAtom (!@loc, TacChange (p,c,cl)) + TacAtom (Loc.tag ~loc:!@loc @@ TacChange (p,c,cl)) ] ] ; END;; diff --git a/plugins/ltac/pptactic.ml b/plugins/ltac/pptactic.ml index 26ac3c53e..bdafbdc78 100644 --- a/plugins/ltac/pptactic.ml +++ b/plugins/ltac/pptactic.ml @@ -352,7 +352,7 @@ type 'a extra_genarg_printer = let pr_ltac_or_var pr = function | ArgArg x -> pr x - | ArgVar (loc,id) -> pr_with_comments ~loc (pr_id id) + | ArgVar (loc,id) -> pr_with_comments ?loc (pr_id id) let pr_ltac_constant kn = if !Flags.in_debugger then pr_kn kn @@ -507,7 +507,7 @@ type 'a extra_genarg_printer = let pr_core_destruction_arg prc prlc = function | ElimOnConstr c -> pr_with_bindings prc prlc c - | ElimOnIdent (loc,id) -> pr_with_comments ~loc (pr_id id) + | ElimOnIdent (loc,id) -> pr_with_comments ?loc (pr_id id) | ElimOnAnonHyp n -> int n let pr_destruction_arg prc prlc (clear_flag,h) = @@ -1037,7 +1037,7 @@ type 'a extra_genarg_printer = | TacId l -> keyword "idtac" ++ prlist (pr_arg (pr_message_token pr.pr_name)) l, latom | TacAtom (loc,t) -> - pr_with_comments ~loc (hov 1 (pr_atom pr strip_prod_binders tag_atom t)), ltatom + pr_with_comments ?loc (hov 1 (pr_atom pr strip_prod_binders tag_atom t)), ltatom | TacArg(_,Tacexp e) -> pr.pr_tactic (latom,E) e, latom | TacArg(_,ConstrMayEval (ConstrTerm c)) -> @@ -1051,16 +1051,16 @@ type 'a extra_genarg_printer = | TacArg(_,TacCall(loc,(f,[]))) -> pr.pr_reference f, latom | TacArg(_,TacCall(loc,(f,l))) -> - pr_with_comments ~loc (hov 1 ( + pr_with_comments ?loc (hov 1 ( pr.pr_reference f ++ spc () ++ prlist_with_sep spc pr_tacarg l)), lcall | TacArg (_,a) -> pr_tacarg a, latom | TacML (loc,(s,l)) -> - pr_with_comments ~loc (pr.pr_extend 1 s l), lcall + pr_with_comments ?loc (pr.pr_extend 1 s l), lcall | TacAlias (loc,(kn,l)) -> - pr_with_comments ~loc (pr.pr_alias (level_of inherited) kn l), latom + pr_with_comments ?loc (pr.pr_alias (level_of inherited) kn l), latom ) in if prec_less prec inherited then strm diff --git a/plugins/ltac/taccoerce.ml b/plugins/ltac/taccoerce.ml index b76009c99..e037bb4b2 100644 --- a/plugins/ltac/taccoerce.ml +++ b/plugins/ltac/taccoerce.ml @@ -268,11 +268,11 @@ let coerce_to_constr_list env v = List.map map l | None -> raise (CannotCoerceTo "a term list") -let coerce_to_intro_pattern_list loc env sigma v = +let coerce_to_intro_pattern_list ?loc env sigma v = match Value.to_list v with | None -> raise (CannotCoerceTo "an intro pattern list") | Some l -> - let map v = (loc, coerce_to_intro_pattern env sigma v) in + let map v = Loc.tag ?loc @@ coerce_to_intro_pattern env sigma v in List.map map l let coerce_to_hyp env sigma v = diff --git a/plugins/ltac/taccoerce.mli b/plugins/ltac/taccoerce.mli index 9c4ac5265..b09672a12 100644 --- a/plugins/ltac/taccoerce.mli +++ b/plugins/ltac/taccoerce.mli @@ -76,7 +76,7 @@ val coerce_to_evaluable_ref : val coerce_to_constr_list : Environ.env -> Value.t -> constr list val coerce_to_intro_pattern_list : - Loc.t -> Environ.env -> Evd.evar_map -> Value.t -> Tacexpr.intro_patterns + ?loc:Loc.t -> Environ.env -> Evd.evar_map -> Value.t -> Tacexpr.intro_patterns val coerce_to_hyp : Environ.env -> Evd.evar_map -> Value.t -> Id.t diff --git a/plugins/ltac/tacentries.ml b/plugins/ltac/tacentries.ml index f608515aa..a19dbd327 100644 --- a/plugins/ltac/tacentries.ml +++ b/plugins/ltac/tacentries.ml @@ -170,7 +170,7 @@ let add_tactic_entry (kn, ml, tg) state = TacGeneric arg in let l = List.map map l in - (TacAlias (loc,(kn,l)):raw_tactic_expr) + (TacAlias (Loc.tag ~loc (kn,l)):raw_tactic_expr) in let () = if Int.equal tg.tacgram_level 0 && not (head_is_ident tg) then @@ -180,6 +180,7 @@ let add_tactic_entry (kn, ml, tg) state = | TacTerm s -> GramTerminal s | TacNonTerm (loc, (s, _)) -> let EntryName (typ, e) = prod_item_of_symbol tg.tacgram_level s in + let loc = Option.default Loc.internal_ghost loc in GramNonTerminal (loc, typ, e) in let prods = List.map map tg.tacgram_prods in @@ -405,7 +406,7 @@ let create_ltac_quotation name cast (e, l) = entry), Atoken (CLexer.terminal ")")) in - let action _ v _ _ _ loc = cast (loc, v) in + let action _ v _ _ _ loc = cast (Some loc, v) in let gram = (level, assoc, [Rule (rule, action)]) in Pcoq.grammar_extend Pltac.tactic_arg None (None, [gram]) @@ -431,7 +432,7 @@ let register_ltac local tacl = let kn = Lib.make_kn id in let id_pp = pr_id id in let () = if is_defined_tac kn then - CErrors.user_err ~loc + CErrors.user_err ?loc (str "There is already an Ltac named " ++ id_pp ++ str".") in let is_shadowed = @@ -448,7 +449,7 @@ let register_ltac local tacl = let kn = try Nametab.locate_tactic (snd (qualid_of_reference ident)) with Not_found -> - CErrors.user_err ~loc + CErrors.user_err ?loc (str "There is no Ltac named " ++ pr_reference ident ++ str ".") in UpdateTac kn, body diff --git a/plugins/ltac/tacintern.ml b/plugins/ltac/tacintern.ml index 787a5f50b..44ea3ff1d 100644 --- a/plugins/ltac/tacintern.ml +++ b/plugins/ltac/tacintern.ml @@ -72,7 +72,7 @@ let intern_name l ist = function let strict_check = ref false -let adjust_loc loc = if !strict_check then None else Some loc +let adjust_loc loc = if !strict_check then None else loc (* Globalize a name which must be bound -- actually just check it is bound *) let intern_hyp ist (loc,id as locid) = @@ -81,7 +81,7 @@ let intern_hyp ist (loc,id as locid) = else if find_ident id ist then Loc.tag id else - Pretype_errors.error_var_not_found ~loc id + Pretype_errors.error_var_not_found ?loc id let intern_or_var f ist = function | ArgVar locid -> ArgVar (intern_hyp ist locid) @@ -126,7 +126,7 @@ let intern_move_location ist = function let intern_isolated_global_tactic_reference r = let (loc,qid) = qualid_of_reference r in - TacCall (Loc.tag ~loc (ArgArg (loc,locate_tactic qid),[])) + TacCall (Loc.tag ?loc (ArgArg (loc,locate_tactic qid),[])) let intern_isolated_tactic_reference strict ist r = (* An ltac reference *) @@ -293,7 +293,7 @@ let intern_evaluable_reference_or_by_notation ist = function | AN r -> intern_evaluable_global_reference ist r | ByNotation (loc,(ntn,sc)) -> evaluable_of_global_reference ist.genv - (Notation.interp_notation_as_global_reference ~loc + (Notation.interp_notation_as_global_reference ?loc (function ConstRef _ | VarRef _ -> true | _ -> false) ntn sc) (* Globalize a reduction expression *) @@ -372,13 +372,13 @@ let intern_typed_pattern_or_ref_with_occurrences ist (l,p) = let dump_glob_red_expr = function | Unfold occs -> List.iter (fun (_, r) -> try - Dumpglob.add_glob (loc_of_or_by_notation Libnames.loc_of_reference r) + Dumpglob.add_glob ?loc:(loc_of_or_by_notation Libnames.loc_of_reference r) (Smartlocate.smart_global r) with e when CErrors.noncritical e -> ()) occs | Cbv grf | Lazy grf -> List.iter (fun r -> try - Dumpglob.add_glob (loc_of_or_by_notation Libnames.loc_of_reference r) + Dumpglob.add_glob ?loc:(loc_of_or_by_notation Libnames.loc_of_reference r) (Smartlocate.smart_global r) with e when CErrors.noncritical e -> ()) grf.rConst | _ -> () @@ -459,7 +459,7 @@ let rec intern_match_goal_hyps ist ?(as_type=false) lfun = function (* Utilities *) let extract_let_names lrc = let fold accu ((loc, name), _) = - if Id.Set.mem name accu then user_err ~loc + if Id.Set.mem name accu then user_err ?loc ~hdr:"glob_tactic" (str "This variable is bound several times.") else Id.Set.add name accu in @@ -626,7 +626,7 @@ and intern_tactic_seq onlytac ist = function (* For extensions *) | TacAlias (loc,(s,l)) -> let l = List.map (intern_tacarg !strict_check false ist) l in - ist.ltacvars, TacAlias (Loc.tag ~loc (s,l)) + ist.ltacvars, TacAlias (Loc.tag ?loc (s,l)) | TacML (loc,(opn,l)) -> let _ignore = Tacenv.interp_ml_tactic opn in ist.ltacvars, TacML (loc, (opn,List.map (intern_tacarg !strict_check false ist) l)) @@ -637,7 +637,7 @@ and intern_tactic_as_arg loc onlytac ist a = | TacGeneric _ as a -> TacArg (loc,a) | Tacexp a -> a | ConstrMayEval _ | TacFreshId _ | TacPretype _ | TacNumgoals as a -> - if onlytac then error_tactic_expected ~loc else TacArg (loc,a) + if onlytac then error_tactic_expected ?loc else TacArg (loc,a) and intern_tactic_or_tacarg ist = intern_tactic false ist @@ -652,7 +652,7 @@ and intern_tacarg strict onlytac ist = function | ConstrMayEval c -> ConstrMayEval (intern_constr_may_eval ist c) | TacCall (loc,(f,[])) -> intern_isolated_tactic_reference strict ist f | TacCall (loc,(f,l)) -> - TacCall (Loc.tag ~loc ( + TacCall (Loc.tag ?loc ( intern_applied_tactic_reference ist f, List.map (intern_tacarg !strict_check false ist) l)) | TacFreshId x -> TacFreshId (List.map (intern_string_or_var ist) x) diff --git a/plugins/ltac/tacinterp.ml b/plugins/ltac/tacinterp.ml index e969b86f6..4d5b84455 100644 --- a/plugins/ltac/tacinterp.ml +++ b/plugins/ltac/tacinterp.ml @@ -313,7 +313,7 @@ let append_trace trace v = (* Dynamically check that an argument is a tactic *) let coerce_to_tactic loc id v = let v = Value.normalize v in - let fail () = user_err ~loc + let fail () = user_err ?loc (str "Variable " ++ pr_id id ++ str " should be bound to a tactic.") in let v = Value.normalize v in @@ -376,7 +376,7 @@ let error_ltac_variable ?loc id env v s = (* Raise Not_found if not in interpretation sign *) let try_interp_ltac_var coerce ist env (loc,id) = let v = Id.Map.find id ist.lfun in - try coerce v with CannotCoerceTo s -> error_ltac_variable ~loc id env v s + try coerce v with CannotCoerceTo s -> error_ltac_variable ?loc id env v s let interp_ltac_var coerce ist env locid = try try_interp_ltac_var coerce ist env locid @@ -402,7 +402,7 @@ let interp_intro_pattern_naming_var loc ist env sigma id = let interp_int ist locid = try try_interp_ltac_var coerce_to_int ist None locid with Not_found -> - user_err ~loc:(fst locid) ~hdr:"interp_int" + user_err ?loc:(fst locid) ~hdr:"interp_int" (str "Unbound variable " ++ pr_id (snd locid) ++ str".") let interp_int_or_var ist = function @@ -425,7 +425,7 @@ let interp_hyp ist env sigma (loc,id as locid) = with Not_found -> (* Then look if bound in the proof context at calling time *) if is_variable env id then id - else Loc.raise ~loc (Logic.RefinerError (Logic.NoSuchHyp id)) + else Loc.raise ?loc (Logic.RefinerError (Logic.NoSuchHyp id)) let interp_hyp_list_as_list ist env sigma (loc,id as x) = try coerce_to_hyp_list env sigma (Id.Map.find id ist.lfun) @@ -447,7 +447,7 @@ let interp_reference ist env sigma = function with Not_found -> try VarRef (get_id (Environ.lookup_named id env)) - with Not_found -> error_global_not_found ~loc (qualid_of_ident id) + with Not_found -> error_global_not_found ?loc (qualid_of_ident id) let try_interp_evaluable env (loc, id) = let v = Environ.lookup_named id env in @@ -463,14 +463,14 @@ let interp_evaluable ist env sigma = function with Not_found -> match r with | EvalConstRef _ -> r - | _ -> error_global_not_found ~loc (qualid_of_ident id) + | _ -> error_global_not_found ?loc (qualid_of_ident id) end | ArgArg (r,None) -> r | ArgVar (loc, id) -> try try_interp_ltac_var (coerce_to_evaluable_ref env sigma) ist (Some (env,sigma)) (loc, id) with Not_found -> try try_interp_evaluable env (loc, id) - with Not_found -> error_global_not_found ~loc (qualid_of_ident id) + with Not_found -> error_global_not_found ?loc (qualid_of_ident id) (* Interprets an hypothesis name *) let interp_occurrences ist occs = @@ -739,7 +739,7 @@ let interp_closed_typed_pattern_with_occurrences ist env sigma (occs, a) = Inr (pattern_of_constr env sigma (EConstr.to_constr sigma c)) in (try try_interp_ltac_var coerce_eval_ref_or_constr ist (Some (env,sigma)) (loc,id) with Not_found -> - error_global_not_found ~loc (qualid_of_ident id)) + error_global_not_found ?loc (qualid_of_ident id)) | Inl (ArgArg _ as b) -> Inl (interp_evaluable ist env sigma b) | Inr c -> Inr (interp_typed_pattern ist env sigma c) in interp_occurrences ist occs, p @@ -797,7 +797,7 @@ let interp_may_eval f ist env sigma = function !evdref , c with | Not_found -> - user_err ~loc ~hdr:"interp_may_eval" + user_err ?loc ~hdr:"interp_may_eval" (str "Unbound context identifier" ++ pr_id s ++ str".")) | ConstrTypeOf c -> let (sigma,c_interp) = f ist env sigma c in @@ -939,7 +939,7 @@ and interp_or_and_intro_pattern ist env sigma = function and interp_intro_pattern_list_as_list ist env sigma = function | [loc,IntroNaming (IntroIdentifier id)] as l -> - (try sigma, coerce_to_intro_pattern_list loc env sigma (Id.Map.find id ist.lfun) + (try sigma, coerce_to_intro_pattern_list ?loc env sigma (Id.Map.find id ist.lfun) with Not_found | CannotCoerceTo _ -> List.fold_map (interp_intro_pattern ist env) sigma l) | l -> List.fold_map (interp_intro_pattern ist env) sigma l @@ -954,7 +954,7 @@ let interp_or_and_intro_pattern_option ist env sigma = function (match coerce_to_intro_pattern env sigma (Id.Map.find id ist.lfun) with | IntroAction (IntroOrAndPattern l) -> sigma, Some (loc,l) | _ -> - user_err ~loc (str "Cannot coerce to a disjunctive/conjunctive pattern.")) + user_err ?loc (str "Cannot coerce to a disjunctive/conjunctive pattern.")) | Some (ArgArg (loc,l)) -> let sigma,l = interp_or_and_intro_pattern ist env sigma l in sigma, Some (loc,l) @@ -1011,13 +1011,13 @@ let interp_open_constr_with_bindings ist env sigma (c,bl) = let loc_of_bindings = function | NoBindings -> None -| ImplicitBindings l -> Some (loc_of_glob_constr (fst (List.last l))) -| ExplicitBindings l -> Some (fst (List.last l)) +| ImplicitBindings l -> loc_of_glob_constr (fst (List.last l)) +| ExplicitBindings l -> fst (List.last l) let interp_open_constr_with_bindings_loc ist ((c,_),bl as cb) = let loc1 = loc_of_glob_constr c in let loc2 = loc_of_bindings bl in - let loc = Loc.opt_merge loc1 loc2 in + let loc = Loc.merge_opt loc1 loc2 in let f = { delayed = fun env sigma -> let sigma = Sigma.to_evar_map sigma in let (sigma, c) = interp_open_constr_with_bindings ist env sigma cb in @@ -1035,7 +1035,7 @@ let interp_destruction_arg ist gl arg = } | keep,ElimOnAnonHyp n as x -> x | keep,ElimOnIdent (loc,id) -> - let error () = user_err ~loc + let error () = user_err ?loc (strbrk "Cannot coerce " ++ pr_id id ++ strbrk " neither to a quantified hypothesis nor to a term.") in @@ -1046,7 +1046,7 @@ let interp_destruction_arg ist gl arg = (keep, ElimOnConstr { delayed = begin fun env sigma -> try Sigma.here (constr_of_id env id', NoBindings) sigma with Not_found -> - user_err ~loc ~hdr:"interp_destruction_arg" ( + user_err ?loc ~hdr:"interp_destruction_arg" ( pr_id id ++ strbrk " binds to " ++ pr_id id' ++ strbrk " which is neither a declared nor a quantified hypothesis.") end }) in @@ -1072,7 +1072,7 @@ let interp_destruction_arg ist gl arg = if Tactics.is_quantified_hypothesis id gl then keep,ElimOnIdent (loc,id) else - let c = (Loc.tag ~loc @@ GVar id,Some (Loc.tag @@ CRef (Ident (loc,id),None))) in + let c = (Loc.tag ?loc @@ GVar id,Some (Loc.tag @@ CRef (Ident (loc,id),None))) in let f = { delayed = fun env sigma -> let sigma = Sigma.to_evar_map sigma in let (sigma,c) = interp_open_constr ist env sigma c in @@ -1287,7 +1287,7 @@ and eval_tactic ist tac : unit Proofview.tactic = match tac with Ftactic.run tac (fun () -> Proofview.tclUNIT ()) | TacML (loc,(opn,l)) -> - push_trace (Loc.tag ~loc @@ LtacMLCall tac) ist >>= fun trace -> + push_trace (Loc.tag ?loc @@ LtacMLCall tac) ist >>= fun trace -> let ist = { ist with extra = TacStore.set ist.extra f_trace trace; } in let tac = Tacenv.interp_ml_tactic opn in let args = Ftactic.List.map_right (fun a -> interp_tacarg ist a) l in diff --git a/plugins/ltac/tacsubst.ml b/plugins/ltac/tacsubst.ml index 0ee6e8a85..4390ff08b 100644 --- a/plugins/ltac/tacsubst.ml +++ b/plugins/ltac/tacsubst.ml @@ -242,7 +242,7 @@ and subst_tacarg subst = function | Reference r -> Reference (subst_reference subst r) | ConstrMayEval c -> ConstrMayEval (subst_raw_may_eval subst c) | TacCall (loc,(f,l)) -> - TacCall (Loc.tag ~loc (subst_reference subst f, List.map (subst_tacarg subst) l)) + TacCall (Loc.tag ?loc (subst_reference subst f, List.map (subst_tacarg subst) l)) | TacFreshId _ as x -> x | TacPretype c -> TacPretype (subst_glob_constr subst c) | TacNumgoals -> TacNumgoals diff --git a/plugins/ltac/tactic_debug.ml b/plugins/ltac/tactic_debug.ml index ac8534bdc..f04495c61 100644 --- a/plugins/ltac/tactic_debug.ml +++ b/plugins/ltac/tactic_debug.ml @@ -387,12 +387,11 @@ let skip_extensions trace = | [] -> [] in List.rev (aux (List.rev trace)) -let finer_loc loc1 loc2 = Loc.merge loc1 loc2 = loc2 +let finer_loc loc1 loc2 = Loc.merge_opt loc1 loc2 = loc2 let extract_ltac_trace ?loc trace = let trace = skip_extensions trace in let (tloc,c),tail = List.sep_last trace in - let loc = Option.default tloc loc in if is_defined_ltac trace then (* We entered a user-defined tactic, we display the trace with location of the call *) @@ -405,8 +404,8 @@ let extract_ltac_trace ?loc trace = (* trace is with innermost call coming first *) let rec aux best_loc = function | (loc,_)::tail -> - if Loc.is_ghost best_loc || - not (Loc.is_ghost loc) && finer_loc loc best_loc + if Option.is_empty best_loc || + not (Option.is_empty loc) && finer_loc loc best_loc then aux loc tail else diff --git a/plugins/ssrmatching/ssrmatching.ml4 b/plugins/ssrmatching/ssrmatching.ml4 index a6a6bd6f9..0bc3f502c 100644 --- a/plugins/ssrmatching/ssrmatching.ml4 +++ b/plugins/ssrmatching/ssrmatching.ml4 @@ -63,7 +63,7 @@ DECLARE PLUGIN "ssrmatching_plugin" type loc = Loc.t let errorstrm = CErrors.user_err ~hdr:"ssrmatching" -let loc_error loc msg = CErrors.user_err ~loc ~hdr:msg (str msg) +let loc_error loc msg = CErrors.user_err ?loc ~hdr:msg (str msg) let ppnl = Feedback.msg_info (* 0 cost pp function. Active only if env variable SSRDEBUG is set *) @@ -1071,7 +1071,7 @@ GEXTEND Gram GLOBAL: cpattern; cpattern: [[ k = ssrtermkind; c = constr -> let pattern = mk_term k c in - if loc_ofCG pattern <> !@loc && k = '(' then mk_term 'x' c else pattern ]]; + if loc_ofCG pattern <> Some !@loc && k = '(' then mk_term 'x' c else pattern ]]; END ARGUMENT EXTEND lcpattern @@ -1088,7 +1088,7 @@ GEXTEND Gram GLOBAL: lcpattern; lcpattern: [[ k = ssrtermkind; c = lconstr -> let pattern = mk_term k c in - if loc_ofCG pattern <> !@loc && k = '(' then mk_term 'x' c else pattern ]]; + if loc_ofCG pattern <> Some !@loc && k = '(' then mk_term 'x' c else pattern ]]; END let thin id sigma goal = @@ -1187,16 +1187,16 @@ let interp_pattern ?wit_ssrpatternarg ist gl red redty = pp(lazy(str"decoded as: " ++ pr_pattern_w_ids red)); let red = match redty with None -> red | Some ty -> let ty = ' ', ty in match red with - | T t -> T (combineCG t ty (mkCCast ~loc:(loc_ofCG t)) mkRCast) + | T t -> T (combineCG t ty (mkCCast ?loc:(loc_ofCG t)) mkRCast) | X_In_T (x,t) -> let ty = pf_intern_term ist gl ty in E_As_X_In_T (mkG (mkRCast mkRHole ty), x, t) | E_In_X_In_T (e,x,t) -> let ty = mkG (pf_intern_term ist gl ty) in - E_In_X_In_T (combineCG e ty (mkCCast ~loc:(loc_ofCG t)) mkRCast, x, t) + E_In_X_In_T (combineCG e ty (mkCCast ?loc:(loc_ofCG t)) mkRCast, x, t) | E_As_X_In_T (e,x,t) -> let ty = mkG (pf_intern_term ist gl ty) in - E_As_X_In_T (combineCG e ty (mkCCast ~loc:(loc_ofCG t)) mkRCast, x, t) + E_As_X_In_T (combineCG e ty (mkCCast ?loc:(loc_ofCG t)) mkRCast, x, t) | red -> red in pp(lazy(str"typed as: " ++ pr_pattern_w_ids red)); let mkXLetIn ?loc x (a,(g,c)) = match c with diff --git a/plugins/ssrmatching/ssrmatching.mli b/plugins/ssrmatching/ssrmatching.mli index 894cdb943..a8862d264 100644 --- a/plugins/ssrmatching/ssrmatching.mli +++ b/plugins/ssrmatching/ssrmatching.mli @@ -222,7 +222,7 @@ val pf_unify_HO : goal sigma -> EConstr.constr -> EConstr.constr -> goal sigma (** Some more low level functions needed to implement the full SSR language on top of the former APIs *) val tag_of_cpattern : cpattern -> char -val loc_of_cpattern : cpattern -> Loc.t +val loc_of_cpattern : cpattern -> Loc.t option val id_of_pattern : pattern -> Names.variable option val is_wildcard : cpattern -> bool val cpattern_of_id : Names.variable -> cpattern diff --git a/pretyping/cases.ml b/pretyping/cases.ml index f5e2e52a1..eb0d01718 100644 --- a/pretyping/cases.ml +++ b/pretyping/cases.ml @@ -349,15 +349,15 @@ let find_tomatch_tycon evdref env loc = function let coerce_row typing_fun evdref env pats (tomatch,(_,indopt)) = let loc = loc_of_glob_constr tomatch in - let tycon,realnames = find_tomatch_tycon evdref env (Some loc) indopt in + let tycon,realnames = find_tomatch_tycon evdref env loc indopt in let j = typing_fun tycon env evdref tomatch in - let evd, j = Coercion.inh_coerce_to_base ~loc:(loc_of_glob_constr tomatch) env !evdref j in + let evd, j = Coercion.inh_coerce_to_base ?loc:(loc_of_glob_constr tomatch) env !evdref j in evdref := evd; let typ = nf_evar !evdref j.uj_type in let t = try try_find_ind env !evdref typ realnames with Not_found -> - unify_tomatch_with_patterns evdref env (Some loc) typ pats realnames in + unify_tomatch_with_patterns evdref env loc typ pats realnames in (j.uj_val,t) let coerce_to_indtype typing_fun evdref env matx tomatchl = @@ -427,7 +427,7 @@ let current_pattern eqn = | pat::_ -> pat | [] -> anomaly (Pp.str "Empty list of patterns") -let alias_of_pat = Loc.with_loc (fun ~loc -> function +let alias_of_pat = Loc.with_loc (fun ?loc -> function | PatVar name -> name | PatCstr(_,_,name) -> name ) @@ -489,23 +489,23 @@ let check_and_adjust_constructor env ind cstrs = function if Int.equal (List.length args) nb_args_constr then pat else try - let args' = adjust_local_defs ~loc (args, List.rev ci.cs_args) - in Loc.tag ~loc @@ PatCstr (cstr, args', alias) + let args' = adjust_local_defs ?loc (args, List.rev ci.cs_args) + in Loc.tag ?loc @@ PatCstr (cstr, args', alias) with NotAdjustable -> - error_wrong_numarg_constructor ~loc env cstr nb_args_constr + error_wrong_numarg_constructor ?loc env cstr nb_args_constr else (* Try to insert a coercion *) try - Coercion.inh_pattern_coerce_to ~loc env pat ind' ind + Coercion.inh_pattern_coerce_to ?loc env pat ind' ind with Not_found -> - error_bad_constructor ~loc env cstr ind + error_bad_constructor ?loc env cstr ind let check_all_variables env sigma typ mat = List.iter (fun eqn -> match current_pattern eqn with | _, PatVar id -> () | loc, PatCstr (cstr_sp,_,_) -> - error_bad_pattern ~loc env sigma cstr_sp typ) + error_bad_pattern ?loc env sigma cstr_sp typ) mat let check_unused_pattern env eqn = @@ -1545,7 +1545,7 @@ let matx_of_eqns env eqns = it = Some initial_rhs } in { patterns = initial_lpat; alias_stack = []; - eqn_loc = Some loc; + eqn_loc = loc; used = ref false; rhs = rhs } in List.map build_eqn eqns @@ -1853,7 +1853,7 @@ let extract_arity_signature ?(dolift=true) env0 tomatchl tmsign = | None -> [LocalAssum (na, lift n typ)] | Some b -> [LocalDef (na, lift n b, lift n typ)]) | Some (loc,_) -> - user_err ~loc + user_err ?loc (str"Unexpected type annotation for a term of non inductive type.")) | IsInd (term,IndType(indf,realargs),_) -> let indf' = if dolift then lift_inductive_family n indf else indf in @@ -1865,7 +1865,7 @@ let extract_arity_signature ?(dolift=true) env0 tomatchl tmsign = match t with | Some (loc,(ind',realnal)) -> if not (eq_ind ind ind') then - user_err ~loc (str "Wrong inductive type."); + user_err ?loc (str "Wrong inductive type."); if not (Int.equal nrealargs_ctxt (List.length realnal)) then anomaly (Pp.str "Ill-formed 'in' clause in cases"); List.rev realnal @@ -2073,7 +2073,7 @@ let constr_of_pat env evdref arsign pat avoid = let previd, id = prime avoid (Name (Id.of_string "wildcard")) in Name id, id :: avoid in - ((Loc.tag ~loc @@ PatVar name), [LocalAssum (name, ty)] @ realargs, mkRel 1, ty, + ((Loc.tag ?loc @@ PatVar name), [LocalAssum (name, ty)] @ realargs, mkRel 1, ty, (List.map (fun x -> mkRel 1) realargs), 1, avoid) | PatCstr (((_, i) as cstr),args,alias) -> let cind = inductive_of_constructor cstr in @@ -2084,7 +2084,7 @@ let constr_of_pat env evdref arsign pat avoid = in let (ind,u), params = dest_ind_family indf in let params = List.map EConstr.of_constr params in - if not (eq_ind ind cind) then error_bad_constructor ~loc env cstr ind; + if not (eq_ind ind cind) then error_bad_constructor ?loc env cstr ind; let cstrs = get_constructors env indf in let ci = cstrs.(i-1) in let nb_args_constr = ci.cs_nargs in @@ -2104,7 +2104,7 @@ let constr_of_pat env evdref arsign pat avoid = in let args = List.rev args in let patargs = List.rev patargs in - let pat' = Loc.tag ~loc @@ PatCstr (cstr, patargs, alias) in + let pat' = Loc.tag ?loc @@ PatCstr (cstr, patargs, alias) in let cstr = mkConstructU (on_snd EInstance.make ci.cs_cstr) in let app = applist (cstr, List.map (lift (List.length sign)) params) in let app = applist (app, args) in diff --git a/pretyping/detyping.ml b/pretyping/detyping.ml index 3079d1052..721d1d749 100644 --- a/pretyping/detyping.ml +++ b/pretyping/detyping.ml @@ -68,14 +68,14 @@ let isomorphic_to_tuple lc = Int.equal (Array.length lc) 1 let encode_bool r = let (x,lc) = encode_inductive r in if not (has_two_constructors lc) then - user_err ~loc:(loc_of_reference r) ~hdr:"encode_if" + user_err ?loc:(loc_of_reference r) ~hdr:"encode_if" (str "This type has not exactly two constructors."); x let encode_tuple r = let (x,lc) = encode_inductive r in if not (isomorphic_to_tuple lc) then - user_err ~loc:(loc_of_reference r) ~hdr:"encode_tuple" + user_err ?loc:(loc_of_reference r) ~hdr:"encode_tuple" (str "This type cannot be seen as a tuple type."); x @@ -793,7 +793,7 @@ let detype_closed_glob ?lax isgoal avoid env sigma t = (**********************************************************************) (* Module substitution: relies on detyping *) -let rec subst_cases_pattern subst (loc, pat) = Loc.tag ~loc @@ +let rec subst_cases_pattern subst (loc, pat) = Loc.tag ?loc @@ match pat with | PatVar _ -> pat | PatCstr (((kn,i),j),cpl,n) -> @@ -804,7 +804,7 @@ let rec subst_cases_pattern subst (loc, pat) = Loc.tag ~loc @@ let (f_subst_genarg, subst_genarg_hook) = Hook.make () -let rec subst_glob_constr subst (loc, raw) = Loc.tag ~loc @@ +let rec subst_glob_constr subst (loc, raw) = Loc.tag ?loc @@ match raw with | GRef (ref,u) -> let ref',t = subst_global subst ref in diff --git a/pretyping/evarconv.ml b/pretyping/evarconv.ml index 95680e18a..1318942c5 100644 --- a/pretyping/evarconv.ml +++ b/pretyping/evarconv.ml @@ -1268,7 +1268,7 @@ let solve_unconstrained_impossible_cases env evd = match ev_info.evar_source with | loc,Evar_kinds.ImpossibleCase -> let j, ctx = coq_unit_judge () in - let evd' = Evd.merge_context_set Evd.univ_flexible_alg ~loc evd' ctx in + let evd' = Evd.merge_context_set Evd.univ_flexible_alg ?loc evd' ctx in let ty = j_type j in let conv_algo = evar_conv_x full_transparent_state in let evd' = check_evar_instance evd' evk ty conv_algo in diff --git a/pretyping/evardefine.ml b/pretyping/evardefine.ml index c5ae684e3..c2b267dd8 100644 --- a/pretyping/evardefine.ml +++ b/pretyping/evardefine.ml @@ -181,7 +181,7 @@ let define_evar_as_sort env evd (ev,args) = constraint on its domain and codomain. If the input constraint is an evar instantiate it with the product of 2 new evars. *) -let split_tycon loc env evd tycon = +let split_tycon ?loc env evd tycon = let rec real_split evd c = let t = Reductionops.whd_all env evd c in match EConstr.kind evd t with @@ -193,7 +193,7 @@ let split_tycon loc env evd tycon = | App (c,args) when isEvar evd c -> let (evd',lam) = define_evar_as_lambda env evd (destEvar evd c) in real_split evd' (mkApp (lam,args)) - | _ -> error_not_product ~loc env evd c + | _ -> error_not_product ?loc env evd c in match tycon with | None -> evd,(Anonymous,None,None) diff --git a/pretyping/evardefine.mli b/pretyping/evardefine.mli index 2f7ac4efb..b8134a28c 100644 --- a/pretyping/evardefine.mli +++ b/pretyping/evardefine.mli @@ -31,7 +31,7 @@ val evar_absorb_arguments : env -> evar_map -> existential -> constr list -> evar_map * existential val split_tycon : - Loc.t -> env -> evar_map -> type_constraint -> + ?loc:Loc.t -> env -> evar_map -> type_constraint -> evar_map * (Name.t * type_constraint * type_constraint) val valcon_of_tycon : type_constraint -> val_constraint diff --git a/pretyping/glob_ops.ml b/pretyping/glob_ops.ml index cba1780ba..7e6b063d0 100644 --- a/pretyping/glob_ops.ml +++ b/pretyping/glob_ops.ml @@ -354,7 +354,7 @@ let add_and_check_ident id set = Id.Set.add id set let bound_glob_vars = - let rec vars bound = Loc.with_loc (fun ~loc -> function + let rec vars bound = Loc.with_loc (fun ?loc -> function | GLambda (na,_,_,_) | GProd (na,_,_,_) | GLetIn (na,_,_,_) as c -> let bound = name_fold add_and_check_ident na bound in fold_glob_constr vars bound (loc, c) @@ -495,7 +495,7 @@ let rename_var l id = if List.exists (fun (_,id') -> Id.equal id id') l then raise UnsoundRenaming else id -let rec rename_glob_vars l = Loc.map_with_loc (fun ~loc -> function +let rec rename_glob_vars l = Loc.map_with_loc (fun ?loc -> function | GVar id as r -> let id' = rename_var l id in if id == id' then r else GVar id' @@ -558,10 +558,10 @@ let rec cases_pattern_of_glob_constr na = Loc.map (function ) (* Turn a closed cases pattern into a glob_constr *) -let rec glob_constr_of_closed_cases_pattern_aux x = Loc.map_with_loc (fun ~loc -> function +let rec glob_constr_of_closed_cases_pattern_aux x = Loc.map_with_loc (fun ?loc -> function | PatCstr (cstr,[],Anonymous) -> GRef (ConstructRef cstr,None) | PatCstr (cstr,l,Anonymous) -> - let ref = Loc.tag ~loc @@ GRef (ConstructRef cstr,None) in + let ref = Loc.tag ?loc @@ GRef (ConstructRef cstr,None) in GApp (ref, List.map glob_constr_of_closed_cases_pattern_aux l) | _ -> raise Not_found ) x diff --git a/pretyping/glob_ops.mli b/pretyping/glob_ops.mli index ac2118ba7..dae662747 100644 --- a/pretyping/glob_ops.mli +++ b/pretyping/glob_ops.mli @@ -20,7 +20,7 @@ val glob_constr_eq : glob_constr -> glob_constr -> bool (** Operations on [glob_constr] *) -val cases_pattern_loc : cases_pattern -> Loc.t +val cases_pattern_loc : cases_pattern -> Loc.t option val cases_predicate_names : tomatch_tuples -> Name.t list @@ -41,7 +41,8 @@ val iter_glob_constr : (glob_constr -> unit) -> glob_constr -> unit val occur_glob_constr : Id.t -> glob_constr -> bool val free_glob_vars : glob_constr -> Id.t list val bound_glob_vars : glob_constr -> Id.Set.t -val loc_of_glob_constr : glob_constr -> Loc.t +(* Obsolete *) +val loc_of_glob_constr : glob_constr -> Loc.t option val glob_visible_short_qualid : glob_constr -> Id.t list (* Renaming free variables using a renaming map; fails with diff --git a/pretyping/patternops.ml b/pretyping/patternops.ml index 6696e174b..84d846afd 100644 --- a/pretyping/patternops.ml +++ b/pretyping/patternops.ml @@ -324,7 +324,7 @@ let warn_cast_in_pattern = CWarnings.create ~name:"cast-in-pattern" ~category:"automation" (fun () -> Pp.strbrk "Casts are ignored in patterns") -let rec pat_of_raw metas vars = Loc.with_loc (fun ~loc -> function +let rec pat_of_raw metas vars = Loc.with_loc (fun ?loc -> function | GVar id -> (try PRel (List.index Name.equal (Name id) vars) with Not_found -> PVar id) @@ -362,7 +362,7 @@ let rec pat_of_raw metas vars = Loc.with_loc (fun ~loc -> function PIf (pat_of_raw metas vars c, pat_of_raw metas vars b1,pat_of_raw metas vars b2) | GLetTuple (nal,(_,None),b,c) -> - let mkGLambda c na = Loc.tag ~loc @@ + let mkGLambda c na = Loc.tag ?loc @@ GLambda (na,Explicit, Loc.tag @@ GHole (Evar_kinds.InternalHole, IntroAnonymous, None),c) in let c = List.fold_left mkGLambda c nal in let cip = @@ -391,7 +391,7 @@ let rec pat_of_raw metas vars = Loc.with_loc (fun ~loc -> function rev_it_mkPLambda nal (mkPLambda na (pat_of_raw metas nvars p)) | (None | Some (_, GHole _)), _ -> PMeta None | Some p, None -> - user_err ~loc (strbrk "Clause \"in\" expected in patterns over \"match\" expressions with an explicit \"return\" clause.") + user_err ?loc (strbrk "Clause \"in\" expected in patterns over \"match\" expressions with an explicit \"return\" clause.") in let info = { cip_style = sty; @@ -404,7 +404,7 @@ let rec pat_of_raw metas vars = Loc.with_loc (fun ~loc -> function one non-trivial branch. These facts are used in [Constrextern]. *) PCase (info, pred, pat_of_raw metas vars c, brs) - | r -> err ~loc (Pp.str "Non supported pattern.") + | r -> err ?loc (Pp.str "Non supported pattern.") ) and pats_of_glob_branches loc metas vars ind brs = @@ -412,7 +412,7 @@ and pats_of_glob_branches loc metas vars ind brs = | _, PatVar na -> name_iter (fun n -> metas := n::!metas) na; na - | loc, PatCstr(_,_,_) -> err ~loc (Pp.str "Non supported pattern.") + | loc, PatCstr(_,_,_) -> err ?loc (Pp.str "Non supported pattern.") in let rec get_pat indexes = function | [] -> false, [] @@ -421,10 +421,10 @@ and pats_of_glob_branches loc metas vars ind brs = let () = match ind with | Some sp when eq_ind sp indsp -> () | _ -> - err ~loc (Pp.str "All constructors must be in the same inductive type.") + err ?loc (Pp.str "All constructors must be in the same inductive type.") in if Int.Set.mem (j-1) indexes then - err ~loc + err ?loc (str "No unique branch for " ++ int j ++ str"-th constructor."); let lna = List.map get_arg lv in let vars' = List.rev lna @ vars in @@ -432,7 +432,7 @@ and pats_of_glob_branches loc metas vars ind brs = let ext,pats = get_pat (Int.Set.add (j-1) indexes) brs in let tags = List.map (fun _ -> false) lv (* approximation, w/o let-in *) in ext, ((j-1, tags, pat) :: pats) - | (loc,(_,_,_)) :: _ -> err ~loc (Pp.str "Non supported pattern.") + | (loc,(_,_,_)) :: _ -> err ?loc (Pp.str "Non supported pattern.") in get_pat Int.Set.empty brs diff --git a/pretyping/pretype_errors.ml b/pretyping/pretype_errors.ml index 24f6d1689..ef187e7a6 100644 --- a/pretyping/pretype_errors.ml +++ b/pretyping/pretype_errors.ml @@ -177,7 +177,7 @@ let unsatisfiable_constraints env evd ev comp = | Some ev -> let loc, kind = Evd.evar_source ev evd in let err = UnsatisfiableConstraints (Some (ev, kind), comp) in - Loc.raise ~loc (PretypeError (env,evd,err)) + Loc.raise ?loc (PretypeError (env,evd,err)) let unsatisfiable_exception exn = match exn with diff --git a/pretyping/pretyping.ml b/pretyping/pretyping.ml index fe15cb490..a256eaa5d 100644 --- a/pretyping/pretyping.ml +++ b/pretyping/pretyping.ml @@ -216,8 +216,8 @@ let interp_universe_level_name evd (loc,s) = evd, snd (Idmap.find id names) with Not_found -> if not (is_strict_universe_declarations ()) then - new_univ_level_variable ~loc ~name:s univ_rigid evd - else user_err ~loc ~hdr:"interp_universe_level_name" + new_univ_level_variable ?loc ~name:s univ_rigid evd + else user_err ?loc ~hdr:"interp_universe_level_name" (Pp.(str "Undeclared universe: " ++ str s)) let interp_universe ?loc evd = function @@ -229,8 +229,8 @@ let interp_universe ?loc evd = function (evd', Univ.sup u (Univ.Universe.make l))) (evd, Univ.Universe.type0m) l -let interp_level_info loc evd : Misctypes.level_info -> _ = function - | None -> new_univ_level_variable ~loc univ_rigid evd +let interp_level_info ?loc evd : Misctypes.level_info -> _ = function + | None -> new_univ_level_variable ?loc univ_rigid evd | Some (loc,s) -> interp_universe_level_name evd (loc,s) let interp_sort ?loc evd = function @@ -337,7 +337,7 @@ let check_extra_evars_are_solved env current_sigma frozen = match frozen with match k with | Evar_kinds.ImplicitArg (gr, (i, id), false) -> () | _ -> - error_unsolvable_implicit ~loc env current_sigma evk None) pending + error_unsolvable_implicit ?loc env current_sigma evk None) pending (* [check_evars] fails if some unresolved evar remains *) @@ -349,7 +349,7 @@ let check_evars env initial_sigma sigma c = let (loc,k) = evar_source evk sigma in begin match k with | Evar_kinds.ImplicitArg (gr, (i, id), false) -> () - | _ -> Pretype_errors.error_unsolvable_implicit ~loc env sigma evk None + | _ -> Pretype_errors.error_unsolvable_implicit ?loc env sigma evk None end | _ -> EConstr.iter sigma proc_rec c in proc_rec c @@ -393,9 +393,9 @@ let check_instance loc subst = function | [] -> () | (id,_) :: _ -> if List.mem_assoc id subst then - user_err ~loc (pr_id id ++ str "appears more than once.") + user_err ?loc (pr_id id ++ str "appears more than once.") else - user_err ~loc (str "No such variable in the signature of the existential variable: " ++ pr_id id ++ str ".") + user_err ?loc (str "No such variable in the signature of the existential variable: " ++ pr_id id ++ str ".") (* used to enforce a name in Lambda when the type constraints itself is named, hence possibly dependent *) @@ -475,7 +475,7 @@ let pretype_id pretype k0 loc env evdref lvar id = (* and build a nice error message *) if Id.Map.mem id lvar.ltac_genargs then begin let Geninterp.Val.Dyn (typ, _) = Id.Map.find id lvar.ltac_genargs in - user_err ~loc + user_err ?loc (str "Variable " ++ pr_id id ++ str " should be bound to a term but is \ bound to a " ++ Geninterp.Val.pr typ ++ str ".") end; @@ -484,47 +484,47 @@ let pretype_id pretype k0 loc env evdref lvar id = { uj_val = mkVar id; uj_type = NamedDecl.get_type (lookup_named id env) } with Not_found -> (* [id] not found, standard error message *) - error_var_not_found ~loc id + error_var_not_found ?loc id (*************************************************************************) (* Main pretyping function *) -let interp_glob_level loc evd : Misctypes.glob_level -> _ = function +let interp_glob_level ?loc evd : Misctypes.glob_level -> _ = function | GProp -> evd, Univ.Level.prop | GSet -> evd, Univ.Level.set - | GType s -> interp_level_info loc evd s + | GType s -> interp_level_info ?loc evd s -let interp_instance loc evd ~len l = +let interp_instance ?loc evd ~len l = if len != List.length l then - user_err ~loc ~hdr:"pretype" + user_err ?loc ~hdr:"pretype" (str "Universe instance should have length " ++ int len) else let evd, l' = List.fold_left (fun (evd, univs) l -> - let evd, l = interp_glob_level loc evd l in + let evd, l = interp_glob_level ?loc evd l in (evd, l :: univs)) (evd, []) l in if List.exists (fun l -> Univ.Level.is_prop l) l' then - user_err ~loc ~hdr:"pretype" + user_err ?loc ~hdr:"pretype" (str "Universe instances cannot contain Prop, polymorphic" ++ str " universe instances must be greater or equal to Set."); evd, Some (Univ.Instance.of_array (Array.of_list (List.rev l'))) -let pretype_global loc rigid env evd gr us = +let pretype_global ?loc rigid env evd gr us = let evd, instance = match us with | None -> evd, None | Some l -> let _, ctx = Universes.unsafe_constr_of_global gr in let len = Univ.UContext.size ctx in - interp_instance loc evd ~len l + interp_instance ?loc evd ~len l in - let (sigma, c) = Evd.fresh_global ~loc ~rigid ?names:instance env.ExtraEnv.env evd gr in + let (sigma, c) = Evd.fresh_global ?loc ~rigid ?names:instance env.ExtraEnv.env evd gr in (sigma, EConstr.of_constr c) -let pretype_ref loc evdref env ref us = +let pretype_ref ?loc evdref env ref us = match ref with | VarRef id -> (* Section variable *) @@ -533,15 +533,15 @@ let pretype_ref loc evdref env ref us = (* This may happen if env is a goal env and section variables have been cleared - section variables should be different from goal variables *) - Pretype_errors.error_var_not_found ~loc id) + Pretype_errors.error_var_not_found ?loc id) | ref -> - let evd, c = pretype_global loc univ_flexible env !evdref ref us in + let evd, c = pretype_global ?loc univ_flexible env !evdref ref us in let () = evdref := evd in let ty = unsafe_type_of env.ExtraEnv.env evd c in make_judge c ty let judge_of_Type loc evd s = - let evd, s = interp_universe ~loc evd s in + let evd, s = interp_universe ?loc evd s in let judge = { uj_val = mkSort (Type s); uj_type = mkSort (Type (Univ.super s)) } in @@ -574,12 +574,12 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre let open Context.Rel.Declaration in match t with | GRef (ref,u) -> - inh_conv_coerce_to_tycon ~loc env evdref - (pretype_ref loc evdref env ref u) + inh_conv_coerce_to_tycon ?loc env evdref + (pretype_ref ?loc evdref env ref u) tycon | GVar id -> - inh_conv_coerce_to_tycon ~loc env evdref + inh_conv_coerce_to_tycon ?loc env evdref (pretype_id (fun e r l t -> pretype tycon e r l t) k0 loc env evdref lvar id) tycon @@ -589,12 +589,12 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre let evk = try Evd.evar_key id !evdref with Not_found -> - user_err ~loc (str "Unknown existential variable.") in + user_err ?loc (str "Unknown existential variable.") in let hyps = evar_filtered_context (Evd.find !evdref evk) in let args = pretype_instance k0 resolve_tc env evdref lvar loc hyps evk inst in let c = mkEvar (evk, args) in let j = (Retyping.get_judgment_of env.ExtraEnv.env !evdref c) in - inh_conv_coerce_to_tycon ~loc env evdref j tycon + inh_conv_coerce_to_tycon ?loc env evdref j tycon | GPatVar (someta,n) -> let env = ltac_interp_name_env k0 lvar env !evdref in @@ -674,7 +674,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre { uj_val = it_mkLambda_or_LetIn j.uj_val ctxt; uj_type = it_mkProd_or_LetIn j.uj_type ctxt }) ctxtv vdef in - Typing.check_type_fixpoint ~loc env.ExtraEnv.env evdref names ftys vdefj; + Typing.check_type_fixpoint ?loc env.ExtraEnv.env evdref names ftys vdefj; let nf c = nf_evar !evdref c in let ftys = Array.map nf ftys in (** FIXME *) let fdefs = Array.map (fun x -> nf (j_val x)) vdefj in @@ -696,7 +696,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre let fixdecls = (names,ftys,fdefs) in let indexes = search_guard - ~loc env.ExtraEnv.env possible_indexes (nf_fix !evdref fixdecls) + ?loc env.ExtraEnv.env possible_indexes (nf_fix !evdref fixdecls) in make_judge (mkFix ((indexes,i),fixdecls)) ftys.(i) | GCoFix i -> @@ -705,15 +705,15 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre (try check_cofix env.ExtraEnv.env (i, nf_fix !evdref fixdecls) with reraise -> let (e, info) = CErrors.push reraise in - let info = Loc.add_loc info loc in + let info = Option.cata (Loc.add_loc info) info loc in iraise (e, info)); make_judge (mkCoFix cofix) ftys.(i) in - inh_conv_coerce_to_tycon ~loc env evdref fixj tycon + inh_conv_coerce_to_tycon ?loc env evdref fixj tycon | GSort s -> let j = pretype_sort loc evdref s in - inh_conv_coerce_to_tycon ~loc env evdref j tycon + inh_conv_coerce_to_tycon ?loc env evdref j tycon | GApp (f,args) -> let fj = pretype empty_tycon env evdref lvar f in @@ -775,7 +775,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre | _ -> let hj = pretype empty_tycon env evdref lvar c in error_cant_apply_not_functional - ~loc:(Loc.merge floc argloc) env.ExtraEnv.env !evdref + ?loc:(Loc.merge_opt floc argloc) env.ExtraEnv.env !evdref resj [|hj|] in let resj = apply_rec env 1 fj candargs args in @@ -792,7 +792,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre else resj | _ -> resj in - inh_conv_coerce_to_tycon ~loc env evdref resj tycon + inh_conv_coerce_to_tycon ?loc env evdref resj tycon | GLambda(name,bk,c1,c2) -> let tycon' = evd_comb1 @@ -800,11 +800,11 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre match tycon with | None -> evd, tycon | Some ty -> - let evd, ty' = Coercion.inh_coerce_to_prod ~loc env.ExtraEnv.env evd ty in + let evd, ty' = Coercion.inh_coerce_to_prod ?loc env.ExtraEnv.env evd ty in evd, Some ty') evdref tycon in - let (name',dom,rng) = evd_comb1 (split_tycon loc env.ExtraEnv.env) evdref tycon' in + let (name',dom,rng) = evd_comb1 (split_tycon ?loc env.ExtraEnv.env) evdref tycon' in let dom_valcon = valcon_of_tycon dom in let j = pretype_type dom_valcon env evdref lvar c1 in (* The name specified by ltac is used also to create bindings. So @@ -814,7 +814,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre let j' = pretype rng (push_rel !evdref var env) evdref lvar c2 in let name = ltac_interp_name lvar name in let resj = judge_of_abstraction env.ExtraEnv.env (orelse_name name name') j j' in - inh_conv_coerce_to_tycon ~loc env evdref resj tycon + inh_conv_coerce_to_tycon ?loc env evdref resj tycon | GProd(name,bk,c1,c2) -> let j = pretype_type empty_valcon env evdref lvar c1 in @@ -836,9 +836,9 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre judge_of_product env.ExtraEnv.env name j j' with TypeError _ as e -> let (e, info) = CErrors.push e in - let info = Loc.add_loc info loc in + let info = Option.cata (Loc.add_loc info) info loc in iraise (e, info) in - inh_conv_coerce_to_tycon ~loc env evdref resj tycon + inh_conv_coerce_to_tycon ?loc env evdref resj tycon | GLetIn(name,c1,t,c2) -> let tycon1 = @@ -867,15 +867,15 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre try find_rectype env.ExtraEnv.env !evdref cj.uj_type with Not_found -> let cloc = loc_of_glob_constr c in - error_case_not_inductive ~loc:cloc env.ExtraEnv.env !evdref cj + error_case_not_inductive ?loc:cloc env.ExtraEnv.env !evdref cj in let cstrs = get_constructors env.ExtraEnv.env indf in if not (Int.equal (Array.length cstrs) 1) then - user_err ~loc (str "Destructing let is only for inductive types" ++ + user_err ?loc (str "Destructing let is only for inductive types" ++ str " with one constructor."); let cs = cstrs.(0) in if not (Int.equal (List.length nal) cs.cs_nargs) then - user_err ~loc:loc (str "Destructing let on this type expects " ++ + user_err ?loc:loc (str "Destructing let on this type expects " ++ int cs.cs_nargs ++ str " variables."); let fsign, record = let set_name na d = set_name na (map_rel_decl EConstr.of_constr d) in @@ -944,7 +944,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre if noccur_between !evdref 1 cs.cs_nargs ccl then lift (- cs.cs_nargs) ccl else - error_cant_find_case_type ~loc env.ExtraEnv.env !evdref + error_cant_find_case_type ?loc env.ExtraEnv.env !evdref cj.uj_val in (* let ccl = refresh_universes ccl in *) let p = it_mkLambda_or_LetIn (lift (nar+1) ccl) psign in @@ -960,10 +960,10 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre try find_rectype env.ExtraEnv.env !evdref cj.uj_type with Not_found -> let cloc = loc_of_glob_constr c in - error_case_not_inductive ~loc:cloc env.ExtraEnv.env !evdref cj in + error_case_not_inductive ?loc:cloc env.ExtraEnv.env !evdref cj in let cstrs = get_constructors env.ExtraEnv.env indf in if not (Int.equal (Array.length cstrs) 2) then - user_err ~loc + user_err ?loc (str "If is only for inductive types with two constructors."); let arsgn = @@ -1020,10 +1020,10 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre mkCase (ci, pred, cj.uj_val, [|b1;b2|]) in let cj = { uj_val = v; uj_type = p } in - inh_conv_coerce_to_tycon ~loc env evdref cj tycon + inh_conv_coerce_to_tycon ?loc env evdref cj tycon | GCases (sty,po,tml,eqns) -> - Cases.compile_cases ~loc sty + Cases.compile_cases ?loc sty ((fun vtyc env evdref -> pretype vtyc (make_env env !evdref) evdref lvar),evdref) tycon env.ExtraEnv.env (* loc *) (po,tml,eqns) @@ -1032,7 +1032,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre match k with | CastCoerce -> let cj = pretype empty_tycon env evdref lvar c in - evd_comb1 (Coercion.inh_coerce_to_base ~loc env.ExtraEnv.env) evdref cj + evd_comb1 (Coercion.inh_coerce_to_base ?loc env.ExtraEnv.env) evdref cj | CastConv t | CastVM t | CastNative t -> let k = (match k with CastVM _ -> VMcast | CastNative _ -> NATIVEcast | _ -> DEFAULTcast) in let tj = pretype_type empty_valcon env evdref lvar t in @@ -1048,9 +1048,9 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre let (evd,b) = Reductionops.vm_infer_conv env.ExtraEnv.env !evdref cty tval in if b then (evdref := evd; cj, tval) else - error_actual_type ~loc env.ExtraEnv.env !evdref cj tval + error_actual_type ?loc env.ExtraEnv.env !evdref cj tval (ConversionFailed (env.ExtraEnv.env,cty,tval)) - else user_err ~loc (str "Cannot check cast with vm: " ++ + else user_err ?loc (str "Cannot check cast with vm: " ++ str "unresolved arguments remain.") | NATIVEcast -> let cj = pretype empty_tycon env evdref lvar c in @@ -1059,7 +1059,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre let (evd,b) = Nativenorm.native_infer_conv env.ExtraEnv.env !evdref cty tval in if b then (evdref := evd; cj, tval) else - error_actual_type ~loc env.ExtraEnv.env !evdref cj tval + error_actual_type ?loc env.ExtraEnv.env !evdref cj tval (ConversionFailed (env.ExtraEnv.env,cty,tval)) end | _ -> @@ -1067,7 +1067,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre in let v = mkCast (cj.uj_val, k, tval) in { uj_val = v; uj_type = tval } - in inh_conv_coerce_to_tycon ~loc env evdref cj tycon + in inh_conv_coerce_to_tycon ?loc env evdref cj tycon and pretype_instance k0 resolve_tc env evdref lvar loc hyps evk update = let f decl (subst,update) = @@ -1087,7 +1087,7 @@ and pretype_instance k0 resolve_tc env evdref lvar loc hyps evk update = let t' = env |> lookup_named id |> NamedDecl.get_type in if is_conv env.ExtraEnv.env !evdref t t' then mkVar id, update else raise Not_found with Not_found -> - user_err ~loc (str "Cannot interpret " ++ + user_err ?loc (str "Cannot interpret " ++ pr_existential_key !evdref evk ++ str " in current context: no binding for " ++ pr_id id ++ str ".") in ((id,c)::subst, update) in @@ -1128,14 +1128,14 @@ and pretype_type k0 resolve_tc valcon (env : ExtraEnv.t) evdref lvar = function | c -> let j = pretype k0 resolve_tc empty_tycon env evdref lvar c in let loc = loc_of_glob_constr c in - let tj = evd_comb1 (Coercion.inh_coerce_to_sort ~loc env.ExtraEnv.env) evdref j in + let tj = evd_comb1 (Coercion.inh_coerce_to_sort ?loc env.ExtraEnv.env) evdref j in match valcon with | None -> tj | Some v -> if e_cumul env.ExtraEnv.env evdref v tj.utj_val then tj else error_unexpected_type - ~loc:(loc_of_glob_constr c) env.ExtraEnv.env !evdref tj.utj_val v + ?loc:(loc_of_glob_constr c) env.ExtraEnv.env !evdref tj.utj_val v let ise_pretype_gen flags env sigma lvar kind c = let env = make_env env sigma in diff --git a/printing/ppconstr.ml b/printing/ppconstr.ml index fddd54a9f..900bf2daf 100644 --- a/printing/ppconstr.ml +++ b/printing/ppconstr.ml @@ -147,7 +147,7 @@ let tag_var = tag Tag.variable let pr_with_comments ?loc pp = pr_located (fun x -> x) (Loc.tag ?loc pp) - let pr_sep_com sep f c = pr_with_comments ~loc:(constr_loc c) (sep() ++ f c) + let pr_sep_com sep f c = pr_with_comments ?loc:(constr_loc c) (sep() ++ f c) let pr_univ l = match l with @@ -221,11 +221,10 @@ let tag_var = tag Tag.variable | t -> str " :" ++ pr_sep_com (fun()->brk(1,2)) (pr ltop) t let pr_lident (loc,id) = - if not (Loc.is_ghost loc) then - let (b,_) = Loc.unloc loc in - pr_located pr_id (Loc.make_loc (b,b + String.length (Id.to_string id)), id) - else - pr_id id + match loc with + | None -> pr_id id + | Some loc -> let (b,_) = Loc.unloc loc in + pr_located pr_id @@ Loc.tag ~loc:(Loc.make_loc (b,b + String.length (Id.to_string id))) id let pr_lname = function | (loc,Name id) -> pr_lident (loc,id) @@ -302,7 +301,7 @@ let tag_var = tag Tag.variable assert false in let loc = cases_pattern_expr_loc (loc, p) in - pr_with_comments ~loc + pr_with_comments ?loc (sep() ++ if prec_less prec inh then strm else surround strm) let pr_patt = pr_patt mt @@ -310,16 +309,18 @@ let tag_var = tag Tag.variable let pr_eqn pr (loc,(pl,rhs)) = let pl = List.map snd pl in spc() ++ hov 4 - (pr_with_comments ~loc + (pr_with_comments ?loc (str "| " ++ hov 0 (prlist_with_sep pr_bar (prlist_with_sep sep_v (pr_patt ltop)) pl ++ str " =>") ++ pr_sep_com spc (pr ltop) rhs)) - let begin_of_binder = function - | CLocalDef((loc,_),_,_) -> fst (Loc.unloc loc) - | CLocalAssum((loc,_)::_,_,_) -> fst (Loc.unloc loc) - | CLocalPattern(loc,(_,_)) -> fst (Loc.unloc loc) + let begin_of_binder l_bi = + let b_loc l = fst (Option.cata Loc.unloc (0,0) l) in + match l_bi with + | CLocalDef((loc,_),_,_) -> b_loc loc + | CLocalAssum((loc,_)::_,_,_) -> b_loc loc + | CLocalPattern(loc,(_,_)) -> b_loc loc | _ -> assert false let begin_of_binders = function @@ -420,12 +421,12 @@ let tag_var = tag Tag.variable | CLambdaN ((nal,bk,t)::bl,c) -> let bl,c = extract_lam_binders (loc, CLambdaN(bl,c)) in CLocalAssum (nal,bk,t) :: bl, c - | c -> [], Loc.tag ~loc c + | c -> [], Loc.tag ?loc c - let split_lambda = Loc.with_loc (fun ~loc -> function + let split_lambda = Loc.with_loc (fun ?loc -> function | CLambdaN ([[na],bk,t],c) -> (na,t,c) - | CLambdaN (([na],bk,t)::bl,c) -> (na,t,Loc.tag ~loc @@ CLambdaN(bl,c)) - | CLambdaN ((na::nal,bk,t)::bl,c) -> (na,t,Loc.tag ~loc @@ CLambdaN((nal,bk,t)::bl,c)) + | CLambdaN (([na],bk,t)::bl,c) -> (na,t,Loc.tag ?loc @@ CLambdaN(bl,c)) + | CLambdaN ((na::nal,bk,t)::bl,c) -> (na,t,Loc.tag ?loc @@ CLambdaN((nal,bk,t)::bl,c)) | _ -> anomaly (Pp.str "ill-formed fixpoint body") ) @@ -436,11 +437,11 @@ let tag_var = tag Tag.variable | (_,Name id), (_,Anonymous) -> (na,t,c) | _ -> (na',t,c) - let split_product na' = Loc.with_loc (fun ~loc -> function + let split_product na' = Loc.with_loc (fun ?loc -> function | CProdN ([[na],bk,t],c) -> rename na na' t c - | CProdN (([na],bk,t)::bl,c) -> rename na na' t (Loc.tag ~loc @@ CProdN(bl,c)) + | CProdN (([na],bk,t)::bl,c) -> rename na na' t (Loc.tag ?loc @@ CProdN(bl,c)) | CProdN ((na::nal,bk,t)::bl,c) -> - rename na na' t (Loc.tag ~loc @@ CProdN((nal,bk,t)::bl,c)) + rename na na' t (Loc.tag ?loc @@ CProdN((nal,bk,t)::bl,c)) | _ -> anomaly (Pp.str "ill-formed fixpoint body") ) @@ -730,7 +731,7 @@ let tag_var = tag Tag.variable return (pr_delimiters sc (pr mt (ldelim,E) a), ldelim) in let loc = constr_loc a in - pr_with_comments ~loc + pr_with_comments ?loc (sep() ++ if prec_less prec inherited then strm else surround strm) type term_pr = { diff --git a/printing/pputils.ml b/printing/pputils.ml index 4ae5035a2..d1ba1a4a1 100644 --- a/printing/pputils.ml +++ b/printing/pputils.ml @@ -15,14 +15,15 @@ open Locus open Genredexpr let pr_located pr (loc, x) = - if !Flags.beautify && not (Loc.is_ghost loc) then + match loc with + | Some loc when !Flags.beautify -> let (b, e) = Loc.unloc loc in (* Side-effect: order matters *) let before = Pp.comment (CLexer.extract_comments b) in let x = pr x in let after = Pp.comment (CLexer.extract_comments e) in before ++ x ++ after - else pr x + | _ -> pr x let pr_or_var pr = function | ArgArg x -> pr x diff --git a/printing/ppvernac.ml b/printing/ppvernac.ml index 27faa7c5c..5b6c5580a 100644 --- a/printing/ppvernac.ml +++ b/printing/ppvernac.ml @@ -32,11 +32,10 @@ open Decl_kinds let pr_spc_lconstr = pr_sep_com spc pr_lconstr_expr let pr_lident (loc,id) = - if Loc.is_ghost loc then - let (b,_) = Loc.unloc loc in - pr_located pr_id (Loc.make_loc (b,b + String.length(Id.to_string id)),id) - else - pr_id id + match loc with + | None -> pr_id id + | Some loc -> let (b,_) = Loc.unloc loc in + pr_located pr_id @@ Loc.tag ~loc:(Loc.make_loc (b,b + String.length (Id.to_string id))) id let pr_plident (lid, l) = pr_lident lid ++ @@ -50,11 +49,10 @@ open Decl_kinds let pr_fqid fqid = str (string_of_fqid fqid) let pr_lfqid (loc,fqid) = - if Loc.is_ghost loc then - let (b,_) = Loc.unloc loc in - pr_located pr_fqid (Loc.make_loc (b,b + String.length(string_of_fqid fqid)),fqid) - else - pr_fqid fqid + match loc with + | None -> pr_fqid fqid + | Some loc -> let (b,_) = Loc.unloc loc in + pr_located pr_fqid @@ Loc.tag ~loc:(Loc.make_loc (b,b + String.length (string_of_fqid fqid))) fqid let pr_lname = function | (loc,Name id) -> pr_lident (loc,id) @@ -296,7 +294,7 @@ open Decl_kinds let begin_of_inductive = function | [] -> 0 - | (_,((loc,_),_))::_ -> fst (Loc.unloc loc) + | (_,((loc,_),_))::_ -> Option.cata (fun loc -> fst (Loc.unloc loc)) 0 loc let pr_class_rawexpr = function | FunClass -> keyword "Funclass" diff --git a/printing/prettyp.ml b/printing/prettyp.ml index 025514902..eb1124e73 100644 --- a/printing/prettyp.ml +++ b/printing/prettyp.ml @@ -711,7 +711,7 @@ let read_sec_context r = let dir = try Nametab.locate_section qid with Not_found -> - user_err ~loc ~hdr:"read_sec_context" (str "Unknown section.") in + user_err ?loc ~hdr:"read_sec_context" (str "Unknown section.") in let rec get_cxt in_cxt = function | (_,Lib.OpenedSection ((dir',_),_) as hd)::rest -> if DirPath.equal dir dir' then (hd::in_cxt) else get_cxt (hd::in_cxt) rest @@ -752,7 +752,7 @@ let print_any_name = function let print_name = function | ByNotation (loc,(ntn,sc)) -> print_any_name - (Term (Notation.interp_notation_as_global_reference ~loc (fun _ -> true) + (Term (Notation.interp_notation_as_global_reference ?loc (fun _ -> true) ntn sc)) | AN ref -> print_any_name (locate_any_name ref) @@ -776,11 +776,11 @@ let print_opaque_name qid = | VarRef id -> env |> lookup_named id |> NamedDecl.set_id id |> print_named_decl -let print_about_any loc k = +let print_about_any ?loc k = match k with | Term ref -> let rb = Reductionops.ReductionBehaviour.print ref in - Dumpglob.add_glob loc ref; + Dumpglob.add_glob ?loc ref; pr_infos_list (print_ref false ref :: blankline :: print_name_infos ref @ @@ -789,7 +789,7 @@ let print_about_any loc k = [hov 0 (str "Expands to: " ++ pr_located_qualid k)]) | Syntactic kn -> let () = match Syntax_def.search_syntactic_definition kn with - | [],Notation_term.NRef ref -> Dumpglob.add_glob loc ref + | [],Notation_term.NRef ref -> Dumpglob.add_glob ?loc ref | _ -> () in v 0 ( print_syntactic_def kn ++ fnl () ++ @@ -799,11 +799,11 @@ let print_about_any loc k = let print_about = function | ByNotation (loc,(ntn,sc)) -> - print_about_any loc - (Term (Notation.interp_notation_as_global_reference ~loc (fun _ -> true) + print_about_any ?loc + (Term (Notation.interp_notation_as_global_reference ?loc (fun _ -> true) ntn sc)) | AN ref -> - print_about_any (loc_of_reference ref) (locate_any_name ref) + print_about_any ?loc:(loc_of_reference ref) (locate_any_name ref) (* for debug *) let inspect depth = diff --git a/proofs/evar_refiner.ml b/proofs/evar_refiner.ml index 8367c09b8..b9c925796 100644 --- a/proofs/evar_refiner.ml +++ b/proofs/evar_refiner.ml @@ -56,7 +56,7 @@ let w_refine (evk,evi) (ltac_var,rawc) sigma = env sigma ltac_var (Pretyping.OfType (EConstr.of_constr evi.evar_concl)) rawc with e when CErrors.noncritical e -> let loc = Glob_ops.loc_of_glob_constr rawc in - user_err ~loc + user_err ?loc (str "Instance is not well-typed in the environment of " ++ Termops.pr_existential_key sigma evk ++ str ".") in diff --git a/proofs/pfedit.mli b/proofs/pfedit.mli index f9fb0b76d..89aa32643 100644 --- a/proofs/pfedit.mli +++ b/proofs/pfedit.mli @@ -130,7 +130,7 @@ val set_end_tac : Genarg.glob_generic_argument -> unit (** [set_used_variables l] declares that section variables [l] will be used in the proof *) val set_used_variables : - Id.t list -> Context.Named.t * (Loc.t * Names.Id.t) list + Id.t list -> Context.Named.t * Names.Id.t Loc.located list val get_used_variables : unit -> Context.Named.t option (** {6 Universe binders } *) diff --git a/proofs/proof_global.ml b/proofs/proof_global.ml index d659eba24..32eb9a3c1 100644 --- a/proofs/proof_global.ml +++ b/proofs/proof_global.ml @@ -207,7 +207,7 @@ let discard (loc,id) = let n = List.length !pstates in discard_gen id; if Int.equal (List.length !pstates) n then - CErrors.user_err ~loc + CErrors.user_err ?loc ~hdr:"Pfedit.delete_proof" (str"No such proof" ++ msg_proofs ()) let discard_current () = diff --git a/proofs/proof_global.mli b/proofs/proof_global.mli index 6bb6f5e2c..b0b0eba08 100644 --- a/proofs/proof_global.mli +++ b/proofs/proof_global.mli @@ -140,7 +140,7 @@ val set_endline_tactic : Genarg.glob_generic_argument -> unit * (w.r.t. type dependencies and let-ins covered by it) + a list of * ids to be cleared *) val set_used_variables : - Names.Id.t list -> Context.Named.t * (Loc.t * Names.Id.t) list + Names.Id.t list -> Context.Named.t * Names.Id.t Loc.located list val get_used_variables : unit -> Context.Named.t option val get_universe_binders : unit -> universe_binders option diff --git a/stm/stm.ml b/stm/stm.ml index 38415c1dd..d41e0c94f 100644 --- a/stm/stm.ml +++ b/stm/stm.ml @@ -1111,7 +1111,7 @@ end (* }}} *) let hints = ref Aux_file.empty_aux_file let set_compilation_hints file = hints := Aux_file.load_aux_file_for file -let get_hint_ctx ?loc = +let get_hint_ctx loc = let s = Aux_file.get ?loc !hints "context_used" in match Str.split (Str.regexp ";") s with | ids :: _ -> @@ -2027,7 +2027,7 @@ let collect_proof keep cur hd brkind id = !Flags.compilation_mode = Flags.BuildVio -> assert (VCS.Branch.equal hd hd'||VCS.Branch.equal hd VCS.edit_branch); (try - let name, hint = name ids, get_hint_ctx ?loc in + let name, hint = name ids, get_hint_ctx loc in let t, v = proof_no_using last in v.expr <- VernacProof(t, Some hint); `ASync (parent last,proof_using_ast last,accn,name,delegate name) @@ -2697,8 +2697,8 @@ let parse_sentence sid pa = Flags.with_option Flags.we_are_parsing (fun () -> try match Pcoq.Gram.entry_parse Pcoq.main_entry pa with - | None -> raise End_of_input - | Some com -> com + | None -> raise End_of_input + | Some (loc, cmd) -> Loc.tag ~loc cmd with e when CErrors.noncritical e -> let (e, info) = CErrors.push e in Exninfo.iraise (e, info)) @@ -2723,7 +2723,7 @@ let ind_len_loc_of_id sid = Note, this could maybe improved by handling more cases in compute_indentation. *) -let compute_indentation sid loc = +let compute_indentation ?loc sid = Option.cata (fun loc -> let open Loc in (* The effective lenght is the lenght on the last line *) let len = loc.ep - loc.bp in @@ -2737,20 +2737,20 @@ let compute_indentation sid loc = eff_indent + prev_indent, len else eff_indent, len - + ) (0, 0) loc let add ~ontop ?newtip verb (loc, ast) = let cur_tip = VCS.cur_tip () in if not (Stateid.equal ontop cur_tip) then - user_err ~hdr:"Stm.add" + user_err ?loc ~hdr:"Stm.add" (str "Stm.add called for a different state (" ++ str (Stateid.to_string ontop) ++ str ") than the tip: " ++ str (Stateid.to_string cur_tip) ++ str "." ++ fnl () ++ str "This is not supported yet, sorry."); - let indentation, strlen = compute_indentation ontop loc in + let indentation, strlen = compute_indentation ?loc ontop in CWarnings.set_current_loc loc; (* XXX: Classifiy vernac should be moved inside process transaction *) let clas = classify_vernac ast in - let aast = { verbose = verb; indentation; strlen; loc = Some loc; expr = ast } in + let aast = { verbose = verb; indentation; strlen; loc; expr = ast } in match process_transaction ?newtip aast clas with | `Ok -> VCS.cur_tip (), `NewTip | `Unfocus qed_id -> qed_id, `Unfocus (VCS.cur_tip ()) @@ -2768,10 +2768,10 @@ let query ~at ?(report_with=(Stateid.dummy,default_route)) s = if Stateid.equal at Stateid.dummy then finish () else Reach.known_state ~cache:`Yes at; let loc, ast = parse_sentence at s in - let indentation, strlen = compute_indentation at loc in + let indentation, strlen = compute_indentation ?loc at in CWarnings.set_current_loc loc; let clas = classify_vernac ast in - let aast = { verbose = true; indentation; strlen; loc = Some loc; expr = ast } in + let aast = { verbose = true; indentation; strlen; loc; expr = ast } in match clas with | VtStm (w,_), _ -> ignore(process_transaction aast (VtStm (w,false), VtNow)) diff --git a/tactics/autorewrite.ml b/tactics/autorewrite.ml index e58ec5a31..ff8b54ebf 100644 --- a/tactics/autorewrite.ml +++ b/tactics/autorewrite.ml @@ -86,7 +86,7 @@ let print_rewrite_hintdb bas = Pputils.pr_glb_generic (Global.env()) tac) (mt ()) h.rew_tac) (find_rewrites bas)) -type raw_rew_rule = Loc.t * constr Univ.in_universe_context_set * bool * Genarg.raw_generic_argument option +type raw_rew_rule = (constr Univ.in_universe_context_set * bool * Genarg.raw_generic_argument option) Loc.located (* Applies all the rules of one base *) let one_base general_rewrite_maybe_in tac_main bas = @@ -291,12 +291,12 @@ let decompose_applied_relation metas env sigma c ctype left2right = | Some c -> Some c | None -> None -let find_applied_relation metas loc env sigma c left2right = +let find_applied_relation ?loc metas env sigma c left2right = let ctype = Typing.unsafe_type_of env sigma (EConstr.of_constr c) in match decompose_applied_relation metas env sigma c ctype left2right with | Some c -> c | None -> - user_err ~loc ~hdr:"decompose_applied_relation" + user_err ?loc ~hdr:"decompose_applied_relation" (str"The type" ++ spc () ++ Printer.pr_econstr_env env sigma ctype ++ spc () ++ str"of this term does not end with an applied relation.") @@ -309,9 +309,9 @@ let add_rew_rules base lrul = let intern tac = snd (Genintern.generic_intern ist tac) in let lrul = List.fold_left - (fun dn (loc,(c,ctx),b,t) -> + (fun dn (loc,((c,ctx),b,t)) -> let sigma = Evd.merge_context_set Evd.univ_rigid sigma ctx in - let info = find_applied_relation false loc env sigma c b in + let info = find_applied_relation ?loc false env sigma c b in let pat = if b then info.hyp_left else info.hyp_right in let rul = { rew_lemma = c; rew_type = info.hyp_ty; rew_pat = pat; rew_ctx = ctx; rew_l2r = b; diff --git a/tactics/autorewrite.mli b/tactics/autorewrite.mli index 49e8588da..f765318d0 100644 --- a/tactics/autorewrite.mli +++ b/tactics/autorewrite.mli @@ -12,7 +12,7 @@ open Term open Equality (** Rewriting rules before tactic interpretation *) -type raw_rew_rule = Loc.t * constr Univ.in_universe_context_set * bool * Genarg.raw_generic_argument option +type raw_rew_rule = (constr Univ.in_universe_context_set * bool * Genarg.raw_generic_argument option) Loc.located (** To add rewriting rules to a base *) val add_rew_rules : string -> raw_rew_rule list -> unit @@ -56,7 +56,7 @@ type hypinfo = { hyp_right : constr; } -val find_applied_relation : bool -> - Loc.t -> +val find_applied_relation : + ?loc:Loc.t -> bool -> Environ.env -> Evd.evar_map -> Term.constr -> bool -> hypinfo diff --git a/tactics/hints.ml b/tactics/hints.ml index bcc068d3d..52e3dd09f 100644 --- a/tactics/hints.ml +++ b/tactics/hints.ml @@ -1282,7 +1282,7 @@ let interp_hints poly = prepare_hint true (poly,false) (Global.env()) Evd.empty (evd,c) in let fref r = let gr = global_with_alias r in - Dumpglob.add_glob (loc_of_reference r) gr; + Dumpglob.add_glob ?loc:(loc_of_reference r) gr; gr in let fr r = evaluable_of_global_reference (Global.env()) (fref r) @@ -1311,7 +1311,7 @@ let interp_hints poly = let constr_hints_of_ind qid = let ind = global_inductive_with_alias qid in let mib,_ = Global.lookup_inductive ind in - Dumpglob.dump_reference (fst (qualid_of_reference qid)) "<>" (string_of_reference qid) "ind"; + Dumpglob.dump_reference ?loc:(fst (qualid_of_reference qid)) "<>" (string_of_reference qid) "ind"; List.init (nconstructors ind) (fun i -> let c = (ind,i+1) in let gr = ConstructRef c in diff --git a/tactics/inv.ml b/tactics/inv.ml index f2147db40..a50ac31e1 100644 --- a/tactics/inv.ml +++ b/tactics/inv.ml @@ -283,9 +283,9 @@ let generalizeRewriteIntros as_mode tac depids id = end } let error_too_many_names pats = - let loc = Loc.merge (fst (List.hd pats)) (fst (List.last pats)) in + let loc = Loc.merge_opt (fst (List.hd pats)) (fst (List.last pats)) in Proofview.tclENV >>= fun env -> - tclZEROMSG ~loc ( + tclZEROMSG ?loc ( str "Unexpected " ++ str (String.plural (List.length pats) "introduction pattern") ++ str ": " ++ pr_enum (Miscprint.pr_intro_pattern (fun c -> Printer.pr_constr (EConstr.Unsafe.to_constr (fst (run_delayed env Evd.empty c))))) pats ++ diff --git a/tactics/tacticals.ml b/tactics/tacticals.ml index 074c8b9de..5c89331b8 100644 --- a/tactics/tacticals.ml +++ b/tactics/tacticals.ml @@ -187,7 +187,7 @@ let check_or_and_pattern_size ?loc check_and names branchsigns = | IntroAndPattern l -> if not (Int.equal n 1) then errn n; let l' = List.filter (function _,IntroForthcoming _ -> true | _,IntroNaming _ | _,IntroAction _ -> false) l in - if l' != [] then errforthcoming ~loc:(fst (List.hd l')); + if l' != [] then errforthcoming ?loc:(fst (List.hd l')); if check_and then let p1 = List.count (fun x -> x) branchsigns.(0) in let p2 = List.length branchsigns.(0) in @@ -221,7 +221,7 @@ let compute_induction_names_gen check_and branchletsigns = function Array.make (Array.length branchletsigns) [] | Some (loc,names) -> let names = fix_empty_or_and_pattern (Array.length branchletsigns) names in - get_and_check_or_and_pattern_gen check_and ~loc names branchletsigns + get_and_check_or_and_pattern_gen check_and ?loc names branchletsigns let compute_induction_names = compute_induction_names_gen true @@ -492,7 +492,7 @@ module New = struct | [] -> () | (evk,evi) :: _ -> let (loc,_) = evi.Evd.evar_source in - Pretype_errors.error_unsolvable_implicit ~loc env sigma evk None + Pretype_errors.error_unsolvable_implicit ?loc env sigma evk None let tclWITHHOLES accept_unresolved_holes tac sigma = tclEVARMAP >>= fun sigma_initial -> diff --git a/tactics/tactics.ml b/tactics/tactics.ml index 7dad90242..d265269b8 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -443,7 +443,7 @@ let find_name mayrepl decl naming gl = match naming with let ids_of_hyps = Tacmach.New.pf_ids_of_hyps gl in let id' = next_ident_away id ids_of_hyps in if not mayrepl && not (Id.equal id' id) then - user_err ~loc (pr_id id ++ str" is already used."); + user_err ?loc (pr_id id ++ str" is already used."); id (**************************************************************) @@ -1734,7 +1734,7 @@ let general_apply with_delta with_destruct with_evars clear_flag (loc,(c,lbind : with Redelimination -> (* Last chance: if the head is a variable, apply may try second order unification *) - let info = Loc.add_loc info loc in + let info = Option.cata (fun loc -> Loc.add_loc info loc) info loc in let tac = if with_destruct then descend_in_conjunctions [] @@ -1838,8 +1838,8 @@ let progress_with_clause flags innerclause clause = try List.find_map f ordered_metas with Not_found -> raise UnableToApply -let explain_unable_to_apply_lemma loc env sigma thm innerclause = - user_err ~loc (hov 0 +let explain_unable_to_apply_lemma ?loc env sigma thm innerclause = + user_err ?loc (hov 0 (Pp.str "Unable to apply lemma of type" ++ brk(1,1) ++ Pp.quote (Printer.pr_leconstr_env env sigma thm) ++ spc() ++ str "on hypothesis of type" ++ brk(1,1) ++ @@ -1855,7 +1855,7 @@ let apply_in_once_main flags innerclause env sigma (loc,d,lbind) = try aux (clenv_push_prod clause) with NotExtensibleClause -> match e with - | UnableToApply -> explain_unable_to_apply_lemma loc env sigma thm innerclause + | UnableToApply -> explain_unable_to_apply_lemma ?loc env sigma thm innerclause | _ -> iraise e' in aux (make_clenv_binding env sigma (d,thm) lbind) @@ -2281,7 +2281,7 @@ let error_unexpected_extra_pattern loc bound pat = | IntroNaming (IntroIdentifier _) -> "name", (String.plural nb " introduction pattern"), "no" | _ -> "introduction pattern", "", "none" in - user_err ~loc (str "Unexpected " ++ str s1 ++ str " (" ++ + user_err ?loc (str "Unexpected " ++ str s1 ++ str " (" ++ (if Int.equal nb 0 then (str s3 ++ str s2) else (str "at most " ++ int nb ++ str s2)) ++ spc () ++ str (if Int.equal nb 1 then "was" else "were") ++ @@ -2461,7 +2461,7 @@ let rec intro_patterns_core with_evars b avoid ids thin destopt bound n tac = | IntroAction pat -> intro_then_gen (make_tmp_naming avoid l pat) destopt true false - (intro_pattern_action ~loc with_evars (b || not (List.is_empty l)) false + (intro_pattern_action ?loc with_evars (b || not (List.is_empty l)) false pat thin destopt (fun thin bound' -> intro_patterns_core with_evars b avoid ids thin destopt bound' 0 (fun ids thin -> @@ -2498,9 +2498,9 @@ and intro_pattern_action ?loc with_evars b style pat thin destopt tac id = rewrite_hyp_then style with_evars thin l2r id (fun thin -> tac thin None []) | IntroApplyOn ((loc',f),(loc,pat)) -> let naming,tac_ipat = - prepare_intros ~loc with_evars (IntroIdentifier id) destopt pat in + prepare_intros ?loc with_evars (IntroIdentifier id) destopt pat in let doclear = - if naming = NamingMustBe (Loc.tag ~loc id) then + if naming = NamingMustBe (Loc.tag ?loc id) then Proofview.tclUNIT () (* apply_in_once do a replacement *) else clear [id] in @@ -2549,7 +2549,7 @@ let intros_patterns with_evars = function let prepare_intros_opt with_evars dft destopt = function | None -> prepare_naming dft, (fun _id -> Proofview.tclUNIT ()) - | Some (loc,ipat) -> prepare_intros ~loc with_evars dft destopt ipat + | Some (loc,ipat) -> prepare_intros ?loc with_evars dft destopt ipat let ipat_of_name = function | Anonymous -> None @@ -2692,7 +2692,7 @@ let mkletin_goal env sigma store with_eq dep (id,lastlhyp,ccl,c) ty = | IntroFresh heq_base -> fresh_id_in_env [id] heq_base env | IntroIdentifier id -> if List.mem id (ids_of_named_context (named_context env)) then - user_err ~loc (pr_id id ++ str" is already used."); + user_err ?loc (pr_id id ++ str" is already used."); id in let eqdata = build_coq_eq_data () in let args = if lr then [t;mkVar id;c] else [t;c;mkVar id]in @@ -2973,7 +2973,7 @@ let specialize (c,lbind) ipat = | Some (loc,ipat) -> (* Like pose proof with extra support for "with" bindings *) (* even though the "with" bindings forces full application *) - let naming,tac = prepare_intros ~loc false IntroAnonymous MoveLast ipat in + let naming,tac = prepare_intros ?loc false IntroAnonymous MoveLast ipat in Tacticals.New.tclTHENFIRST (assert_before_then_gen false naming typ tac) (exact_no_check term) @@ -4467,7 +4467,7 @@ let induction_gen_l isrec with_evars elim names lc = let lc = List.map (function | (c,None) -> c | (c,Some(loc,eqname)) -> - user_err ~loc (str "Do not know what to do with " ++ + user_err ?loc (str "Do not know what to do with " ++ Miscprint.pr_intro_pattern_naming eqname)) lc in let rec atomize_list l = match l with diff --git a/toplevel/coqloop.ml b/toplevel/coqloop.ml index 4641a2bc8..ab360f98d 100644 --- a/toplevel/coqloop.ml +++ b/toplevel/coqloop.ml @@ -146,13 +146,10 @@ let print_highlight_location ib loc = highlight_lines let valid_buffer_loc ib loc = - not (Loc.is_ghost loc) && let (b,e) = Loc.unloc loc in b-ib.start >= 0 && e-ib.start < ib.len && b<=e (* This is specific to the toplevel *) -let pr_loc loc = - if Loc.is_ghost loc then str"" - else +let pr_loc ?loc = Option.default (fun loc -> let fname = loc.Loc.fname in if CString.equal fname "" then Loc.(str"Toplevel input, characters " ++ int loc.bp ++ @@ -162,6 +159,7 @@ let pr_loc loc = str", line " ++ int loc.line_nb ++ str", characters " ++ int (loc.bp-loc.bol_pos) ++ str"-" ++ int (loc.ep-loc.bol_pos) ++ str":") + ) loc (* Toplevel error explanation. *) let error_info_for_buffer ?loc buf = diff --git a/toplevel/vernac.ml b/toplevel/vernac.ml index d6bcd2f15..3a67f4cbf 100644 --- a/toplevel/vernac.ml +++ b/toplevel/vernac.ml @@ -21,11 +21,12 @@ open Vernacprop let checknav_simple (loc, cmd) = if is_navigation_vernac cmd && not (is_reset cmd) then - CErrors.user_err ~loc (str "Navigation commands forbidden in files.") + CErrors.user_err ?loc (str "Navigation commands forbidden in files.") let checknav_deep (loc, ast) = if is_deep_navigation_vernac ast then - CErrors.user_err ~loc (str "Navigation commands forbidden in nested commands.") + CErrors.user_err ?loc (str "Navigation commands forbidden in nested commands.") + let disable_drop = function | Drop -> CErrors.error "Drop is forbidden." @@ -33,10 +34,12 @@ let disable_drop = function (* Echo from a buffer based on position. XXX: Should move to utility file. *) -let vernac_echo loc in_chan = let open Loc in - let len = loc.ep - loc.bp in - seek_in in_chan loc.bp; - Feedback.msg_notice @@ str @@ really_input_string in_chan len +let vernac_echo ?loc in_chan = let open Loc in + Option.iter (fun loc -> + let len = loc.ep - loc.bp in + seek_in in_chan loc.bp; + 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 *) @@ -49,8 +52,8 @@ let set_formatter_translator ch = Format.set_formatter_output_functions out (fun () -> flush ch); Format.set_max_boxes max_int -let pr_new_syntax_in_context loc chan_beautify ocom = - let loc = Loc.unloc loc in +let pr_new_syntax_in_context ?loc chan_beautify ocom = + let loc = Option.cata Loc.unloc (0,0) loc in if !beautify_file then set_formatter_translator chan_beautify; let fs = States.freeze ~marshallable:`No in (* The content of this is not supposed to fail, but if ever *) @@ -72,14 +75,14 @@ let pr_new_syntax_in_context loc chan_beautify ocom = States.unfreeze fs; Format.set_formatter_out_channel stdout -let pr_new_syntax po loc chan_beautify ocom = +let pr_new_syntax ?loc po chan_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 chan_beautify loc) ocom + Pcoq.Gram.with_parsable po (pr_new_syntax_in_context ?loc chan_beautify) ocom (* For coqtop -time, we display the position in the file, and a glimpse of the executed command *) -let pp_cmd_header loc com = +let pp_cmd_header ?loc com = let shorten s = try (String.sub s 0 30)^"..." with _ -> s in let noblank s = String.map (fun c -> match c with @@ -87,7 +90,7 @@ let pp_cmd_header loc com = | x -> x ) s in - let (start,stop) = Loc.unloc loc in + let (start,stop) = Option.cata Loc.unloc (0,0) loc in let safe_pr_vernac x = try Ppvernac.pr_vernac x with e -> str (Printexc.to_string e) in @@ -98,9 +101,8 @@ let pp_cmd_header loc com = (* This is a special case where we assume we are in console batch mode and take control of the console. *) -(* FIXME *) -let print_cmd_header loc com = - Pp.pp_with !Topfmt.std_ft (pp_cmd_header loc com); +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 () = @@ -141,14 +143,14 @@ let rec interp_vernac sid po (loc,com) = try (* The -time option is only supported from console-based clients due to the way it prints. *) - if !Flags.time then print_cmd_header loc com; + if !Flags.time then print_cmd_header ?loc com; let com = if !Flags.time then VernacTime (loc,com) else com in interp com with reraise -> let (reraise, info) = CErrors.push reraise in let info = begin match Loc.get_loc info with - | None -> Loc.add_loc info loc + | None -> Option.cata (Loc.add_loc info) info loc | Some _ -> info end in iraise (reraise, info) @@ -177,8 +179,8 @@ and load_vernac verbosely sid file = in (* Printing of vernacs *) - if !beautify then pr_new_syntax in_pa chan_beautify loc (Some ast); - Option.iter (vernac_echo loc) in_echo; + if !beautify then pr_new_syntax ?loc in_pa chan_beautify (Some ast); + Option.iter (vernac_echo ?loc) in_echo; checknav_simple (loc, ast); let nsid = Flags.silently (interp_vernac !rsid in_pa) (loc, ast) in @@ -193,7 +195,7 @@ and load_vernac verbosely sid file = | Stm.End_of_input -> (* Is this called so comments at EOF are printed? *) if !beautify then - pr_new_syntax in_pa chan_beautify (Loc.make_loc (max_int,max_int)) None; + pr_new_syntax ~loc:(Loc.make_loc (max_int,max_int)) in_pa chan_beautify None; if !Flags.beautify_file then close_out chan_beautify; !rsid | reraise -> diff --git a/vernac/classes.ml b/vernac/classes.ml index bdaa3ece6..fb300dbc1 100644 --- a/vernac/classes.ml +++ b/vernac/classes.ml @@ -77,7 +77,7 @@ let existing_instance glob g info = match class_of_constr Evd.empty (EConstr.of_constr r) with | Some (_, ((tc,u), _)) -> add_instance (new_instance tc info glob (*FIXME*) (Flags.use_polymorphic_flag ()) c) - | None -> user_err ~loc:(loc_of_reference g) + | None -> user_err ?loc:(loc_of_reference g) ~hdr:"declare_instance" (Pp.str "Constant does not build instances of a declared type class.") @@ -237,7 +237,7 @@ let new_instance ?(abstract=false) ?(global=false) ?(refine= !refine_instance) p let get_id = function | Ident id' -> id' - | Qualid (loc,id') -> (loc, snd (repr_qualid id')) + | Qualid (loc,id') -> (Loc.tag ?loc @@ snd (repr_qualid id')) in let props, rest = List.fold_left @@ -257,7 +257,7 @@ let new_instance ?(abstract=false) ?(global=false) ?(refine= !refine_instance) p let (loc, mid) = get_id loc_mid in List.iter (fun (n, _, x) -> if Name.equal n (Name mid) then - Option.iter (fun x -> Dumpglob.add_glob loc (ConstRef x)) x) + Option.iter (fun x -> Dumpglob.add_glob ?loc (ConstRef x)) x) k.cl_projs; c :: props, rest' with Not_found -> diff --git a/vernac/command.ml b/vernac/command.ml index fbaa09430..82d7b19d7 100644 --- a/vernac/command.ml +++ b/vernac/command.ml @@ -53,16 +53,16 @@ let rec under_binders env sigma f n c = mkLetIn (x,b,t,under_binders (push_rel (LocalDef (x,b,t)) env) sigma f (n-1) c) | _ -> assert false -let rec complete_conclusion a cs = Loc.map_with_loc (fun ~loc -> function +let rec complete_conclusion a cs = Loc.map_with_loc (fun ?loc -> function | CProdN (bl,c) -> CProdN (bl,complete_conclusion a cs c) | CLetIn (na,b,t,c) -> CLetIn (na,b,t,complete_conclusion a cs c) | CHole (k, _, _) -> let (has_no_args,name,params) = a in if not has_no_args then - user_err ~loc + user_err ?loc (strbrk"Cannot infer the non constant arguments of the conclusion of " ++ pr_id cs ++ str "."); - let args = List.map (fun id -> Loc.tag ~loc @@ CRef(Ident(loc,id),None)) params in + let args = List.map (fun id -> Loc.tag ?loc @@ CRef(Ident(loc,id),None)) params in CAppExpl ((None,Ident(loc,name),None),List.rev args) | c -> c ) @@ -344,7 +344,7 @@ let do_assumptions kind nl l = match l with | (Discharge, _, _) when Lib.sections_are_opened () -> let loc = fst id in let msg = Pp.str "Section variables cannot be polymorphic." in - user_err ~loc msg + user_err ?loc msg | _ -> () in do_assumptions_bound_univs coe kind nl id (Some pl) c @@ -356,7 +356,7 @@ let do_assumptions kind nl l = match l with let loc = fst id in let msg = Pp.str "Assumptions with bound universes can only be defined one at a time." in - user_err ~loc msg + user_err ?loc msg in (coe, (List.map map idl, c)) in @@ -452,7 +452,7 @@ let interp_ind_arity env evdref ind = let t, impls = understand_tcc_evars env evdref ~expected_type:IsType c, imps in let pseudo_poly = check_anonymous_type c in let () = if not (Reductionops.is_arity env !evdref t) then - user_err ~loc:(constr_loc ind.ind_arity) (str "Not an arity") + user_err ?loc:(constr_loc ind.ind_arity) (str "Not an arity") in let t = EConstr.Unsafe.to_constr t in t, pseudo_poly, impls @@ -566,7 +566,7 @@ let check_named (loc, na) = match na with | Name _ -> () | Anonymous -> let msg = str "Parameters must be named." in - user_err ~loc msg + user_err ?loc msg let check_param = function @@ -982,7 +982,7 @@ let build_wellfounded (recname,pl,n,bl,arityc,body) poly r measure notation = let relty = Typing.unsafe_type_of env !evdref rel in let relargty = let error () = - user_err ~loc:(constr_loc r) + user_err ?loc:(constr_loc r) ~hdr:"Command.build_wellfounded" (Printer.pr_econstr_env env !evdref rel ++ str " is not an homogeneous binary relation.") in diff --git a/vernac/explainErr.ml b/vernac/explainErr.ml index ed2dd274a..040c86805 100644 --- a/vernac/explainErr.ml +++ b/vernac/explainErr.ml @@ -117,8 +117,9 @@ let process_vernac_interp_error ?(allow_uncaught=true) (exc, info) = try Some (CList.find_map (fun f -> f e) !additional_error_info) with _ -> None in + let add_loc_opt ?loc info = Option.cata (fun l -> Loc.add_loc info l) info loc in match e' with | None -> e - | Some (loc, None) -> (fst e, Loc.add_loc (snd e) loc) + | Some (loc, None) -> (fst e, add_loc_opt ?loc (snd e)) | Some (loc, Some msg) -> - (EvaluatedError (msg, Some (fst e)), Loc.add_loc (snd e) loc) + (EvaluatedError (msg, Some (fst e)), add_loc_opt ?loc (snd e)) diff --git a/vernac/lemmas.ml b/vernac/lemmas.ml index b51792b1e..c6f9f21d8 100644 --- a/vernac/lemmas.ml +++ b/vernac/lemmas.ml @@ -224,7 +224,7 @@ let compute_proof_name locality = function if Nametab.exists_cci (Lib.make_path id) || is_section_variable id || locality == Global && Nametab.exists_cci (Lib.make_path_except_section id) then - user_err ~loc (pr_id id ++ str " already exists."); + user_err ?loc (pr_id id ++ str " already exists."); id, pl | None -> next_global_ident_away default_thm_id (Pfedit.get_all_proof_names ()), None @@ -337,7 +337,7 @@ let get_proof proof do_guard hook opacity = let check_exist = List.iter (fun (loc,id) -> if not (Nametab.exists_cci (Lib.make_path id)) then - user_err ~loc (pr_id id ++ str " does not exist.") + user_err ?loc (pr_id id ++ str " does not exist.") ) let universe_proof_terminator compute_guard hook = diff --git a/vernac/metasyntax.ml b/vernac/metasyntax.ml index a586fe9f2..feacfe915 100644 --- a/vernac/metasyntax.ml +++ b/vernac/metasyntax.ml @@ -192,7 +192,7 @@ let parse_format ((loc, str) : lstring) = error "Empty format." with reraise -> let (e, info) = CErrors.push reraise in - let info = Loc.add_loc info loc in + let info = Option.cata (Loc.add_loc info) info loc in iraise (e, info) (***********************) diff --git a/vernac/obligations.ml b/vernac/obligations.ml index 5e65855ef..f0883e286 100644 --- a/vernac/obligations.ml +++ b/vernac/obligations.ml @@ -38,7 +38,7 @@ let check_evars env evm = | Evar_kinds.QuestionMark _ | Evar_kinds.ImplicitArg (_,_,false) -> () | _ -> - Pretype_errors.error_unsolvable_implicit ~loc env evm key None) + Pretype_errors.error_unsolvable_implicit ?loc env evm key None) (Evd.undefined_map evm) type oblinfo = @@ -991,7 +991,7 @@ and solve_obligation_by_tac prg obls i tac = let (e, _) = CErrors.push e in match e with | Refiner.FailError (_, s) -> - user_err ~loc:(fst obl.obl_location) ~hdr:"solve_obligation" (Lazy.force s) + user_err ?loc:(fst obl.obl_location) ~hdr:"solve_obligation" (Lazy.force s) | e -> None (* FIXME really ? *) and solve_prg_obligations prg ?oblset tac = diff --git a/vernac/record.ml b/vernac/record.ml index 95f5ad7cc..8bd583b81 100644 --- a/vernac/record.ml +++ b/vernac/record.ml @@ -94,7 +94,7 @@ let compute_constructor_level evars env l = let binder_of_decl = function | Vernacexpr.AssumExpr(n,t) -> (n,None,t) | Vernacexpr.DefExpr(n,c,t) -> (n,Some c, match t with Some c -> c - | None -> Loc.tag ~loc:(fst n) @@ CHole (None, Misctypes.IntroAnonymous, None)) + | None -> Loc.tag ?loc:(fst n) @@ CHole (None, Misctypes.IntroAnonymous, None)) let binders_of_decls = List.map binder_of_decl @@ -106,14 +106,14 @@ let typecheck_params_and_fields def id pl t ps nots fs = let error bk (loc, name) = match bk, name with | Default _, Anonymous -> - user_err ~loc ~hdr:"record" (str "Record parameters must be named") + user_err ?loc ~hdr:"record" (str "Record parameters must be named") | _ -> () in List.iter (function CLocalDef (b, _, _) -> error default_binder_kind b | CLocalAssum (ls, bk, ce) -> List.iter (error bk) ls | CLocalPattern (loc,(_,_)) -> - Loc.raise ~loc (Stream.Error "pattern with quote not allowed in record parameters.")) ps + Loc.raise ?loc (Stream.Error "pattern with quote not allowed in record parameters.")) ps in let impls_env, ((env1,newps), imps) = interp_context_evars env0 evars ps in let newps = List.map (fun d -> Termops.map_rel_decl EConstr.Unsafe.to_constr d) newps in @@ -135,7 +135,7 @@ let typecheck_params_and_fields def id pl t ps nots fs = sred, true | None -> s, false else s, false) - | _ -> user_err ~loc:(constr_loc t) (str"Sort expected.")) + | _ -> user_err ?loc:(constr_loc t) (str"Sort expected.")) | None -> let uvarkind = Evd.univ_flexible_alg in mkSort (Evarutil.evd_comb0 (Evd.new_sort_variable uvarkind) evars), true diff --git a/vernac/topfmt.ml b/vernac/topfmt.ml index c25dd55fb..2f01cfb54 100644 --- a/vernac/topfmt.ml +++ b/vernac/topfmt.ml @@ -108,7 +108,7 @@ end type logger = ?loc:Loc.t -> level -> std_ppcmds -> unit -let msgnl_with ?pre_hdr fmt strm = +let msgnl_with fmt strm = pp_with fmt (strm ++ fnl ()); Format.pp_print_flush fmt () diff --git a/vernac/vernacentries.ml b/vernac/vernacentries.ml index 75f403e33..a78350749 100644 --- a/vernac/vernacentries.ml +++ b/vernac/vernacentries.ml @@ -364,43 +364,43 @@ let msg_found_library = function Feedback.msg_info (hov 0 (pr_dirpath fulldir ++ strbrk " is bound to file " ++ str file)) -let err_unmapped_library loc ?from qid = +let err_unmapped_library ?loc ?from qid = let dir = fst (repr_qualid qid) in let prefix = match from with | None -> str "." | Some from -> str " and prefix " ++ pr_dirpath from ++ str "." in - user_err ~loc + user_err ?loc ~hdr:"locate_library" (strbrk "Cannot find a physical path bound to logical path matching suffix " ++ pr_dirpath dir ++ prefix) -let err_notfound_library loc ?from qid = +let err_notfound_library ?loc ?from qid = let prefix = match from with | None -> str "." | Some from -> str " with prefix " ++ pr_dirpath from ++ str "." in - user_err ~loc ~hdr:"locate_library" + user_err ?loc ~hdr:"locate_library" (strbrk "Unable to locate library " ++ pr_qualid qid ++ prefix) let print_located_library r = let (loc,qid) = qualid_of_reference r in try msg_found_library (Library.locate_qualified_library ~warn:false qid) with - | Library.LibUnmappedDir -> err_unmapped_library loc qid - | Library.LibNotFound -> err_notfound_library loc qid + | Library.LibUnmappedDir -> err_unmapped_library ?loc qid + | Library.LibNotFound -> err_notfound_library ?loc qid let smart_global r = let gr = Smartlocate.smart_global r in - Dumpglob.add_glob (Stdarg.loc_of_or_by_notation loc_of_reference r) gr; + Dumpglob.add_glob ?loc:(Stdarg.loc_of_or_by_notation loc_of_reference r) gr; gr let dump_global r = try let gr = Smartlocate.smart_global r in - Dumpglob.add_glob (Stdarg.loc_of_or_by_notation loc_of_reference r) gr + Dumpglob.add_glob ?loc:(Stdarg.loc_of_or_by_notation loc_of_reference r) gr with e when CErrors.noncritical e -> () (**********) (* Syntax *) @@ -608,14 +608,14 @@ let vernac_combined_scheme lid l = let vernac_universe loc poly l = if poly && not (Lib.sections_are_opened ()) then - user_err ~loc ~hdr:"vernac_universe" + user_err ?loc ~hdr:"vernac_universe" (str"Polymorphic universes can only be declared inside sections, " ++ str "use Monomorphic Universe instead"); do_universe poly l let vernac_constraint loc poly l = if poly && not (Lib.sections_are_opened ()) then - user_err ~loc ~hdr:"vernac_constraint" + user_err ?loc ~hdr:"vernac_constraint" (str"Polymorphic universe constraints can only be declared" ++ str " inside sections, use Monomorphic Constraint instead"); do_constraint poly l @@ -641,7 +641,7 @@ let vernac_declare_module export (loc, id) binders_ast mty_ast = Declaremods.declare_module Modintern.interp_module_ast id binders_ast (Enforce mty_ast) [] in - Dumpglob.dump_moddef loc mp "mod"; + Dumpglob.dump_moddef ?loc mp "mod"; if_verbose Feedback.msg_info (str "Module " ++ pr_id id ++ str " is declared"); Option.iter (fun export -> vernac_import export [Ident (Loc.tag id)]) export @@ -662,7 +662,7 @@ let vernac_define_module export (loc, id) binders_ast mty_ast_o mexpr_ast_l = Declaremods.start_module Modintern.interp_module_ast export id binders_ast mty_ast_o in - Dumpglob.dump_moddef loc mp "mod"; + Dumpglob.dump_moddef ?loc mp "mod"; if_verbose Feedback.msg_info (str "Interactive Module " ++ pr_id id ++ str " started"); List.iter @@ -680,7 +680,7 @@ let vernac_define_module export (loc, id) binders_ast mty_ast_o mexpr_ast_l = Declaremods.declare_module Modintern.interp_module_ast id binders_ast mty_ast_o mexpr_ast_l in - Dumpglob.dump_moddef loc mp "mod"; + Dumpglob.dump_moddef ?loc mp "mod"; if_verbose Feedback.msg_info (str "Module " ++ pr_id id ++ str " is defined"); Option.iter (fun export -> vernac_import export [Ident (Loc.tag id)]) @@ -688,7 +688,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"; + Dumpglob.dump_modref ?loc mp "mod"; if_verbose Feedback.msg_info (str "Module " ++ pr_id id ++ str " is defined"); Option.iter (fun export -> vernac_import export [Ident lid]) export @@ -709,7 +709,7 @@ let vernac_declare_module_type (loc,id) binders_ast mty_sign mty_ast_l = Declaremods.start_modtype Modintern.interp_module_ast id binders_ast mty_sign in - Dumpglob.dump_moddef loc mp "modtype"; + Dumpglob.dump_moddef ?loc mp "modtype"; if_verbose Feedback.msg_info (str "Interactive Module Type " ++ pr_id id ++ str " started"); List.iter @@ -728,13 +728,13 @@ let vernac_declare_module_type (loc,id) binders_ast mty_sign mty_ast_l = Declaremods.declare_modtype Modintern.interp_module_ast id binders_ast mty_sign mty_ast_l in - Dumpglob.dump_moddef loc mp "modtype"; + Dumpglob.dump_moddef ?loc mp "modtype"; 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"; + Dumpglob.dump_modref ?loc mp "modtype"; if_verbose Feedback.msg_info (str "Module Type " ++ pr_id id ++ str " is defined") let vernac_include l = @@ -751,7 +751,7 @@ let vernac_begin_section (_, id as lid) = Lib.open_section id let vernac_end_section (loc,_) = - Dumpglob.dump_reference loc + Dumpglob.dump_reference ?loc (DirPath.to_string (Lib.current_dirpath true)) "<>" "sec"; Lib.close_section () @@ -784,12 +784,12 @@ let vernac_require from import qidl = let (_, dir, f) = Library.locate_qualified_library ?root ~warn qid in (dir, f) with - | Library.LibUnmappedDir -> err_unmapped_library loc ?from:root qid - | Library.LibNotFound -> err_notfound_library loc ?from:root qid + | Library.LibUnmappedDir -> err_unmapped_library ?loc ?from:root qid + | Library.LibNotFound -> err_notfound_library ?loc ?from:root qid in let modrefl = List.map locate qidl in if Dumpglob.dump () then - List.iter2 (fun (loc, _) dp -> Dumpglob.dump_libref loc dp "lib") qidl (List.map fst modrefl); + List.iter2 (fun (loc, _) dp -> Dumpglob.dump_libref ?loc dp "lib") qidl (List.map fst modrefl); Library.require_library_from_dirpath modrefl import (* Coercions and canonical structures *) @@ -1179,7 +1179,7 @@ let vernac_arguments locality reference args more_implicits nargs_for_red flags let scopes = List.map (Option.map (fun (loc,k) -> try ignore (Notation.find_scope k); k with UserError _ -> - Notation.find_delimiters_scope ~loc k)) scopes + Notation.find_delimiters_scope ?loc k)) scopes in vernac_arguments_scope locality reference scopes end; @@ -1533,14 +1533,14 @@ let get_current_context_of_args = function | Some n -> get_goal_context n | None -> get_current_context () -let query_command_selector ~loc = function +let query_command_selector ?loc = function | None -> None | Some (SelectNth n) -> Some n - | _ -> user_err ~loc ~hdr:"query_command_selector" + | _ -> user_err ?loc ~hdr:"query_command_selector" (str "Query commands only support the single numbered goal selector.") -let vernac_check_may_eval ~loc redexp glopt rc = - let glopt = query_command_selector ~loc glopt in +let vernac_check_may_eval ?loc redexp glopt rc = + let glopt = query_command_selector ?loc glopt in let (sigma, env) = get_current_context_of_args glopt in let sigma', c = interp_open_constr env sigma rc in let c = EConstr.Unsafe.to_constr c in @@ -1602,10 +1602,10 @@ exception NoHyp (* Printing "About" information of a hypothesis of the current goal. 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 ~loc ref_or_by_not glopt = +let print_about_hyp_globs ?loc ref_or_by_not glopt = let open Context.Named.Declaration in try - let glnumopt = query_command_selector ~loc glopt in + let glnumopt = query_command_selector ?loc glopt in let gl,id = match glnumopt,ref_or_by_not with | None,AN (Ident (_loc,id)) -> (* goal number not given, catch any failure *) @@ -1613,7 +1613,7 @@ let print_about_hyp_globs ~loc ref_or_by_not glopt = | Some n,AN (Ident (_loc,id)) -> (* goal number given, catch if wong *) (try get_nth_goal n,id with - Failure _ -> user_err ~loc ~hdr:"print_about_hyp_globs" + Failure _ -> user_err ?loc ~hdr:"print_about_hyp_globs" (str "No such goal: " ++ int n ++ str ".")) | _ , _ -> raise NoHyp in let hyps = pf_hyps gl in @@ -1627,7 +1627,7 @@ let print_about_hyp_globs ~loc ref_or_by_not glopt = | NoHyp | Not_found -> print_about ref_or_by_not -let vernac_print ~loc = let open Feedback in function +let vernac_print ?loc = 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) @@ -1672,7 +1672,7 @@ let vernac_print ~loc = let open Feedback in function | PrintVisibility s -> msg_notice (Notation.pr_visibility (Constrextern.without_symbols pr_lglob_constr) s) | PrintAbout (ref_or_by_not,glnumopt) -> - msg_notice (print_about_hyp_globs ~loc ref_or_by_not glnumopt) + msg_notice (print_about_hyp_globs ?loc ref_or_by_not glnumopt) | PrintImplicit qid -> dump_global qid; msg_notice (print_impargs qid) | PrintAssumptions (o,t,r) -> @@ -1689,7 +1689,7 @@ let global_module r = let (loc,qid) = qualid_of_reference r in try Nametab.full_name_module qid with Not_found -> - user_err ~loc ~hdr:"global_module" + user_err ?loc ~hdr:"global_module" (str "Module/section " ++ pr_qualid qid ++ str " not found.") let interp_search_restriction = function @@ -1737,8 +1737,8 @@ let _ = optread = (fun () -> !search_output_name_only); optwrite = (:=) search_output_name_only } -let vernac_search ~loc s gopt r = - let gopt = query_command_selector ~loc gopt in +let vernac_search ?loc s gopt r = + let gopt = query_command_selector ?loc gopt in let r = interp_search_restriction r in let env,gopt = match gopt with | None -> @@ -1913,7 +1913,7 @@ let vernac_load interp fname = * 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 = +let interp ?proof ?loc locality poly c = vernac_pperr_endline (fun () -> str "interpreting: " ++ Ppvernac.pr_vernac c); match c with (* The below vernac are candidates for removal from the main type @@ -2045,11 +2045,11 @@ let interp ?proof ~loc locality poly c = | VernacAddOption (key,v) -> vernac_add_option key v | VernacMemOption (key,v) -> vernac_mem_option key v | VernacPrintOption key -> vernac_print_option key - | VernacCheckMayEval (r,g,c) -> vernac_check_may_eval ~loc r g c + | VernacCheckMayEval (r,g,c) -> vernac_check_may_eval ?loc r g c | VernacDeclareReduction (s,r) -> vernac_declare_reduction locality s r | VernacGlobalCheck c -> vernac_global_check c - | VernacPrint p -> vernac_print ~loc p - | VernacSearch (s,g,r) -> vernac_search ~loc s g r + | VernacPrint p -> vernac_print ?loc p + | VernacSearch (s,g,r) -> vernac_search ?loc s g r | VernacLocate l -> vernac_locate l | VernacRegister (id, r) -> vernac_register id r | VernacComments l -> if_verbose Feedback.msg_info (str "Comments ok\n") @@ -2065,15 +2065,15 @@ let interp ?proof ~loc locality poly c = | VernacShow s -> vernac_show s | VernacCheckGuard -> vernac_check_guard () | VernacProof (None, None) -> - Aux_file.record_in_aux_at ~loc "VernacProof" "tac:no using:no" + Aux_file.record_in_aux_at ?loc "VernacProof" "tac:no using:no" | VernacProof (Some tac, None) -> - Aux_file.record_in_aux_at ~loc "VernacProof" "tac:yes using:no"; + Aux_file.record_in_aux_at ?loc "VernacProof" "tac:yes using:no"; vernac_set_end_tac tac | VernacProof (None, Some l) -> - Aux_file.record_in_aux_at ~loc "VernacProof" "tac:no using:yes"; + Aux_file.record_in_aux_at ?loc "VernacProof" "tac:no using:yes"; vernac_set_used_variables l | VernacProof (Some tac, Some l) -> - Aux_file.record_in_aux_at ~loc "VernacProof" "tac:yes using:yes"; + Aux_file.record_in_aux_at ?loc "VernacProof" "tac:yes using:yes"; vernac_set_end_tac tac; vernac_set_used_variables l | VernacProofMode mn -> Proof_global.set_proof_mode mn @@ -2149,10 +2149,10 @@ let vernac_timeout f = let restore_timeout () = current_timeout := None -let locate_if_not_already loc (e, info) = +let locate_if_not_already ?loc (e, info) = match Loc.get_loc info with - | None -> (e, Loc.add_loc info loc) - | Some l -> if Loc.is_ghost l then (e, Loc.add_loc info loc) else (e, info) + | None -> (e, Option.cata (Loc.add_loc info) info loc) + | Some l -> (e, info) exception HasNotFailed exception HasFailed of std_ppcmds @@ -2219,8 +2219,8 @@ let interp ?(verbosely=true) ?proof (loc,c) = try vernac_timeout begin fun () -> if verbosely - then Flags.verbosely (interp ?proof ~loc locality poly) c - else Flags.silently (interp ?proof ~loc locality poly) c; + then Flags.verbosely (interp ?proof ?loc locality poly) c + else Flags.silently (interp ?proof ?loc locality poly) c; if orig_program_mode || not !Flags.program_mode || isprogcmd then Flags.program_mode := orig_program_mode; ignore (Flags.use_polymorphic_flag ()) @@ -2232,7 +2232,7 @@ let interp ?(verbosely=true) ?proof (loc,c) = | e -> CErrors.noncritical e) -> let e = CErrors.push reraise in - let e = locate_if_not_already loc e 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 ()); -- cgit v1.2.3