From 30d3515546cf244837c6340b6b87c5f51e68cbf4 Mon Sep 17 00:00:00 2001 From: Emilio Jesus Gallego Arias Date: Tue, 17 Jan 2017 23:40:35 +0100 Subject: [location] Remove Loc.ghost. Now it is a private field, locations are optional. --- .merlin | 2 + checker/modops.ml | 6 +- dev/top_printers.ml | 30 +-- engine/evarutil.ml | 2 +- engine/evd.ml | 12 +- engine/evd.mli | 2 +- engine/proofview.ml | 2 +- engine/uState.ml | 6 +- grammar/argextend.mlp | 2 +- grammar/tacextend.mlp | 6 +- ide/coqOps.ml | 22 +- ide/ide_slave.ml | 4 +- ide/texmacspp.ml | 449 ++++++++++++++++---------------- ide/texmacspp.mli | 2 +- ide/xmlprotocol.ml | 4 +- interp/constrexpr_ops.ml | 45 ++-- interp/constrexpr_ops.mli | 6 +- interp/constrextern.ml | 112 ++++---- interp/constrextern.mli | 6 +- interp/constrintern.ml | 99 +++---- interp/implicit_quantifiers.ml | 2 +- interp/notation.ml | 64 ++--- interp/notation.mli | 12 +- interp/notation_ops.ml | 74 +++--- interp/notation_ops.mli | 4 +- interp/smartlocate.ml | 4 +- interp/topconstr.ml | 2 +- intf/glob_term.mli | 9 +- lib/aux_file.ml | 20 +- lib/aux_file.mli | 8 +- lib/cWarnings.ml | 19 +- lib/feedback.ml | 2 +- lib/feedback.mli | 2 +- lib/loc.ml | 17 +- lib/loc.mli | 17 +- lib/stateid.ml | 2 +- lib/stateid.mli | 2 +- library/declare.ml | 4 +- parsing/g_constr.ml4 | 10 +- parsing/g_vernac.ml4 | 16 +- plugins/firstorder/g_ground.ml4 | 2 +- plugins/funind/glob_term_to_relation.ml | 22 +- plugins/funind/indfun.ml | 22 +- plugins/funind/indfun_common.ml | 2 +- plugins/funind/invfun.ml | 2 +- plugins/funind/merge.ml | 12 +- plugins/funind/recdef.ml | 4 +- plugins/ltac/coretactics.ml4 | 5 +- plugins/ltac/evar_tactics.ml | 2 +- plugins/ltac/extraargs.ml4 | 6 +- plugins/ltac/extratactics.ml4 | 8 +- plugins/ltac/g_ltac.ml4 | 4 +- plugins/ltac/g_obligations.ml4 | 2 +- plugins/ltac/pptactic.ml | 42 +-- plugins/ltac/pptactic.mli | 2 +- plugins/ltac/profile_ltac.ml | 4 +- plugins/ltac/rewrite.ml | 41 ++- plugins/ltac/tacentries.ml | 22 +- plugins/ltac/tacentries.mli | 2 +- plugins/ltac/tacexpr.mli | 2 +- plugins/ltac/tacintern.ml | 10 +- plugins/ltac/tacinterp.ml | 52 ++-- plugins/ltac/tacinterp.mli | 2 +- plugins/ltac/tacsubst.ml | 2 +- plugins/ltac/tactic_debug.ml | 19 +- plugins/ltac/tactic_debug.mli | 2 +- plugins/ltac/tauto.ml | 11 +- plugins/micromega/coq_micromega.ml | 4 +- plugins/quote/g_quote.ml4 | 5 +- plugins/setoid_ring/newring.ml | 14 +- plugins/ssrmatching/ssrmatching.ml4 | 34 +-- plugins/syntax/ascii_syntax.ml | 12 +- plugins/syntax/nat_syntax.ml | 10 +- plugins/syntax/numbers_syntax.ml | 68 ++--- plugins/syntax/r_syntax.ml | 12 +- plugins/syntax/string_syntax.ml | 8 +- plugins/syntax/z_syntax.ml | 42 +-- pretyping/cases.ml | 66 ++--- pretyping/cases.mli | 8 +- pretyping/coercion.ml | 65 ++--- pretyping/coercion.mli | 14 +- pretyping/detyping.ml | 2 - pretyping/evarconv.ml | 4 +- pretyping/glob_ops.ml | 2 +- pretyping/glob_ops.mli | 2 +- pretyping/pretyping.ml | 44 ++-- pretyping/pretyping.mli | 2 +- pretyping/typing.ml | 6 +- pretyping/typing.mli | 2 +- pretyping/unification.ml | 4 +- printing/ppconstr.ml | 10 +- printing/ppconstr.mli | 2 +- printing/pputils.ml | 2 +- printing/prettyp.ml | 6 +- printing/printer.ml | 6 +- proofs/clenv.ml | 2 +- proofs/goal.ml | 2 +- proofs/proof_global.ml | 4 +- proofs/proof_using.ml | 2 +- proofs/refine.ml | 2 +- stm/stm.ml | 83 +++--- stm/stm.mli | 2 +- stm/vio_checking.ml | 2 +- tactics/contradiction.ml | 2 +- tactics/equality.ml | 2 +- tactics/inv.ml | 2 +- tactics/tacticals.ml | 24 +- tactics/tacticals.mli | 2 +- tactics/tactics.ml | 100 ++++--- toplevel/coqtop.ml | 4 +- toplevel/vernac.ml | 10 +- vernac/auto_ind_decl.ml | 12 +- vernac/classes.ml | 2 +- vernac/command.ml | 11 +- vernac/explainErr.ml | 4 +- vernac/explainErr.mli | 2 +- vernac/indschemes.ml | 2 +- vernac/lemmas.ml | 2 +- vernac/metasyntax.ml | 2 +- vernac/obligations.ml | 5 +- vernac/vernacentries.ml | 22 +- vernac/vernacentries.mli | 2 +- 122 files changed, 1083 insertions(+), 1099 deletions(-) diff --git a/.merlin b/.merlin index 5cae15f5f..5c8c73851 100644 --- a/.merlin +++ b/.merlin @@ -36,6 +36,8 @@ S toplevel B toplevel S vernac B vernac +S plugins/ltac +B plugins/ltac S tools B tools diff --git a/checker/modops.ml b/checker/modops.ml index aba9da2fe..07dda8a06 100644 --- a/checker/modops.ml +++ b/checker/modops.ml @@ -32,10 +32,10 @@ let error_no_such_label_sub l l1 = error ("The field "^ Label.to_string l^" is missing in "^l1^".") -let error_not_a_module_loc loc s = - user_err ~loc (str ("\""^Label.to_string s^"\" is not a module")) +let error_not_a_module_loc ?loc s = + user_err ?loc (str ("\""^Label.to_string s^"\" is not a module")) -let error_not_a_module s = error_not_a_module_loc Loc.ghost s +let error_not_a_module s = error_not_a_module_loc s let error_with_module () = error "Unsupported 'with' constraint in module implementation" diff --git a/dev/top_printers.ml b/dev/top_printers.ml index 52cf8cf97..96cb9443c 100644 --- a/dev/top_printers.ml +++ b/dev/top_printers.ml @@ -521,7 +521,7 @@ let _ = extend_vernac_command_grammar ("PrintConstr", 0) None [GramTerminal "PrintConstr"; GramNonTerminal - (Loc.ghost,rawwit wit_constr,Extend.Aentry Pcoq.Constr.constr)] + (Loc.internal_ghost, rawwit wit_constr,Extend.Aentry Pcoq.Constr.constr)] let _ = try @@ -537,46 +537,46 @@ let _ = extend_vernac_command_grammar ("PrintPureConstr", 0) None [GramTerminal "PrintPureConstr"; GramNonTerminal - (Loc.ghost,rawwit wit_constr,Extend.Aentry Pcoq.Constr.constr)] + (Loc.internal_ghost,rawwit wit_constr,Extend.Aentry Pcoq.Constr.constr)] (* Setting printer of unbound global reference *) open Names open Libnames -let encode_path loc prefix mpdir suffix id = +let encode_path ?loc prefix mpdir suffix id = let dir = match mpdir with | None -> [] | Some (mp,dir) -> (DirPath.repr (dirpath_of_string (string_of_mp mp))@ DirPath.repr dir) in - Qualid (loc, make_qualid + Qualid (Loc.tag ?loc @@ make_qualid (DirPath.make (List.rev (Id.of_string prefix::dir@suffix))) id) -let raw_string_of_ref loc _ = function +let raw_string_of_ref ?loc _ = function | ConstRef cst -> let (mp,dir,id) = repr_con cst in - encode_path loc "CST" (Some (mp,dir)) [] (Label.to_id id) + encode_path ?loc "CST" (Some (mp,dir)) [] (Label.to_id id) | IndRef (kn,i) -> let (mp,dir,id) = repr_mind kn in - encode_path loc "IND" (Some (mp,dir)) [Label.to_id id] + encode_path ?loc "IND" (Some (mp,dir)) [Label.to_id id] (Id.of_string ("_"^string_of_int i)) | ConstructRef ((kn,i),j) -> let (mp,dir,id) = repr_mind kn in - encode_path loc "CSTR" (Some (mp,dir)) + encode_path ?loc "CSTR" (Some (mp,dir)) [Label.to_id id;Id.of_string ("_"^string_of_int i)] (Id.of_string ("_"^string_of_int j)) | VarRef id -> - encode_path loc "SECVAR" None [] id + encode_path ?loc "SECVAR" None [] id -let short_string_of_ref loc _ = function - | VarRef id -> Ident (loc,id) - | ConstRef cst -> Ident (loc,Label.to_id (pi3 (repr_con cst))) - | IndRef (kn,0) -> Ident (loc,Label.to_id (pi3 (repr_mind kn))) +let short_string_of_ref ?loc _ = function + | VarRef id -> Ident (Loc.tag ?loc id) + | ConstRef cst -> Ident (Loc.tag ?loc @@ Label.to_id (pi3 (repr_con cst))) + | IndRef (kn,0) -> Ident (Loc.tag ?loc @@ Label.to_id (pi3 (repr_mind kn))) | IndRef (kn,i) -> - encode_path loc "IND" None [Label.to_id (pi3 (repr_mind kn))] + encode_path ?loc "IND" None [Label.to_id (pi3 (repr_mind kn))] (Id.of_string ("_"^string_of_int i)) | ConstructRef ((kn,i),j) -> - encode_path loc "CSTR" None + encode_path ?loc "CSTR" None [Label.to_id (pi3 (repr_mind kn));Id.of_string ("_"^string_of_int i)] (Id.of_string ("_"^string_of_int j)) diff --git a/engine/evarutil.ml b/engine/evarutil.ml index 1624dc93e..ac64ab834 100644 --- a/engine/evarutil.ml +++ b/engine/evarutil.ml @@ -366,7 +366,7 @@ let push_rel_context_to_named_context env sigma typ = * Entry points to define new evars * *------------------------------------*) -let default_source = (Loc.ghost,Evar_kinds.InternalHole) +let default_source = Loc.tag @@ Evar_kinds.InternalHole let restrict_evar evd evk filter candidates = let evd = Sigma.to_evar_map evd in diff --git a/engine/evd.ml b/engine/evd.ml index 5419a10a8..9e81ccd36 100644 --- a/engine/evd.ml +++ b/engine/evd.ml @@ -151,7 +151,7 @@ let make_evar hyps ccl = { evar_hyps = hyps; evar_body = Evar_empty; evar_filter = Filter.identity; - evar_source = (Loc.ghost,Evar_kinds.InternalHole); + evar_source = Loc.tag @@ Evar_kinds.InternalHole; evar_candidates = None; evar_extra = Store.empty } @@ -704,11 +704,11 @@ 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,_) -> fst (evar_source evk1 evd) + | Evar (evk1,_) -> Some (fst (evar_source evk1 evd)) | _ -> match kind_of_term (fst (decompose_app t2)) with - | Evar (evk2,_) -> fst (evar_source evk2 evd) - | _ -> Loc.ghost + | Evar (evk2,_) -> Some (fst (evar_source evk2 evd)) + | _ -> None (** The following functions return the set of evars immediately contained in the object *) @@ -1086,8 +1086,8 @@ let retract_coercible_metas evd = let evar_source_of_meta mv evd = match meta_name evd mv with - | Anonymous -> (Loc.ghost,Evar_kinds.GoalEvar) - | Name id -> (Loc.ghost,Evar_kinds.VarInstance id) + | Anonymous -> Loc.tag Evar_kinds.GoalEvar + | Name id -> Loc.tag @@ Evar_kinds.VarInstance id let dependent_evar_ident ev evd = let evi = find evd ev in diff --git a/engine/evd.mli b/engine/evd.mli index 9c40c8b71..005332470 100644 --- a/engine/evd.mli +++ b/engine/evd.mli @@ -414,7 +414,7 @@ val extract_changed_conv_pbs : evar_map -> (Evar.Set.t -> evar_constraint -> bool) -> evar_map * evar_constraint list val extract_all_conv_pbs : evar_map -> evar_map * evar_constraint list -val loc_of_conv_pb : evar_map -> evar_constraint -> Loc.t +val loc_of_conv_pb : evar_map -> evar_constraint -> Loc.t option (** The following functions return the set of evars immediately contained in the object; need the term to be evar-normal otherwise diff --git a/engine/proofview.ml b/engine/proofview.ml index f054038e9..84bcecc44 100644 --- a/engine/proofview.ml +++ b/engine/proofview.ml @@ -66,7 +66,7 @@ let dependent_init = for type classes. *) let store = Evd.Store.set Evd.Store.empty typeclass_resolvable () in (* Goals don't have a source location. *) - let src = (Loc.ghost,Evar_kinds.GoalEvar) in + let src = Loc.tag @@ Evar_kinds.GoalEvar in (* Main routine *) let rec aux = function | TNil sigma -> [], { solution = sigma; comb = []; shelf = [] } diff --git a/engine/uState.ml b/engine/uState.ml index c66af02bb..c9653b6cd 100644 --- a/engine/uState.ml +++ b/engine/uState.ml @@ -266,10 +266,10 @@ let universe_context ?names ctx = try let info = Univ.LMap.find (Univ.LSet.choose left) (snd ctx.uctx_names) in - Option.default Loc.ghost info.uloc - with Not_found -> Loc.ghost + info.uloc + with Not_found -> None in - user_err ~loc ~hdr:"universe_context" + user_err ?loc ~hdr:"universe_context" ((str(CString.plural n "Universe") ++ spc () ++ Univ.LSet.pr (pr_uctx_level ctx) left ++ spc () ++ str (CString.conjugate_verb_to_be n) ++ diff --git a/grammar/argextend.mlp b/grammar/argextend.mlp index d00ee4e5d..5ec8086d0 100644 --- a/grammar/argextend.mlp +++ b/grammar/argextend.mlp @@ -9,7 +9,7 @@ open Q_util let loc = Ploc.dummy -let default_loc = <:expr< Loc.ghost >> +let default_loc = <:expr< Loc.internal_ghost >> IFDEF STRICT THEN let ploc_vala x = Ploc.VaVal x diff --git a/grammar/tacextend.mlp b/grammar/tacextend.mlp index 1dd8da12a..aa463ada5 100644 --- a/grammar/tacextend.mlp +++ b/grammar/tacextend.mlp @@ -20,8 +20,6 @@ let make_fun loc cl = let l = cl @ [default_patt loc] in MLast.ExFun (loc, ploc_vala l) (* correspond to <:expr< fun [ $list:l$ ] >> *) -let dloc = <:expr< Loc.ghost >> - let plugin_name = <:expr< __coq_plugin_name >> let mlexpr_of_ident id = @@ -75,7 +73,7 @@ let rec mlexpr_of_symbol = function let make_prod_item = function | ExtTerminal s -> <:expr< Tacentries.TacTerm $str:s$ >> | ExtNonTerminal (g, id) -> - <:expr< Tacentries.TacNonTerm $default_loc$ $mlexpr_of_symbol g$ $mlexpr_of_ident id$ >> + <:expr< Tacentries.TacNonTerm (Loc.tag ( $mlexpr_of_symbol g$ , $mlexpr_of_ident id$ ) ) >> let mlexpr_of_clause cl = mlexpr_of_list (fun (a,_,_) -> mlexpr_of_list make_prod_item a) cl @@ -114,7 +112,7 @@ let declare_tactic loc tacname ~level classification clause = match clause with (** Arguments are not passed directly to the ML tactic in the TacML node, the ML tactic retrieves its arguments in the [ist] environment instead. This is the rĂ´le of the [lift_constr_tac_to_ml_tac] function. *) - let body = <:expr< Tacexpr.TacFun ($vars$, Tacexpr.TacML ($dloc$, $ml$, [])) >> in + let body = <:expr< Tacexpr.TacFun ($vars$, Tacexpr.TacML (Loc.tag ( $ml$ , []))) >> in let name = <:expr< Names.Id.of_string $name$ >> in declare_str_items loc [ <:str_item< do { diff --git a/ide/coqOps.ml b/ide/coqOps.ml index 222b9eed9..7825fb45e 100644 --- a/ide/coqOps.ml +++ b/ide/coqOps.ml @@ -465,20 +465,22 @@ object(self) self#attach_tooltip ~loc sentence (Printf.sprintf "%s %s %s" filepath ident ty) | Message(Error, loc, msg), Some (id,sentence) -> - let uloc = Option.default Loc.ghost loc in log_pp ?id Pp.(str "ErrorMsg" ++ msg); remove_flag sentence `PROCESSING; let rmsg = Pp.string_of_ppcmds msg in - add_flag sentence (`ERROR (uloc, rmsg)); + Option.iter (fun loc -> + add_flag sentence (`ERROR (loc, rmsg)); + ) loc; self#mark_as_needed sentence; - self#attach_tooltip sentence ?loc rmsg; + self#attach_tooltip ?loc sentence rmsg; self#position_tag_at_sentence ?loc Tags.Script.error sentence | Message(Warning, loc, msg), Some (id,sentence) -> - let uloc = Option.default Loc.ghost loc in log_pp ?id Pp.(str "WarningMsg" ++ msg); let rmsg = Pp.string_of_ppcmds msg in - add_flag sentence (`WARNING (uloc, rmsg)); - self#attach_tooltip sentence ?loc rmsg; + Option.iter (fun loc -> + add_flag sentence (`WARNING (loc, rmsg)); + ) loc; + self#attach_tooltip ?loc sentence rmsg; self#position_tag_at_sentence ?loc Tags.Script.warning sentence; messages#push Warning msg | Message(lvl, loc, msg), Some (id,sentence) -> @@ -528,14 +530,14 @@ object(self) let start, stop, phrase = self#get_sentence sentence in self#position_tag_at_iter ?loc start stop tag phrase - method private process_interp_error queue sentence loc msg tip id = + method private process_interp_error ?loc queue sentence msg tip id = Coq.bind (Coq.return ()) (function () -> let start, stop, phrase = self#get_sentence sentence in buffer#remove_tag Tags.Script.to_process ~start ~stop; self#discard_command_queue queue; pop_info (); if Stateid.equal id tip || Stateid.equal id Stateid.dummy then begin - self#position_tag_at_iter ~loc start stop Tags.Script.error phrase; + self#position_tag_at_iter ?loc start stop Tags.Script.error phrase; buffer#place_cursor ~where:stop; messages#clear; messages#push Feedback.Error msg; @@ -649,9 +651,9 @@ object(self) if Queue.is_empty queue then loop tip [] else loop tip (List.rev topstack) | Fail (id, loc, msg) -> - let loc = Option.cata Loc.make_loc Loc.ghost loc in + let loc = Option.map Loc.make_loc loc in let sentence = Doc.pop document in - self#process_interp_error queue sentence loc msg tip id in + self#process_interp_error ?loc queue sentence msg tip id in Coq.bind coq_query handle_answer in let tip = diff --git a/ide/ide_slave.ml b/ide/ide_slave.ml index bf86db08f..7f30a4acc 100644 --- a/ide/ide_slave.ml +++ b/ide/ide_slave.ml @@ -390,8 +390,8 @@ let quit = ref false (** Serializes the output of Stm.get_ast *) let print_ast id = match Stm.get_ast id with - | Some (expr, loc) -> begin - try Texmacspp.tmpp expr loc + | Some (loc, expr) -> begin + 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 f86b260df..0a07a830b 100644 --- a/ide/texmacspp.ml +++ b/ide/texmacspp.ml @@ -17,12 +17,12 @@ open Extend open Libnames open Constrexpr_ops -let unlock loc = - let start, stop = Loc.unloc loc in +let unlock ?loc = + let start, stop = Option.cata Loc.unloc (0,0) loc in (string_of_int start, string_of_int stop) -let xmlWithLoc loc ename attr xml = - let start, stop = unlock loc in +let xmlWithLoc ?loc ename attr xml = + let start, stop = unlock ?loc in Element(ename, [ "begin", start; "end", stop ] @ attr, xml) let get_fst_attr_in_xml_list attr xml_list = @@ -49,32 +49,32 @@ let compare_begin_att xml1 xml2 = | (_, s1), (_, s2) when int_of_string s1 < int_of_string s2 -> -1 | _ -> 0 -let xmlBeginSection loc name = xmlWithLoc loc "beginsection" ["name", name] [] +let xmlBeginSection ?loc name = xmlWithLoc ?loc "beginsection" ["name", name] [] -let xmlEndSegment loc name = xmlWithLoc loc "endsegment" ["name", name] [] +let xmlEndSegment ?loc name = xmlWithLoc ?loc "endsegment" ["name", name] [] -let xmlThm typ name loc xml = - xmlWithLoc loc "theorem" ["type", typ; "name", name] xml +let xmlThm ?loc typ name xml = + xmlWithLoc ?loc "theorem" ["type", typ; "name", name] xml -let xmlDef typ name loc xml = - xmlWithLoc loc "definition" ["type", typ; "name", name] xml +let xmlDef ?loc typ name xml = + xmlWithLoc ?loc "definition" ["type", typ; "name", name] xml -let xmlNotation attr name loc xml = - xmlWithLoc loc "notation" (("name", name) :: attr) xml +let xmlNotation ?loc attr name xml = + xmlWithLoc ?loc "notation" (("name", name) :: attr) xml -let xmlReservedNotation attr name loc = - xmlWithLoc loc "reservednotation" (("name", name) :: attr) [] +let xmlReservedNotation ?loc attr name = + xmlWithLoc ?loc "reservednotation" (("name", name) :: attr) [] -let xmlCst name ?(attr=[]) loc = - xmlWithLoc loc "constant" (("name", name) :: attr) [] +let xmlCst ?loc ?(attr=[]) name = + xmlWithLoc ?loc "constant" (("name", name) :: attr) [] -let xmlOperator name ?(attr=[]) ?(pprules=[]) loc = - xmlWithLoc loc "operator" +let xmlOperator ?loc ?(attr=[]) ?(pprules=[]) name = + xmlWithLoc ?loc "operator" (("name", name) :: List.map (fun (a,b) -> "format"^a,b) pprules @ attr) [] -let xmlApply loc ?(attr=[]) xml = xmlWithLoc loc "apply" attr xml +let xmlApply ?loc ?(attr=[]) xml = xmlWithLoc ?loc "apply" attr xml -let xmlToken loc ?(attr=[]) xml = xmlWithLoc loc "token" attr xml +let xmlToken ?loc ?(attr=[]) xml = xmlWithLoc ?loc "token" attr xml let xmlTyped xml = Element("typed", (backstep_loc xml), xml) @@ -88,23 +88,23 @@ let xmlWith xml = Element("with", [], xml) let xmlAssign id xml = Element("assign", ["target",string_of_id id], [xml]) -let xmlInductive kind loc xml = xmlWithLoc loc "inductive" ["kind",kind] xml +let xmlInductive ?loc kind xml = xmlWithLoc ?loc "inductive" ["kind",kind] xml let xmlCoFixpoint xml = Element("cofixpoint", [], xml) let xmlFixpoint xml = Element("fixpoint", [], xml) -let xmlCheck loc xml = xmlWithLoc loc "check" [] xml +let xmlCheck ?loc xml = xmlWithLoc ?loc "check" [] xml -let xmlAssumption kind loc xml = xmlWithLoc loc "assumption" ["kind",kind] xml +let xmlAssumption ?loc kind xml = xmlWithLoc ?loc "assumption" ["kind",kind] xml -let xmlComment loc xml = xmlWithLoc loc "comment" [] xml +let xmlComment ?loc xml = xmlWithLoc ?loc "comment" [] xml -let xmlCanonicalStructure attr loc = xmlWithLoc loc "canonicalstructure" attr [] +let xmlCanonicalStructure ?loc attr = xmlWithLoc ?loc "canonicalstructure" attr [] -let xmlQed ?(attr=[]) loc = xmlWithLoc loc "qed" attr [] +let xmlQed ?loc ?(attr=[]) = xmlWithLoc ?loc "qed" attr [] -let xmlPatvar id loc = xmlWithLoc loc "patvar" ["id", id] [] +let xmlPatvar ?loc id = xmlWithLoc ?loc "patvar" ["id", id] [] let xmlReference ref = let name = Libnames.string_of_reference ref in @@ -112,38 +112,38 @@ let xmlReference ref = let b, e = string_of_int i, string_of_int j in Element("reference",["name", name; "begin", b; "end", e] ,[]) -let xmlRequire loc ?(attr=[]) xml = xmlWithLoc loc "require" attr xml -let xmlImport loc ?(attr=[]) xml = xmlWithLoc loc "import" attr xml +let xmlRequire ?loc ?(attr=[]) xml = xmlWithLoc ?loc "require" attr xml +let xmlImport ?loc ?(attr=[]) xml = xmlWithLoc ?loc "import" attr xml -let xmlAddLoadPath loc ?(attr=[]) xml = xmlWithLoc loc "addloadpath" attr xml -let xmlRemoveLoadPath loc ?(attr=[]) = xmlWithLoc loc "removeloadpath" attr -let xmlAddMLPath loc ?(attr=[]) = xmlWithLoc loc "addmlpath" attr +let xmlAddLoadPath ?loc ?(attr=[]) xml = xmlWithLoc ?loc "addloadpath" attr xml +let xmlRemoveLoadPath ?loc ?(attr=[]) = xmlWithLoc ?loc "removeloadpath" attr +let xmlAddMLPath ?loc ?(attr=[]) = xmlWithLoc ?loc "addmlpath" attr -let xmlExtend loc xml = xmlWithLoc loc "extend" [] xml +let xmlExtend ?loc xml = xmlWithLoc ?loc "extend" [] xml -let xmlScope loc action ?(attr=[]) name xml = - xmlWithLoc loc "scope" (["name",name;"action",action] @ attr) xml +let xmlScope ?loc ?(attr=[]) action name xml = + xmlWithLoc ?loc "scope" (["name",name;"action",action] @ attr) xml -let xmlProofMode loc name = xmlWithLoc loc "proofmode" ["name",name] [] +let xmlProofMode ?loc name = xmlWithLoc ?loc "proofmode" ["name",name] [] -let xmlProof loc xml = xmlWithLoc loc "proof" [] xml +let xmlProof ?loc xml = xmlWithLoc ?loc "proof" [] xml let xmlSectionSubsetDescr name ssd = Element("sectionsubsetdescr",["name",name], [PCData (Proof_using.to_string ssd)]) -let xmlDeclareMLModule loc s = - xmlWithLoc loc "declarexmlmodule" [] +let xmlDeclareMLModule ?loc s = + xmlWithLoc ?loc "declarexmlmodule" [] (List.map (fun x -> Element("path",["value",x],[])) s) (* tactics *) -let xmlLtac loc xml = xmlWithLoc loc "ltac" [] xml +let xmlLtac ?loc xml = xmlWithLoc ?loc "ltac" [] xml (* toplevel commands *) -let xmlGallina loc xml = xmlWithLoc loc "gallina" [] xml +let xmlGallina ?loc xml = xmlWithLoc ?loc "gallina" [] xml -let xmlTODO loc x = - xmlWithLoc loc "todo" [] [PCData (Pp.string_of_ppcmds (Ppvernac.pr_vernac x))] +let xmlTODO ?loc x = + xmlWithLoc ?loc "todo" [] [PCData (Pp.string_of_ppcmds (Ppvernac.pr_vernac x))] let string_of_name n = match n with @@ -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 (string_of_name name) loc) loc_names) in + xmlCst ~loc (string_of_name name)) loc_names) in match e with | _, CHole _ -> names | _ -> names @ [pp_expr e]) @@ -237,13 +237,13 @@ and pp_local_binder lb = (* don't know what it is for now *) pp_expr ~attr:attrs value | CLocalAssum (namll, _, ce) -> let ppl = - List.map (fun (loc, nam) -> (xmlCst (string_of_name nam) loc)) namll in + List.map (fun (loc, nam) -> (xmlCst ~loc (string_of_name nam))) namll in xmlTyped (ppl @ [pp_expr ce]) | CLocalPattern _ -> assert false and pp_local_decl_expr lde = (* don't know what it is for now *) match lde with - | AssumExpr (_, ce) -> pp_expr ce + | AssumExpr (_, ce) -> pp_expr ce | DefExpr (_, ce, _) -> pp_expr ce and pp_inductive_expr ((_, ((l, id),_)), lbl, ceo, _, cl_or_rdexpr) = (* inductive_expr *) @@ -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,37 +300,37 @@ 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 (Id.to_string id) loc +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 "alias" ~attr:["name", string_of_id id] loc :: + xmlApply ~loc + (xmlOperator ~loc ~attr:["name", string_of_id id] "alias" :: [pp_cases_pattern_expr cpe]) | CPatCstr (ref, None, cpel2) -> - xmlApply loc - (xmlOperator "reference" - ~attr:["name", Libnames.string_of_reference ref] loc :: + 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 "reference" - ~attr:["name", Libnames.string_of_reference ref] loc :: + 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))]) | CPatAtom optr -> let attrs = match optr with | None -> [] | Some r -> ["name", Libnames.string_of_reference r] in - xmlApply loc (xmlOperator "atom" ~attr:attrs loc :: []) + xmlApply ~loc (xmlOperator ~loc "atom" ~attr:attrs :: []) | CPatOr cpel -> - xmlApply loc (xmlOperator "or" loc :: 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 "notation" loc :: - [xmlOperator n loc; + 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 "record" loc :: + 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 "delimiter" ~attr:["name", delim] loc :: + 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] @@ -378,12 +378,12 @@ and pp_branch_expr_list bel = let ppe = [pp_expr e] in xmlCase (ppcepl @ ppe)) bel) -and pp_token loc tok = +and pp_token ?loc tok = let tokstr = match tok with | String s -> PCData s | Numeral n -> PCData (to_string n) in - xmlToken loc [tokstr] + xmlToken ?loc [tokstr] and pp_local_binder_list lbl = let l = (List.map pp_local_binder lbl) in Element ("recurse", (backstep_loc l), l) @@ -393,142 +393,140 @@ and pp_const_expr_list cel = and pp_expr ?(attr=[]) (loc, e) = match e with | CRef (r, _) -> - xmlCst ~attr - (Libnames.string_of_reference r) (Libnames.loc_of_reference r) + xmlCst ~loc:(Libnames.loc_of_reference r) ~attr (Libnames.string_of_reference r) | CProdN (bl, e) -> - xmlApply loc - (xmlOperator "forall" loc :: [pp_bindlist bl] @ [pp_expr e]) + xmlApply ~loc + (xmlOperator ~loc "forall" :: [pp_bindlist bl] @ [pp_expr e]) | CApp ((_, hd), args) -> - xmlApply ~attr loc (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 ~attr loc - (xmlCst (Libnames.string_of_reference r) - (Libnames.loc_of_reference r) :: List.map pp_expr args) + xmlApply ~loc ~attr + (xmlCst ~loc:(Libnames.loc_of_reference r) (Libnames.string_of_reference r) + :: List.map pp_expr args) | CNotation (notation, ([],[],[])) -> - xmlOperator notation loc + xmlOperator ~loc notation | CNotation (notation, (args, cell, lbll)) -> let fmts = Notation.find_notation_extra_printing_rules notation in - let oper = xmlOperator notation loc ~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 (string_of_glob_sort s) loc + xmlOperator ~loc (string_of_glob_sort s) | CDelimiters (scope, ce) -> - xmlApply loc (xmlOperator "delimiter" ~attr:["name", scope] loc :: + 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 "generalization" ~attr:["kind", kind] loc :: [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 "evar" loc ~attr:["id", string_of_id ek] :: + xmlApply ~loc + (xmlOperator ~loc "evar" ~attr:["id", string_of_id ek] :: ppcel) - | CPatVar id -> xmlPatvar (string_of_id id) loc - | CHole (_, _, _) -> xmlCst ~attr "_" loc + | 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 "if" loc :: + 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 "lettuple" loc :: + xmlApply ~loc + (xmlOperator ~loc "lettuple" :: return @ - (List.map (fun (loc, var) -> xmlCst (string_of_name var) loc) 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 "match" loc ~attr:["style", (string_of_case_style sty)] :: + 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])) | CRecord _ -> assert false - | 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)) | None -> value in - xmlApply loc - (xmlOperator "let" loc :: - [xmlCst (string_of_name var) varloc; 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 "lambda" loc :: [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 "fix" loc :: + 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)) -let pp_comment (c) = +let pp_comment c = match c with | CommentConstr e -> [pp_expr e] | CommentString s -> [Element ("string", [], [PCData s])] - | CommentInt i -> [PCData (string_of_int i)] + | CommentInt i -> [PCData (string_of_int i)] -let rec tmpp v loc = +let rec tmpp ?loc v = match v with (* Control *) | VernacLoad (verbose,f) -> - xmlWithLoc loc "load" ["verbose",string_of_bool verbose;"file",f] [] + xmlWithLoc ?loc "load" ["verbose",string_of_bool verbose;"file",f] [] | VernacTime (loc,e) -> - xmlApply loc (Element("time",[],[]) :: - [tmpp e loc]) + xmlApply ~loc (Element("time",[],[]) :: + [tmpp ~loc e]) | VernacRedirect (s, (loc,e)) -> - xmlApply loc (Element("redirect",["path", s],[]) :: - [tmpp e loc]) + xmlApply ~loc (Element("redirect",["path", s],[]) :: + [tmpp ~loc e]) | VernacTimeout (s,e) -> - xmlApply loc (Element("timeout",["val",string_of_int s],[]) :: - [tmpp e loc]) - | VernacFail e -> xmlApply loc (Element("fail",[],[]) :: [tmpp e loc]) + xmlApply ?loc (Element("timeout",["val",string_of_int s],[]) :: + [tmpp ?loc e]) + | VernacFail e -> xmlApply ?loc (Element("fail",[],[]) :: [tmpp ?loc e]) (* Syntax *) | VernacSyntaxExtension (_, ((_, name), sml)) -> let attrs = List.flatten (List.map attribute_of_syntax_modifier sml) in - xmlReservedNotation attrs name loc + xmlReservedNotation ?loc attrs name - | VernacOpenCloseScope (_,(true,name)) -> xmlScope loc "open" name [] - | VernacOpenCloseScope (_,(false,name)) -> xmlScope loc "close" name [] + | VernacOpenCloseScope (_,(true,name)) -> xmlScope ?loc "open" name [] + | VernacOpenCloseScope (_,(false,name)) -> xmlScope ?loc "close" name [] | VernacDelimiters (name,Some tag) -> - xmlScope loc "delimit" name ~attr:["delimiter",tag] [] + xmlScope ?loc "delimit" name ~attr:["delimiter",tag] [] | VernacDelimiters (name,None) -> - xmlScope loc "undelimit" name ~attr:[] [] + xmlScope ?loc "undelimit" name ~attr:[] [] | VernacInfix (_,((_,name),sml),ce,sn) -> let attrs = List.flatten (List.map attribute_of_syntax_modifier sml) in let sc_attr = match sn with | Some scope -> ["scope", scope] | None -> [] in - xmlNotation (sc_attr @ attrs) name loc [pp_expr ce] + xmlNotation ?loc (sc_attr @ attrs) name [pp_expr ce] | VernacNotation (_, ce, (lstr, sml), sn) -> let name = snd lstr in let attrs = List.flatten (List.map attribute_of_syntax_modifier sml) in @@ -536,12 +534,12 @@ let rec tmpp v loc = match sn with | Some scope -> ["scope", scope] | None -> [] in - xmlNotation (sc_attr @ attrs) name loc [pp_expr ce] - | VernacBindScope _ as x -> xmlTODO loc x - | VernacNotationAddFormat _ as x -> xmlTODO loc x + xmlNotation ?loc (sc_attr @ attrs) name [pp_expr ce] + | VernacBindScope _ as x -> xmlTODO ?loc x + | VernacNotationAddFormat _ as x -> xmlTODO ?loc x | VernacUniverse _ | VernacConstraint _ - | VernacPolymorphic (_, _) as x -> xmlTODO loc x + | VernacPolymorphic (_, _) as x -> xmlTODO ?loc x (* Gallina *) | VernacDefinition (ldk, ((_,id),_), de) -> let l, dk = @@ -557,26 +555,26 @@ let rec tmpp v loc = | DefineBody (_, None , ce, Some _) -> ce in let str_dk = Kindops.string_of_definition_kind (l, false, dk) in let str_id = Id.to_string id in - (xmlDef str_dk str_id loc [pp_expr e]) + (xmlDef ?loc str_dk str_id [pp_expr e]) | VernacStartTheoremProof (tk, [ Some ((_,id),_), ([], statement, None) ], b) -> let str_tk = Kindops.string_of_theorem_kind tk in let str_id = Id.to_string id in - (xmlThm str_tk str_id loc [pp_expr statement]) - | VernacStartTheoremProof _ as x -> xmlTODO loc x + (xmlThm ?loc str_tk str_id [pp_expr statement]) + | VernacStartTheoremProof _ as x -> xmlTODO ?loc x | VernacEndProof pe -> begin match pe with - | Admitted -> xmlQed loc + | Admitted -> xmlQed ?loc ?attr:None | Proved (_, Some ((_, id), Some tk)) -> let nam = Id.to_string id in let typ = Kindops.string_of_theorem_kind tk in - xmlQed ~attr:["name", nam; "type", typ] loc + xmlQed ?loc ~attr:["name", nam; "type", typ] | Proved (_, Some ((_, id), None)) -> let nam = Id.to_string id in - xmlQed ~attr:["name", nam] loc - | Proved _ -> xmlQed loc + xmlQed ?loc ~attr:["name", nam] + | Proved _ -> xmlQed ?loc ?attr:None end - | VernacExactProof _ as x -> xmlTODO loc x + | VernacExactProof _ as x -> xmlTODO ?loc x | VernacAssumption ((l, a), _, sbwcl) -> let binders = List.map (fun (_, (id, c)) -> (List.map fst id, c)) sbwcl in let many = @@ -585,7 +583,7 @@ let rec tmpp v loc = List.flatten (List.map pp_simple_binder binders) in let l = match l with Some x -> x | None -> Decl_kinds.Global in let kind = string_of_assumption_kind l a many in - xmlAssumption kind loc exprs + xmlAssumption ?loc kind exprs | VernacInductive (_, _, iednll) -> let kind = let (_, _, _, k, _),_ = List.hd iednll in @@ -603,7 +601,7 @@ let rec tmpp v loc = (List.map (fun (ie, dnl) -> (pp_inductive_expr ie) @ (List.map pp_decl_notation dnl)) iednll) in - xmlInductive kind loc exprs + xmlInductive ?loc kind exprs | VernacFixpoint (_, fednll) -> let exprs = List.flatten (* should probably not be flattened *) @@ -619,13 +617,13 @@ let rec tmpp v loc = (fun (cfe, dnl) -> (pp_cofixpoint_expr cfe) @ (List.map pp_decl_notation dnl)) cfednll) in xmlCoFixpoint exprs - | VernacScheme _ as x -> xmlTODO loc x - | VernacCombinedScheme _ as x -> xmlTODO loc x + | VernacScheme _ as x -> xmlTODO ?loc x + | VernacCombinedScheme _ as x -> xmlTODO ?loc x (* Gallina extensions *) - | VernacBeginSection (_, id) -> xmlBeginSection loc (Id.to_string id) - | VernacEndSegment (_, id) -> xmlEndSegment loc (Id.to_string id) - | VernacNameSectionHypSet _ as x -> xmlTODO loc x + | VernacBeginSection (_, id) -> xmlBeginSection ?loc (Id.to_string id) + | VernacEndSegment (_, id) -> xmlEndSegment ?loc (Id.to_string id) + | VernacNameSectionHypSet _ as x -> xmlTODO ?loc x | VernacRequire (from, import, l) -> let import = match import with | None -> [] @@ -636,137 +634,136 @@ let rec tmpp v loc = | None -> [] | Some r -> ["from", Libnames.string_of_reference r] in - xmlRequire loc ~attr:(from @ import) (List.map (fun ref -> + xmlRequire ?loc ~attr:(from @ import) (List.map (fun ref -> xmlReference ref) l) | VernacImport (true,l) -> - xmlImport loc ~attr:["export","true"] (List.map (fun ref -> + xmlImport ?loc ~attr:["export","true"] (List.map (fun ref -> xmlReference ref) l) | VernacImport (false,l) -> - xmlImport loc (List.map (fun ref -> - xmlReference ref) l) + xmlImport ?loc (List.map (fun ref -> xmlReference ref) l) | VernacCanonical r -> let attr = match r with | AN (Qualid (_, q)) -> ["qualid", string_of_qualid q] | AN (Ident (_, id)) -> ["id", Id.to_string id] | ByNotation (_, (s, _)) -> ["notation", s] in - xmlCanonicalStructure attr loc - | VernacCoercion _ as x -> xmlTODO loc x - | VernacIdentityCoercion _ as x -> xmlTODO loc x + xmlCanonicalStructure ?loc attr + | VernacCoercion _ as x -> xmlTODO ?loc x + | VernacIdentityCoercion _ as x -> xmlTODO ?loc x (* Type classes *) - | VernacInstance _ as x -> xmlTODO loc x + | VernacInstance _ as x -> xmlTODO ?loc x - | VernacContext _ as x -> xmlTODO loc x + | VernacContext _ as x -> xmlTODO ?loc x - | VernacDeclareInstances _ as x -> xmlTODO loc x + | VernacDeclareInstances _ as x -> xmlTODO ?loc x - | VernacDeclareClass _ as x -> xmlTODO loc x + | VernacDeclareClass _ as x -> xmlTODO ?loc x (* Modules and Module Types *) - | VernacDeclareModule _ as x -> xmlTODO loc x - | VernacDefineModule _ as x -> xmlTODO loc x - | VernacDeclareModuleType _ as x -> xmlTODO loc x - | VernacInclude _ as x -> xmlTODO loc x + | VernacDeclareModule _ as x -> xmlTODO ?loc x + | VernacDefineModule _ as x -> xmlTODO ?loc x + | VernacDeclareModuleType _ as x -> xmlTODO ?loc x + | VernacInclude _ as x -> xmlTODO ?loc x (* Solving *) | (VernacSolveExistential _) as x -> - xmlLtac loc [PCData (Pp.string_of_ppcmds (Ppvernac.pr_vernac x))] + xmlLtac ?loc [PCData (Pp.string_of_ppcmds (Ppvernac.pr_vernac x))] (* Auxiliary file and library management *) | VernacAddLoadPath (recf,name,None) -> - xmlAddLoadPath loc ~attr:["rec",string_of_bool recf;"path",name] [] + xmlAddLoadPath ?loc ~attr:["rec",string_of_bool recf;"path",name] [] | VernacAddLoadPath (recf,name,Some dp) -> - xmlAddLoadPath loc ~attr:["rec",string_of_bool recf;"path",name] + xmlAddLoadPath ?loc ~attr:["rec",string_of_bool recf;"path",name] [PCData (Names.DirPath.to_string dp)] - | VernacRemoveLoadPath name -> xmlRemoveLoadPath loc ~attr:["path",name] [] + | VernacRemoveLoadPath name -> xmlRemoveLoadPath ?loc ~attr:["path",name] [] | VernacAddMLPath (recf,name) -> - xmlAddMLPath loc ~attr:["rec",string_of_bool recf;"path",name] [] - | VernacDeclareMLModule sl -> xmlDeclareMLModule loc sl - | VernacChdir _ as x -> xmlTODO loc x + xmlAddMLPath ?loc ~attr:["rec",string_of_bool recf;"path",name] [] + | VernacDeclareMLModule sl -> xmlDeclareMLModule ?loc sl + | VernacChdir _ as x -> xmlTODO ?loc x (* State management *) - | VernacWriteState _ as x -> xmlTODO loc x - | VernacRestoreState _ as x -> xmlTODO loc x + | VernacWriteState _ as x -> xmlTODO ?loc x + | VernacRestoreState _ as x -> xmlTODO ?loc x (* Resetting *) - | VernacResetName _ as x -> xmlTODO loc x - | VernacResetInitial as x -> xmlTODO loc x - | VernacBack _ as x -> xmlTODO loc x + | VernacResetName _ as x -> xmlTODO ?loc x + | VernacResetInitial as x -> xmlTODO ?loc x + | VernacBack _ as x -> xmlTODO ?loc x | VernacBackTo _ -> PCData "VernacBackTo" (* Commands *) - | VernacCreateHintDb _ as x -> xmlTODO loc x - | VernacRemoveHints _ as x -> xmlTODO loc x - | VernacHints _ as x -> xmlTODO loc x + | VernacCreateHintDb _ as x -> xmlTODO ?loc x + | VernacRemoveHints _ as x -> xmlTODO ?loc x + | VernacHints _ as x -> xmlTODO ?loc x | VernacSyntacticDefinition ((_, name), (idl, ce), _, _) -> let name = Id.to_string name in let attrs = List.map (fun id -> ("id", Id.to_string id)) idl in - xmlNotation attrs name loc [pp_expr ce] - | VernacDeclareImplicits _ as x -> xmlTODO loc x - | VernacArguments _ as x -> xmlTODO loc x - | VernacArgumentsScope _ as x -> xmlTODO loc x - | VernacReserve _ as x -> xmlTODO loc x - | VernacGeneralizable _ as x -> xmlTODO loc x - | VernacSetOpacity _ as x -> xmlTODO loc x - | VernacSetStrategy _ as x -> xmlTODO loc x - | VernacUnsetOption _ as x -> xmlTODO loc x - | VernacSetOption _ as x -> xmlTODO loc x - | VernacSetAppendOption _ as x -> xmlTODO loc x - | VernacAddOption _ as x -> xmlTODO loc x - | VernacRemoveOption _ as x -> xmlTODO loc x - | VernacMemOption _ as x -> xmlTODO loc x - | VernacPrintOption _ as x -> xmlTODO loc x - | VernacCheckMayEval (_,_,e) -> xmlCheck loc [pp_expr e] - | VernacGlobalCheck _ as x -> xmlTODO loc x - | VernacDeclareReduction _ as x -> xmlTODO loc x - | VernacPrint _ as x -> xmlTODO loc x - | VernacSearch _ as x -> xmlTODO loc x - | VernacLocate _ as x -> xmlTODO loc x - | VernacRegister _ as x -> xmlTODO loc x + xmlNotation ?loc attrs name [pp_expr ce] + | VernacDeclareImplicits _ as x -> xmlTODO ?loc x + | VernacArguments _ as x -> xmlTODO ?loc x + | VernacArgumentsScope _ as x -> xmlTODO ?loc x + | VernacReserve _ as x -> xmlTODO ?loc x + | VernacGeneralizable _ as x -> xmlTODO ?loc x + | VernacSetOpacity _ as x -> xmlTODO ?loc x + | VernacSetStrategy _ as x -> xmlTODO ?loc x + | VernacUnsetOption _ as x -> xmlTODO ?loc x + | VernacSetOption _ as x -> xmlTODO ?loc x + | VernacSetAppendOption _ as x -> xmlTODO ?loc x + | VernacAddOption _ as x -> xmlTODO ?loc x + | VernacRemoveOption _ as x -> xmlTODO ?loc x + | VernacMemOption _ as x -> xmlTODO ?loc x + | VernacPrintOption _ as x -> xmlTODO ?loc x + | VernacCheckMayEval (_,_,e) -> xmlCheck ?loc [pp_expr e] + | VernacGlobalCheck _ as x -> xmlTODO ?loc x + | VernacDeclareReduction _ as x -> xmlTODO ?loc x + | VernacPrint _ as x -> xmlTODO ?loc x + | VernacSearch _ as x -> xmlTODO ?loc x + | VernacLocate _ as x -> xmlTODO ?loc x + | VernacRegister _ as x -> xmlTODO ?loc x | VernacComments (cl) -> - xmlComment loc (List.flatten (List.map pp_comment cl)) + xmlComment ?loc (List.flatten (List.map pp_comment cl)) (* Stm backdoor *) - | VernacStm _ as x -> xmlTODO loc x + | VernacStm _ as x -> xmlTODO ?loc x (* Proof management *) - | VernacGoal _ as x -> xmlTODO loc x - | VernacAbort _ as x -> xmlTODO loc x + | VernacGoal _ as x -> xmlTODO ?loc x + | VernacAbort _ as x -> xmlTODO ?loc x | VernacAbortAll -> PCData "VernacAbortAll" - | VernacRestart as x -> xmlTODO loc x - | VernacUndo _ as x -> xmlTODO loc x - | VernacUndoTo _ as x -> xmlTODO loc x - | VernacBacktrack _ as x -> xmlTODO loc x - | VernacFocus _ as x -> xmlTODO loc x - | VernacUnfocus as x -> xmlTODO loc x - | VernacUnfocused as x -> xmlTODO loc x - | VernacBullet _ as x -> xmlTODO loc x - | VernacSubproof _ as x -> xmlTODO loc x - | VernacEndSubproof as x -> xmlTODO loc x - | VernacShow _ as x -> xmlTODO loc x - | VernacCheckGuard as x -> xmlTODO loc x + | VernacRestart as x -> xmlTODO ?loc x + | VernacUndo _ as x -> xmlTODO ?loc x + | VernacUndoTo _ as x -> xmlTODO ?loc x + | VernacBacktrack _ as x -> xmlTODO ?loc x + | VernacFocus _ as x -> xmlTODO ?loc x + | VernacUnfocus as x -> xmlTODO ?loc x + | VernacUnfocused as x -> xmlTODO ?loc x + | VernacBullet _ as x -> xmlTODO ?loc x + | VernacSubproof _ as x -> xmlTODO ?loc x + | VernacEndSubproof as x -> xmlTODO ?loc x + | VernacShow _ as x -> xmlTODO ?loc x + | VernacCheckGuard as x -> xmlTODO ?loc x | VernacProof (tac,using) -> let tac = None (** FIXME *) in let using = Option.map (xmlSectionSubsetDescr "using") using in - xmlProof loc (Option.List.(cons tac (cons using []))) - | VernacProofMode name -> xmlProofMode loc name + xmlProof ?loc (Option.List.(cons tac (cons using []))) + | VernacProofMode name -> xmlProofMode ?loc name (* Toplevel control *) - | VernacToplevelControl _ as x -> xmlTODO loc x + | VernacToplevelControl _ as x -> xmlTODO ?loc x (* For extension *) | VernacExtend _ as x -> - xmlExtend loc [PCData (Pp.string_of_ppcmds (Ppvernac.pr_vernac x))] + xmlExtend ?loc [PCData (Pp.string_of_ppcmds (Ppvernac.pr_vernac x))] (* Flags *) - | VernacProgram e -> xmlApply loc (Element("program",[],[]) :: [tmpp e loc]) + | VernacProgram e -> xmlApply ?loc (Element("program",[],[]) :: [tmpp ?loc e]) | VernacLocal (b,e) -> - xmlApply loc (Element("local",["flag",string_of_bool b],[]) :: - [tmpp e loc]) + xmlApply ?loc (Element("local",["flag",string_of_bool b],[]) :: + [tmpp ?loc e]) -let tmpp v loc = - match tmpp v loc with +let tmpp ?loc v = + match tmpp ?loc v with | Element("ltac",_,_) as x -> x - | xml -> xmlGallina loc [xml] + | xml -> xmlGallina ?loc [xml] diff --git a/ide/texmacspp.mli b/ide/texmacspp.mli index 858847fb6..c1086a633 100644 --- a/ide/texmacspp.mli +++ b/ide/texmacspp.mli @@ -9,4 +9,4 @@ open Xml_datatype open Vernacexpr -val tmpp : vernac_expr -> Loc.t -> xml +val tmpp : ?loc:Loc.t -> vernac_expr -> xml diff --git a/ide/xmlprotocol.ml b/ide/xmlprotocol.ml index bf52b0b52..53eb1a95f 100644 --- a/ide/xmlprotocol.ml +++ b/ide/xmlprotocol.ml @@ -863,7 +863,7 @@ let to_feedback_content = do_match "feedback_content" (fun s a -> match s,a with | "workerstatus", [ns] -> let n, s = to_pair to_string to_string ns in WorkerStatus(n,s) - | "custom", [loc;name;x]-> Custom (to_loc loc, to_string name, x) + | "custom", [loc;name;x]-> Custom (to_option to_loc loc, to_string name, x) | "filedependency", [from; dep] -> FileDependency (to_option to_string from, to_string dep) | "fileloaded", [dirpath; filename] -> @@ -896,7 +896,7 @@ let of_feedback_content = function constructor "feedback_content" "workerstatus" [of_pair of_string of_string (n,s)] | Custom (loc, name, x) -> - constructor "feedback_content" "custom" [of_loc loc; of_string name; x] + constructor "feedback_content" "custom" [of_option of_loc loc; of_string name; x] | FileDependency (from, depends_on) -> constructor "feedback_content" "filedependency" [ of_option of_string from; diff --git a/interp/constrexpr_ops.ml b/interp/constrexpr_ops.ml index 61115c00b..4b61ab494 100644 --- a/interp/constrexpr_ops.ml +++ b/interp/constrexpr_ops.ml @@ -242,13 +242,12 @@ let local_binder_loc = function | CLocalPattern (loc,_) -> loc let local_binders_loc bll = match bll with - | [] -> Loc.ghost - | h :: l -> - Loc.merge (local_binder_loc h) (local_binder_loc (List.last bll)) + | [] -> None + | h :: l -> Some (Loc.merge (local_binder_loc h) (local_binder_loc (List.last bll))) (** Pseudo-constructors *) -let mkIdentC id = Loc.tag @@ CRef (Ident (Loc.ghost, id),None) +let mkIdentC id = Loc.tag @@ CRef (Ident (Loc.tag id),None) let mkRefC r = Loc.tag @@ CRef (r,None) let mkCastC (a,k) = Loc.tag @@ CCast (a,k) let mkLambdaC (idl,bk,a,b) = Loc.tag @@ CLambdaN ([idl,bk,a],b) @@ -268,23 +267,23 @@ let add_name_in_env env n = let (fresh_var, fresh_var_hook) = Hook.make ~default:(fun _ _ -> assert false) () -let expand_binders ~loc mkC bl c = - let rec loop ~loc bl c = +let expand_binders ?loc mkC bl c = + let rec loop ?loc bl c = match bl with | [] -> ([], c) | b :: bl -> match b with | CLocalDef ((loc1,_) as n, oty, b) -> - let env, c = loop ~loc:(Loc.merge loc1 loc) bl c in + let env, c = loop ~loc:(Loc.opt_merge loc1 loc) bl c in let env = add_name_in_env env n in - (env, Loc.tag ~loc @@ CLetIn (n,oty,b,c)) + (env, Loc.tag ?loc @@ CLetIn (n,oty,b,c)) | CLocalAssum ((loc1,_)::_ as nl, bk, t) -> - let env, c = loop ~loc:(Loc.merge loc1 loc) bl c in + let env, c = loop ~loc:(Loc.opt_merge 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 + (env, mkC ?loc (nl,bk,t) c) + | CLocalAssum ([],_,_) -> loop ?loc bl c | CLocalPattern (loc1, (p, ty)) -> - let env, c = loop ~loc:(Loc.merge loc1 loc) bl c in + let env, c = loop ~loc:(Loc.opt_merge 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 @@ -292,27 +291,27 @@ let expand_binders ~loc mkC bl c = | 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 @@ + let c = Loc.tag ?loc @@ CCases (LetPatternStyle, None, [(e,None,None)], [(Loc.tag ~loc:loc1 ([(loc1,[p])], c))]) in - (ni :: env, mkC ~loc ([id],Default Explicit,ty) c) + (ni :: env, mkC ?loc ([id],Default Explicit,ty) c) in - let (_, c) = loop loc bl c in + let (_, c) = loop ?loc bl c in c -let mkCProdN ~loc bll c = - let mk ~loc b c = Loc.tag ~loc @@ CProdN ([b],c) in - expand_binders ~loc mk bll c +let mkCProdN ?loc bll c = + let mk ?loc b c = Loc.tag ?loc @@ CProdN ([b],c) in + expand_binders ?loc mk bll c -let mkCLambdaN ~loc bll c = - let mk ~loc b c = Loc.tag ~loc @@ CLambdaN ([b],c) in - expand_binders ~loc mk bll c +let mkCLambdaN ?loc bll c = + let mk ?loc b c = Loc.tag ?loc @@ CLambdaN ([b],c) in + expand_binders ?loc mk bll c (* Deprecated *) -let abstract_constr_expr c bl = mkCLambdaN (local_binders_loc bl) bl c -let prod_constr_expr c bl = mkCProdN (local_binders_loc bl) bl c +let abstract_constr_expr c bl = mkCLambdaN ?loc:(local_binders_loc bl) bl c +let prod_constr_expr c bl = mkCProdN ?loc:(local_binders_loc bl) bl c let coerce_reference_to_id = function | Ident (_,id) -> id diff --git a/interp/constrexpr_ops.mli b/interp/constrexpr_ops.mli index ae5ec2be5..82e4f54b0 100644 --- a/interp/constrexpr_ops.mli +++ b/interp/constrexpr_ops.mli @@ -37,7 +37,7 @@ val binder_kind_eq : binder_kind -> binder_kind -> bool 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 local_binders_loc : local_binder_expr list -> Loc.t +val local_binders_loc : local_binder_expr list -> Loc.t option (** {6 Constructors}*) @@ -49,10 +49,10 @@ val mkLambdaC : Name.t located list * binder_kind * constr_expr * constr_expr -> val mkLetInC : Name.t located * constr_expr * constr_expr option * constr_expr -> constr_expr val mkProdC : Name.t located list * binder_kind * constr_expr * constr_expr -> constr_expr -val mkCLambdaN : loc:Loc.t -> local_binder_expr list -> constr_expr -> constr_expr +val mkCLambdaN : ?loc:Loc.t -> local_binder_expr list -> constr_expr -> constr_expr (** Same as [abstract_constr_expr], with location *) -val mkCProdN : loc:Loc.t -> local_binder_expr list -> constr_expr -> constr_expr +val mkCProdN : ?loc:Loc.t -> local_binder_expr list -> constr_expr -> constr_expr (** Same as [prod_constr_expr], with location *) (** @deprecated variant of mkCLambdaN *) diff --git a/interp/constrextern.ml b/interp/constrextern.ml index 8d9f8552d..5960a6baa 100644 --- a/interp/constrextern.ml +++ b/interp/constrextern.ml @@ -146,13 +146,13 @@ let insert_delimiters e = function | None -> e | Some sc -> Loc.tag @@ CDelimiters (sc,e) -let insert_pat_delimiters loc p = function +let insert_pat_delimiters ?loc p = function | None -> p - | Some sc -> Loc.tag ~loc @@ CPatDelimiters (sc,p) + | Some sc -> Loc.tag ?loc @@ CPatDelimiters (sc,p) -let insert_pat_alias loc p = function +let insert_pat_alias ?loc p = function | Anonymous -> p - | Name id -> Loc.tag ~loc @@ CPatAlias (p,id) + | Name id -> Loc.tag ?loc @@ CPatAlias (p,id) (**********************************************************************) (* conversion of references *) @@ -163,15 +163,15 @@ let extern_evar n l = CEvar (n,l) For instance, in the debugger the tables of global references may be inaccurate *) -let default_extern_reference loc vars r = - Qualid (loc,shortest_qualid_of_global vars r) +let default_extern_reference ?loc vars r = + Qualid (Loc.tag ?loc @@ shortest_qualid_of_global vars r) let my_extern_reference = ref default_extern_reference let set_extern_reference f = my_extern_reference := f let get_extern_reference () = !my_extern_reference -let extern_reference loc vars l = !my_extern_reference loc vars l +let extern_reference ?loc vars l = !my_extern_reference ?loc vars l (**********************************************************************) (* mapping patterns to cases_pattern_expr *) @@ -266,16 +266,16 @@ let make_notation loc ntn (terms,termlists,binders as subst) = (fun (loc,p) -> Loc.tag ~loc @@ CPrim p) destPrim terms -let make_pat_notation loc ntn (terms,termlists as subst) args = - if not (List.is_empty termlists) then (loc, CPatNotation (ntn,subst,args)) else +let make_pat_notation ?loc ntn (terms,termlists as subst) args = + if not (List.is_empty termlists) then (Loc.tag ?loc @@ CPatNotation (ntn,subst,args)) else make_notation_gen loc ntn - (fun (loc,ntn,l) -> Loc.tag ~loc @@ CPatNotation (ntn,(l,[]),args)) - (fun (loc,p) -> Loc.tag ~loc @@ CPatPrim p) + (fun (loc,ntn,l) -> Loc.tag ?loc @@ CPatNotation (ntn,(l,[]),args)) + (fun (loc,p) -> Loc.tag ?loc @@ CPatPrim p) destPatPrim terms -let mkPat loc qid l = +let mkPat ?loc qid l = Loc.tag ?loc @@ (* Normally irrelevant test with v8 syntax, but let's do it anyway *) - if List.is_empty l then Loc.tag ~loc @@ CPatAtom (Some qid) else Loc.tag ~loc @@ CPatCstr (qid,None,l) + if List.is_empty l then CPatAtom (Some qid) else CPatCstr (qid,None,l) let pattern_printable_in_both_syntax (ind,_ as c) = let impl_st = extract_impargs_data (implicits_of_global (ConstructRef c)) in @@ -293,7 +293,7 @@ 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), []) | _ -> @@ -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; @@ -330,12 +330,12 @@ 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 [])) 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) @@ -345,8 +345,8 @@ let rec extern_cases_pattern_in_scope (scopes:local_scopes) vars pat = 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 -and apply_notation_to_pattern loc gr ((subst,substlist),(nb_to_drop,more_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 | NotationRule (sc,ntn) -> @@ -373,11 +373,11 @@ and apply_notation_to_pattern loc gr ((subst,substlist),(nb_to_drop,more_args)) |Some true_args -> true_args |None -> raise No_match in - insert_pat_delimiters loc - (make_pat_notation loc ntn (l,ll) l2') key + insert_pat_delimiters ?loc + (make_pat_notation ?loc ntn (l,ll) l2') key end | SynDefRule kn -> - let qid = Qualid (loc, shortest_qualid_of_syndef vars kn) in + let qid = Qualid (Loc.tag ?loc @@ shortest_qualid_of_syndef vars kn) in let l1 = List.rev_map (fun (c,(scopt,scl)) -> extern_cases_pattern_in_scope (scopt,scl@scopes) vars c) @@ -390,7 +390,7 @@ and apply_notation_to_pattern loc gr ((subst,substlist),(nb_to_drop,more_args)) |None -> raise No_match in assert (List.is_empty substlist); - mkPat loc qid (List.rev_append l1 l2') + mkPat ?loc qid (List.rev_append l1 l2') and extern_notation_pattern (tmp_scope,scopes as allscopes) vars (loc, t) = function | [] -> raise No_match | (keyrule,pat,n as _rule)::rules -> @@ -398,9 +398,9 @@ 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 + 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 @@ -411,7 +411,7 @@ let rec extern_notation_ind_pattern allscopes vars ind args = function | (keyrule,pat,n as _rule)::rules -> try if List.mem keyrule !print_non_active_notations then raise No_match; - apply_notation_to_pattern Loc.ghost (IndRef ind) + apply_notation_to_pattern (IndRef ind) (match_notation_constr_ind_pattern ind args pat) allscopes vars keyrule with No_match -> extern_notation_ind_pattern allscopes vars ind args rules @@ -420,7 +420,7 @@ let extern_ind_pattern_in_scope (scopes:local_scopes) vars ind args = (* pboutill: There are letins in pat which is incompatible with notations and not explicit application. *) if !Flags.in_debugger||Inductiveops.inductive_has_local_defs ind then - let c = extern_reference Loc.ghost vars (IndRef ind) in + let c = extern_reference vars (IndRef ind) in let args = List.map (extern_cases_pattern_in_scope scopes vars) args in Loc.tag @@ CPatCstr (c, Some (add_patt_for_params ind args), []) else @@ -430,14 +430,14 @@ let extern_ind_pattern_in_scope (scopes:local_scopes) vars ind args = match availability_of_prim_token p sc scopes with | None -> raise No_match | Some key -> - insert_pat_delimiters Loc.ghost (Loc.tag @@ CPatPrim p) key + insert_pat_delimiters (Loc.tag @@ CPatPrim p) key with No_match -> try if !Flags.raw_print || !print_no_symbol then raise No_match; extern_notation_ind_pattern scopes vars ind args (uninterp_ind_pattern_notations ind) with No_match -> - let c = extern_reference Loc.ghost vars (IndRef ind) in + let c = extern_reference vars (IndRef ind) in let args = List.map (extern_cases_pattern_in_scope scopes vars) args in match drop_implicits_in_patt (IndRef ind) 0 args with |Some true_args -> Loc.tag @@ CPatCstr (c, None, true_args) @@ -490,7 +490,7 @@ let explicitize inctx impl (cf,f) args = is_significant_implicit (Lazy.force a)) in if visible then - (Lazy.force a,Some (Loc.ghost, ExplByName (name_of_implicit imp))) :: tail + (Lazy.force a,Some (Loc.tag @@ ExplByName (name_of_implicit imp))) :: tail else tail | a::args, _::impl -> (Lazy.force a,None) :: exprec (q+1) (args,impl) @@ -615,9 +615,11 @@ let extern_optimal_prim_token scopes r r' = (* mapping decl *) let extended_glob_local_binder_of_decl loc = function - | (p,bk,None,t) -> GLocalAssum (loc,p,bk,t) - | (p,bk,Some x,(_,GHole ( _, Misctypes.IntroAnonymous, None))) -> GLocalDef (loc,p,bk,x,None) - | (p,bk,Some x,t) -> GLocalDef (loc,p,bk,x,Some t) + | (p,bk,None,t) -> GLocalAssum (p,bk,t) + | (p,bk,Some x,(_,GHole ( _, Misctypes.IntroAnonymous, None))) -> GLocalDef (p,bk,x,None) + | (p,bk,Some x,t) -> GLocalDef (p,bk,x,Some t) + +let extended_glob_local_binder_of_decl ?loc u = Loc.tag ?loc (extended_glob_local_binder_of_decl loc u) (**********************************************************************) (* mapping glob_constr to constr_expr *) @@ -645,7 +647,7 @@ let rec extern inctx scopes vars r = 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) @@ -699,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 @@ -707,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 rloc vars ref) (extern_universes us) args + (Some ref,extern_reference ~loc:rloc vars ref) (extern_universes us) args end | _ -> @@ -722,12 +724,12 @@ let rec extern inctx scopes vars r = | GProd (na,bk,t,c) -> let t = extern_typ scopes vars t in let (idl,c) = factorize_prod scopes (add_vname vars na) na bk t c in - CProdN ([(Loc.ghost,na)::idl,Default bk,t],c) + CProdN ([(Loc.tag na)::idl,Default bk,t],c) | GLambda (na,bk,t,c) -> let t = extern_typ scopes vars t in let (idl,c) = factorize_lambda inctx scopes (add_vname vars na) na bk t c in - CLambdaN ([(Loc.ghost,na)::idl,Default bk,t],c) + CLambdaN ([(Loc.tag na)::idl,Default bk,t],c) | GCases (sty,rtntypopt,tml,eqns) -> let vars' = @@ -741,12 +743,12 @@ let rec extern inctx scopes vars r = | None -> None | Some ntn -> if occur_glob_constr id ntn then - Some (Loc.ghost, Anonymous) + Some (Loc.tag Anonymous) else None end | Anonymous, _ -> None | Name id, (_, GVar id') when Id.equal id id' -> None - | Name _, _ -> Some (Loc.ghost,na) in + | Name _, _ -> Some (Loc.tag na) in (sub_extern false scopes vars tm, na', Option.map (fun (loc,(ind,nal)) -> @@ -760,15 +762,15 @@ let rec extern inctx scopes vars r = CCases (sty,rtntypopt',tml,eqns) | GLetTuple (nal,(na,typopt),tm,b) -> - CLetTuple (List.map (fun na -> (Loc.ghost,na)) nal, - (Option.map (fun _ -> (Loc.ghost,na)) typopt, + CLetTuple (List.map (fun na -> (Loc.tag na)) nal, + (Option.map (fun _ -> (Loc.tag na)) typopt, Option.map (extern_typ scopes (add_vname vars na)) typopt), sub_extern false scopes vars tm, extern inctx scopes (List.fold_left add_vname vars nal) b) | GIf (c,(na,typopt),b1,b2) -> CIf (sub_extern false scopes vars c, - (Option.map (fun _ -> (Loc.ghost,na)) typopt, + (Option.map (fun _ -> (Loc.tag na)) typopt, Option.map (extern_typ scopes (add_vname vars na)) typopt), sub_extern inctx scopes vars b1, sub_extern inctx scopes vars b2) @@ -779,28 +781,28 @@ 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 let n = match fst nv.(i) with | None -> None - | Some x -> Some (Loc.ghost, out_name (List.nth assums x)) + | Some x -> Some (Loc.tag @@ out_name (List.nth assums x)) in let ro = extern_recursion_order scopes vars (snd nv.(i)) in - ((Loc.ghost, fi), (n, ro), bl, extern_typ scopes vars0 ty, + ((Loc.tag fi), (n, ro), bl, extern_typ scopes vars0 ty, extern false scopes vars1 def)) idv in CFix ((loc,idv.(n)),Array.to_list listdecl) | 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 - ((Loc.ghost, fi),bl,extern_typ scopes vars0 tyv.(i), + ((Loc.tag fi),bl,extern_typ scopes vars0 tyv.(i), sub_extern false scopes vars1 bv.(i))) idv in CCoFix ((loc,idv.(n)),Array.to_list listdecl)) @@ -841,14 +843,14 @@ and factorize_lambda inctx scopes vars na bk aty c = and extern_local_binder scopes vars = function [] -> ([],[],[]) - | GLocalDef (_,na,bk,bd,ty)::l -> + | (_, GLocalDef (na,bk,bd,ty))::l -> let (assums,ids,l) = extern_local_binder scopes (name_fold Id.Set.add na vars) l in (assums,na::ids, - CLocalDef((Loc.ghost,na), extern false scopes vars bd, + CLocalDef((Loc.tag na), extern false scopes vars bd, Option.map (extern false scopes vars) ty) :: l) - | GLocalAssum (_,na,bk,ty)::l -> + | (_, GLocalAssum (na,bk,ty))::l -> let ty = extern_typ scopes vars ty in (match extern_local_binder scopes (name_fold Id.Set.add na vars) l with (assums,ids,CLocalAssum(nal,k,ty')::l) @@ -856,12 +858,12 @@ and extern_local_binder scopes vars = function match na with Name id -> not (occur_var_constr_expr id ty') | _ -> true -> (na::assums,na::ids, - CLocalAssum((Loc.ghost,na)::nal,k,ty')::l) + CLocalAssum((Loc.tag na)::nal,k,ty')::l) | (assums,ids,l) -> (na::assums,na::ids, - CLocalAssum([(Loc.ghost,na)],Default bk,ty) :: l)) + CLocalAssum([(Loc.tag na)],Default bk,ty) :: l)) - | GLocalPattern (_,(p,_),_,bk,ty)::l -> + | (_, GLocalPattern ((p,_),_,bk,ty))::l -> let ty = if !Flags.raw_print then Some (extern_typ scopes vars ty) else None in let p = extern_cases_pattern vars p in @@ -1078,5 +1080,5 @@ let extern_rel_context where env sigma sign = let where = Option.map EConstr.of_constr where in let a = detype_rel_context where [] (names_of_rel_context env,env) sigma sign in let vars = vars_of_env env in - let a = List.map (extended_glob_local_binder_of_decl Loc.ghost) a in + let a = List.map (extended_glob_local_binder_of_decl) a in pi3 (extern_local_binder (None,[]) vars a) diff --git a/interp/constrextern.mli b/interp/constrextern.mli index b39339450..ea627cff1 100644 --- a/interp/constrextern.mli +++ b/interp/constrextern.mli @@ -37,7 +37,7 @@ val extern_closed_glob : ?lax:bool -> bool -> env -> Evd.evar_map -> closed_glob val extern_constr : ?lax:bool -> bool -> env -> Evd.evar_map -> constr -> constr_expr val extern_constr_in_scope : bool -> scope_name -> env -> Evd.evar_map -> constr -> constr_expr -val extern_reference : Loc.t -> Id.Set.t -> global_reference -> reference +val extern_reference : ?loc:Loc.t -> Id.Set.t -> global_reference -> reference val extern_type : bool -> env -> Evd.evar_map -> types -> constr_expr val extern_sort : Evd.evar_map -> sorts -> glob_sort val extern_rel_context : constr option -> env -> Evd.evar_map -> @@ -55,9 +55,9 @@ val print_projections : bool ref (** Customization of the global_reference printer *) val set_extern_reference : - (Loc.t -> Id.Set.t -> global_reference -> reference) -> unit + (?loc:Loc.t -> Id.Set.t -> global_reference -> reference) -> unit val get_extern_reference : - unit -> (Loc.t -> Id.Set.t -> global_reference -> reference) + unit -> (?loc:Loc.t -> Id.Set.t -> global_reference -> reference) (** This governs printing of implicit arguments. If [with_implicits] is on and not [with_arguments] then implicit args are printed prefixed diff --git a/interp/constrintern.ml b/interp/constrintern.ml index d1b931a22..585f03808 100644 --- a/interp/constrintern.ml +++ b/interp/constrintern.ml @@ -345,13 +345,13 @@ let rec check_capture ty = function | [] -> () -let locate_if_hole loc na = function +let locate_if_hole ?loc na = function | _, GHole (_,naming,arg) -> (try match na with - | Name id -> glob_constr_of_notation_constr loc + | Name id -> glob_constr_of_notation_constr ?loc (Reserve.find_reserved_type id) | Anonymous -> raise Not_found - with Not_found -> Loc.tag ~loc @@ GHole (Evar_kinds.BinderType na, naming, arg)) + with Not_found -> Loc.tag ?loc @@ GHole (Evar_kinds.BinderType na, naming, arg)) | x -> x let reset_hidden_inductive_implicit_test env = @@ -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,(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 @@ -454,27 +454,28 @@ let intern_local_pattern intern lvar env p = env) env (free_vars_of_pat [] p) -let glob_local_binder_of_extended = function - | GLocalAssum (loc,na,bk,t) -> (na,bk,None,t) - | GLocalDef (loc,na,bk,c,Some t) -> (na,bk,Some c,t) - | GLocalDef (loc,na,bk,c,None) -> +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 (na,bk,Some c,t) - | GLocalPattern (loc,_,_,_,_) -> + | GLocalPattern (_,_,_,_) -> Loc.raise ~loc (Stream.Error "pattern with quote not allowed here.") + ) 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)) -> GLocalAssum (loc,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, - GLocalDef (loc,na,Explicit,term,ty) :: bl) + (Loc.tag ~loc @@ GLocalDef (na,Explicit,term,ty)) :: bl) | CLocalPattern (loc,(p,ty)) -> let tyc = match ty with @@ -494,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, GLocalPattern(loc,(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 @@ -582,13 +583,13 @@ let make_letins = 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 *) - | GLocalDef (loc,na,_,b,t)::l -> + | (loc, GLocalDef (na,_,b,t))::l -> subordinate_letins (LPLetIn (loc,(na,b,t))::letins) l - | GLocalAssum (loc,na,bk,t)::l -> + | (loc, GLocalAssum (na,bk,t))::l -> let letins',rest = subordinate_letins [] l in letins',((loc,(na,bk,t)),letins)::rest - | GLocalPattern (loc,u,id,bk,t) :: l -> - subordinate_letins (LPCases (loc,u,id)::letins) ([GLocalAssum (loc,Name id,bk,t)] @ l) + | (loc, GLocalPattern (u,id,bk,t)) :: l -> + subordinate_letins (LPCases (loc,u,id)::letins) ([Loc.tag ~loc @@ GLocalAssum (Name id,bk,t)] @ l) | [] -> letins,[] @@ -602,11 +603,11 @@ let terms_of_binders bl = 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 - | GLocalAssum (loc,Name id,_,_)::l -> (Loc.tag ~loc @@ CRef (Ident (loc,id), None)) :: extract_variables l - | GLocalDef (loc,Name id,_,_,_)::l -> extract_variables l - | GLocalDef (loc,Anonymous,_,_,_)::l - | GLocalAssum (loc,Anonymous,_,_)::l -> error "Cannot turn \"_\" into a term." - | GLocalPattern (loc,(u,_),_,_,_) :: l -> term_of_pat u :: 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." + | (loc, GLocalPattern ((u,_),_,_,_)) :: l -> term_of_pat u :: extract_variables l | [] -> [] in extract_variables bl @@ -697,7 +698,7 @@ let instantiate_notation_constr loc intern ntnvars subst infos 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 *) @@ -728,7 +729,7 @@ 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 + 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 @@ -809,8 +810,8 @@ let find_appl_head_data c = List.skipn_at_least n scopes,[] | _ -> c,[],[],[] -let error_not_enough_arguments loc = - user_err ~loc (str "Abbreviation is not applied enough.") +let error_not_enough_arguments ?loc = + user_err ?loc (str "Abbreviation is not applied enough.") let check_no_explicitation l = let is_unset (a, b) = match b with None -> false | Some _ -> true in @@ -843,7 +844,7 @@ let intern_qualid loc qid intern env lvar us 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 @@ -893,7 +894,7 @@ let intern_applied_reference intern env namedctx (_, ntnvars as lvar) us args = let interp_reference vars r = let (r,_,_,_),_ = - intern_applied_reference (fun _ -> error_not_enough_arguments Loc.ghost) + intern_applied_reference (fun _ -> error_not_enough_arguments ?loc:None) {ids = Id.Set.empty; unb = false ; tmp_scope = None; scopes = []; impls = empty_internalization_env} [] (vars, Id.Map.empty) None [] r @@ -990,10 +991,10 @@ let add_implicits_check_length fail nargs nargs_with_letin impls_st len_pl1 pl2 else Int.equal args_len nargs_with_letin || (fst (fail (nargs - List.length impl_list + i)))) ,l) |imp::q as il,[] -> if is_status_implicit imp && maximal_insertion_of imp - then let (b,out) = aux i (q,[]) in (b,(Loc.ghost,RCPatAtom(None))::out) + then let (b,out) = aux i (q,[]) in (b,(Loc.tag @@ RCPatAtom(None))::out) else fail (remaining_args (len_pl1+i) il) |imp::q,(hh::tt as l) -> if is_status_implicit imp - then let (b,out) = aux i (q,l) in (b,(Loc.ghost, RCPatAtom(None))::out) + then let (b,out) = aux i (q,l) in (b,(Loc.tag @@ RCPatAtom(None))::out) else let (b,out) = aux (succ i) (q,tt) in (b,hh::out) in aux 0 (impl_list,pl2) @@ -1239,7 +1240,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 @@ -1288,20 +1289,20 @@ let drop_notations_pattern looked_for = 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; 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 @@ -1540,7 +1541,9 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = let before, after = split_at_annot bl n in let (env',rbefore) = List.fold_left intern_local_binder (env,[]) before in let ro = f (intern env') in - let n' = Option.map (fun _ -> List.count (function GLocalAssum _ -> true | _ -> false (* remove let-ins *)) rbefore) n in + let n' = Option.map (fun _ -> List.count (function | _, GLocalAssum _ -> true + | _ -> false (* remove let-ins *)) + rbefore) n in n', ro, List.fold_left intern_local_binder (env',rbefore) after in let n, ro, (env',rbl) = @@ -1559,7 +1562,7 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = let (_,bli,tyi,_) = idl_temp.(i) in let fix_args = (List.map (fun (na, bk, _, _) -> (build_impls bk na)) bli) in push_name_env ntnvars (impls_type_list ~args:fix_args tyi) - en (Loc.ghost, Name name)) 0 env' lf in + 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 @@ GRec (GFix @@ -1586,7 +1589,7 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = let (bli,tyi,_) = idl_tmp.(i) in let cofix_args = List.map (fun (na, bk, _, _) -> (build_impls bk na)) bli in push_name_env ntnvars (impls_type_list ~args:cofix_args tyi) - en (Loc.ghost, Name name)) 0 env' lf in + 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 @@ GRec (GCoFix n, @@ -1617,10 +1620,10 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = | 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 @@ -1679,7 +1682,7 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = (tm,ind)::inds, Option.fold_right Id.Set.add extra_id ex_ids, List.rev_append match_td matchs) tms ([],Id.Set.empty,[]) in let env' = Id.Set.fold - (fun var bli -> push_name_env ntnvars (Variable,[],[],[]) bli (Loc.ghost,Name var)) + (fun var bli -> push_name_env ntnvars (Variable,[],[],[]) bli (Loc.tag @@ Name var)) (Id.Set.union ex_ids as_in_vars) (reset_hidden_inductive_implicit_test env) in (* PatVars before a real pattern do not need to be matched *) let stripped_match_from_in = @@ -1715,7 +1718,7 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = let ((b',(na',_)),_,_) = intern_case_item env' Id.Set.empty (b,na,None) in let p' = Option.map (fun u -> let env'' = push_name_env ntnvars (Variable,[],[],[]) (reset_hidden_inductive_implicit_test env') - (Loc.ghost,na') in + (Loc.tag na') in intern_type env'' u) po in Loc.tag ~loc @@ GLetTuple (List.map snd nal, (na', p'), b', @@ -1725,7 +1728,7 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = let ((c',(na',_)),_,_) = intern_case_item env' Id.Set.empty (c,na,None) in (* no "in" no match to ad too *) let p' = Option.map (fun p -> let env'' = push_name_env ntnvars (Variable,[],[],[]) (reset_hidden_inductive_implicit_test env) - (Loc.ghost,na') in + (Loc.tag na') in intern_type env'' p) po in Loc.tag ~loc @@ GIf (c', (na', p'), intern env b1, intern env b2) @@ -1779,7 +1782,7 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = ) and intern_type env = intern (set_type_scope env) - and intern_local_binder env bind = + and intern_local_binder env bind : intern_env * Glob_term.extended_glob_local_binder list = intern_local_binder_aux intern ntnvars env bind (* Expands a multiple pattern into a disjunction of multiple patterns *) @@ -1815,7 +1818,7 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = let extra_id,na = match tm', na with | (loc , GVar id), None when not (Id.Map.mem id (snd lvar)) -> Some id,(loc,Name id) | (loc, GRef (VarRef id, _)), None -> Some id,(loc,Name id) - | _, None -> None,(Loc.ghost,Anonymous) + | _, None -> None,(Loc.tag Anonymous) | _, Some (loc,na) -> None,(loc,na) in (* the "in" part *) let match_td,typ = match t with @@ -1837,7 +1840,7 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = match case_rel_ctxt,arg_pats with (* LetIn in the rel_context *) | LocalDef _ :: t, l when not with_letin -> - canonize_args t l forbidden_names match_acc ((Loc.ghost,Anonymous)::var_acc) + canonize_args t l forbidden_names match_acc ((Loc.tag Anonymous)::var_acc) | [],[] -> (add_name match_acc na, var_acc) | _::t, (loc, PatVar x)::tt -> @@ -2052,12 +2055,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_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_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 @@ -2084,7 +2087,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_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/implicit_quantifiers.ml b/interp/implicit_quantifiers.ml index 51152bb24..fa7712bdc 100644 --- a/interp/implicit_quantifiers.ml +++ b/interp/implicit_quantifiers.ml @@ -252,7 +252,7 @@ let combine_params avoid fn applied needed = let combine_params_freevar = fun avoid (_, decl) -> let id' = next_name_away_from (RelDecl.get_name decl) avoid in - (Loc.tag @@ CRef (Ident (Loc.ghost, id'),None), Id.Set.add id' avoid) + (Loc.tag @@ CRef (Ident (Loc.tag id'),None), Id.Set.add id' avoid) let destClassApp (loc, cl) = match cl with diff --git a/interp/notation.ml b/interp/notation.ml index 3bcec3001..150be040f 100644 --- a/interp/notation.ml +++ b/interp/notation.ml @@ -220,10 +220,10 @@ let remove_delimiters scope = with Not_found -> assert false (* A delimiter for scope [scope] should exist *) -let find_delimiters_scope loc key = +let find_delimiters_scope ?loc key = try String.Map.find key !delimiters_map with Not_found -> - user_err ~loc ~hdr:"find_delimiters" + user_err ?loc ~hdr:"find_delimiters" (str "Unknown scope delimiting key " ++ str key ++ str ".") (* Uninterpretation tables *) @@ -291,7 +291,7 @@ let notation_constr_key = function (* Rem: NApp(NRef ref,[]) stands for @ref *) type required_module = full_path * string list type 'a prim_token_interpreter = - Loc.t -> 'a -> glob_constr + ?loc:Loc.t -> 'a -> glob_constr type cases_pattern_status = bool (* true = use prim token in patterns *) @@ -299,7 +299,7 @@ type 'a prim_token_uninterpreter = glob_constr list * (glob_constr -> 'a option) * cases_pattern_status type internal_prim_token_interpreter = - Loc.t -> prim_token -> required_module * (unit -> glob_constr) + ?loc:Loc.t -> prim_token -> required_module * (unit -> glob_constr) let prim_token_interpreter_tab = (Hashtbl.create 7 : (scope_name, internal_prim_token_interpreter) Hashtbl.t) @@ -309,7 +309,7 @@ let add_prim_token_interpreter sc interp = let cont = Hashtbl.find prim_token_interpreter_tab sc in Hashtbl.replace prim_token_interpreter_tab sc (interp cont) with Not_found -> - let cont = (fun _loc _p -> raise Not_found) in + let cont = (fun ?loc _p -> raise Not_found) in Hashtbl.add prim_token_interpreter_tab sc (interp cont) let declare_prim_token_interpreter sc interp (patl,uninterp,b) = @@ -325,22 +325,22 @@ let mkString = function | None -> None | Some s -> if Unicode.is_utf8 s then Some (String s) else None -let delay dir int loc x = (dir, (fun () -> int loc x)) +let delay dir int ?loc x = (dir, (fun () -> int ?loc x)) let declare_numeral_interpreter sc dir interp (patl,uninterp,inpat) = declare_prim_token_interpreter sc - (fun cont loc -> function Numeral n-> delay dir interp loc n | p -> cont loc p) + (fun cont ?loc -> function Numeral n-> delay dir interp ?loc n | p -> cont ?loc p) (patl, (fun r -> Option.map mkNumeral (uninterp r)), inpat) let declare_string_interpreter sc dir interp (patl,uninterp,inpat) = declare_prim_token_interpreter sc - (fun cont loc -> function String s -> delay dir interp loc s | p -> cont loc p) + (fun cont ?loc -> function String s -> delay dir interp ?loc s | p -> cont ?loc p) (patl, (fun r -> mkString (uninterp r)), inpat) -let check_required_module loc sc (sp,d) = +let check_required_module ?loc sc (sp,d) = try let _ = Nametab.global_of_path sp in () with Not_found -> - user_err ~loc ~hdr:"prim_token_interpreter" + user_err ?loc ~hdr:"prim_token_interpreter" (str "Cannot interpret in " ++ str sc ++ str " without requiring first module " ++ str (List.last d) ++ str ".") (* Look if some notation or numeral printer in [scope] can be used in @@ -445,23 +445,23 @@ let notation_of_prim_token = function | Numeral n -> "- "^(to_string (neg n)) | String _ -> raise Not_found -let find_prim_token g loc p sc = +let find_prim_token ?loc g p sc = (* Try for a user-defined numerical notation *) try let (_,c),df = find_notation (notation_of_prim_token p) sc in - g (Notation_ops.glob_constr_of_notation_constr loc c),df + g (Notation_ops.glob_constr_of_notation_constr ?loc c),df with Not_found -> (* Try for a primitive numerical notation *) - let (spdir,interp) = Hashtbl.find prim_token_interpreter_tab sc loc p in - check_required_module loc sc spdir; + let (spdir,interp) = (Hashtbl.find prim_token_interpreter_tab sc) ?loc p in + check_required_module ?loc sc spdir; g (interp ()), ((dirpath (fst spdir),DirPath.empty),"") -let interp_prim_token_gen g loc p local_scopes = +let interp_prim_token_gen g ?loc p local_scopes = let scopes = make_current_scopes local_scopes in let p_as_ntn = try notation_of_prim_token p with Not_found -> "" in - try find_interpretation p_as_ntn (find_prim_token g loc p) scopes + try find_interpretation p_as_ntn (find_prim_token ?loc g p) scopes with Not_found -> - user_err ~loc ~hdr:"interp_prim_token" + user_err ?loc ~hdr:"interp_prim_token" ((match p with | Numeral n -> str "No interpretation for numeral " ++ str (to_string n) | String s -> str "No interpretation for string " ++ qs s) ++ str ".") @@ -480,14 +480,14 @@ let rec rcp_of_glob looked_for gt = Loc.map (function | _ -> raise Not_found ) gt -let interp_prim_token_cases_pattern_expr loc looked_for p = - interp_prim_token_gen (rcp_of_glob looked_for) loc p +let interp_prim_token_cases_pattern_expr ?loc looked_for p = + interp_prim_token_gen (rcp_of_glob looked_for) ?loc p -let interp_notation loc ntn local_scopes = +let interp_notation ?loc ntn local_scopes = let scopes = make_current_scopes local_scopes in try find_interpretation ntn (find_notation ntn) scopes with Not_found -> - user_err ~loc + user_err ?loc (str "Unknown interpretation for notation \"" ++ str ntn ++ str "\".") let uninterp_notations c = @@ -541,7 +541,7 @@ let uninterp_prim_token_cases_pattern c = let availability_of_prim_token n printer_scope local_scopes = let f scope = - try ignore (Hashtbl.find prim_token_interpreter_tab scope Loc.ghost n); true + try ignore ((Hashtbl.find prim_token_interpreter_tab scope) n); true with Not_found -> false in let scopes = make_current_scopes local_scopes in Option.map snd (find_without_delimiters f (Some printer_scope,None) scopes) @@ -823,7 +823,7 @@ let pr_scope_classes sc = let pr_notation_info prglob ntn c = str "\"" ++ str ntn ++ str "\" := " ++ - prglob (Notation_ops.glob_constr_of_notation_constr Loc.ghost c) + prglob (Notation_ops.glob_constr_of_notation_constr c) let pr_named_scope prglob scope sc = (if String.equal scope default_scope then @@ -891,25 +891,25 @@ let global_reference_of_notation test (ntn,(sc,c,_)) = Some (ntn,sc,ref) | _ -> None -let error_ambiguous_notation loc _ntn = - user_err ~loc (str "Ambiguous notation.") +let error_ambiguous_notation ?loc _ntn = + user_err ?loc (str "Ambiguous notation.") -let error_notation_not_reference loc ntn = - user_err ~loc +let error_notation_not_reference ?loc ntn = + user_err ?loc (str "Unable to interpret " ++ quote (str ntn) ++ str " as a reference.") -let interp_notation_as_global_reference loc test ntn sc = +let interp_notation_as_global_reference ?loc test ntn sc = let scopes = match sc with | Some sc -> - let scope = find_scope (find_delimiters_scope Loc.ghost sc) in + let scope = find_scope (find_delimiters_scope sc) in String.Map.add sc scope String.Map.empty | None -> !scope_map in let ntns = browse_notation true ntn scopes in let refs = List.map (global_reference_of_notation test) ntns in match Option.List.flatten refs with | [_,_,ref] -> ref - | [] -> error_notation_not_reference loc ntn + | [] -> error_notation_not_reference ?loc ntn | refs -> let f (ntn,sc,ref) = let def = find_default ntn !scope_stack in @@ -919,8 +919,8 @@ let interp_notation_as_global_reference loc test ntn sc = in match List.filter f refs with | [_,_,ref] -> ref - | [] -> error_notation_not_reference loc ntn - | _ -> error_ambiguous_notation loc ntn + | [] -> error_notation_not_reference ?loc ntn + | _ -> error_ambiguous_notation ?loc ntn let locate_notation prglob ntn scope = let ntns = factorize_entries (browse_notation false ntn !scope_map) in diff --git a/interp/notation.mli b/interp/notation.mli index 2e92a00a8..10c7b85e4 100644 --- a/interp/notation.mli +++ b/interp/notation.mli @@ -55,7 +55,7 @@ val find_scope : scope_name -> scope val declare_delimiters : scope_name -> delimiters -> unit val remove_delimiters : scope_name -> unit -val find_delimiters_scope : Loc.t -> delimiters -> scope_name +val find_delimiters_scope : ?loc:Loc.t -> delimiters -> scope_name (** {6 Declare and uses back and forth an interpretation of primitive token } *) @@ -69,7 +69,7 @@ type required_module = full_path * string list type cases_pattern_status = bool (** true = use prim token in patterns *) type 'a prim_token_interpreter = - Loc.t -> 'a -> glob_constr + ?loc:Loc.t -> 'a -> glob_constr type 'a prim_token_uninterpreter = glob_constr list * (glob_constr -> 'a option) * cases_pattern_status @@ -83,9 +83,9 @@ val declare_string_interpreter : scope_name -> required_module -> (** Return the [term]/[cases_pattern] bound to a primitive token in a given scope context*) -val interp_prim_token : Loc.t -> prim_token -> local_scopes -> +val interp_prim_token : ?loc:Loc.t -> prim_token -> local_scopes -> glob_constr * (notation_location * scope_name option) -val interp_prim_token_cases_pattern_expr : Loc.t -> (global_reference -> unit) -> prim_token -> +val interp_prim_token_cases_pattern_expr : ?loc:Loc.t -> (global_reference -> unit) -> prim_token -> local_scopes -> raw_cases_pattern_expr * (notation_location * scope_name option) (** Return the primitive token associated to a [term]/[cases_pattern]; @@ -114,7 +114,7 @@ val declare_notation_interpretation : notation -> scope_name option -> val declare_uninterpretation : interp_rule -> interpretation -> unit (** Return the interpretation bound to a notation *) -val interp_notation : Loc.t -> notation -> local_scopes -> +val interp_notation : ?loc:Loc.t -> notation -> local_scopes -> interpretation * (notation_location * scope_name option) type notation_rule = interp_rule * interpretation * int option @@ -137,7 +137,7 @@ val level_of_notation : notation -> level (** raise [Not_found] if no level *) (** {6 Miscellaneous} *) -val interp_notation_as_global_reference : Loc.t -> (global_reference -> bool) -> +val interp_notation_as_global_reference : ?loc:Loc.t -> (global_reference -> bool) -> notation -> delimiters option -> global_reference (** Checks for already existing notations *) diff --git a/interp/notation_ops.ml b/interp/notation_ops.ml index 32c900504..32c564156 100644 --- a/interp/notation_ops.ml +++ b/interp/notation_ops.ml @@ -117,13 +117,13 @@ let name_to_ident = function let to_id g e id = let e,na = g e (Name id) in e,name_to_ident na -let rec cases_pattern_fold_map loc g e = Loc.with_unloc (function +let rec cases_pattern_fold_map ?loc g e = Loc.with_unloc (function | PatVar na -> - let e',na' = g e na in e', Loc.tag ~loc @@ PatVar na' + let e',na' = g e na in e', Loc.tag ?loc @@ PatVar na' | PatCstr (cstr,patl,na) -> let e',na' = g e na in - let e',patl' = List.fold_map (cases_pattern_fold_map loc g) e patl in - e', Loc.tag ~loc @@ PatCstr (cstr,patl',na') + let e',patl' = List.fold_map (cases_pattern_fold_map ?loc g) e patl in + e', Loc.tag ?loc @@ PatCstr (cstr,patl',na') ) let subst_binder_type_vars l = function @@ -152,8 +152,8 @@ let rec subst_glob_vars l gc = Loc.map (function let ldots_var = Id.of_string ".." -let glob_constr_of_notation_constr_with_binders loc g f e nc = - let lt x = Loc.tag ~loc x in lt @@ match nc with +let glob_constr_of_notation_constr_with_binders ?loc g f e nc = + let lt x = Loc.tag ?loc x in lt @@ match nc with | NVar id -> GVar id | NApp (a,args) -> GApp (f e a, List.map (f e) args) | NList (x,y,iter,tail,swap) -> @@ -181,13 +181,13 @@ let glob_constr_of_notation_constr_with_binders loc g f e nc = | Some (ind,nal) -> let e',nal' = List.fold_right (fun na (e',nal) -> let e',na' = g e' na in e',na'::nal) nal (e',[]) in - e',Some (loc,(ind,nal')) in + e',Some (Loc.tag ?loc (ind,nal')) in let e',na' = g e' na in (e',(f e tm,(na',t'))::tml')) tml (e,[]) in let fold (idl,e) na = let (e,na) = g e na in ((name_cons na idl,e),na) in let eqnl' = List.map (fun (patl,rhs) -> let ((idl,e),patl) = - List.fold_map (cases_pattern_fold_map loc fold) ([],e) patl in + List.fold_map (cases_pattern_fold_map ?loc fold) ([],e) patl in lt (idl,patl,f e rhs)) eqnl in GCases (sty,Option.map (f e') rtntypopt,tml',eqnl') | NLetTuple (nal,(na,po),b,c) -> @@ -208,9 +208,9 @@ let glob_constr_of_notation_constr_with_binders loc g f e nc = | NHole (x, naming, arg) -> GHole (x, naming, arg) | NRef x -> GRef (x,None) -let glob_constr_of_notation_constr loc x = +let glob_constr_of_notation_constr ?loc x = let rec aux () x = - glob_constr_of_notation_constr_with_binders loc (fun () id -> ((),id)) aux () x + glob_constr_of_notation_constr_with_binders ?loc (fun () id -> ((),id)) aux () x in aux () x (******************************************************************************) @@ -795,17 +795,17 @@ let bind_bindinglist_env alp (terms,onlybinders,termlists,binderlists as sigma) | (Some _ as x), None | None, (Some _ as x) -> x | None, None -> None in let unify_binding_kind bk bk' = if bk == bk' then bk' else raise No_match in - let unify_binder alp b b' = + let unify_binder alp (loc, b) (loc', b') = match b, b' with - | GLocalAssum (loc,na,bk,t), GLocalAssum (_,na',bk',t') -> + | GLocalAssum (na,bk,t), GLocalAssum (na',bk',t') -> let alp, na = unify_name alp na na' in - alp, GLocalAssum (loc, na, unify_binding_kind bk bk', unify_term alp t t') - | GLocalDef (loc,na,bk,c,t), GLocalDef (_,na',bk',c',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, GLocalDef (loc, na, unify_binding_kind bk bk', unify_term alp c c', unify_opt_term alp t t') - | GLocalPattern (loc,(p,ids),id,bk,t), GLocalPattern (_,(p',_),_,bk',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, GLocalPattern (loc, (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,18 +832,18 @@ 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 b' = + let unify_term_binder c (loc, b') = Loc.tag ~loc @@ match c, b' with - | (_, GVar id), GLocalAssum (loc, na', bk', t') -> - GLocalAssum (loc, unify_id id na', bk', t') - | c, GLocalPattern (loc, (p',ids), id, bk', t') -> + | (_, GVar id), GLocalAssum (na', bk', t') -> + GLocalAssum (unify_id id na', bk', t') + | c, GLocalPattern ((p',ids), id, bk', t') -> let p = pat_binder_of_term c in - GLocalPattern (loc, (unify_pat p p',ids), id, bk', t') + GLocalPattern ((unify_pat p p',ids), id, bk', t') | _ -> raise No_match in let rec unify cl bl' = match cl, bl' with | [], [] -> [] - | c :: cl, GLocalDef (_, _, _, _, t) :: bl' -> unify cl bl' + | c :: cl, (_loc, GLocalDef ( _, _, _, t)) :: bl' -> unify cl bl' | c :: cl, b' :: bl' -> unify_term_binder c b' :: unify cl bl' | _ -> raise No_match in let bl = unify cl bl' in @@ -898,17 +898,17 @@ let glue_letin_with_decls = true 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 (GLocalPattern(loc,(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 (GLocalAssum(loc,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 (GLocalPattern(loc,(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 (GLocalAssum(loc,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 - (GLocalDef (loc,na,Explicit (*?*), c,t)::decls) b + ((Loc.tag ~loc @@ GLocalDef (na,Explicit (*?*), c,t))::decls) b | b -> (decls, Loc.tag ~loc b) ) bi @@ -989,13 +989,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 [GLocalPattern(loc,(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 [GLocalAssum (loc,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 +1003,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 [GLocalPattern (loc,(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 [GLocalAssum (loc,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 +1021,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 [GLocalPattern (loc,(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 [GLocalAssum (loc,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 [GLocalAssum (loc,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 *) @@ -1121,10 +1121,10 @@ let rec match_ inner u alp metas sigma a1 a2 = | _ -> assert false in let (alp,sigma) = if is_bindinglist_meta id metas then - bind_bindinglist_env alp sigma id [GLocalAssum (Loc.ghost,Name id',Explicit,t1)] + bind_bindinglist_env alp sigma id [Loc.tag @@ GLocalAssum (Name id',Explicit,t1)] else match_names metas (alp,sigma) (Name id') na in - match_in u alp metas sigma (mkGApp Loc.ghost a1 (Loc.tag @@ GVar id')) b2 + match_in u alp metas sigma (mkGApp a1 (Loc.tag @@ GVar id')) b2 | (GRec _ | GEvar _), _ | _,_ -> raise No_match diff --git a/interp/notation_ops.mli b/interp/notation_ops.mli index a61ba172e..64f811dc2 100644 --- a/interp/notation_ops.mli +++ b/interp/notation_ops.mli @@ -33,12 +33,12 @@ val notation_constr_of_glob_constr : notation_interp_env -> (** Re-interpret a notation as a [glob_constr], taking care of binders *) -val glob_constr_of_notation_constr_with_binders : Loc.t -> +val glob_constr_of_notation_constr_with_binders : ?loc:Loc.t -> ('a -> Name.t -> 'a * Name.t) -> ('a -> notation_constr -> glob_constr) -> 'a -> notation_constr -> glob_constr -val glob_constr_of_notation_constr : Loc.t -> notation_constr -> glob_constr +val glob_constr_of_notation_constr : ?loc:Loc.t -> notation_constr -> glob_constr (** {5 Matching a notation pattern against a [glob_constr]} *) diff --git a/interp/smartlocate.ml b/interp/smartlocate.ml index 64d260cc1..fd9599ec0 100644 --- a/interp/smartlocate.ml +++ b/interp/smartlocate.ml @@ -67,14 +67,14 @@ 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/topconstr.ml b/interp/topconstr.ml index c8fbdaf28..2ffeb1f83 100644 --- a/interp/topconstr.ml +++ b/interp/topconstr.ml @@ -282,7 +282,7 @@ let locs_of_notation loc locs ntn = let ntn_loc loc (args,argslist,binderslist) = locs_of_notation loc (List.map constr_loc (args@List.flatten argslist)@ - List.map local_binders_loc binderslist) + List.map_filter local_binders_loc binderslist) let patntn_loc loc (args,argslist) = locs_of_notation loc diff --git a/intf/glob_term.mli b/intf/glob_term.mli index a8311e093..5aa5bdeeb 100644 --- a/intf/glob_term.mli +++ b/intf/glob_term.mli @@ -79,10 +79,11 @@ and cases_clause = (Id.t list * cases_pattern list * glob_constr) Loc.located of [t] are members of [il]. *) and cases_clauses = cases_clause list -type extended_glob_local_binder = - | GLocalAssum of Loc.t * Name.t * binding_kind * glob_constr - | GLocalDef of Loc.t * Name.t * binding_kind * glob_constr * glob_constr option - | GLocalPattern of Loc.t * (cases_pattern * Id.t list) * Id.t * binding_kind * glob_constr +type extended_glob_local_binder_r = + | GLocalAssum of Name.t * binding_kind * glob_constr + | GLocalDef of Name.t * binding_kind * glob_constr * glob_constr option + | GLocalPattern of (cases_pattern * Id.t list) * Id.t * binding_kind * glob_constr +and extended_glob_local_binder = extended_glob_local_binder_r Loc.located (** A globalised term together with a closure representing the value of its free variables. Intended for use when these variables are taken diff --git a/lib/aux_file.ml b/lib/aux_file.ml index 1b6651a55..fe6e87388 100644 --- a/lib/aux_file.ml +++ b/lib/aux_file.ml @@ -46,19 +46,21 @@ let contents x = x let empty_aux_file = H.empty -let get aux loc key = M.find key (H.find (Loc.unloc loc) aux) +let get ?loc aux key = M.find key (H.find (Option.cata Loc.unloc (0,0) loc) aux) -let record_in_aux_at loc key v = +let record_in_aux_at ?loc key v = Option.iter (fun oc -> - let i, j = Loc.unloc loc in - Printf.fprintf oc "%d %d %s %S\n" i j key v) - !oc + match loc with + | Some loc -> let i, j = Loc.unloc loc in + Printf.fprintf oc "%d %d %s %S\n" i j key v + | None -> Printf.fprintf oc "--- %s %S\n" key v + ) !oc -let current_loc = ref Loc.ghost +let current_loc : Loc.t option ref = ref None -let record_in_aux_set_at loc = current_loc := loc +let record_in_aux_set_at ?loc = current_loc := loc -let record_in_aux key v = record_in_aux_at !current_loc key v +let record_in_aux key v = record_in_aux_at ?loc:!current_loc key v let set h loc k v = let m = try H.find loc h with Not_found -> M.empty in @@ -91,4 +93,4 @@ let load_aux_file_for vfile = Flags.if_verbose Feedback.msg_info Pp.(str"Loading file "++str aux_fname++str": "++str s); empty_aux_file -let set h loc k v = set h (Loc.unloc loc) k v +let set ?loc h k v = set h (Option.cata Loc.unloc (0,0) loc) k v diff --git a/lib/aux_file.mli b/lib/aux_file.mli index 86e322b71..41b7b0855 100644 --- a/lib/aux_file.mli +++ b/lib/aux_file.mli @@ -9,9 +9,9 @@ type aux_file val load_aux_file_for : string -> aux_file -val get : aux_file -> Loc.t -> string -> string val empty_aux_file : aux_file -val set : aux_file -> Loc.t -> string -> string -> aux_file +val get : ?loc:Loc.t -> aux_file -> string -> string +val set : ?loc:Loc.t -> aux_file -> string -> string -> aux_file module H : Map.S with type key = int * int module M : Map.S with type key = string @@ -22,6 +22,6 @@ val start_aux_file : aux_file:string -> v_file:string -> unit val stop_aux_file : unit -> unit val recording : unit -> bool -val record_in_aux_at : Loc.t -> string -> string -> unit +val record_in_aux_at : ?loc:Loc.t -> string -> string -> unit val record_in_aux : string -> string -> unit -val record_in_aux_set_at : Loc.t -> unit +val record_in_aux_set_at : ?loc:Loc.t -> unit diff --git a/lib/cWarnings.ml b/lib/cWarnings.ml index 2f569d284..4e692af36 100644 --- a/lib/cWarnings.ml +++ b/lib/cWarnings.ml @@ -20,10 +20,10 @@ type t = { let warnings : (string, t) Hashtbl.t = Hashtbl.create 97 let categories : (string, string list) Hashtbl.t = Hashtbl.create 97 -let current_loc = ref Loc.ghost +let current_loc = ref None let flags = ref "" -let set_current_loc = (:=) current_loc +let set_current_loc loc = current_loc := Some loc let get_flags () = !flags @@ -35,29 +35,22 @@ let add_warning_in_category ~name ~category = in Hashtbl.replace categories category (name::ws) -let refine_loc = function - | None when not (Loc.is_ghost !current_loc) -> Some !current_loc - | loc -> loc - let create ~name ~category ?(default=Enabled) pp = Hashtbl.add warnings name { default; category; status = default }; add_warning_in_category ~name ~category; if default <> Disabled then add_warning_in_category ~name ~category:"default"; - fun ?loc x -> let w = Hashtbl.find warnings name in + fun ?loc x -> + let w = Hashtbl.find warnings name in + let loc = Option.append loc !current_loc in match w.status with | Disabled -> () - | AsError -> - begin match refine_loc loc with - | Some loc -> CErrors.user_err ~loc (pp x) - | None -> CErrors.user_err (pp x) - end + | AsError -> CErrors.user_err ?loc (pp x) | Enabled -> let msg = pp x ++ spc () ++ str "[" ++ str name ++ str "," ++ str category ++ str "]" in - let loc = refine_loc loc in Feedback.msg_warning ?loc msg let warn_unknown_warning = diff --git a/lib/feedback.ml b/lib/feedback.ml index df6fe3a62..3552fa4a0 100644 --- a/lib/feedback.ml +++ b/lib/feedback.ml @@ -30,7 +30,7 @@ type feedback_content = | FileDependency of string option * string | FileLoaded of string * string (* Extra metadata *) - | Custom of Loc.t * string * xml + | Custom of Loc.t option * string * xml (* Generic messages *) | Message of level * Loc.t option * Pp.std_ppcmds diff --git a/lib/feedback.mli b/lib/feedback.mli index bdd236ac7..dc104132a 100644 --- a/lib/feedback.mli +++ b/lib/feedback.mli @@ -38,7 +38,7 @@ type feedback_content = | FileDependency of string option * string | FileLoaded of string * string (* Extra metadata *) - | Custom of Loc.t * string * xml + | Custom of Loc.t option * string * xml (* Generic messages *) | Message of level * Loc.t option * Pp.std_ppcmds diff --git a/lib/loc.ml b/lib/loc.ml index 2a785fac4..3051ca7b9 100644 --- a/lib/loc.ml +++ b/lib/loc.ml @@ -26,12 +26,6 @@ let make_loc (bp, ep) = { fname = ""; line_nb = -1; bol_pos = 0; line_nb_last = -1; bol_pos_last = 0; bp = bp; ep = ep; } -let ghost = { - fname = ""; line_nb = -1; bol_pos = 0; line_nb_last = -1; bol_pos_last = 0; - bp = 0; ep = 0; } - -let is_ghost loc = loc.ep = 0 - let merge loc1 loc2 = if loc1.bp < loc2.bp then if loc1.ep < loc2.ep then { @@ -51,15 +45,24 @@ 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 unloc loc = (loc.bp, loc.ep) -let dummy_loc = ghost let join_loc = merge (** Located type *) type 'a located = t * 'a +let is_ghost loc = loc.ep = 0 + +let ghost = { + fname = ""; line_nb = -1; bol_pos = 0; line_nb_last = -1; bol_pos_last = 0; + bp = 0; ep = 0; } + +let internal_ghost = ghost let to_pair x = x let tag ?loc x = Option.default ghost loc, x let obj (_,x) = x diff --git a/lib/loc.mli b/lib/loc.mli index 10f63a8dd..82792613c 100644 --- a/lib/loc.mli +++ b/lib/loc.mli @@ -32,14 +32,15 @@ val unloc : t -> int * int val make_loc : int * int -> t (** Make a location out of its start and end position *) -val ghost : t -(** Dummy location *) - +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 + (** {5 Located exceptions} *) val add_loc : Exninfo.info -> t -> Exninfo.info @@ -71,15 +72,7 @@ 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 -(* Current not used *) +(* Currently not used *) val located_iter2 : ('a -> 'b -> unit) -> 'a located -> 'b located -> unit (** Projects out a located object *) - -(** {5 Backward compatibility} *) - -val dummy_loc : t -(** Same as [ghost] *) - -val join_loc : t -> t -> t -(** Same as [merge] *) diff --git a/lib/stateid.ml b/lib/stateid.ml index ae25735c5..34d49685b 100644 --- a/lib/stateid.ml +++ b/lib/stateid.ml @@ -41,7 +41,7 @@ type ('a,'b) request = { exn_info : t * t; stop : t; document : 'b; - loc : Loc.t; + loc : Loc.t option; uuid : 'a; name : string } diff --git a/lib/stateid.mli b/lib/stateid.mli index 1d87a343b..d9e75f584 100644 --- a/lib/stateid.mli +++ b/lib/stateid.mli @@ -34,7 +34,7 @@ type ('a,'b) request = { exn_info : t * t; stop : t; document : 'b; - loc : Loc.t; + loc : Loc.t option; uuid : 'a; name : string } diff --git a/library/declare.ml b/library/declare.ml index 31c9c24bc..6b505ac09 100644 --- a/library/declare.ml +++ b/library/declare.ml @@ -512,8 +512,8 @@ let do_constraint poly l = let open Misctypes in let u_of_id x = match x with - | GProp -> Loc.dummy_loc, (false, Univ.Level.prop) - | GSet -> Loc.dummy_loc, (false, Univ.Level.set) + | GProp -> Loc.tag (false, Univ.Level.prop) + | GSet -> Loc.tag (false, Univ.Level.set) | GType None -> user_err ~hdr:"Constraint" (str "Cannot declare constraints on anonymous universes") diff --git a/parsing/g_constr.ml4 b/parsing/g_constr.ml4 index 33e7236f5..e6c82c894 100644 --- a/parsing/g_constr.ml4 +++ b/parsing/g_constr.ml4 @@ -226,19 +226,19 @@ GEXTEND Gram record_field_declaration: [ [ id = global; bl = binders; ":="; c = lconstr -> - (id, mkCLambdaN (!@loc) bl c) ] ] + (id, mkCLambdaN ~loc:!@loc bl c) ] ] ; binder_constr: [ [ "forall"; bl = open_binders; ","; c = operconstr LEVEL "200" -> - mkCProdN (!@loc) bl c + mkCProdN ~loc:!@loc bl c | "fun"; bl = open_binders; "=>"; c = operconstr LEVEL "200" -> - mkCLambdaN (!@loc) bl c + mkCLambdaN ~loc:!@loc bl c | "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 *) | _, _ -> ty, c1 in - Loc.tag ~loc:!@loc @@ CLetIn(id,mkCLambdaN (constr_loc c1) bl c1, + 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 @@ -411,7 +411,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.join_loc (fst na) !@loc) @@ CHole (Some (Evar_kinds.BinderType (snd na)), IntroAnonymous, None))) + Loc.tag ~loc:(Loc.merge (fst na) !@loc) @@ CHole (Some (Evar_kinds.BinderType (snd na)), IntroAnonymous, None))) | ":"; c=lconstr; "}" -> (fun na -> CLocalAssum ([na],Default Implicit,c)) ] ] diff --git a/parsing/g_vernac.ml4 b/parsing/g_vernac.ml4 index 3d0322b10..7a3a2ace1 100644 --- a/parsing/g_vernac.ml4 +++ b/parsing/g_vernac.ml4 @@ -349,14 +349,14 @@ GEXTEND Gram ; record_binder_body: [ [ l = binders; oc = of_type_with_opt_coercion; - t = lconstr -> fun id -> (oc,AssumExpr (id,mkCProdN (!@loc) l t)) + t = lconstr -> fun id -> (oc,AssumExpr (id,mkCProdN ~loc:!@loc l t)) | l = binders; oc = of_type_with_opt_coercion; t = lconstr; ":="; b = lconstr -> fun id -> - (oc,DefExpr (id,mkCLambdaN ~loc:!@loc l b,Some (mkCProdN (!@loc) l t))) + (oc,DefExpr (id,mkCLambdaN ~loc:!@loc l b,Some (mkCProdN ~loc:!@loc l t))) | l = binders; ":="; b = lconstr -> fun id -> match snd b with | CCast(b', (CastConv t|CastVM t|CastNative t)) -> - (None,DefExpr(id,mkCLambdaN ~loc:!@loc l b',Some (mkCProdN (!@loc) l t))) + (None,DefExpr(id,mkCLambdaN ~loc:!@loc l b',Some (mkCProdN ~loc:!@loc l t))) | _ -> (None,DefExpr(id,mkCLambdaN ~loc:!@loc l b,None)) ] ] ; @@ -378,9 +378,9 @@ GEXTEND Gram constructor_type: [[ l = binders; t= [ coe = of_type_with_opt_coercion; c = lconstr -> - fun l id -> (not (Option.is_empty coe),(id,mkCProdN (!@loc) l c)) + fun l id -> (not (Option.is_empty coe),(id,mkCProdN ~loc:!@loc l c)) | -> - fun l id -> (false,(id,mkCProdN (!@loc) l (Loc.tag ~loc:!@loc @@ CHole (None, Misctypes.IntroAnonymous, None)))) ] + fun l id -> (false,(id,mkCProdN ~loc:!@loc l (Loc.tag ~loc:!@loc @@ CHole (None, Misctypes.IntroAnonymous, None)))) ] -> t l ]] ; @@ -592,15 +592,15 @@ GEXTEND Gram d = def_body -> let s = coerce_reference_to_id qid in VernacDefinition - ((Some Global,CanonicalStructure),((Loc.ghost,s),None),d) + ((Some Global,CanonicalStructure),((Loc.tag s),None),d) (* Coercions *) | IDENT "Coercion"; qid = global; d = def_body -> let s = coerce_reference_to_id qid in - VernacDefinition ((None,Coercion),((Loc.ghost,s),None),d) + VernacDefinition ((None,Coercion),((Loc.tag s),None),d) | IDENT "Coercion"; IDENT "Local"; qid = global; d = def_body -> let s = coerce_reference_to_id qid in - VernacDefinition ((Some Decl_kinds.Local,Coercion),((Loc.ghost,s),None),d) + VernacDefinition ((Some Decl_kinds.Local,Coercion),((Loc.tag s),None),d) | IDENT "Identity"; IDENT "Coercion"; IDENT "Local"; f = identref; ":"; s = class_rawexpr; ">->"; t = class_rawexpr -> VernacIdentityCoercion (true, f, s, t) diff --git a/plugins/firstorder/g_ground.ml4 b/plugins/firstorder/g_ground.ml4 index 3c0319319..70748f056 100644 --- a/plugins/firstorder/g_ground.ml4 +++ b/plugins/firstorder/g_ground.ml4 @@ -61,7 +61,7 @@ let default_intuition_tac = let name = { Tacexpr.mltac_plugin = "ground_plugin"; mltac_tactic = "auto_with"; } in let entry = { Tacexpr.mltac_name = name; mltac_index = 0 } in Tacenv.register_ml_tactic name [| tac |]; - Tacexpr.TacML (Loc.ghost, entry, []) + Tacexpr.TacML (Loc.tag (entry, [])) let (set_default_solver, default_solver, print_default_solver) = Tactic_option.declare_tactic_option ~default:default_intuition_tac "Firstorder default solver" diff --git a/plugins/funind/glob_term_to_relation.ml b/plugins/funind/glob_term_to_relation.ml index 07a0d5a50..946ee55d4 100644 --- a/plugins/funind/glob_term_to_relation.ml +++ b/plugins/funind/glob_term_to_relation.ml @@ -672,9 +672,7 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return = in let case_pats = build_constructors_of_type (fst ind) nal_as_glob_constr in assert (Int.equal (Array.length case_pats) 1); - let br = - (Loc.ghost,([],[case_pats.(0)],e)) - in + let br = Loc.tag ([],[case_pats.(0)],e) in let match_expr = mkGCases(None,[b,(Anonymous,None)],[br]) in build_entry_lc env funnames avoid match_expr @@ -1254,7 +1252,7 @@ let rec rebuild_return_type (loc, rt) = 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.ghost,Anonymous], + | _ -> Loc.tag ~loc @@ Constrexpr.CProdN([[Loc.tag Anonymous], Constrexpr.Default Decl_kinds.Explicit,Loc.tag ~loc rt], Loc.tag @@ Constrexpr.CSort(GType [])) @@ -1307,12 +1305,12 @@ let do_build_inductive (fun (n,t,typ) acc -> match typ with | Some typ -> - Loc.tag @@ Constrexpr.CLetIn((Loc.ghost, n),with_full_print (Constrextern.extern_glob_constr Id.Set.empty) t, + Loc.tag @@ Constrexpr.CLetIn((Loc.tag n),with_full_print (Constrextern.extern_glob_constr Id.Set.empty) t, Some (with_full_print (Constrextern.extern_glob_constr Id.Set.empty) typ), acc) | None -> Loc.tag @@ Constrexpr.CProdN - ([[(Loc.ghost,n)],Constrexpr_ops.default_binder_kind,with_full_print (Constrextern.extern_glob_constr Id.Set.empty) t], + ([[(Loc.tag n)],Constrexpr_ops.default_binder_kind,with_full_print (Constrextern.extern_glob_constr Id.Set.empty) t], acc ) ) @@ -1374,12 +1372,12 @@ let do_build_inductive (fun (n,t,typ) acc -> match typ with | Some typ -> - Loc.tag @@ Constrexpr.CLetIn((Loc.ghost, n),with_full_print (Constrextern.extern_glob_constr Id.Set.empty) t, + Loc.tag @@ Constrexpr.CLetIn((Loc.tag n),with_full_print (Constrextern.extern_glob_constr Id.Set.empty) t, Some (with_full_print (Constrextern.extern_glob_constr Id.Set.empty) typ), acc) | None -> Loc.tag @@ Constrexpr.CProdN - ([[(Loc.ghost,n)],Constrexpr_ops.default_binder_kind,with_full_print (Constrextern.extern_glob_constr Id.Set.empty) t], + ([[(Loc.tag n)],Constrexpr_ops.default_binder_kind,with_full_print (Constrextern.extern_glob_constr Id.Set.empty) t], acc ) ) @@ -1406,18 +1404,18 @@ let do_build_inductive (fun (n,t,typ) -> match typ with | Some typ -> - Constrexpr.CLocalDef((Loc.ghost,n), Constrextern.extern_glob_constr Id.Set.empty t, + Constrexpr.CLocalDef((Loc.tag n), Constrextern.extern_glob_constr Id.Set.empty t, Some (with_full_print (Constrextern.extern_glob_constr Id.Set.empty) typ)) | None -> Constrexpr.CLocalAssum - ([(Loc.ghost,n)], Constrexpr_ops.default_binder_kind, Constrextern.extern_glob_constr Id.Set.empty t) + ([(Loc.tag n)], Constrexpr_ops.default_binder_kind, Constrextern.extern_glob_constr Id.Set.empty t) ) rels_params in let ext_rels_constructors = Array.map (List.map (fun (id,t) -> - false,((Loc.ghost,id), + false,((Loc.tag id), with_full_print (Constrextern.extern_glob_type Id.Set.empty) ((* zeta_normalize *) (alpha_rt rel_params_ids t)) ) @@ -1425,7 +1423,7 @@ let do_build_inductive (rel_constructors) in let rel_ind i ext_rel_constructors = - (((Loc.ghost,relnames.(i)), None), + (((Loc.tag @@ relnames.(i)), None), rel_params, Some rel_arities.(i), ext_rel_constructors),[] diff --git a/plugins/funind/indfun.ml b/plugins/funind/indfun.ml index cad96946c..1da85c467 100644 --- a/plugins/funind/indfun.ml +++ b/plugins/funind/indfun.ml @@ -156,7 +156,7 @@ let build_newrecursive let (rec_sign,rec_impls) = List.fold_left (fun (env,impls) (((_,recname),_),bl,arityc,_) -> - let arityc = Constrexpr_ops.mkCProdN Loc.ghost bl arityc in + let arityc = Constrexpr_ops.mkCProdN bl arityc in let arity,ctx = Constrintern.interp_type env0 sigma arityc in let evdref = ref (Evd.from_env env0) in let _, (_, impls') = Constrintern.interp_context_evars env evdref bl in @@ -355,7 +355,7 @@ let generate_principle (evd:Evd.evar_map ref) pconstants on_error (*i The next call to mk_rel_id is valid since we have just construct the graph Ensures by : do_built i*) - let f_R_mut = Ident (Loc.ghost,mk_rel_id (List.nth names 0)) in + let f_R_mut = Ident (Loc.tag @@ mk_rel_id (List.nth names 0)) in let ind_kn = fst (locate_with_msg (pr_reference f_R_mut++str ": Not an inductive type!") @@ -453,7 +453,7 @@ let generate_correction_proof_wf f_ref tcc_lemma_ref let register_wf ?(is_mes=false) fname rec_impls wf_rel_expr wf_arg using_lemmas args ret_type body pre_hook = - let type_of_f = Constrexpr_ops.mkCProdN Loc.ghost args ret_type in + let type_of_f = Constrexpr_ops.mkCProdN args ret_type in let rec_arg_num = let names = List.map @@ -470,7 +470,7 @@ let register_wf ?(is_mes=false) fname rec_impls wf_rel_expr wf_arg using_lemmas let unbounded_eq = let f_app_args = Loc.tag @@ Constrexpr.CAppExpl( - (None,(Ident (Loc.ghost,fname)),None) , + (None,(Ident (Loc.tag fname)),None) , (List.map (function | _,Anonymous -> assert false @@ -480,10 +480,10 @@ let register_wf ?(is_mes=false) fname rec_impls wf_rel_expr wf_arg using_lemmas ) ) in - Loc.tag @@ Constrexpr.CApp ((None,Constrexpr_ops.mkRefC (Qualid (Loc.ghost,(qualid_of_string "Logic.eq")))), + Loc.tag @@ Constrexpr.CApp ((None,Constrexpr_ops.mkRefC (Qualid (Loc.tag (qualid_of_string "Logic.eq")))), [(f_app_args,None);(body,None)]) in - let eq = Constrexpr_ops.mkCProdN Loc.ghost args unbounded_eq in + let eq = Constrexpr_ops.mkCProdN args unbounded_eq in let hook ((f_ref,_) as fconst) tcc_lemma_ref (functional_ref,_) (eq_ref,_) rec_arg_num rec_arg_type nb_args relation = try @@ -537,13 +537,13 @@ let register_mes fname rec_impls wf_mes_expr wf_rel_expr_opt wf_arg using_lemmas | None -> let ltof = let make_dir l = DirPath.make (List.rev_map Id.of_string l) in - Libnames.Qualid (Loc.ghost,Libnames.qualid_of_path + Libnames.Qualid (Loc.tag @@ Libnames.qualid_of_path (Libnames.make_path (make_dir ["Arith";"Wf_nat"]) (Id.of_string "ltof"))) in let fun_from_mes = let applied_mes = Constrexpr_ops.mkAppC(wf_mes_expr,[Constrexpr_ops.mkIdentC wf_arg]) in - Constrexpr_ops.mkLambdaC ([(Loc.ghost,Name wf_arg)],Constrexpr_ops.default_binder_kind,wf_arg_type,applied_mes) + Constrexpr_ops.mkLambdaC ([(Loc.tag @@ Name wf_arg)],Constrexpr_ops.default_binder_kind,wf_arg_type,applied_mes) in let wf_rel_from_mes = Constrexpr_ops.mkAppC(Constrexpr_ops.mkRefC ltof,[wf_arg_type;fun_from_mes]) @@ -554,7 +554,7 @@ let register_mes fname rec_impls wf_mes_expr wf_rel_expr_opt wf_arg using_lemmas let a = Names.Id.of_string "___a" in let b = Names.Id.of_string "___b" in Constrexpr_ops.mkLambdaC( - [Loc.ghost,Name a;Loc.ghost,Name b], + [Loc.tag @@ Name a;Loc.tag @@ Name b], Constrexpr.Default Explicit, wf_arg_type, Constrexpr_ops.mkAppC(wf_rel_expr, @@ -891,14 +891,14 @@ let make_graph (f_ref:global_reference) = ) in let b' = add_args (snd id) new_args b in - ((((id,None), ( Some (Loc.ghost,rec_id),CStructRec),nal_tas@bl,t,Some b'),[]):(Vernacexpr.fixpoint_expr * Vernacexpr.decl_notation list)) + ((((id,None), ( Some (Loc.tag rec_id),CStructRec),nal_tas@bl,t,Some b'),[]):(Vernacexpr.fixpoint_expr * Vernacexpr.decl_notation list)) ) fixexprl in l | _ -> let id = Label.to_id (con_label c) in - [(((Loc.ghost,id),None),(None,Constrexpr.CStructRec),nal_tas,t,Some b),[]] + [(((Loc.tag id),None),(None,Constrexpr.CStructRec),nal_tas,t,Some b),[]] in let mp,dp,_ = repr_con c in do_generate_principle [c,Univ.Instance.empty] error_error false false expr_list; diff --git a/plugins/funind/indfun_common.ml b/plugins/funind/indfun_common.ml index 8f0c98624..de8dc53f1 100644 --- a/plugins/funind/indfun_common.ml +++ b/plugins/funind/indfun_common.ml @@ -106,7 +106,7 @@ let list_add_set_eq eq_fun x l = let const_of_id id = let _,princ_ref = - qualid_of_reference (Libnames.Ident (Loc.ghost,id)) + qualid_of_reference (Libnames.Ident (Loc.tag id)) in try Constrintern.locate_reference princ_ref with Not_found -> diff --git a/plugins/funind/invfun.ml b/plugins/funind/invfun.ml index 8da4f9233..e4536278c 100644 --- a/plugins/funind/invfun.ml +++ b/plugins/funind/invfun.ml @@ -274,7 +274,7 @@ let prove_fun_correct evd functional_induction funs_constr graphs_constr schemes List.map (fun decl -> List.map - (fun id -> Loc.ghost, IntroNaming (IntroIdentifier id)) + (fun id -> Loc.tag @@ IntroNaming (IntroIdentifier id)) (generate_fresh_id (Id.of_string "y") ids (List.length (fst (decompose_prod_assum evd (RelDecl.get_type decl))))) ) branches diff --git a/plugins/funind/merge.ml b/plugins/funind/merge.ml index 9dc1d48df..f2654b5de 100644 --- a/plugins/funind/merge.ml +++ b/plugins/funind/merge.ml @@ -73,7 +73,7 @@ let isVarf f x = in global environment. *) let ident_global_exist id = try - let ans = Loc.tag @@ CRef (Libnames.Ident (Loc.ghost,id), None) in + let ans = Loc.tag @@ CRef (Libnames.Ident (Loc.tag id), None) in let _ = ignore (Constrintern.intern_constr (Global.env()) ans) in true with e when CErrors.noncritical e -> false @@ -825,7 +825,7 @@ let merge_rec_params_and_arity prms1 prms2 shift (concl:constr) = let _ = prNamedRConstr (string_of_name nme) tp in let _ = prstr " ; " in let typ = glob_constr_to_constr_expr tp in - CLocalAssum ([(Loc.ghost,nme)], Constrexpr_ops.default_binder_kind, typ) :: acc) + CLocalAssum ([(Loc.tag nme)], Constrexpr_ops.default_binder_kind, typ) :: acc) [] params in let concl = Constrextern.extern_constr false (Global.env()) Evd.empty concl in let arity,_ = @@ -835,7 +835,7 @@ let merge_rec_params_and_arity prms1 prms2 shift (concl:constr) = let c = RelDecl.get_type decl in let typ = Constrextern.extern_constr false env Evd.empty c in let newenv = Environ.push_rel (LocalAssum (nm,c)) env in - Loc.tag @@ CProdN ([[(Loc.ghost,nm)],Constrexpr_ops.default_binder_kind,typ] , acc) , newenv) + Loc.tag @@ CProdN ([[(Loc.tag nm)],Constrexpr_ops.default_binder_kind,typ] , acc) , newenv) (concl,Global.env()) (shift.funresprms2 @ shift.funresprms1 @ shift.args2 @ shift.args1 @ shift.otherprms2 @ shift.otherprms1) in @@ -849,12 +849,12 @@ let merge_rec_params_and_arity prms1 prms2 shift (concl:constr) = FIXME: params et cstr_expr (arity) *) let glob_constr_list_to_inductive_expr prms1 prms2 mib1 mib2 shift (rawlist:(Id.t * glob_constr) list) = - let lident = (Loc.ghost, shift.ident), None in + let lident = (Loc.tag shift.ident), None in let bindlist , cstr_expr = (* params , arities *) merge_rec_params_and_arity prms1 prms2 shift mkSet in let lcstor_expr : (bool * (lident * constr_expr)) list = List.map (* zeta_normalize t ? *) - (fun (id,t) -> false, ((Loc.ghost,id),glob_constr_to_constr_expr t)) + (fun (id,t) -> false, ((Loc.tag id),glob_constr_to_constr_expr t)) rawlist in lident , bindlist , Some cstr_expr , lcstor_expr @@ -902,7 +902,7 @@ let merge_inductive (ind1: inductive) (ind2: inductive) (* Find infos on identifier id. *) let find_Function_infos_safe (id:Id.t): Indfun_common.function_info = let kn_of_id x = - let f_ref = Libnames.Ident (Loc.ghost,x) in + let f_ref = Libnames.Ident (Loc.tag x) in locate_with_msg (str "Don't know what to do with " ++ Libnames.pr_reference f_ref) locate_constant f_ref in try find_Function_infos (kn_of_id id) diff --git a/plugins/funind/recdef.ml b/plugins/funind/recdef.ml index 25daedd0e..9e469c756 100644 --- a/plugins/funind/recdef.ml +++ b/plugins/funind/recdef.ml @@ -1306,7 +1306,7 @@ let open_new_goal build_proof sigma using_lemmas ref_ goal_name (gls_type,decomp CErrors.error "\"abstract\" cannot handle existentials"; let hook _ _ = let opacity = - let na_ref = Libnames.Ident (Loc.ghost,na) in + let na_ref = Libnames.Ident (Loc.tag na) in let na_global = Smartlocate.global_with_alias na_ref in match na_global with ConstRef c -> is_opaque_constant c @@ -1556,7 +1556,7 @@ let recursive_definition is_mes function_name rec_impls type_of_f r rec_arg_num let hook _ _ = let term_ref = Nametab.locate (qualid_of_ident term_id) in let f_ref = declare_f function_name (IsProof Lemma) arg_types term_ref in - let _ = Extraction_plugin.Table.extraction_inline true [Ident (Loc.ghost,term_id)] in + let _ = Extraction_plugin.Table.extraction_inline true [Ident (Loc.tag term_id)] in (* message "start second proof"; *) let stop = try com_eqn (List.length res_vars) equation_id functional_ref f_ref term_ref (subst_var function_name equation_lemma_type); diff --git a/plugins/ltac/coretactics.ml4 b/plugins/ltac/coretactics.ml4 index 28ff6df83..6890b3198 100644 --- a/plugins/ltac/coretactics.ml4 +++ b/plugins/ltac/coretactics.ml4 @@ -305,10 +305,9 @@ END open Tacexpr let initial_atomic () = - let dloc = Loc.ghost in let nocl = {onhyps=Some[];concl_occs=AllOccurrences} in let iter (s, t) = - let body = TacAtom (dloc, t) in + let body = TacAtom (Loc.tag t) in Tacenv.register_ltac false false (Id.of_string s) body in let () = List.iter iter @@ -323,7 +322,7 @@ let initial_atomic () = List.iter iter [ "idtac",TacId []; "fail", TacFail(TacLocal,ArgArg 0,[]); - "fresh", TacArg(dloc,TacFreshId []) + "fresh", TacArg(Loc.tag @@ TacFreshId []) ] let () = Mltop.declare_cache_obj initial_atomic "coretactics" diff --git a/plugins/ltac/evar_tactics.ml b/plugins/ltac/evar_tactics.ml index 5d3f6df03..a5d9697ae 100644 --- a/plugins/ltac/evar_tactics.ml +++ b/plugins/ltac/evar_tactics.ml @@ -82,7 +82,7 @@ let instantiate_tac_by_name id c = end let let_evar name typ = - let src = (Loc.ghost,Evar_kinds.GoalEvar) in + let src = (Loc.tag Evar_kinds.GoalEvar) in Proofview.Goal.s_enter { s_enter = begin fun gl -> let sigma = Tacmach.New.project gl in let env = Proofview.Goal.env gl in diff --git a/plugins/ltac/extraargs.ml4 b/plugins/ltac/extraargs.ml4 index 53b726432..aa81f148e 100644 --- a/plugins/ltac/extraargs.ml4 +++ b/plugins/ltac/extraargs.ml4 @@ -228,11 +228,11 @@ ARGUMENT EXTEND hloc | [ "in" "|-" "*" ] -> [ ConclLocation () ] | [ "in" ident(id) ] -> - [ HypLocation ((Loc.ghost,id),InHyp) ] + [ HypLocation ((Loc.tag id),InHyp) ] | [ "in" "(" "Type" "of" ident(id) ")" ] -> - [ HypLocation ((Loc.ghost,id),InHypTypeOnly) ] + [ HypLocation ((Loc.tag id),InHypTypeOnly) ] | [ "in" "(" "Value" "of" ident(id) ")" ] -> - [ HypLocation ((Loc.ghost,id),InHypValueOnly) ] + [ HypLocation ((Loc.tag id),InHypValueOnly) ] END diff --git a/plugins/ltac/extratactics.ml4 b/plugins/ltac/extratactics.ml4 index 232bd851f..5123d6c40 100644 --- a/plugins/ltac/extratactics.ml4 +++ b/plugins/ltac/extratactics.ml4 @@ -76,7 +76,7 @@ END let induction_arg_of_quantified_hyp = function | AnonHyp n -> None,ElimOnAnonHyp n - | NamedHyp id -> None,ElimOnIdent (Loc.ghost,id) + | NamedHyp id -> None,ElimOnIdent (Loc.tag id) (* Versions *_main must come first!! so that "1" is interpreted as a ElimOnAnonHyp and not as a "constr", and "id" is interpreted as a @@ -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 -> Loc.ghost | Some loc -> loc in - resolve_hole (subst_hole_with_term (fst (Loc.unloc loc)) c_raw t_hole) + 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) 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 @@ -784,7 +784,7 @@ let case_eq_intros_rewrite x = let rec find_a_destructable_match sigma t = let cl = induction_arg_of_quantified_hyp (NamedHyp (Id.of_string "x")) in let cl = [cl, (None, None), None], None in - let dest = TacAtom (Loc.ghost, TacInductionDestruct(false, false, cl)) in + let dest = TacAtom (Loc.tag @@ TacInductionDestruct(false, false, cl)) in match EConstr.kind sigma t with | Case (_,_,x,_) when closed0 sigma x -> if isVar sigma x then diff --git a/plugins/ltac/g_ltac.ml4 b/plugins/ltac/g_ltac.ml4 index d6cbd8db1..23643c5d3 100644 --- a/plugins/ltac/g_ltac.ml4 +++ b/plugins/ltac/g_ltac.ml4 @@ -460,7 +460,7 @@ END let pr_ltac_production_item = function | Tacentries.TacTerm s -> quote (str s) -| Tacentries.TacNonTerm (_, (arg, sep), id) -> +| Tacentries.TacNonTerm (_, ((arg, sep), id)) -> let sep = match sep with | None -> mt () | Some sep -> str "," ++ spc () ++ quote (str sep) @@ -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, ((Names.Id.to_string nt, sep), p)) ] END VERNAC COMMAND EXTEND VernacTacticNotation diff --git a/plugins/ltac/g_obligations.ml4 b/plugins/ltac/g_obligations.ml4 index 3e6e2db60..7d4075836 100644 --- a/plugins/ltac/g_obligations.ml4 +++ b/plugins/ltac/g_obligations.ml4 @@ -50,7 +50,7 @@ module Tactic = Pltac open Pcoq -let sigref = mkRefC (Qualid (Loc.ghost, Libnames.qualid_of_string "Coq.Init.Specif.sig")) +let sigref = mkRefC (Qualid (Loc.tag @@ Libnames.qualid_of_string "Coq.Init.Specif.sig")) type 'a withtac_argtype = (Tacexpr.raw_tactic_expr option, 'a) Genarg.abstract_argument_type diff --git a/plugins/ltac/pptactic.ml b/plugins/ltac/pptactic.ml index ae596c411..26ac3c53e 100644 --- a/plugins/ltac/pptactic.ml +++ b/plugins/ltac/pptactic.ml @@ -51,7 +51,7 @@ let pr_global x = Nametab.pr_global_env Id.Set.empty x type 'a grammar_tactic_prod_item_expr = | TacTerm of string -| TacNonTerm of Loc.t * 'a * Names.Id.t +| TacNonTerm of ('a * Names.Id.t) Loc.located type grammar_terminals = Genarg.ArgT.any Extend.user_symbol grammar_tactic_prod_item_expr list @@ -212,7 +212,7 @@ type 'a extra_genarg_printer = let rec tacarg_using_rule_token pr_gen = function | [] -> [] | TacTerm s :: l -> keyword s :: tacarg_using_rule_token pr_gen l - | TacNonTerm (_, (symb, arg), _) :: l -> + | TacNonTerm (_, ((symb, arg), _)) :: l -> pr_gen symb arg :: tacarg_using_rule_token pr_gen l let pr_tacarg_using_rule pr_gen l = @@ -252,7 +252,7 @@ type 'a extra_genarg_printer = let prods = (KNmap.find key !prnotation_tab).pptac_prods in let rec pr = function | TacTerm s -> primitive s - | TacNonTerm (_, symb, _) -> str (Printf.sprintf "(%s)" (pr_user_symbol symb)) + | TacNonTerm (_, (symb, _)) -> str (Printf.sprintf "(%s)" (pr_user_symbol symb)) in pr_sequence pr prods with Not_found -> @@ -264,8 +264,8 @@ type 'a extra_genarg_printer = let rec pack prods args = match prods, args with | [], [] -> [] | TacTerm s :: prods, args -> TacTerm s :: pack prods args - | TacNonTerm (loc, symb, id) :: prods, arg :: args -> - TacNonTerm (loc, (symb, arg), id) :: pack prods args + | TacNonTerm (loc, (symb, id)) :: prods, arg :: args -> + TacNonTerm (loc, ((symb, arg), id)) :: pack prods args | _ -> raise Not_found in let prods = pack pp.pptac_prods l in @@ -275,7 +275,7 @@ type 'a extra_genarg_printer = let pr arg = str "_" in KerName.print key ++ spc() ++ pr_sequence pr l ++ str" (* Generic printer *)" - let pr_farg prtac arg = prtac (1, Any) (TacArg (Loc.ghost, arg)) + let pr_farg prtac arg = prtac (1, Any) (TacArg (Loc.tag arg)) let is_genarg tag wit = let ArgT.Any tag = tag in @@ -331,9 +331,9 @@ type 'a extra_genarg_printer = pr_extend_gen (pr_farg prtac) let pr_raw_alias prc prlc prtac prpat lev key args = - pr_alias_gen (pr_targ (fun l a -> prtac l (TacArg (Loc.ghost, a)))) lev key args + pr_alias_gen (pr_targ (fun l a -> prtac l (TacArg (Loc.tag a)))) lev key args let pr_glob_alias prc prlc prtac prpat lev key args = - pr_alias_gen (pr_targ (fun l a -> prtac l (TacArg (Loc.ghost, a)))) lev key args + pr_alias_gen (pr_targ (fun l a -> prtac l (TacArg (Loc.tag a)))) lev key args (**********************************************************************) (* The tactic printer *) @@ -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 @@ -416,7 +416,7 @@ type 'a extra_genarg_printer = let pr_as_name = function | Anonymous -> mt () - | Name id -> spc () ++ keyword "as" ++ spc () ++ pr_lident (Loc.ghost,id) + | Name id -> spc () ++ keyword "as" ++ spc () ++ pr_lident (Loc.tag id) let pr_pose_as_style prc na c = spc() ++ prc c ++ pr_as_name na @@ -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) = @@ -574,7 +574,7 @@ type 'a extra_genarg_printer = let pr_let_clause k pr (id,(bl,t)) = hov 0 (keyword k ++ spc () ++ pr_lident id ++ prlist pr_funvar bl ++ - str " :=" ++ brk (1,1) ++ pr (TacArg (Loc.ghost,t))) + str " :=" ++ brk (1,1) ++ pr (TacArg (Loc.tag t))) let pr_let_clauses recflag pr = function | hd::tl -> @@ -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 + | TacML (loc,(s,l)) -> + 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 @@ -1078,7 +1078,7 @@ type 'a extra_genarg_printer = | TacNumgoals -> keyword "numgoals" | (TacCall _|Tacexp _ | TacGeneric _) as a -> - hov 0 (keyword "ltac:" ++ surround (pr_tac ltop (TacArg (Loc.ghost,a)))) + hov 0 (keyword "ltac:" ++ surround (pr_tac ltop (TacArg (Loc.tag a)))) in pr_tac @@ -1087,7 +1087,7 @@ type 'a extra_genarg_printer = if Int.equal n 0 then (List.rev acc, (ty,None)) else match Loc.obj ty with Glob_term.GProd(na,Explicit,a,b) -> - strip_ty (([Loc.ghost,na],(a,None))::acc) (n-1) b + strip_ty (([Loc.tag na],(a,None))::acc) (n-1) b | _ -> error "Cannot translate fix tactic: not enough products" in strip_ty [] n ty @@ -1158,7 +1158,7 @@ type 'a extra_genarg_printer = if n=0 then (List.rev acc, EConstr.of_constr ty) else match Term.kind_of_term ty with Term.Prod(na,a,b) -> - strip_ty (([Loc.ghost,na],EConstr.of_constr a)::acc) (n-1) b + strip_ty (([Loc.tag na],EConstr.of_constr a)::acc) (n-1) b | _ -> error "Cannot translate fix tactic: not enough products" in strip_ty [] n ty @@ -1253,7 +1253,7 @@ let () = wit_clause_dft_concl (pr_clauses (Some true) pr_lident) (pr_clauses (Some true) pr_lident) - (pr_clauses (Some true) (fun id -> pr_lident (Loc.ghost,id))) + (pr_clauses (Some true) (fun id -> pr_lident (Loc.tag id))) ; Genprint.register_print0 wit_constr diff --git a/plugins/ltac/pptactic.mli b/plugins/ltac/pptactic.mli index 729338fb9..23570392d 100644 --- a/plugins/ltac/pptactic.mli +++ b/plugins/ltac/pptactic.mli @@ -21,7 +21,7 @@ open Ppextend type 'a grammar_tactic_prod_item_expr = | TacTerm of string -| TacNonTerm of Loc.t * 'a * Names.Id.t +| TacNonTerm of ('a * Names.Id.t) Loc.located type 'a raw_extra_genarg_printer = (constr_expr -> std_ppcmds) -> diff --git a/plugins/ltac/profile_ltac.ml b/plugins/ltac/profile_ltac.ml index bcb28f77c..eb97a0e70 100644 --- a/plugins/ltac/profile_ltac.ml +++ b/plugins/ltac/profile_ltac.ml @@ -129,7 +129,7 @@ let to_ltacprof_results xml = let feedback_results results = Feedback.(feedback - (Custom (Loc.dummy_loc, "ltacprof_results", of_ltacprof_results results))) + (Custom (None, "ltacprof_results", of_ltacprof_results results))) (* ************** pretty printing ************************************* *) @@ -250,7 +250,7 @@ let string_of_call ck = | Tacexpr.LtacVarCall (id, t) -> Nameops.pr_id id | Tacexpr.LtacAtomCall te -> (Pptactic.pr_glob_tactic (Global.env ()) - (Tacexpr.TacAtom (Loc.ghost, te))) + (Tacexpr.TacAtom (Loc.tag te))) | Tacexpr.LtacConstrInterp (c, _) -> pr_glob_constr_env (Global.env ()) c | Tacexpr.LtacMLCall te -> diff --git a/plugins/ltac/rewrite.ml b/plugins/ltac/rewrite.ml index 40484bfea..19c2eaf0a 100644 --- a/plugins/ltac/rewrite.ml +++ b/plugins/ltac/rewrite.ml @@ -1787,11 +1787,11 @@ let rec strategy_of_ast = function (* By default the strategy for "rewrite_db" is top-down *) -let mkappc s l = Loc.tag @@ CAppExpl ((None,(Libnames.Ident (Loc.ghost,Id.of_string s)),None),l) +let mkappc s l = Loc.tag @@ CAppExpl ((None,(Libnames.Ident (Loc.tag @@ Id.of_string s)),None),l) let declare_an_instance n s args = - (((Loc.ghost,Name n),None), Explicit, - Loc.tag @@ CAppExpl ((None, Qualid (Loc.ghost, qualid_of_string s),None), + (((Loc.tag @@ Name n),None), Explicit, + Loc.tag @@ CAppExpl ((None, Qualid (Loc.tag @@ qualid_of_string s),None), args)) let declare_instance a aeq n s = declare_an_instance n s [a;aeq] @@ -1804,17 +1804,17 @@ let anew_instance global binders instance fields = let declare_instance_refl global binders a aeq n lemma = let instance = declare_instance a aeq (add_suffix n "_Reflexive") "Coq.Classes.RelationClasses.Reflexive" in anew_instance global binders instance - [(Ident (Loc.ghost,Id.of_string "reflexivity"),lemma)] + [(Ident (Loc.tag @@ Id.of_string "reflexivity"),lemma)] let declare_instance_sym global binders a aeq n lemma = let instance = declare_instance a aeq (add_suffix n "_Symmetric") "Coq.Classes.RelationClasses.Symmetric" in anew_instance global binders instance - [(Ident (Loc.ghost,Id.of_string "symmetry"),lemma)] + [(Ident (Loc.tag @@ Id.of_string "symmetry"),lemma)] let declare_instance_trans global binders a aeq n lemma = let instance = declare_instance a aeq (add_suffix n "_Transitive") "Coq.Classes.RelationClasses.Transitive" in anew_instance global binders instance - [(Ident (Loc.ghost,Id.of_string "transitivity"),lemma)] + [(Ident (Loc.tag @@ Id.of_string "transitivity"),lemma)] let declare_relation ?(binders=[]) a aeq n refl symm trans = init_setoid (); @@ -1838,16 +1838,16 @@ let declare_relation ?(binders=[]) a aeq n refl symm trans = let instance = declare_instance a aeq n "Coq.Classes.RelationClasses.PreOrder" in ignore( anew_instance global binders instance - [(Ident (Loc.ghost,Id.of_string "PreOrder_Reflexive"), lemma1); - (Ident (Loc.ghost,Id.of_string "PreOrder_Transitive"),lemma3)]) + [(Ident (Loc.tag @@ Id.of_string "PreOrder_Reflexive"), lemma1); + (Ident (Loc.tag @@ Id.of_string "PreOrder_Transitive"),lemma3)]) | (None, Some lemma2, Some lemma3) -> let _lemma_sym = declare_instance_sym global binders a aeq n lemma2 in let _lemma_trans = declare_instance_trans global binders a aeq n lemma3 in let instance = declare_instance a aeq n "Coq.Classes.RelationClasses.PER" in ignore( anew_instance global binders instance - [(Ident (Loc.ghost,Id.of_string "PER_Symmetric"), lemma2); - (Ident (Loc.ghost,Id.of_string "PER_Transitive"),lemma3)]) + [(Ident (Loc.tag @@ Id.of_string "PER_Symmetric"), lemma2); + (Ident (Loc.tag @@ Id.of_string "PER_Transitive"),lemma3)]) | (Some lemma1, Some lemma2, Some lemma3) -> let _lemma_refl = declare_instance_refl global binders a aeq n lemma1 in let _lemma_sym = declare_instance_sym global binders a aeq n lemma2 in @@ -1855,9 +1855,9 @@ let declare_relation ?(binders=[]) a aeq n refl symm trans = let instance = declare_instance a aeq n "Coq.Classes.RelationClasses.Equivalence" in ignore( anew_instance global binders instance - [(Ident (Loc.ghost,Id.of_string "Equivalence_Reflexive"), lemma1); - (Ident (Loc.ghost,Id.of_string "Equivalence_Symmetric"), lemma2); - (Ident (Loc.ghost,Id.of_string "Equivalence_Transitive"), lemma3)]) + [(Ident (Loc.tag @@ Id.of_string "Equivalence_Reflexive"), lemma1); + (Ident (Loc.tag @@ Id.of_string "Equivalence_Symmetric"), lemma2); + (Ident (Loc.tag @@ Id.of_string "Equivalence_Transitive"), lemma3)]) let cHole = Loc.tag @@ CHole (None, Misctypes.IntroAnonymous, None) @@ -1959,17 +1959,16 @@ let add_setoid global binders a aeq t n = let instance = declare_instance a aeq n "Coq.Classes.RelationClasses.Equivalence" in ignore( anew_instance global binders instance - [(Ident (Loc.ghost,Id.of_string "Equivalence_Reflexive"), mkappc "Seq_refl" [a;aeq;t]); - (Ident (Loc.ghost,Id.of_string "Equivalence_Symmetric"), mkappc "Seq_sym" [a;aeq;t]); - (Ident (Loc.ghost,Id.of_string "Equivalence_Transitive"), mkappc "Seq_trans" [a;aeq;t])]) + [(Ident (Loc.tag @@ Id.of_string "Equivalence_Reflexive"), mkappc "Seq_refl" [a;aeq;t]); + (Ident (Loc.tag @@ Id.of_string "Equivalence_Symmetric"), mkappc "Seq_sym" [a;aeq;t]); + (Ident (Loc.tag @@ Id.of_string "Equivalence_Transitive"), mkappc "Seq_trans" [a;aeq;t])]) let make_tactic name = let open Tacexpr in - let loc = Loc.ghost in let tacpath = Libnames.qualid_of_string name in - let tacname = Qualid (loc, tacpath) in - TacArg (loc, TacCall (Loc.tag ~loc (tacname, []))) + let tacname = Qualid (Loc.tag tacpath) in + TacArg (Loc.tag @@ TacCall (Loc.tag (tacname, []))) let add_morphism_infer glob m n = init_setoid (); @@ -2012,9 +2011,9 @@ let add_morphism glob binders m s n = let poly = Flags.is_universe_polymorphism () in let instance_id = add_suffix n "_Proper" in let instance = - (((Loc.ghost,Name instance_id),None), Explicit, + (((Loc.tag @@ Name instance_id),None), Explicit, Loc.tag @@ CAppExpl ( - (None, Qualid (Loc.ghost, Libnames.qualid_of_string "Coq.Classes.Morphisms.Proper"),None), + (None, Qualid (Loc.tag @@ Libnames.qualid_of_string "Coq.Classes.Morphisms.Proper"),None), [cHole; s; m])) in let tac = Tacinterp.interp (make_tactic "add_morphism_tactic") in diff --git a/plugins/ltac/tacentries.ml b/plugins/ltac/tacentries.ml index 39121ac94..f608515aa 100644 --- a/plugins/ltac/tacentries.ml +++ b/plugins/ltac/tacentries.ml @@ -22,7 +22,7 @@ open Nameops type 'a grammar_tactic_prod_item_expr = 'a Pptactic.grammar_tactic_prod_item_expr = | TacTerm of string -| TacNonTerm of Loc.t * 'a * Names.Id.t +| TacNonTerm of ('a * Names.Id.t) Loc.located type raw_argument = string * string option type argument = Genarg.ArgT.any Extend.user_symbol @@ -178,7 +178,7 @@ let add_tactic_entry (kn, ml, tg) state = in let map = function | TacTerm s -> GramTerminal s - | TacNonTerm (loc, s, _) -> + | TacNonTerm (loc, (s, _)) -> let EntryName (typ, e) = prod_item_of_symbol tg.tacgram_level s in GramNonTerminal (loc, typ, e) in @@ -206,7 +206,7 @@ let register_tactic_notation_entry name entry = let interp_prod_item = function | TacTerm s -> TacTerm s - | TacNonTerm (loc, (nt, sep), id) -> + | TacNonTerm (loc, ((nt, sep), id)) -> let symbol = parse_user_entry nt sep in let interp s = function | None -> @@ -224,7 +224,7 @@ let interp_prod_item = function end in let symbol = interp_entry_name interp symbol in - TacNonTerm (loc, symbol, id) + TacNonTerm (loc, (symbol, id)) let make_fresh_key = let id = Summary.ref ~name:"TACTIC-NOTATION-COUNTER" 0 in @@ -300,7 +300,7 @@ let inTacticGrammar : tactic_grammar_obj -> obj = let cons_production_parameter = function | TacTerm _ -> None -| TacNonTerm (_, _, id) -> Some id +| TacNonTerm (_, (_, id)) -> Some id let add_glob_tactic_notation local ~level prods forml ids tac = let parule = { @@ -338,10 +338,10 @@ let extend_atomic_tactic name entries = in let empty_value = function | TacTerm s -> raise NonEmptyArgument - | TacNonTerm (_, symb, _) -> + | TacNonTerm (_, (symb, _)) -> let EntryName (typ, e) = prod_item_of_symbol 0 symb in let Genarg.Rawwit wit = typ in - let inj x = TacArg (Loc.ghost, TacGeneric (Genarg.in_gen typ x)) in + let inj x = TacArg (Loc.tag @@ TacGeneric (Genarg.in_gen typ x)) in let default = epsilon_value inj e in match default with | None -> raise NonEmptyArgument @@ -355,7 +355,7 @@ let extend_atomic_tactic name entries = | Some (id, args) -> let args = List.map (fun a -> Tacexp a) args in let entry = { mltac_name = name; mltac_index = i } in - let body = TacML (Loc.ghost, entry, args) in + let body = TacML (Loc.tag (entry, args)) in Tacenv.register_ltac false false (Names.Id.of_string id) body in List.iteri add_atomic entries @@ -366,12 +366,12 @@ let add_ml_tactic_notation name ~level prods = let open Tacexpr in let get_id = function | TacTerm s -> None - | TacNonTerm (_, _, id) -> Some id + | TacNonTerm (_, (_, id)) -> Some id in let ids = List.map_filter get_id prods in let entry = { mltac_name = name; mltac_index = len - i - 1 } in - let map id = Reference (Misctypes.ArgVar (Loc.ghost, id)) in - let tac = TacML (Loc.ghost, entry, List.map map ids) in + let map id = Reference (Misctypes.ArgVar (Loc.tag id)) in + let tac = TacML (Loc.tag (entry, List.map map ids)) in add_glob_tactic_notation false ~level prods true ids tac in List.iteri iter (List.rev prods); diff --git a/plugins/ltac/tacentries.mli b/plugins/ltac/tacentries.mli index 069504473..48598f7f4 100644 --- a/plugins/ltac/tacentries.mli +++ b/plugins/ltac/tacentries.mli @@ -20,7 +20,7 @@ val register_ltac : locality_flag -> Tacexpr.tacdef_body list -> unit type 'a grammar_tactic_prod_item_expr = 'a Pptactic.grammar_tactic_prod_item_expr = | TacTerm of string -| TacNonTerm of Loc.t * 'a * Names.Id.t +| TacNonTerm of ('a * Names.Id.t) Loc.located type raw_argument = string * string option (** An argument type as provided in Tactic notations, i.e. a string like diff --git a/plugins/ltac/tacexpr.mli b/plugins/ltac/tacexpr.mli index e8e99fcfe..bf760e7bb 100644 --- a/plugins/ltac/tacexpr.mli +++ b/plugins/ltac/tacexpr.mli @@ -265,7 +265,7 @@ and 'a gen_tactic_expr = | TacArg of 'a gen_tactic_arg located | TacSelect of goal_selector * 'a gen_tactic_expr (* For ML extensions *) - | TacML of Loc.t * ml_tactic_entry * 'a gen_tactic_arg list + | TacML of (ml_tactic_entry * 'a gen_tactic_arg list) Loc.located (* For syntax extensions *) | TacAlias of (KerName.t * 'a gen_tactic_arg list) Loc.located diff --git a/plugins/ltac/tacintern.ml b/plugins/ltac/tacintern.ml index df845b1a4..787a5f50b 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 Loc.ghost else loc +let adjust_loc loc = if !strict_check then None else Some loc (* Globalize a name which must be bound -- actually just check it is bound *) let intern_hyp ist (loc,id as locid) = @@ -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 *) @@ -550,7 +550,7 @@ and intern_tactic_seq onlytac ist = function | TacAtom (loc,t) -> let lf = ref ist.ltacvars in let t = intern_atomic lf ist t in - !lf, TacAtom (adjust_loc loc, t) + !lf, TacAtom (Loc.tag ?loc:(adjust_loc loc) t) | TacFun tacfun -> ist.ltacvars, TacFun (intern_tactic_fun ist tacfun) | TacLetIn (isrec,l,u) -> let ltacvars = Id.Set.union (extract_let_names l) ist.ltacvars in @@ -627,9 +627,9 @@ and intern_tactic_seq onlytac ist = function | TacAlias (loc,(s,l)) -> let l = List.map (intern_tacarg !strict_check false ist) l in ist.ltacvars, TacAlias (Loc.tag ~loc (s,l)) - | TacML (loc,opn,l) -> + | TacML (loc,(opn,l)) -> let _ignore = Tacenv.interp_ml_tactic opn in - ist.ltacvars, TacML (adjust_loc loc, opn,List.map (intern_tacarg !strict_check false ist) l) + ist.ltacvars, TacML (loc, (opn,List.map (intern_tacarg !strict_check false ist) l)) and intern_tactic_as_arg loc onlytac ist a = match intern_tacarg !strict_check onlytac ist a with diff --git a/plugins/ltac/tacinterp.ml b/plugins/ltac/tacinterp.ml index 64eda28ed..e969b86f6 100644 --- a/plugins/ltac/tacinterp.ml +++ b/plugins/ltac/tacinterp.ml @@ -201,8 +201,6 @@ end let print_top_val env v = Pptactic.pr_value Pptactic.ltop v -let dloc = Loc.ghost - let catching_error call_trace fail (e, info) = let inner_trace = Option.default [] (Exninfo.get info ltac_trace_info) @@ -326,7 +324,7 @@ let coerce_to_tactic loc id v = | _ -> fail () else fail () -let intro_pattern_of_ident id = (Loc.ghost, IntroNaming (IntroIdentifier id)) +let intro_pattern_of_ident id = (Loc.tag @@ IntroNaming (IntroIdentifier id)) let value_of_ident id = in_gen (topwit wit_intro_pattern) (intro_pattern_of_ident id) @@ -370,22 +368,22 @@ let debugging_exception_step ist signal_anomaly e pp = debugging_step ist (fun () -> pp() ++ spc() ++ str "raised the exception" ++ fnl() ++ explain_exc e) -let error_ltac_variable loc id env v s = - user_err ~loc (str "Ltac variable " ++ pr_id id ++ +let error_ltac_variable ?loc id env v s = + user_err ?loc (str "Ltac variable " ++ pr_id id ++ strbrk " is bound to" ++ spc () ++ pr_value env v ++ spc () ++ strbrk "which cannot be coerced to " ++ str s ++ str".") (* 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 with Not_found -> anomaly (str "Detected '" ++ Id.print (snd locid) ++ str "' as ltac var at interning time") let interp_ident ist env sigma id = - try try_interp_ltac_var (coerce_var_to_ident false env sigma) ist (Some (env,sigma)) (dloc,id) + try try_interp_ltac_var (coerce_var_to_ident false env sigma) ist (Some (env,sigma)) (Loc.tag id) with Not_found -> id (* Interprets an optional identifier, bound or fresh *) @@ -531,7 +529,7 @@ let extract_ids ids lfun = if has_type v (topwit wit_intro_pattern) then let (_, ipat) = out_gen (topwit wit_intro_pattern) v in if Id.List.mem id ids then accu - else accu @ intropattern_ids (dloc, ipat) + else accu @ intropattern_ids (Loc.tag ipat) else accu in Id.Map.fold fold lfun [] @@ -541,7 +539,7 @@ let default_fresh_id = Id.of_string "H" let interp_fresh_id ist env sigma l = let extract_ident ist env sigma id = try try_interp_ltac_var (coerce_to_ident_not_fresh env sigma) - ist (Some (env,sigma)) (dloc,id) + ist (Some (env,sigma)) (Loc.tag id) with Not_found -> id in let ids = List.map_filter (function ArgVar (_, id) -> Some id | _ -> None) l in let avoid = match TacStore.get ist.extra f_avoid_ids with @@ -977,14 +975,14 @@ let interp_binding_name ist sigma = function (* If a name is bound, it has to be a quantified hypothesis *) (* user has to use other names for variables if these ones clash with *) (* a name intented to be used as a (non-variable) identifier *) - try try_interp_ltac_var (coerce_to_quantified_hypothesis sigma) ist None(dloc,id) + try try_interp_ltac_var (coerce_to_quantified_hypothesis sigma) ist None(Loc.tag id) with Not_found -> NamedHyp id let interp_declared_or_quantified_hypothesis ist env sigma = function | AnonHyp n -> AnonHyp n | NamedHyp id -> try try_interp_ltac_var - (coerce_to_decl_or_quant_hyp env sigma) ist (Some (env,sigma)) (dloc,id) + (coerce_to_decl_or_quant_hyp env sigma) ist (Some (env,sigma)) (Loc.tag id) with Not_found -> NamedHyp id let interp_binding ist env sigma (loc,(b,c)) = @@ -1012,14 +1010,14 @@ let interp_open_constr_with_bindings ist env sigma (c,bl) = sigma, (c, bl) let loc_of_bindings = function -| NoBindings -> Loc.ghost -| ImplicitBindings l -> loc_of_glob_constr (fst (List.last l)) -| ExplicitBindings l -> fst (List.last l) +| NoBindings -> None +| ImplicitBindings l -> Some (loc_of_glob_constr (fst (List.last l))) +| ExplicitBindings l -> Some (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 = if Loc.is_ghost loc2 then loc1 else Loc.merge loc1 loc2 in + let loc = Loc.opt_merge 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 @@ -1288,8 +1286,8 @@ and eval_tactic ist tac : unit Proofview.tactic = match tac with in Ftactic.run tac (fun () -> Proofview.tclUNIT ()) - | TacML (loc,opn,l) -> - push_trace (loc,LtacMLCall tac) ist >>= fun trace -> + | TacML (loc,(opn,l)) -> + 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 @@ -1308,7 +1306,7 @@ and force_vrec ist v : Val.t Ftactic.t = | v -> Ftactic.return (of_tacvalue v) else Ftactic.return v -and interp_ltac_reference loc' mustbetac ist r : Val.t Ftactic.t = +and interp_ltac_reference ?loc' mustbetac ist r : Val.t Ftactic.t = match r with | ArgVar (loc,id) -> let v = @@ -1322,7 +1320,7 @@ and interp_ltac_reference loc' mustbetac ist r : Val.t Ftactic.t = end | ArgArg (loc,r) -> let ids = extract_ids [] ist.lfun in - let loc_info = ((if Loc.is_ghost loc' then loc else loc'),LtacNameCall r) in + let loc_info = (Option.default loc loc',LtacNameCall r) in let extra = TacStore.set ist.extra f_avoid_ids ids in push_trace loc_info ist >>= fun trace -> let extra = TacStore.set extra f_trace trace in @@ -1333,7 +1331,7 @@ and interp_ltac_reference loc' mustbetac ist r : Val.t Ftactic.t = and interp_tacarg ist arg : Val.t Ftactic.t = match arg with | TacGeneric arg -> interp_genarg ist arg - | Reference r -> interp_ltac_reference dloc false ist r + | Reference r -> interp_ltac_reference false ist r | ConstrMayEval c -> Ftactic.s_enter { s_enter = begin fun gl -> let sigma = project gl in @@ -1342,16 +1340,16 @@ and interp_tacarg ist arg : Val.t Ftactic.t = Sigma.Unsafe.of_pair (Ftactic.return (Value.of_constr c_interp), sigma) end } | TacCall (loc,(r,[])) -> - interp_ltac_reference loc true ist r + interp_ltac_reference true ist r | TacCall (loc,(f,l)) -> let (>>=) = Ftactic.bind in - interp_ltac_reference loc true ist f >>= fun fv -> + interp_ltac_reference true ist f >>= fun fv -> Ftactic.List.map (fun a -> interp_tacarg ist a) l >>= fun largs -> interp_app loc ist fv largs | TacFreshId l -> Ftactic.enter { enter = begin fun gl -> let id = interp_fresh_id ist (pf_env gl) (project gl) l in - Ftactic.return (in_gen (topwit wit_intro_pattern) (dloc, IntroNaming (IntroIdentifier id))) + Ftactic.return (in_gen (topwit wit_intro_pattern) (Loc.tag @@ IntroNaming (IntroIdentifier id))) end } | TacPretype c -> Ftactic.s_enter { s_enter = begin fun gl -> @@ -1442,7 +1440,7 @@ and interp_letrec ist llc u = Proofview.tclUNIT () >>= fun () -> (* delay for the effects of [lref], just in case. *) let lref = ref ist.lfun in let fold accu ((_, id), b) = - let v = of_tacvalue (VRec (lref, TacArg (dloc, b))) in + let v = of_tacvalue (VRec (lref, TacArg (Loc.tag b))) in Id.Map.add id v accu in let lfun = List.fold_left fold ist.lfun llc in @@ -1768,7 +1766,7 @@ and interp_atomic ist tac : unit Proofview.tactic = (* We try to fully-typecheck the term *) let (sigma,c_interp) = interp_constr ist env sigma c in let let_tac b na c cl eqpat = - let id = Option.default (Loc.ghost,IntroAnonymous) eqpat in + let id = Option.default (Loc.tag IntroAnonymous) eqpat in let with_eq = if b then None else Some (true,id) in Tactics.letin_tac with_eq na c None cl in @@ -1780,7 +1778,7 @@ and interp_atomic ist tac : unit Proofview.tactic = else (* We try to keep the pattern structure as much as possible *) let let_pat_tac b na c cl eqpat = - let id = Option.default (Loc.ghost,IntroAnonymous) eqpat in + let id = Option.default (Loc.tag IntroAnonymous) eqpat in let with_eq = if b then None else Some (true,id) in Tactics.letin_pat_tac with_eq na c cl in @@ -2132,7 +2130,7 @@ let lift_constr_tac_to_ml_tac vars tac = let c = Id.Map.find id ist.lfun in try Some (coerce_to_closed_constr env c) with CannotCoerceTo ty -> - error_ltac_variable Loc.ghost dummy_id (Some (env,sigma)) c ty + error_ltac_variable dummy_id (Some (env,sigma)) c ty in let args = List.map_filter map vars in tac args ist diff --git a/plugins/ltac/tacinterp.mli b/plugins/ltac/tacinterp.mli index 1e5f6bd42..6cd5a63b3 100644 --- a/plugins/ltac/tacinterp.mli +++ b/plugins/ltac/tacinterp.mli @@ -111,7 +111,7 @@ val interp_int : interp_sign -> Id.t Loc.located -> int val interp_int_or_var : interp_sign -> int or_var -> int -val error_ltac_variable : Loc.t -> Id.t -> +val error_ltac_variable : ?loc:Loc.t -> Id.t -> (Environ.env * Evd.evar_map) option -> value -> string -> 'a (** Transforms a constr-expecting tactic into a tactic finding its arguments in diff --git a/plugins/ltac/tacsubst.ml b/plugins/ltac/tacsubst.ml index c92dd23a0..0ee6e8a85 100644 --- a/plugins/ltac/tacsubst.ml +++ b/plugins/ltac/tacsubst.ml @@ -234,7 +234,7 @@ and subst_tactic subst (t:glob_tactic_expr) = match t with | TacAlias (_,(s,l)) -> let s = subst_kn subst s in TacAlias (Loc.tag (s,List.map (subst_tacarg subst) l)) - | TacML (_loc,opn,l) -> TacML (_loc, opn,List.map (subst_tacarg subst) l) + | TacML (loc,(opn,l)) -> TacML (loc, (opn,List.map (subst_tacarg subst) l)) and subst_tactic_fun subst (var,body) = (var,subst_tactic subst body) diff --git a/plugins/ltac/tactic_debug.ml b/plugins/ltac/tactic_debug.ml index dffeade29..ac8534bdc 100644 --- a/plugins/ltac/tactic_debug.ml +++ b/plugins/ltac/tactic_debug.ml @@ -352,7 +352,7 @@ let explain_ltac_call_trace last trace loc = Pptactic.pr_glob_tactic (Global.env()) t ++ str ")" | Tacexpr.LtacAtomCall te -> quote (Pptactic.pr_glob_tactic (Global.env()) - (Tacexpr.TacAtom (Loc.ghost,te))) + (Tacexpr.TacAtom (Loc.tag te))) | Tacexpr.LtacConstrInterp (c, { Pretyping.ltac_constrs = vars }) -> quote (Printer.pr_glob_constr_env (Global.env()) c) ++ (if not (Id.Map.is_empty vars) then @@ -389,14 +389,15 @@ let skip_extensions trace = let finer_loc loc1 loc2 = Loc.merge loc1 loc2 = loc2 -let extract_ltac_trace trace eloc = +let extract_ltac_trace ?loc trace = let trace = skip_extensions trace in - let (loc,c),tail = List.sep_last 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 *) - let msg = hov 0 (explain_ltac_call_trace c tail eloc ++ fnl()) in - Some msg, if finer_loc eloc loc then eloc else loc + let msg = hov 0 (explain_ltac_call_trace c tail loc ++ fnl()) in + (if finer_loc loc tloc then loc else tloc), Some msg else (* We entered a primitive tactic, we don't display trace but report on the finest location *) @@ -411,14 +412,14 @@ let extract_ltac_trace trace eloc = else aux best_loc tail | [] -> best_loc in - aux eloc trace in - None, best_loc + aux loc trace in + best_loc, None let get_ltac_trace (_, info) = let ltac_trace = Exninfo.get info ltac_trace_info in - let loc = Option.default Loc.ghost (Loc.get_loc info) in + let loc = Loc.get_loc info in match ltac_trace with | None -> None - | Some trace -> Some (extract_ltac_trace trace loc) + | Some trace -> Some (extract_ltac_trace ?loc trace) let () = ExplainErr.register_additional_error_info get_ltac_trace diff --git a/plugins/ltac/tactic_debug.mli b/plugins/ltac/tactic_debug.mli index 7745d9b7b..38d8caca6 100644 --- a/plugins/ltac/tactic_debug.mli +++ b/plugins/ltac/tactic_debug.mli @@ -78,4 +78,4 @@ val db_breakpoint : debug_info -> Id.t Loc.located message_token list -> unit Proofview.NonLogical.t val extract_ltac_trace : - Tacexpr.ltac_trace -> Loc.t -> Pp.std_ppcmds option * Loc.t + ?loc:Loc.t -> Tacexpr.ltac_trace -> Pp.std_ppcmds option Loc.located diff --git a/plugins/ltac/tauto.ml b/plugins/ltac/tauto.ml index 30eed08d7..df186cc46 100644 --- a/plugins/ltac/tauto.ml +++ b/plugins/ltac/tauto.ml @@ -89,7 +89,6 @@ let _ = (** Base tactics *) -let loc = Loc.ghost let idtac = Proofview.tclUNIT () let fail = Proofview.tclINDEPENDENT (tclFAIL 0 (Pp.mt ())) @@ -207,7 +206,7 @@ let u_iff = make_unfold "iff" let u_not = make_unfold "not" let reduction_not_iff _ ist = - let make_reduce c = TacAtom (loc, TacReduce (Genredexpr.Unfold c, Locusops.allHypsAndConcl)) in + let make_reduce c = TacAtom (Loc.tag @@ TacReduce (Genredexpr.Unfold c, Locusops.allHypsAndConcl)) in let tac = match !negation_unfolding, unfold_iff () with | true, true -> make_reduce [u_not; u_iff] | true, false -> make_reduce [u_not] @@ -260,11 +259,11 @@ let tauto_power_flags = { } let with_flags flags _ ist = - let f = (loc, Id.of_string "f") in - let x = (loc, Id.of_string "x") in + let f = (Loc.tag @@ Id.of_string "f") in + let x = (Loc.tag @@ Id.of_string "x") in let arg = Val.Dyn (tag_tauto_flags, flags) in let ist = { ist with lfun = Id.Map.add (snd x) arg ist.lfun } in - eval_tactic_ist ist (TacArg (loc, TacCall (loc, (ArgVar f, [Reference (ArgVar x)])))) + eval_tactic_ist ist (TacArg (Loc.tag @@ TacCall (Loc.tag (ArgVar f, [Reference (ArgVar x)])))) let register_tauto_tactic tac name0 args = let ids = List.map (fun id -> Id.of_string id) args in @@ -272,7 +271,7 @@ let register_tauto_tactic tac name0 args = let name = { mltac_plugin = tauto_plugin; mltac_tactic = name0; } in let entry = { mltac_name = name; mltac_index = 0 } in let () = Tacenv.register_ml_tactic name [| tac |] in - let tac = TacFun (ids, TacML (loc, entry, [])) in + let tac = TacFun (ids, TacML (Loc.tag (entry, []))) in let obj () = Tacenv.register_ltac true true (Id.of_string name0) tac in Mltop.declare_cache_obj obj tauto_plugin diff --git a/plugins/micromega/coq_micromega.ml b/plugins/micromega/coq_micromega.ml index 4b87e6e2e..79d17db25 100644 --- a/plugins/micromega/coq_micromega.ml +++ b/plugins/micromega/coq_micromega.ml @@ -1997,7 +1997,7 @@ let micromega_gen let intro_vars = Tacticals.New.tclTHENLIST (List.map intro vars) in let intro_props = Tacticals.New.tclTHENLIST (List.map intro props) in - let ipat_of_name id = Some (Loc.ghost, Misctypes.IntroNaming (Misctypes.IntroIdentifier id)) in + let ipat_of_name id = Some (Loc.tag @@ Misctypes.IntroNaming (Misctypes.IntroIdentifier id)) in let goal_name = Tactics.fresh_id [] (Names.Id.of_string "__arith") gl in let env' = List.map (fun (id,i) -> Term.mkVar id,i) vars in @@ -2113,7 +2113,7 @@ let micromega_genr prover tac = let intro_vars = Tacticals.New.tclTHENLIST (List.map intro vars) in let intro_props = Tacticals.New.tclTHENLIST (List.map intro props) in - let ipat_of_name id = Some (Loc.ghost, Misctypes.IntroNaming (Misctypes.IntroIdentifier id)) in + let ipat_of_name id = Some (Loc.tag @@ Misctypes.IntroNaming (Misctypes.IntroIdentifier id)) in let goal_name = Tactics.fresh_id [] (Names.Id.of_string "__arith") gl in let env' = List.map (fun (id,i) -> Term.mkVar id,i) vars in diff --git a/plugins/quote/g_quote.ml4 b/plugins/quote/g_quote.ml4 index 432625e4d..980f03db3 100644 --- a/plugins/quote/g_quote.ml4 +++ b/plugins/quote/g_quote.ml4 @@ -19,15 +19,14 @@ open Tacarg DECLARE PLUGIN "quote_plugin" -let loc = Loc.ghost let cont = Id.of_string "cont" let x = Id.of_string "x" let make_cont (k : Val.t) (c : EConstr.t) = let c = Tacinterp.Value.of_constr c in - let tac = TacCall (loc, (ArgVar (loc, cont), [Reference (ArgVar (loc, x))])) in + let tac = TacCall (Loc.tag (ArgVar (Loc.tag cont), [Reference (ArgVar (Loc.tag x))])) in let ist = { lfun = Id.Map.add cont k (Id.Map.singleton x c); extra = TacStore.empty; } in - Tacinterp.eval_tactic_ist ist (TacArg (loc, tac)) + Tacinterp.eval_tactic_ist ist (TacArg (Loc.tag tac)) TACTIC EXTEND quote [ "quote" ident(f) ] -> [ quote f [] ] diff --git a/plugins/setoid_ring/newring.ml b/plugins/setoid_ring/newring.ml index e3e14bcf3..e2a6ad55c 100644 --- a/plugins/setoid_ring/newring.ml +++ b/plugins/setoid_ring/newring.ml @@ -128,11 +128,11 @@ let closed_term_ast l = mltac_name = tacname; mltac_index = 0; } in - let l = List.map (fun gr -> ArgArg(Loc.ghost,gr)) l in + let l = List.map (fun gr -> ArgArg(Loc.tag gr)) l in TacFun([Name(Id.of_string"t")], - TacML(Loc.ghost,tacname, + TacML(Loc.tag (tacname, [TacGeneric (Genarg.in_gen (Genarg.glbwit Stdarg.wit_constr) (Loc.tag @@ GVar(Id.of_string"t"),None)); - TacGeneric (Genarg.in_gen (Genarg.glbwit (Genarg.wit_list Stdarg.wit_ref)) l)])) + TacGeneric (Genarg.in_gen (Genarg.glbwit (Genarg.wit_list Stdarg.wit_ref)) l)]))) (* let _ = add_tacdef false ((Loc.ghost,Id.of_string"ring_closed_term" *) @@ -170,7 +170,7 @@ let ltac_lcall tac args = let ltac_apply (f : Value.t) (args: Tacinterp.Value.t list) = let fold arg (i, vars, lfun) = let id = Id.of_string ("x" ^ string_of_int i) in - let x = Reference (ArgVar (Loc.ghost, id)) in + let x = Reference (ArgVar (Loc.tag id)) in (succ i, x :: vars, Id.Map.add id arg lfun) in let (_, args, lfun) = List.fold_right fold args (0, [], Id.Map.empty) in @@ -205,7 +205,7 @@ let get_res = let exec_tactic env evd n f args = let fold arg (i, vars, lfun) = let id = Id.of_string ("x" ^ string_of_int i) in - let x = Reference (ArgVar (Loc.ghost, id)) in + let x = Reference (ArgVar (Loc.tag id)) in (succ i, x :: vars, Id.Map.add id (Value.of_constr arg) lfun) in let (_, args, lfun) = List.fold_right fold args (0, [], Id.Map.empty) in @@ -213,7 +213,7 @@ let exec_tactic env evd n f args = (** Build the getter *) let lid = List.init n (fun i -> Id.of_string("x"^string_of_int i)) in let n = Genarg.in_gen (Genarg.glbwit Stdarg.wit_int) n in - let get_res = TacML (Loc.ghost, get_res, [TacGeneric n]) in + let get_res = TacML (Loc.tag (get_res, [TacGeneric n])) in let getter = Tacexp (TacFun (List.map (fun n -> Name n) lid, get_res)) in (** Evaluate the whole result *) let gl = dummy_goal env evd in @@ -603,7 +603,7 @@ let interp_power env evd pow = match pow with | None -> let t = ArgArg(Loc.tag (Lazy.force ltac_inv_morph_nothing)) in - (TacArg(Loc.ghost,TacCall(Loc.tag (t,[]))), plapp evd coq_None [|carrier|]) + (TacArg(Loc.tag (TacCall(Loc.tag (t,[])))), plapp evd coq_None [|carrier|]) | Some (tac, spec) -> let tac = match tac with diff --git a/plugins/ssrmatching/ssrmatching.ml4 b/plugins/ssrmatching/ssrmatching.ml4 index f8cccf714..a6a6bd6f9 100644 --- a/plugins/ssrmatching/ssrmatching.ml4 +++ b/plugins/ssrmatching/ssrmatching.ml4 @@ -151,12 +151,12 @@ let dC t = CastConv t let isCVar = function _loc, CRef (Ident _, _) -> true | _ -> false let destCVar = function _loc, CRef (Ident (_, id), _) -> id | _ -> CErrors.anomaly (str"not a CRef") -let mkCHole loc = Loc.tag ~loc @@ CHole (None, IntroAnonymous, None) -let mkCLambda loc name ty t = Loc.tag ~loc @@ - CLambdaN ([[loc, name], Default Explicit, ty], t) -let mkCLetIn loc name bo t = Loc.tag ~loc @@ - CLetIn ((loc, name), bo, None, t) -let mkCCast loc t ty = Loc.tag ~loc @@ CCast (t, dC ty) +let mkCHole ~loc = Loc.tag ?loc @@ CHole (None, IntroAnonymous, None) +let mkCLambda ?loc name ty t = Loc.tag ?loc @@ + CLambdaN ([[Loc.tag ?loc name], Default Explicit, ty], t) +let mkCLetIn ?loc name bo t = Loc.tag ?loc @@ + CLetIn ((Loc.tag ?loc name), bo, None, t) +let mkCCast ?loc t ty = Loc.tag ?loc @@ CCast (t, dC ty) (** Constructors for rawconstr *) let mkRHole = Loc.tag @@ GHole (InternalHole, IntroAnonymous, None) let mkRApp f args = if args = [] then f else Loc.tag @@ GApp (f, args) @@ -943,8 +943,8 @@ let glob_cpattern gs p = let name = Name (id_of_string ("_ssrpat_" ^ s)) in k, (mkRCast mkRHole (mkRLambda name mkRHole (mkRApp mkRHole l)), None) in let bind_in t1 t2 = - let d = Loc.ghost in let n = Name (destCVar t1) in - fst (glob (mkCCast d (mkCHole d) (mkCLambda d n (mkCHole d) t2))) in + let mkCHole = mkCHole ~loc:None in let n = Name (destCVar t1) in + fst (glob (mkCCast mkCHole (mkCLambda n mkCHole t2))) in let check_var t2 = if not (isCVar t2) then loc_error (constr_loc t2) "Only identifiers are allowed here" in match p with @@ -1187,27 +1187,27 @@ 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_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_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_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 - | Some b -> a,(g,Some (mkCLetIn loc x (mkCHole loc) b)) - | None -> a,(Loc.tag ~loc @@ GLetIn (x, Loc.tag ~loc @@ GHole (BinderType x, IntroAnonymous, None), None, g), None) in + let mkXLetIn ?loc x (a,(g,c)) = match c with + | Some b -> a,(g,Some (mkCLetIn ?loc x (mkCHole ?loc) b)) + | None -> a,(Loc.tag ?loc @@ GLetIn (x, Loc.tag ?loc @@ GHole (BinderType x, IntroAnonymous, None), None, g), None) in match red with | T t -> let sigma, t = interp_term ist gl t in sigma, T t | In_T t -> let sigma, t = interp_term ist gl t in sigma, In_T t | X_In_T (x, rp) | In_X_In_T (x, rp) -> let mk x p = match red with X_In_T _ -> X_In_T(x,p) | _ -> In_X_In_T(x,p) in - let rp = mkXLetIn Loc.ghost (Name x) rp in + let rp = mkXLetIn (Name x) rp in let sigma, rp = interp_term ist gl rp in let _, h, _, rp = destLetIn rp in let sigma = cleanup_XinE h x rp sigma in @@ -1216,7 +1216,7 @@ let interp_pattern ?wit_ssrpatternarg ist gl red redty = | E_In_X_In_T(e, x, rp) | E_As_X_In_T (e, x, rp) -> let mk e x p = match red with E_In_X_In_T _ ->E_In_X_In_T(e,x,p)|_->E_As_X_In_T(e,x,p) in - let rp = mkXLetIn Loc.ghost (Name x) rp in + let rp = mkXLetIn (Name x) rp in let sigma, rp = interp_term ist gl rp in let _, h, _, rp = destLetIn rp in let sigma = cleanup_XinE h x rp sigma in @@ -1426,7 +1426,7 @@ let () = let () = Tacenv.register_ml_tactic name [|mltac|] in let tac = TacFun ([Name (Id.of_string "pattern")], - TacML (Loc.ghost, { mltac_name = name; mltac_index = 0 }, [])) in + TacML (Loc.tag ({ mltac_name = name; mltac_index = 0 }, []))) in let obj () = Tacenv.register_ltac true false (Id.of_string "ssrpattern") tac in Mltop.declare_cache_obj obj "ssrmatching_plugin" diff --git a/plugins/syntax/ascii_syntax.ml b/plugins/syntax/ascii_syntax.ml index dc0b87793..ed977c416 100644 --- a/plugins/syntax/ascii_syntax.ml +++ b/plugins/syntax/ascii_syntax.ml @@ -37,24 +37,24 @@ let glob_Ascii = lazy (make_reference "Ascii") open Lazy -let interp_ascii loc p = +let interp_ascii ?loc p = let rec aux n p = if Int.equal n 0 then [] else let mp = p mod 2 in - (Loc.tag ~loc @@ GRef ((if Int.equal mp 0 then glob_false else glob_true),None)) + (Loc.tag ?loc @@ GRef ((if Int.equal mp 0 then glob_false else glob_true),None)) :: (aux (n-1) (p/2)) in - Loc.tag ~loc @@ GApp (Loc.tag ~loc @@ GRef(force glob_Ascii,None), aux 8 p) + Loc.tag ?loc @@ GApp (Loc.tag ?loc @@ GRef(force glob_Ascii,None), aux 8 p) -let interp_ascii_string dloc s = +let interp_ascii_string ?loc s = let p = if Int.equal (String.length s) 1 then int_of_char s.[0] else if Int.equal (String.length s) 3 && is_digit s.[0] && is_digit s.[1] && is_digit s.[2] then int_of_string s else - user_err ~loc:dloc ~hdr:"interp_ascii_string" + user_err ?loc ~hdr:"interp_ascii_string" (str "Expects a single character or a three-digits ascii code.") in - interp_ascii dloc p + interp_ascii ?loc p let uninterp_ascii r = let rec uninterp_bool_list n = function diff --git a/plugins/syntax/nat_syntax.ml b/plugins/syntax/nat_syntax.ml index 90d643b7f..5cdd82024 100644 --- a/plugins/syntax/nat_syntax.ml +++ b/plugins/syntax/nat_syntax.ml @@ -33,21 +33,21 @@ let warn_large_nat = strbrk "may vary from 5000 to 70000 depending on your system " ++ strbrk "limits and on the command executed).") -let nat_of_int loc n = +let nat_of_int ?loc n = if is_pos_or_zero n then begin if less_than threshold n then warn_large_nat (); - let ref_O = Loc.tag ~loc @@ GRef (glob_O, None) in - let ref_S = Loc.tag ~loc @@ GRef (glob_S, None) in + let ref_O = Loc.tag ?loc @@ GRef (glob_O, None) in + let ref_S = Loc.tag ?loc @@ GRef (glob_S, None) in let rec mk_nat acc n = if n <> zero then - mk_nat (Loc.tag ~loc @@ GApp (ref_S, [acc])) (sub_1 n) + mk_nat (Loc.tag ?loc @@ GApp (ref_S, [acc])) (sub_1 n) else acc in mk_nat ref_O n end else - user_err ~hdr:"nat_of_int" + user_err ?loc ~hdr:"nat_of_int" (str "Cannot interpret a negative number as a number of type nat") (************************************************************************) diff --git a/plugins/syntax/numbers_syntax.ml b/plugins/syntax/numbers_syntax.ml index 8876d464a..3ee64ba7e 100644 --- a/plugins/syntax/numbers_syntax.ml +++ b/plugins/syntax/numbers_syntax.ml @@ -86,10 +86,10 @@ exception Non_closed (* parses a *non-negative* integer (from bigint.ml) into an int31 wraps modulo 2^31 *) -let int31_of_pos_bigint loc n = - let ref_construct = Loc.tag ~loc @@ GRef (int31_construct, None) in - let ref_0 = Loc.tag ~loc @@ GRef (int31_0, None) in - let ref_1 = Loc.tag ~loc @@ GRef (int31_1, None) in +let int31_of_pos_bigint ?loc n = + let ref_construct = Loc.tag ?loc @@ GRef (int31_construct, None) in + let ref_0 = Loc.tag ?loc @@ GRef (int31_0, None) in + let ref_1 = Loc.tag ?loc @@ GRef (int31_1, None) in let rec args counter n = if counter <= 0 then [] @@ -97,16 +97,16 @@ let int31_of_pos_bigint loc n = let (q,r) = div2_with_rest n in (if r then ref_1 else ref_0)::(args (counter-1) q) in - Loc.tag ~loc @@ GApp (ref_construct, List.rev (args 31 n)) + Loc.tag ?loc @@ GApp (ref_construct, List.rev (args 31 n)) -let error_negative dloc = - CErrors.user_err ~loc:dloc ~hdr:"interp_int31" (Pp.str "int31 are only non-negative numbers.") +let error_negative ?loc = + CErrors.user_err ?loc ~hdr:"interp_int31" (Pp.str "int31 are only non-negative numbers.") -let interp_int31 dloc n = +let interp_int31 ?loc n = if is_pos_or_zero n then - int31_of_pos_bigint dloc n + int31_of_pos_bigint ?loc n else - error_negative dloc + error_negative ?loc (* Pretty prints an int31 *) @@ -162,40 +162,40 @@ let height bi = in hght 0 base (* n must be a non-negative integer (from bigint.ml) *) -let word_of_pos_bigint loc hght n = - let ref_W0 = Loc.tag ~loc @@ GRef (zn2z_W0, None) in - let ref_WW = Loc.tag ~loc @@ GRef (zn2z_WW, None) in +let word_of_pos_bigint ?loc hght n = + let ref_W0 = Loc.tag ?loc @@ GRef (zn2z_W0, None) in + let ref_WW = Loc.tag ?loc @@ GRef (zn2z_WW, None) in let rec decomp hgt n = if hgt <= 0 then - int31_of_pos_bigint loc n + int31_of_pos_bigint ?loc n else if equal n zero then - Loc.tag ~loc @@ GApp (ref_W0, [Loc.tag ~loc @@ GHole (Evar_kinds.InternalHole, Misctypes.IntroAnonymous, None)]) + Loc.tag ?loc @@ GApp (ref_W0, [Loc.tag ?loc @@ GHole (Evar_kinds.InternalHole, Misctypes.IntroAnonymous, None)]) else let (h,l) = split_at hgt n in - Loc.tag ~loc @@ GApp (ref_WW, [Loc.tag ~loc @@ GHole (Evar_kinds.InternalHole, Misctypes.IntroAnonymous, None); + Loc.tag ?loc @@ GApp (ref_WW, [Loc.tag ?loc @@ GHole (Evar_kinds.InternalHole, Misctypes.IntroAnonymous, None); decomp (hgt-1) h; decomp (hgt-1) l]) in decomp hght n -let bigN_of_pos_bigint loc n = +let bigN_of_pos_bigint ?loc n = let h = height n in - let ref_constructor = Loc.tag ~loc @@ GRef (bigN_constructor h, None) in - let word = word_of_pos_bigint loc h n in + let ref_constructor = Loc.tag ?loc @@ GRef (bigN_constructor h, None) in + let word = word_of_pos_bigint ?loc h n in let args = if h < n_inlined then [word] - else [Nat_syntax_plugin.Nat_syntax.nat_of_int loc (of_int (h-n_inlined));word] + else [Nat_syntax_plugin.Nat_syntax.nat_of_int ?loc (of_int (h-n_inlined));word] in - Loc.tag ~loc @@ GApp (ref_constructor, args) + Loc.tag ?loc @@ GApp (ref_constructor, args) -let bigN_error_negative loc = - CErrors.user_err ~loc ~hdr:"interp_bigN" (Pp.str "bigN are only non-negative numbers.") +let bigN_error_negative ?loc = + CErrors.user_err ?loc ~hdr:"interp_bigN" (Pp.str "bigN are only non-negative numbers.") -let interp_bigN dloc n = +let interp_bigN ?loc n = if is_pos_or_zero n then - bigN_of_pos_bigint dloc n + bigN_of_pos_bigint ?loc n else - bigN_error_negative dloc + bigN_error_negative ?loc (* Pretty prints a bigN *) @@ -256,13 +256,13 @@ let _ = Notation.declare_numeral_interpreter bigN_scope (*** Parsing for bigZ in digital notation ***) -let interp_bigZ loc n = - let ref_pos = Loc.tag ~loc @@ GRef (bigZ_pos, None) in - let ref_neg = Loc.tag ~loc @@ GRef (bigZ_neg, None) in +let interp_bigZ ?loc n = + let ref_pos = Loc.tag ?loc @@ GRef (bigZ_pos, None) in + let ref_neg = Loc.tag ?loc @@ GRef (bigZ_neg, None) in if is_pos_or_zero n then - Loc.tag ~loc @@ GApp (ref_pos, [bigN_of_pos_bigint loc n]) + Loc.tag ?loc @@ GApp (ref_pos, [bigN_of_pos_bigint ?loc n]) else - Loc.tag ~loc @@ GApp (ref_neg, [bigN_of_pos_bigint loc (neg n)]) + Loc.tag ?loc @@ GApp (ref_neg, [bigN_of_pos_bigint ?loc (neg n)]) (* pretty printing functions for bigZ *) let bigint_of_bigZ = function @@ -292,9 +292,9 @@ let _ = Notation.declare_numeral_interpreter bigZ_scope true) (*** Parsing for bigQ in digital notation ***) -let interp_bigQ loc n = - let ref_z = Loc.tag ~loc @@ GRef (bigQ_z, None) in - Loc.tag ~loc @@ GApp (ref_z, [interp_bigZ loc n]) +let interp_bigQ ?loc n = + let ref_z = Loc.tag ?loc @@ GRef (bigQ_z, None) in + Loc.tag ?loc @@ GApp (ref_z, [interp_bigZ ?loc n]) let uninterp_bigQ rc = try match rc with diff --git a/plugins/syntax/r_syntax.ml b/plugins/syntax/r_syntax.ml index 1af3f6c5b..b7041d045 100644 --- a/plugins/syntax/r_syntax.ml +++ b/plugins/syntax/r_syntax.ml @@ -41,7 +41,7 @@ let glob_xI = ConstructRef path_of_xI let glob_xO = ConstructRef path_of_xO let glob_xH = ConstructRef path_of_xH -let pos_of_bignat dloc x = +let pos_of_bignat ?loc x = let ref_xI = Loc.tag @@ GRef (glob_xI, None) in let ref_xH = Loc.tag @@ GRef (glob_xH, None) in let ref_xO = Loc.tag @@ GRef (glob_xO, None) in @@ -77,11 +77,11 @@ let glob_ZERO = ConstructRef path_of_ZERO let glob_POS = ConstructRef path_of_POS let glob_NEG = ConstructRef path_of_NEG -let z_of_int dloc n = +let z_of_int ?loc n = if not (Bigint.equal n zero) then let sgn, n = if is_pos_or_zero n then glob_POS, n else glob_NEG, Bigint.neg n in - Loc.tag @@ GApp(Loc.tag @@ GRef (sgn,None), [pos_of_bignat dloc n]) + Loc.tag @@ GApp(Loc.tag @@ GRef (sgn,None), [pos_of_bignat ?loc n]) else Loc.tag @@ GRef (glob_ZERO, None) @@ -107,8 +107,8 @@ let make_path dir id = Globnames.encode_con dir (Id.of_string id) let glob_IZR = ConstRef (make_path (make_dir rdefinitions) "IZR") -let r_of_int dloc z = - Loc.tag @@ GApp (Loc.tag @@ GRef(glob_IZR,None), [z_of_int dloc z]) +let r_of_int ?loc z = + Loc.tag @@ GApp (Loc.tag @@ GRef(glob_IZR,None), [z_of_int ?loc z]) (**********************************************************************) (* Printing R via scopes *) @@ -128,6 +128,6 @@ let uninterp_r p = let _ = Notation.declare_numeral_interpreter "R_scope" (r_path,["Coq";"Reals";"Rdefinitions"]) r_of_int - ([Loc.tag @@ GRef (glob_IZR,None)], + ([Loc.tag @@ GRef (glob_IZR, None)], uninterp_r, false) diff --git a/plugins/syntax/string_syntax.ml b/plugins/syntax/string_syntax.ml index 539670722..49cb9355c 100644 --- a/plugins/syntax/string_syntax.ml +++ b/plugins/syntax/string_syntax.ml @@ -33,12 +33,12 @@ let glob_EmptyString = lazy (make_reference "EmptyString") open Lazy -let interp_string loc s = +let interp_string ?loc s = let le = String.length s in let rec aux n = - if n = le then Loc.tag ~loc @@ GRef (force glob_EmptyString, None) else - Loc.tag ~loc @@ GApp (Loc.tag ~loc @@ GRef (force glob_String, None), - [interp_ascii loc (int_of_char s.[n]); aux (n+1)]) + if n = le then Loc.tag ?loc @@ GRef (force glob_EmptyString, None) else + Loc.tag ?loc @@ GApp (Loc.tag ?loc @@ GRef (force glob_String, None), + [interp_ascii ?loc (int_of_char s.[n]); aux (n+1)]) in aux 0 let uninterp_string r = diff --git a/plugins/syntax/z_syntax.ml b/plugins/syntax/z_syntax.ml index a00525f91..96c1f3e39 100644 --- a/plugins/syntax/z_syntax.ml +++ b/plugins/syntax/z_syntax.ml @@ -44,25 +44,25 @@ let glob_xI = ConstructRef path_of_xI let glob_xO = ConstructRef path_of_xO let glob_xH = ConstructRef path_of_xH -let pos_of_bignat loc x = - let ref_xI = Loc.tag ~loc @@ GRef (glob_xI, None) in - let ref_xH = Loc.tag ~loc @@ GRef (glob_xH, None) in - let ref_xO = Loc.tag ~loc @@ GRef (glob_xO, None) in +let pos_of_bignat ?loc x = + let ref_xI = Loc.tag ?loc @@ GRef (glob_xI, None) in + let ref_xH = Loc.tag ?loc @@ GRef (glob_xH, None) in + let ref_xO = Loc.tag ?loc @@ GRef (glob_xO, None) in let rec pos_of x = match div2_with_rest x with - | (q,false) -> Loc.tag ~loc @@ GApp (ref_xO,[pos_of q]) - | (q,true) when not (Bigint.equal q zero) -> Loc.tag ~loc @@ GApp (ref_xI,[pos_of q]) + | (q,false) -> Loc.tag ?loc @@ GApp (ref_xO,[pos_of q]) + | (q,true) when not (Bigint.equal q zero) -> Loc.tag ?loc @@ GApp (ref_xI,[pos_of q]) | (q,true) -> ref_xH in pos_of x -let error_non_positive dloc = - user_err ~loc:dloc ~hdr:"interp_positive" +let error_non_positive ?loc = + user_err ?loc ~hdr:"interp_positive" (str "Only strictly positive numbers in type \"positive\".") -let interp_positive dloc n = - if is_strictly_pos n then pos_of_bignat dloc n - else error_non_positive dloc +let interp_positive ?loc n = + if is_strictly_pos n then pos_of_bignat ?loc n + else error_non_positive ?loc (**********************************************************************) (* Printing positive via scopes *) @@ -106,18 +106,18 @@ let glob_Npos = ConstructRef path_of_Npos let n_path = make_path binnums "N" -let n_of_binnat loc pos_or_neg n = Loc.tag ~loc @@ +let n_of_binnat ?loc pos_or_neg n = Loc.tag ?loc @@ if not (Bigint.equal n zero) then - GApp(Loc.tag @@ GRef (glob_Npos,None), [pos_of_bignat loc n]) + GApp(Loc.tag @@ GRef (glob_Npos,None), [pos_of_bignat ?loc n]) else GRef(glob_N0, None) -let error_negative dloc = - user_err ~loc:dloc ~hdr:"interp_N" (str "No negative numbers in type \"N\".") +let error_negative ?loc = + user_err ?loc ~hdr:"interp_N" (str "No negative numbers in type \"N\".") -let n_of_int dloc n = - if is_pos_or_zero n then n_of_binnat dloc true n - else error_negative dloc +let n_of_int ?loc n = + if is_pos_or_zero n then n_of_binnat ?loc true n + else error_negative ?loc (**********************************************************************) (* Printing N via scopes *) @@ -157,13 +157,13 @@ let glob_ZERO = ConstructRef path_of_ZERO let glob_POS = ConstructRef path_of_POS let glob_NEG = ConstructRef path_of_NEG -let z_of_int loc n = +let z_of_int ?loc n = if not (Bigint.equal n zero) then let sgn, n = if is_pos_or_zero n then glob_POS, n else glob_NEG, Bigint.neg n in - Loc.tag ~loc @@ GApp(Loc.tag ~loc @@ GRef(sgn,None), [pos_of_bignat loc n]) + Loc.tag ?loc @@ GApp(Loc.tag ?loc @@ GRef(sgn,None), [pos_of_bignat ?loc n]) else - Loc.tag ~loc @@ GRef(glob_ZERO, None) + Loc.tag ?loc @@ GRef(glob_ZERO, None) (**********************************************************************) (* Printing Z via scopes *) diff --git a/pretyping/cases.ml b/pretyping/cases.ml index 347c49f44..f5e2e52a1 100644 --- a/pretyping/cases.ml +++ b/pretyping/cases.ml @@ -122,7 +122,7 @@ type 'a equation = { patterns : cases_pattern list; rhs : 'a rhs; alias_stack : Name.t list; - eqn_loc : Loc.t; + eqn_loc : Loc.t option; used : bool ref } type 'a matrix = 'a equation list @@ -251,7 +251,7 @@ type 'a pattern_matching_problem = tomatch : tomatch_stack; history : pattern_continuation; mat : 'a matrix; - caseloc : Loc.t; + caseloc : Loc.t option; casestyle : case_style; typing_function: type_constraint -> env -> evar_map ref -> 'a option -> unsafe_judgment } @@ -281,8 +281,8 @@ let inductive_template evdref env tmloc ind = let arsign = inductive_alldecls_env env indu in let indu = on_snd EInstance.make indu in let hole_source i = match tmloc with - | Some loc -> (loc, Evar_kinds.TomatchTypeParameter (ind,i)) - | None -> (Loc.ghost, Evar_kinds.TomatchTypeParameter (ind,i)) in + | Some loc -> Loc.tag ~loc @@ Evar_kinds.TomatchTypeParameter (ind,i) + | None -> Loc.tag @@ Evar_kinds.TomatchTypeParameter (ind,i) in let (_,evarl,_) = List.fold_right (fun decl (subst,evarl,n) -> @@ -351,7 +351,7 @@ 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 j = typing_fun tycon env evdref tomatch in - let evd, j = Coercion.inh_coerce_to_base (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 = @@ -370,7 +370,7 @@ let coerce_to_indtype typing_fun evdref env matx tomatchl = (************************************************************************) (* Utils *) -let mkExistential env ?(src=(Loc.ghost,Evar_kinds.InternalHole)) evdref = +let mkExistential env ?(src=(Loc.tag Evar_kinds.InternalHole)) evdref = let e, u = e_new_type_evar env evdref univ_flexible_alg ~src:src in e let evd_comb2 f evdref x y = @@ -402,7 +402,7 @@ let adjust_tomatch_to_pattern pb ((current,typ),deps,dep) = let _ = e_cumul pb.env pb.evdref indt typ in current else - (evd_comb2 (Coercion.inh_conv_coerce_to true Loc.ghost pb.env) + (evd_comb2 (Coercion.inh_conv_coerce_to true pb.env) pb.evdref (make_judge current typ) indt).uj_val in let sigma = !(pb.evdref) in (current,try_find_ind pb.env sigma indt names)) @@ -469,11 +469,11 @@ let prepend_pattern tms eqn = {eqn with patterns = tms@eqn.patterns } exception NotAdjustable -let rec adjust_local_defs loc = function +let rec adjust_local_defs ?loc = function | (pat :: pats, LocalAssum _ :: decls) -> - pat :: adjust_local_defs loc (pats,decls) + pat :: adjust_local_defs ?loc (pats,decls) | (pats, LocalDef _ :: decls) -> - (Loc.tag ~loc @@ PatVar Anonymous) :: adjust_local_defs loc (pats,decls) + (Loc.tag ?loc @@ PatVar Anonymous) :: adjust_local_defs ?loc (pats,decls) | [], [] -> [] | _ -> raise NotAdjustable @@ -489,14 +489,14 @@ 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) + 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 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 @@ -510,7 +510,7 @@ let check_all_variables env sigma typ mat = let check_unused_pattern env eqn = if not !(eqn.used) then - raise_pattern_matching_error ~loc:eqn.eqn_loc (env, Evd.empty, UnusedClause eqn.patterns) + raise_pattern_matching_error ?loc:eqn.eqn_loc (env, Evd.empty, UnusedClause eqn.patterns) let set_used_pattern eqn = eqn.used := true @@ -978,7 +978,7 @@ let add_assert_false_case pb tomatch = avoid_ids = []; it = None }; alias_stack = Anonymous::aliasnames; - eqn_loc = Loc.ghost; + eqn_loc = None; used = ref false } ] let adjust_impossible_cases pb pred tomatch submat = @@ -1545,7 +1545,7 @@ let matx_of_eqns env eqns = it = Some initial_rhs } in { patterns = initial_lpat; alias_stack = []; - eqn_loc = loc; + eqn_loc = Some loc; used = ref false; rhs = rhs } in List.map build_eqn eqns @@ -1629,11 +1629,11 @@ let rec list_assoc_in_triple x = function * similarly for each ti. *) -let abstract_tycon loc env evdref subst tycon extenv t = +let abstract_tycon ?loc env evdref subst tycon extenv t = let t = nf_betaiota !evdref t in (* it helps in some cases to remove K-redex*) let src = match EConstr.kind !evdref t with - | Evar (evk,_) -> (loc,Evar_kinds.SubEvar evk) - | _ -> (loc,Evar_kinds.CasesType true) in + | Evar (evk,_) -> (Loc.tag ?loc @@ Evar_kinds.SubEvar evk) + | _ -> (Loc.tag ?loc @@ Evar_kinds.CasesType true) in let subst0,t0 = adjust_to_extended_env_and_remove_deps env extenv !evdref subst t in (* We traverse the type T of the original problem Xi looking for subterms that match the non-constructor part of the constraints (this part @@ -1687,7 +1687,7 @@ let abstract_tycon loc env evdref subst tycon extenv t = in aux (0,extenv,subst0) t0 -let build_tycon loc env tycon_env s subst tycon extenv evdref t = +let build_tycon ?loc env tycon_env s subst tycon extenv evdref t = let t,tt = match t with | None -> (* This is the situation we are building a return predicate and @@ -1695,10 +1695,10 @@ let build_tycon loc env tycon_env s subst tycon extenv evdref t = let n = Context.Rel.length (rel_context env) in let n' = Context.Rel.length (rel_context tycon_env) in let impossible_case_type, u = - e_new_type_evar (reset_context env) evdref univ_flexible_alg ~src:(loc,Evar_kinds.ImpossibleCase) in + e_new_type_evar (reset_context env) evdref univ_flexible_alg ~src:(Loc.tag ?loc Evar_kinds.ImpossibleCase) in (lift (n'-n) impossible_case_type, mkSort u) | Some t -> - let t = abstract_tycon loc tycon_env evdref subst tycon extenv t in + let t = abstract_tycon ?loc tycon_env evdref subst tycon extenv t in let evd,tt = Typing.type_of extenv !evdref t in evdref := evd; (t,tt) in @@ -1786,7 +1786,7 @@ let build_inversion_problem loc env sigma tms t = let main_eqn = { patterns = patl; alias_stack = []; - eqn_loc = Loc.ghost; + eqn_loc = None; used = ref false; rhs = { rhs_env = pb_env; (* we assume all vars are used; in practice we discard dependent @@ -1806,7 +1806,7 @@ let build_inversion_problem loc env sigma tms t = else [ { patterns = List.map (fun _ -> Loc.tag @@ PatVar Anonymous) patl; alias_stack = []; - eqn_loc = Loc.ghost; + eqn_loc = None; used = ref false; rhs = { rhs_env = pb_env; rhs_vars = []; @@ -1827,7 +1827,7 @@ let build_inversion_problem loc env sigma tms t = mat = main_eqn :: catch_all_eqn; caseloc = loc; casestyle = RegularStyle; - typing_function = build_tycon loc env pb_env s subst} in + typing_function = build_tycon ?loc env pb_env s subst} in let pred = (compile pb).uj_val in (!evdref,pred) @@ -1880,10 +1880,10 @@ let extract_arity_signature ?(dolift=true) env0 tomatchl tmsign = | _ -> assert false in List.rev (buildrec 0 (tomatchl,tmsign)) -let inh_conv_coerce_to_tycon loc env evdref j tycon = +let inh_conv_coerce_to_tycon ?loc env evdref j tycon = match tycon with | Some p -> - let (evd',j) = Coercion.inh_conv_coerce_to true loc env !evdref j p in + let (evd',j) = Coercion.inh_conv_coerce_to ?loc true env !evdref j p in evdref := evd'; j | None -> j @@ -1966,7 +1966,7 @@ let noccur_with_meta sigma n m term = in try (occur_rec n term; true) with LocalOccur -> false -let prepare_predicate loc typing_fun env sigma tomatchs arsign tycon pred = +let prepare_predicate ?loc typing_fun env sigma tomatchs arsign tycon pred = let refresh_tycon sigma t = (** If we put the typing constraint in the term, it has to be refreshed to preserve the invariant that no algebraic universe @@ -1997,7 +1997,7 @@ let prepare_predicate loc typing_fun env sigma tomatchs arsign tycon pred = | None -> let sigma = Sigma.Unsafe.of_evar_map sigma in let Sigma ((t, _), sigma, _) = - new_type_evar env sigma univ_flexible_alg ~src:(loc, Evar_kinds.CasesType false) in + new_type_evar env sigma univ_flexible_alg ~src:(Loc.tag ?loc @@ Evar_kinds.CasesType false) in let sigma = Sigma.to_evar_map sigma in sigma, t in @@ -2438,7 +2438,7 @@ let context_of_arsign l = l ([], 0) in x -let compile_program_cases loc style (typing_function, evdref) tycon env +let compile_program_cases ?loc style (typing_function, evdref) tycon env (predopt, tomatchl, eqns) = let typing_fun tycon env = function | Some t -> typing_function tycon env evdref t @@ -2545,9 +2545,9 @@ let compile_program_cases loc style (typing_function, evdref) tycon env (**************************************************************************) (* Main entry of the matching compilation *) -let compile_cases loc style (typing_fun, evdref) tycon env (predopt, tomatchl, eqns) = +let compile_cases ?loc style (typing_fun, evdref) tycon env (predopt, tomatchl, eqns) = if predopt == None && Flags.is_program_mode () && Program.is_program_cases () then - compile_program_cases loc style (typing_fun, evdref) + compile_program_cases ?loc style (typing_fun, evdref) tycon env (predopt, tomatchl, eqns) else @@ -2564,7 +2564,7 @@ let compile_cases loc style (typing_fun, evdref) tycon env (predopt, tomatchl, e with the type of arguments to match; if none is provided, we build alternative possible predicates *) let arsign = extract_arity_signature env tomatchs tomatchl in - let preds = prepare_predicate loc typing_fun env !evdref tomatchs arsign tycon predopt in + let preds = prepare_predicate ?loc typing_fun env !evdref tomatchs arsign tycon predopt in let compile_for_one_predicate (sigma,nal,pred) = (* We push the initial terms to match and push their alias to rhs' envs *) @@ -2614,7 +2614,7 @@ let compile_cases loc style (typing_fun, evdref) tycon env (predopt, tomatchl, e let j = compile pb in (* We coerce to the tycon (if an elim predicate was provided) *) - let j = inh_conv_coerce_to_tycon loc env myevdref j tycon in + let j = inh_conv_coerce_to_tycon ?loc env myevdref j tycon in evdref := !myevdref; j in diff --git a/pretyping/cases.mli b/pretyping/cases.mli index 6c2b5bf68..b16342db4 100644 --- a/pretyping/cases.mli +++ b/pretyping/cases.mli @@ -38,7 +38,7 @@ val irrefutable : env -> cases_pattern -> bool (** {6 Compilation primitive. } *) val compile_cases : - Loc.t -> case_style -> + ?loc:Loc.t -> case_style -> (type_constraint -> env -> evar_map ref -> glob_constr -> unsafe_judgment) * evar_map ref -> type_constraint -> env -> glob_constr option * tomatch_tuples * cases_clauses -> @@ -65,7 +65,7 @@ type 'a equation = { patterns : cases_pattern list; rhs : 'a rhs; alias_stack : Name.t list; - eqn_loc : Loc.t; + eqn_loc : Loc.t option; used : bool ref } type 'a matrix = 'a equation list @@ -106,14 +106,14 @@ type 'a pattern_matching_problem = tomatch : tomatch_stack; history : pattern_continuation; mat : 'a matrix; - caseloc : Loc.t; + caseloc : Loc.t option; casestyle : case_style; typing_function: type_constraint -> env -> evar_map ref -> 'a option -> unsafe_judgment } val compile : 'a pattern_matching_problem -> unsafe_judgment -val prepare_predicate : Loc.t -> +val prepare_predicate : ?loc:Loc.t -> (Evarutil.type_constraint -> Environ.env -> Evd.evar_map ref -> glob_constr -> unsafe_judgment) -> Environ.env -> diff --git a/pretyping/coercion.ml b/pretyping/coercion.ml index b2c1d0116..acccfc125 100644 --- a/pretyping/coercion.ml +++ b/pretyping/coercion.ml @@ -75,25 +75,25 @@ let apply_coercion_args env evd check isproj argl funj = !evdref, res (* appliquer le chemin de coercions de patterns p *) -let apply_pattern_coercion loc pat p = +let apply_pattern_coercion ?loc pat p = List.fold_left (fun pat (co,n) -> let f i = - if i (None, v) in aux t -and coerce loc env evdref (x : EConstr.constr) (y : EConstr.constr) +and coerce ?loc env evdref (x : EConstr.constr) (y : EConstr.constr) : (EConstr.constr -> EConstr.constr) option = let open Context.Rel.Declaration in @@ -183,7 +183,7 @@ and coerce loc env evdref (x : EConstr.constr) (y : EConstr.constr) let args = List.rev (restargs @ mkRel 1 :: List.map (lift 1) tele) in let pred = mkLambda (n, eqT, applist (lift 1 c, args)) in let eq = papp evdref coq_eq_ind [| eqT; hdx; hdy |] in - let evar = make_existential loc env evdref eq in + let evar = make_existential ?loc env evdref eq in let eq_app x = papp evdref coq_eq_rect [| eqT; hdx; pred; x; hdy; evar|] in @@ -326,7 +326,7 @@ and coerce loc env evdref (x : EConstr.constr) (y : EConstr.constr) Some (fun x -> let cx = app_opt env evdref c x in - let evar = make_existential loc env evdref (mkApp (p, [| cx |])) + let evar = make_existential ?loc env evdref (mkApp (p, [| cx |])) in (papp evdref sig_intro [| u; p; cx; evar |])) | None -> @@ -340,9 +340,9 @@ let app_coercion env evdref coercion v = let v' = Typing.e_solve_evars env evdref (f v) in whd_betaiota !evdref v' -let coerce_itf loc env evd v t c1 = +let coerce_itf ?loc env evd v t c1 = let evdref = ref evd in - let coercion = coerce loc env evdref t c1 in + let coercion = coerce ?loc env evdref t c1 in let t = Option.map (app_coercion env evdref coercion) v in !evdref, t @@ -410,16 +410,16 @@ let type_judgment env sigma j = | Sort s -> {utj_val = j.uj_val; utj_type = ESorts.kind sigma s } | _ -> error_not_a_type env sigma j -let inh_tosort_force loc env evd j = +let inh_tosort_force ?loc env evd j = try let t,p = lookup_path_to_sort_from env evd j.uj_type in let evd,j1 = apply_coercion env evd p j t in let j2 = on_judgment_type (whd_evar evd) j1 in (evd,type_judgment env evd j2) with Not_found | NoCoercion -> - error_not_a_type ~loc env evd j + error_not_a_type ?loc env evd j -let inh_coerce_to_sort loc env evd j = +let inh_coerce_to_sort ?loc env evd j = let typ = whd_all env evd j.uj_type in match EConstr.kind evd typ with | Sort s -> (evd,{ utj_val = j.uj_val; utj_type = ESorts.kind evd s }) @@ -427,9 +427,9 @@ let inh_coerce_to_sort loc env evd j = let (evd',s) = Evardefine.define_evar_as_sort env evd ev in (evd',{ utj_val = j.uj_val; utj_type = s }) | _ -> - inh_tosort_force loc env evd j + inh_tosort_force ?loc env evd j -let inh_coerce_to_base loc env evd j = +let inh_coerce_to_base ?loc env evd j = if Flags.is_program_mode () then let evdref = ref evd in let ct, typ' = mu env evdref j.uj_type in @@ -439,7 +439,7 @@ let inh_coerce_to_base loc env evd j = in !evdref, res else (evd, j) -let inh_coerce_to_prod loc env evd t = +let inh_coerce_to_prod ?loc env evd t = if Flags.is_program_mode () then let evdref = ref evd in let _, typ' = mu env evdref t in @@ -466,7 +466,7 @@ let inh_coerce_to_fail env evd rigidonly v t c1 = try (the_conv_x_leq env t' c1 evd, v') with UnableToUnify _ -> raise NoCoercion -let rec inh_conv_coerce_to_fail loc env evd rigidonly v t c1 = +let rec inh_conv_coerce_to_fail ?loc env evd rigidonly v t c1 = try (the_conv_x_leq env t c1 evd, v) with UnableToUnify (best_failed_evd,e) -> try inh_coerce_to_fail env evd rigidonly v t c1 @@ -488,49 +488,50 @@ let rec inh_conv_coerce_to_fail loc env evd rigidonly v t c1 = let open Context.Rel.Declaration in let env1 = push_rel (LocalAssum (name,u1)) env in let (evd', v1) = - inh_conv_coerce_to_fail loc env1 evd rigidonly + inh_conv_coerce_to_fail ?loc env1 evd rigidonly (Some (mkRel 1)) (lift 1 u1) (lift 1 t1) in let v1 = Option.get v1 in let v2 = Option.map (fun v -> beta_applist evd' (lift 1 v,[v1])) v in let t2 = match v2 with | None -> subst_term evd' v1 t2 | Some v2 -> Retyping.get_type_of env1 evd' v2 in - let (evd'',v2') = inh_conv_coerce_to_fail loc env1 evd' rigidonly v2 t2 u2 in + let (evd'',v2') = inh_conv_coerce_to_fail ?loc env1 evd' rigidonly v2 t2 u2 in (evd'', Option.map (fun v2' -> mkLambda (name, u1, v2')) v2') | _ -> raise (NoCoercionNoUnifier (best_failed_evd,e)) (* Look for cj' obtained from cj by inserting coercions, s.t. cj'.typ = t *) -let inh_conv_coerce_to_gen resolve_tc rigidonly loc env evd cj t = +let inh_conv_coerce_to_gen ?loc resolve_tc rigidonly env evd cj t = let (evd', val') = try - inh_conv_coerce_to_fail loc env evd rigidonly (Some cj.uj_val) cj.uj_type t + inh_conv_coerce_to_fail ?loc env evd rigidonly (Some cj.uj_val) cj.uj_type t with NoCoercionNoUnifier (best_failed_evd,e) -> try if Flags.is_program_mode () then - coerce_itf loc env evd (Some cj.uj_val) cj.uj_type t + coerce_itf ?loc env evd (Some cj.uj_val) cj.uj_type t else raise NoSubtacCoercion with | NoSubtacCoercion when not resolve_tc || not !use_typeclasses_for_conversion -> - error_actual_type ~loc env best_failed_evd cj t e + error_actual_type ?loc env best_failed_evd cj t e | NoSubtacCoercion -> let evd' = saturate_evd env evd in try if evd' == evd then - error_actual_type ~loc env best_failed_evd cj t e + error_actual_type ?loc env best_failed_evd cj t e else - inh_conv_coerce_to_fail loc env evd' rigidonly (Some cj.uj_val) cj.uj_type t + inh_conv_coerce_to_fail ?loc env evd' rigidonly (Some cj.uj_val) cj.uj_type t with NoCoercionNoUnifier (_evd,_error) -> - error_actual_type ~loc env best_failed_evd cj t e + error_actual_type ?loc env best_failed_evd cj t e in let val' = match val' with Some v -> v | None -> assert(false) in (evd',{ uj_val = val'; uj_type = t }) -let inh_conv_coerce_to resolve_tc = inh_conv_coerce_to_gen resolve_tc false -let inh_conv_coerce_rigid_to resolve_tc = inh_conv_coerce_to_gen resolve_tc true +let inh_conv_coerce_to ?loc resolve_tc = inh_conv_coerce_to_gen ?loc resolve_tc false -let inh_conv_coerces_to loc env evd t t' = +let inh_conv_coerce_rigid_to ?loc resolve_tc = inh_conv_coerce_to_gen resolve_tc ?loc true + +let inh_conv_coerces_to ?loc env evd t t' = try - fst (inh_conv_coerce_to_fail loc env evd true None t t') + fst (inh_conv_coerce_to_fail ?loc env evd true None t t') with NoCoercion -> evd (* Maybe not enough information to unify *) diff --git a/pretyping/coercion.mli b/pretyping/coercion.mli index bc63d092d..25ee82bbf 100644 --- a/pretyping/coercion.mli +++ b/pretyping/coercion.mli @@ -26,17 +26,17 @@ val inh_app_fun : bool -> (** [inh_coerce_to_sort env isevars j] coerces [j] to a type; i.e. it inserts a coercion into [j], if needed, in such a way it gets as type a sort; it fails if no coercion is applicable *) -val inh_coerce_to_sort : Loc.t -> +val inh_coerce_to_sort : ?loc:Loc.t -> env -> evar_map -> unsafe_judgment -> evar_map * unsafe_type_judgment (** [inh_coerce_to_base env isevars j] coerces [j] to its base type; i.e. it inserts a coercion into [j], if needed, in such a way it gets as type its base type (the notion depends on the coercion system) *) -val inh_coerce_to_base : Loc.t -> +val inh_coerce_to_base : ?loc:Loc.t -> env -> evar_map -> unsafe_judgment -> evar_map * unsafe_judgment (** [inh_coerce_to_prod env isevars t] coerces [t] to a product type *) -val inh_coerce_to_prod : Loc.t -> +val inh_coerce_to_prod : ?loc:Loc.t -> env -> evar_map -> types -> evar_map * types (** [inh_conv_coerce_to resolve_tc Loc.t env isevars j t] coerces [j] to an @@ -44,20 +44,20 @@ val inh_coerce_to_prod : Loc.t -> a way [t] and [j.uj_type] are convertible; it fails if no coercion is applicable. resolve_tc=false disables resolving type classes (as the last resort before failing) *) -val inh_conv_coerce_to : bool -> Loc.t -> +val inh_conv_coerce_to : ?loc:Loc.t -> bool -> env -> evar_map -> unsafe_judgment -> types -> evar_map * unsafe_judgment -val inh_conv_coerce_rigid_to : bool -> Loc.t -> +val inh_conv_coerce_rigid_to : ?loc:Loc.t -> bool -> env -> evar_map -> unsafe_judgment -> types -> evar_map * unsafe_judgment (** [inh_conv_coerces_to loc env isevars t t'] checks if an object of type [t] is coercible to an object of type [t'] adding evar constraints if needed; it fails if no coercion exists *) -val inh_conv_coerces_to : Loc.t -> +val inh_conv_coerces_to : ?loc:Loc.t -> env -> evar_map -> types -> types -> evar_map (** [inh_pattern_coerce_to loc env isevars pat ind1 ind2] coerces the Cases pattern [pat] typed in [ind1] into a pattern typed in [ind2]; raises [Not_found] if no coercion found *) val inh_pattern_coerce_to : - Loc.t -> env -> cases_pattern -> inductive -> inductive -> cases_pattern + ?loc:Loc.t -> env -> cases_pattern -> inductive -> inductive -> cases_pattern diff --git a/pretyping/detyping.ml b/pretyping/detyping.ml index 05d6a1ad4..3079d1052 100644 --- a/pretyping/detyping.ml +++ b/pretyping/detyping.ml @@ -29,8 +29,6 @@ open Misctypes open Decl_kinds open Context.Named.Declaration -let dl = Loc.ghost - (** Should we keep details of universes during detyping ? *) let print_universes = Flags.univ_print diff --git a/pretyping/evarconv.ml b/pretyping/evarconv.ml index 4bb66b8e9..95680e18a 100644 --- a/pretyping/evarconv.ml +++ b/pretyping/evarconv.ml @@ -888,7 +888,7 @@ and conv_record trs env evd (ctx,(h,h2),c,bs,(params,params1),(us,us2),(sk1,sk2) let test i = evar_conv_x trs env i CUMUL ty (substl ks b) in (i,t2::ks, m-1, test) else - let dloc = (Loc.ghost,Evar_kinds.InternalHole) in + let dloc = Loc.tag Evar_kinds.InternalHole in let i = Sigma.Unsafe.of_evar_map i in let Sigma (ev, i', _) = Evarutil.new_evar env i ~src:dloc (substl ks b) in let i' = Sigma.to_evar_map i' in @@ -1214,7 +1214,7 @@ let apply_conversion_problem_heuristic ts env evd pbty t1 t2 = let error_cannot_unify env evd pb ?reason t1 t2 = Pretype_errors.error_cannot_unify - ~loc:(loc_of_conv_pb evd pb) env + ?loc:(loc_of_conv_pb evd pb) env evd ?reason (t1, t2) let check_problems_are_solved env evd = diff --git a/pretyping/glob_ops.ml b/pretyping/glob_ops.ml index 25ece5b8e..cba1780ba 100644 --- a/pretyping/glob_ops.ml +++ b/pretyping/glob_ops.ml @@ -22,7 +22,7 @@ let cases_predicate_names tml = | (tm,(na,None)) -> [na] | (tm,(na,Some (_,(_,nal)))) -> na::nal) tml) -let mkGApp loc p t = Loc.tag ~loc @@ +let mkGApp ?loc p t = Loc.tag ?loc @@ match snd p with | GApp (f,l) -> GApp (f,l@[t]) | _ -> GApp (p,[t]) diff --git a/pretyping/glob_ops.mli b/pretyping/glob_ops.mli index 55e6b6533..ac2118ba7 100644 --- a/pretyping/glob_ops.mli +++ b/pretyping/glob_ops.mli @@ -25,7 +25,7 @@ val cases_pattern_loc : cases_pattern -> Loc.t val cases_predicate_names : tomatch_tuples -> Name.t list (** Apply one argument to a glob_constr *) -val mkGApp : Loc.t -> glob_constr -> glob_constr -> glob_constr +val mkGApp : ?loc:Loc.t -> glob_constr -> glob_constr -> glob_constr val map_glob_constr : (glob_constr -> glob_constr) -> glob_constr -> glob_constr diff --git a/pretyping/pretyping.ml b/pretyping/pretyping.ml index 5f9f4bb08..fe15cb490 100644 --- a/pretyping/pretyping.ml +++ b/pretyping/pretyping.ml @@ -133,7 +133,7 @@ let nf_fix sigma (nas, cs, ts) = let inj c = EConstr.to_constr sigma c in (nas, Array.map inj cs, Array.map inj ts) -let search_guard loc env possible_indexes fixdefs = +let search_guard ?loc env possible_indexes fixdefs = (* Standard situation with only one possibility for each fix. *) (* We treat it separately in order to get proper error msg. *) let is_singleton = function [_] -> true | _ -> false in @@ -143,7 +143,7 @@ let search_guard loc env possible_indexes fixdefs = (try check_fix env fix with reraise -> let (e, info) = CErrors.push reraise in - let info = Loc.add_loc info loc in + let info = Option.cata (fun loc -> Loc.add_loc info loc) info loc in iraise (e, info)); indexes else @@ -166,7 +166,7 @@ let search_guard loc env possible_indexes fixdefs = with TypeError _ -> ()) (List.combinations possible_indexes); let errmsg = "Cannot guess decreasing argument of fix." in - user_err ~loc ~hdr:"search_guard" (Pp.str errmsg) + user_err ?loc ~hdr:"search_guard" (Pp.str errmsg) with Found indexes -> indexes) (* To force universe name declaration before use *) @@ -384,10 +384,10 @@ let process_inference_flags flags env initial_sigma (sigma,c) = let allow_anonymous_refs = ref false (* coerce to tycon if any *) -let inh_conv_coerce_to_tycon resolve_tc loc env evdref j = function +let inh_conv_coerce_to_tycon ?loc resolve_tc env evdref j = function | None -> j | Some t -> - evd_comb2 (Coercion.inh_conv_coerce_to resolve_tc loc env.ExtraEnv.env) evdref j t + evd_comb2 (Coercion.inh_conv_coerce_to ?loc resolve_tc env.ExtraEnv.env) evdref j t let check_instance loc subst = function | [] -> () @@ -568,18 +568,18 @@ let (f_genarg_interp, genarg_interp_hook) = Hook.make () (* the type constraint tycon *) let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdref (lvar : ltac_var_map) (loc, t) = - let inh_conv_coerce_to_tycon = inh_conv_coerce_to_tycon resolve_tc in + let inh_conv_coerce_to_tycon ?loc = inh_conv_coerce_to_tycon ?loc resolve_tc in let pretype_type = pretype_type k0 resolve_tc in let pretype = pretype k0 resolve_tc in let open Context.Rel.Declaration in match t with | GRef (ref,u) -> - inh_conv_coerce_to_tycon loc env evdref + 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 @@ -594,7 +594,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre 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 -> @@ -709,11 +709,11 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre 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 @@ -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,7 +800,7 @@ 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 @@ -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 @@ -838,7 +838,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre let (e, info) = CErrors.push e in let info = Loc.add_loc 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 = @@ -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 @@ -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) = @@ -1128,7 +1128,7 @@ 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 -> diff --git a/pretyping/pretyping.mli b/pretyping/pretyping.mli index f13c10b05..79fafcf1a 100644 --- a/pretyping/pretyping.mli +++ b/pretyping/pretyping.mli @@ -24,7 +24,7 @@ open Misctypes (** An auxiliary function for searching for fixpoint guard indexes *) val search_guard : - Loc.t -> env -> int list list -> rec_declaration -> int array + ?loc:Loc.t -> env -> int list list -> rec_declaration -> int array type typing_constraint = OfType of types | IsType | WithoutTypeConstraint diff --git a/pretyping/typing.ml b/pretyping/typing.ml index c2a030bcd..41d994553 100644 --- a/pretyping/typing.ml +++ b/pretyping/typing.ml @@ -152,13 +152,13 @@ let e_judge_of_case env evdref ci pj cj lfj = { uj_val = mkCase (ci, pj.uj_val, cj.uj_val, Array.map j_val lfj); uj_type = rslty } -let check_type_fixpoint loc env evdref lna lar vdefj = +let check_type_fixpoint ?loc env evdref lna lar vdefj = let lt = Array.length vdefj in if Int.equal (Array.length lar) lt then for i = 0 to lt-1 do if not (Evarconv.e_cumul env evdref (vdefj.(i)).uj_type (lift lt lar.(i))) then - error_ill_typed_rec_body ~loc env !evdref + error_ill_typed_rec_body ?loc env !evdref i lna vdefj lar done @@ -361,7 +361,7 @@ and execute_recdef env evdref (names,lar,vdef) = let env1 = push_rec_types (names,lara,vdef) env in let vdefj = execute_array env1 evdref vdef in let vdefv = Array.map j_val vdefj in - let _ = check_type_fixpoint Loc.ghost env1 evdref names lara vdefj in + let _ = check_type_fixpoint env1 evdref names lara vdefj in (names,lara,vdefv) and execute_array env evdref = Array.map (execute env evdref) diff --git a/pretyping/typing.mli b/pretyping/typing.mli index 91134b499..1f3ba34e5 100644 --- a/pretyping/typing.mli +++ b/pretyping/typing.mli @@ -44,7 +44,7 @@ val check_allowed_sort : env -> evar_map -> pinductive -> constr -> constr -> (** Raise an error message if bodies have types not unifiable with the expected ones *) -val check_type_fixpoint : Loc.t -> env -> evar_map ref -> +val check_type_fixpoint : ?loc:Loc.t -> env -> evar_map ref -> Names.Name.t array -> types array -> unsafe_judgment array -> unit val judge_of_prop : unsafe_judgment diff --git a/pretyping/unification.ml b/pretyping/unification.ml index 532cc8baa..9678e3a42 100644 --- a/pretyping/unification.ml +++ b/pretyping/unification.ml @@ -1250,7 +1250,7 @@ let applyHead env (type r) (evd : r Sigma.t) n c = let sigma = Sigma.to_evar_map evd in match EConstr.kind sigma (whd_all env sigma cty) with | Prod (_,c1,c2) -> - let Sigma (evar, evd', q) = Evarutil.new_evar env evd ~src:(Loc.ghost,Evar_kinds.GoalEvar) c1 in + let Sigma (evar, evd', q) = Evarutil.new_evar env evd ~src:(Loc.tag Evar_kinds.GoalEvar) c1 in apprec (n-1) (mkApp(c,[|evar|])) (subst1 evar c2) (p +> q) evd' | _ -> error "Apply_Head_Then" in @@ -1265,7 +1265,7 @@ let is_mimick_head sigma ts f = let try_to_coerce env evd c cty tycon = let j = make_judge c cty in - let (evd',j') = inh_conv_coerce_rigid_to true Loc.ghost env evd j tycon in + let (evd',j') = inh_conv_coerce_rigid_to true env evd j tycon in let evd' = Evarconv.solve_unif_constraints_with_heuristics env evd' in let evd' = Evd.map_metas_fvalue (fun c -> EConstr.Unsafe.to_constr (nf_evar evd' (EConstr.of_constr c))) evd' in (evd',j'.uj_val) diff --git a/printing/ppconstr.ml b/printing/ppconstr.ml index 117e1900d..fddd54a9f 100644 --- a/printing/ppconstr.ml +++ b/printing/ppconstr.ml @@ -145,9 +145,9 @@ let tag_var = tag Tag.variable if !Flags.beautify && not (Int.equal n 0) then comment (CLexer.extract_comments n) else mt() - let pr_with_comments loc pp = pr_located (fun x -> x) (loc,pp) + 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 (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 @@ -302,7 +302,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,7 +310,7 @@ 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 " =>") ++ @@ -730,7 +730,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/ppconstr.mli b/printing/ppconstr.mli index f92caf426..482c994c2 100644 --- a/printing/ppconstr.mli +++ b/printing/ppconstr.mli @@ -35,7 +35,7 @@ val pr_or_var : ('a -> std_ppcmds) -> 'a or_var -> std_ppcmds val pr_lident : Id.t located -> std_ppcmds val pr_lname : Name.t located -> std_ppcmds -val pr_with_comments : Loc.t -> std_ppcmds -> std_ppcmds +val pr_with_comments : ?loc:Loc.t -> std_ppcmds -> std_ppcmds val pr_com_at : int -> std_ppcmds val pr_sep_com : (unit -> std_ppcmds) -> diff --git a/printing/pputils.ml b/printing/pputils.ml index e90b54685..4ae5035a2 100644 --- a/printing/pputils.ml +++ b/printing/pputils.ml @@ -15,7 +15,7 @@ open Locus open Genredexpr let pr_located pr (loc, x) = - if !Flags.beautify && loc <> Loc.ghost then + if !Flags.beautify && not (Loc.is_ghost loc) then let (b, e) = Loc.unloc loc in (* Side-effect: order matters *) let before = Pp.comment (CLexer.extract_comments b) in diff --git a/printing/prettyp.ml b/printing/prettyp.ml index 96b0f49d4..025514902 100644 --- a/printing/prettyp.ml +++ b/printing/prettyp.ml @@ -539,7 +539,7 @@ let gallina_print_constant_with_infos sp = let gallina_print_syntactic_def kn = let qid = Nametab.shortest_qualid_of_syndef Id.Set.empty kn and (vars,a) = Syntax_def.search_syntactic_definition kn in - let c = Notation_ops.glob_constr_of_notation_constr Loc.ghost a in + let c = Notation_ops.glob_constr_of_notation_constr a in hov 2 (hov 4 (str "Notation " ++ pr_qualid qid ++ @@ -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) @@ -800,7 +800,7 @@ 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) + (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) diff --git a/printing/printer.ml b/printing/printer.ml index 35ddf2e8c..279295a03 100644 --- a/printing/printer.ml +++ b/printing/printer.ml @@ -196,10 +196,10 @@ let qualid_of_global env r = let safe_gen f env sigma c = let orig_extern_ref = Constrextern.get_extern_reference () in - let extern_ref loc vars r = - try orig_extern_ref loc vars r + let extern_ref ?loc vars r = + try orig_extern_ref ?loc vars r with e when CErrors.noncritical e -> - Libnames.Qualid (loc, qualid_of_global env r) + Libnames.Qualid (Loc.tag ?loc @@ qualid_of_global env r) in Constrextern.set_extern_reference extern_ref; try diff --git a/proofs/clenv.ml b/proofs/clenv.ml index 17a9651cd..251e0d27d 100644 --- a/proofs/clenv.ml +++ b/proofs/clenv.ml @@ -681,7 +681,7 @@ let define_with_type sigma env ev c = let t = Retyping.get_type_of env sigma ev in let ty = Retyping.get_type_of env sigma c in let j = Environ.make_judge c ty in - let (sigma, j) = Coercion.inh_conv_coerce_to true (Loc.ghost) env sigma j t in + let (sigma, j) = Coercion.inh_conv_coerce_to true env sigma j t in let (ev, _) = destEvar sigma ev in let sigma = Evd.define ev (EConstr.Unsafe.to_constr j.Environ.uj_val) sigma in sigma diff --git a/proofs/goal.ml b/proofs/goal.ml index 9046f4534..5cc9d0df9 100644 --- a/proofs/goal.ml +++ b/proofs/goal.ml @@ -69,7 +69,7 @@ module V82 = struct Evd.evar_concl = concl; Evd.evar_filter = Evd.Filter.identity; Evd.evar_body = Evd.Evar_empty; - Evd.evar_source = (Loc.ghost,Evar_kinds.GoalEvar); + Evd.evar_source = (Loc.tag Evar_kinds.GoalEvar); Evd.evar_candidates = None; Evd.evar_extra = extra } in diff --git a/proofs/proof_global.ml b/proofs/proof_global.ml index 5f4a7766f..d659eba24 100644 --- a/proofs/proof_global.ml +++ b/proofs/proof_global.ml @@ -287,13 +287,13 @@ let set_used_variables l = match entry with | LocalAssum (x,_) -> if Id.Set.mem x all_safe then orig - else (ctx, all_safe, (Loc.ghost,x)::to_clear) + else (ctx, all_safe, (Loc.tag x)::to_clear) | LocalDef (x,bo, ty) as decl -> if Id.Set.mem x all_safe then orig else let vars = Id.Set.union (vars_of env bo) (vars_of env ty) in if Id.Set.subset vars all_safe then (decl :: ctx, Id.Set.add x all_safe, to_clear) - else (ctx, all_safe, (Loc.ghost,x) :: to_clear) in + else (ctx, all_safe, (Loc.tag x) :: to_clear) in let ctx, _, to_clear = Environ.fold_named_context aux env ~init:(ctx,ctx_set,[]) in let to_clear = if !proof_using_auto_clear then to_clear else [] in diff --git a/proofs/proof_using.ml b/proofs/proof_using.ml index 2c489d6de..f0854e9aa 100644 --- a/proofs/proof_using.ml +++ b/proofs/proof_using.ml @@ -76,7 +76,7 @@ and full_set env = List.fold_right Id.Set.add (List.map NamedDecl.get_id (named_context env)) Id.Set.empty let process_expr env e ty = - let ty_expr = SsSingl(Loc.ghost, Id.of_string "Type") in + let ty_expr = SsSingl(Loc.tag @@ Id.of_string "Type") in let v_ty = process_expr env ty_expr ty in let s = Id.Set.union v_ty (process_expr env e ty) in Id.Set.elements s diff --git a/proofs/refine.ml b/proofs/refine.ml index 1ee6e0ca5..d423a658a 100644 --- a/proofs/refine.ml +++ b/proofs/refine.ml @@ -137,7 +137,7 @@ let with_type env evd c t = let my_type = Retyping.get_type_of env evd c in let j = Environ.make_judge c my_type in let (evd,j') = - Coercion.inh_conv_coerce_to true (Loc.ghost) env evd j t + Coercion.inh_conv_coerce_to true env evd j t in evd , j'.Environ.uj_val diff --git a/stm/stm.ml b/stm/stm.ml index e823373f7..38415c1dd 100644 --- a/stm/stm.ml +++ b/stm/stm.ml @@ -23,9 +23,9 @@ open Feedback open Vernacexpr open Vernac_classifier -let execution_error state_id loc msg = +let execution_error ?loc state_id msg = feedback ~id:state_id - (Message (Error, Some loc, msg)) + (Message (Error, loc, msg)) module Hooks = struct @@ -72,7 +72,7 @@ let async_proofs_workers_extra_env = ref [||] type aast = { verbose : bool; - loc : Loc.t; + loc : Loc.t option; indentation : int; strlen : int; mutable expr : vernac_expr; (* mutable: Proof using hinted by aux file *) @@ -801,9 +801,9 @@ end = struct (* {{{ *) match Stateid.get info with | Some _ -> (e, info) | None -> - let loc = Option.default Loc.ghost (Loc.get_loc info) in + let loc = Loc.get_loc info in let (e, info) = Hooks.(call_process_error_once (e, info)) in - execution_error id loc (iprint (e, info)); + execution_error ?loc id (iprint (e, info)); (e, Stateid.add info ~valid id) let same_env { system = s1 } { system = s2 } = @@ -949,7 +949,7 @@ let stm_vernac_interp ?proof id ?route { verbose; loc; expr } = the whole document state, such as backtrack, etc... so we start to design the stm command interpreter now *) set_id_for_feedback ?route id; - Aux_file.record_in_aux_set_at loc; + Aux_file.record_in_aux_set_at ?loc; (* We need to check if a command should be filtered from * vernac_entries, as it cannot handle it. This should go away in * future refactorings. @@ -968,7 +968,7 @@ let stm_vernac_interp ?proof id ?route { verbose; loc; expr } = | VernacShow ShowScript -> ShowScript.show_script () | expr -> stm_pperr_endline Pp.(fun () -> str "interpreting " ++ Ppvernac.pr_vernac expr); - try Vernacentries.interp ?verbosely:(Some verbose) ?proof (loc, expr) + try Vernacentries.interp ?verbosely:(Some verbose) ?proof (Loc.tag ?loc expr) with e -> let e = CErrors.push e in Exninfo.iraise Hooks.(call_process_error_once e) @@ -1111,12 +1111,12 @@ 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 s = Aux_file.get !hints loc "context_used" in +let get_hint_ctx ?loc = + let s = Aux_file.get ?loc !hints "context_used" in match Str.split (Str.regexp ";") s with | ids :: _ -> let ids = List.map Names.Id.of_string (Str.split (Str.regexp " ") ids) in - let ids = List.map (fun id -> Loc.ghost, id) ids in + let ids = List.map (fun id -> Loc.tag id) ids in begin match ids with | [] -> SsEmpty | x :: xs -> @@ -1125,15 +1125,15 @@ let get_hint_ctx loc = | _ -> raise Not_found let get_hint_bp_time proof_name = - try float_of_string (Aux_file.get !hints Loc.ghost proof_name) + try float_of_string (Aux_file.get !hints proof_name) with Not_found -> 1.0 -let record_pb_time proof_name loc time = +let record_pb_time ?loc proof_name time = let proof_build_time = Printf.sprintf "%.3f" time in - Aux_file.record_in_aux_at loc "proof_build_time" proof_build_time; + Aux_file.record_in_aux_at ?loc "proof_build_time" proof_build_time; if proof_name <> "" then begin - Aux_file.record_in_aux_at Loc.ghost proof_name proof_build_time; - hints := Aux_file.set !hints Loc.ghost proof_name proof_build_time + Aux_file.record_in_aux_at proof_name proof_build_time; + hints := Aux_file.set !hints proof_name proof_build_time end exception RemoteException of Pp.std_ppcmds @@ -1222,7 +1222,7 @@ module rec ProofTask : sig t_drop : bool; t_states : competence; t_assign : Proof_global.closed_proof_output Future.assignement -> unit; - t_loc : Loc.t; + t_loc : Loc.t option; t_uuid : Future.UUID.t; t_name : string } @@ -1240,8 +1240,9 @@ module rec ProofTask : sig and type request := request val build_proof_here : + ?loc:Loc.t -> drop_pt:bool -> - Stateid.t * Stateid.t -> Loc.t -> Stateid.t -> + Stateid.t * Stateid.t -> Stateid.t -> Proof_global.closed_proof_output Future.computation (* If set, only tasks overlapping with this list are processed *) @@ -1259,7 +1260,7 @@ end = struct (* {{{ *) t_drop : bool; t_states : competence; t_assign : Proof_global.closed_proof_output Future.assignement -> unit; - t_loc : Loc.t; + t_loc : Loc.t option; t_uuid : Future.UUID.t; t_name : string } @@ -1328,7 +1329,7 @@ end = struct (* {{{ *) RespBuiltProof (pl, time) -> feedback (InProgress ~-1); t_assign (`Val pl); - record_pb_time t_name t_loc time; + record_pb_time ?loc:t_loc t_name time; if !Flags.async_proofs_full || t_drop then `Stay(t_states,[States t_states]) else `End @@ -1349,16 +1350,16 @@ end = struct (* {{{ *) let info = Stateid.add ~valid:start Exninfo.null start in let e = (RemoteException (Pp.strbrk s), info) in t_assign (`Exn e); - execution_error start Loc.ghost (Pp.strbrk s); + execution_error start (Pp.strbrk s); feedback (InProgress ~-1) - let build_proof_here ~drop_pt (id,valid) loc eop = + let build_proof_here ?loc ~drop_pt (id,valid) eop = Future.create (State.exn_on id ~valid) (fun () -> let wall_clock1 = Unix.gettimeofday () in if !Flags.batch_mode then Reach.known_state ~cache:`No eop else Reach.known_state ~cache:`Shallow eop; let wall_clock2 = Unix.gettimeofday () in - Aux_file.record_in_aux_at loc "proof_build_time" + Aux_file.record_in_aux_at ?loc "proof_build_time" (Printf.sprintf "%.3f" (wall_clock2 -. wall_clock1)); let p = Proof_global.return_proof ~allow_partial:drop_pt () in if drop_pt then feedback ~id Complete; @@ -1370,7 +1371,7 @@ end = struct (* {{{ *) VCS.print (); let proof, future_proof, time = let wall_clock = Unix.gettimeofday () in - let fp = build_proof_here ~drop_pt:drop exn_info loc stop in + let fp = build_proof_here ?loc ~drop_pt:drop exn_info stop in let proof = Future.force fp in proof, fp, Unix.gettimeofday () -. wall_clock in (* We typecheck the proof with the kernel (in the worker) to spot @@ -1452,7 +1453,7 @@ end = struct (* {{{ *) msg_error(Pp.strbrk("Marshalling error: "^s^". "^ "The system state could not be sent to the worker process. "^ "Falling back to local, lazy, evaluation.")); - t_assign(`Comp(build_proof_here ~drop_pt t_exn_info t_loc t_stop)); + t_assign(`Comp(build_proof_here ?loc:t_loc ~drop_pt t_exn_info t_stop)); feedback (InProgress ~-1) end (* }}} *) @@ -1462,7 +1463,7 @@ and Slaves : sig (* (eventually) remote calls *) val build_proof : - loc:Loc.t -> drop_pt:bool -> + ?loc:Loc.t -> drop_pt:bool -> exn_info:(Stateid.t * Stateid.t) -> block_start:Stateid.t -> block_stop:Stateid.t -> name:string -> future_proof * cancel_switch @@ -1542,7 +1543,7 @@ end = struct (* {{{ *) | { step = `Fork (( { loc }, _, _, _), _) } | { step = `Qed ( { qast = { loc } }, _) } | { step = `Sideff (`Ast ( { loc }, _)) } -> - let start, stop = Loc.unloc loc in + let start, stop = Option.cata Loc.unloc (0,0) loc in msg_error Pp.( str"File " ++ str name ++ str ": proof of " ++ str r_name ++ str ": chars " ++ int start ++ str "-" ++ int stop ++ @@ -1598,10 +1599,10 @@ end = struct (* {{{ *) let info_tasks l = CList.map_i (fun i ({ Stateid.loc; name }, _) -> let time1 = - try float_of_string (Aux_file.get !hints loc "proof_build_time") + try float_of_string (Aux_file.get ?loc !hints "proof_build_time") with Not_found -> 0.0 in let time2 = - try float_of_string (Aux_file.get !hints loc "proof_check_time") + try float_of_string (Aux_file.get ?loc !hints "proof_check_time") with Not_found -> 0.0 in name, max (time1 +. time2) 0.0001,i) 0 l @@ -1622,7 +1623,7 @@ end = struct (* {{{ *) BuildProof { t_states = s2 } -> overlap_rel s1 s2 | _ -> 0) - let build_proof ~loc ~drop_pt ~exn_info ~block_start ~block_stop ~name:pname = + let build_proof ?loc ~drop_pt ~exn_info ~block_start ~block_stop ~name:pname = let id, valid as t_exn_info = exn_info in let cancel_switch = ref false in if TaskQueue.n_workers (Option.get !queue) = 0 then @@ -1637,7 +1638,7 @@ end = struct (* {{{ *) TaskQueue.enqueue_task (Option.get !queue) (task,cancel_switch); f, cancel_switch end else - ProofTask.build_proof_here ~drop_pt t_exn_info loc block_stop, cancel_switch + ProofTask.build_proof_here ?loc ~drop_pt t_exn_info block_stop, cancel_switch else let f, t_assign = Future.create_delegate ~name:pname (State.exn_on id ~valid) in let t_uuid = Future.uuid f in @@ -2026,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) @@ -2122,7 +2123,7 @@ let known_state ?(redefine_qed=false) ~cache id = feedback ~id:id Feedback.AddedAxiom; fst (Pfedit.solve Vernacexpr.SelectAll None tac p), ()); Option.iter (fun expr -> stm_vernac_interp id { - verbose = true; loc = Loc.ghost; expr; indentation = 0; + verbose = true; loc = None; expr; indentation = 0; strlen = 0 }) recovery_command | _ -> assert false @@ -2244,7 +2245,7 @@ let known_state ?(redefine_qed=false) ~cache id = ^"the proof's statement to avoid that.")); let fp, cancel = Slaves.build_proof - ~loc ~drop_pt ~exn_info ~block_start ~block_stop ~name in + ?loc ~drop_pt ~exn_info ~block_start ~block_stop ~name in Future.replace ofp fp; qed.fproof <- Some (fp, cancel); (* We don't generate a new state, but we still need @@ -2256,10 +2257,10 @@ let known_state ?(redefine_qed=false) ~cache id = let fp, cancel = if delegate then Slaves.build_proof - ~loc ~drop_pt ~exn_info ~block_start ~block_stop ~name + ?loc ~drop_pt ~exn_info ~block_start ~block_stop ~name else - ProofTask.build_proof_here - ~drop_pt exn_info loc block_stop, ref false + ProofTask.build_proof_here ?loc + ~drop_pt exn_info block_stop, ref false in qed.fproof <- Some (fp, cancel); let proof = @@ -2279,7 +2280,7 @@ let known_state ?(redefine_qed=false) ~cache id = log_processing_sync id name reason; reach eop; let wall_clock = Unix.gettimeofday () in - record_pb_time name x.loc (wall_clock -. !wall_clock_last_fork); + record_pb_time name ?loc:x.loc (wall_clock -. !wall_clock_last_fork); let proof = match keep with | VtDrop -> None @@ -2296,7 +2297,7 @@ let known_state ?(redefine_qed=false) ~cache id = let wall_clock2 = Unix.gettimeofday () in stm_vernac_interp id ?proof x; let wall_clock3 = Unix.gettimeofday () in - Aux_file.record_in_aux_at x.loc "proof_check_time" + Aux_file.record_in_aux_at ?loc:x.loc "proof_check_time" (Printf.sprintf "%.3f" (wall_clock3 -. wall_clock2)); Proof_global.discard_all () ), `Yes, true @@ -2658,7 +2659,7 @@ let get_ast id = | { step = `Cmd { cast = { loc; expr } } } | { step = `Fork (({ loc; expr }, _, _, _), _) } | { step = `Qed ({ qast = { loc; expr } }, _) } -> - Some (expr, loc) + Some (Loc.tag ?loc expr) | _ -> None let stop_worker n = Slaves.cancel_worker n @@ -2749,7 +2750,7 @@ let add ~ontop ?newtip verb (loc, ast) = 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; expr = ast } in + let aast = { verbose = verb; indentation; strlen; loc = Some 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 ()) @@ -2770,7 +2771,7 @@ let query ~at ?(report_with=(Stateid.dummy,default_route)) s = let indentation, strlen = compute_indentation at loc in CWarnings.set_current_loc loc; let clas = classify_vernac ast in - let aast = { verbose = true; indentation; strlen; loc; expr = ast } in + let aast = { verbose = true; indentation; strlen; loc = Some loc; expr = ast } in match clas with | VtStm (w,_), _ -> ignore(process_transaction aast (VtStm (w,false), VtNow)) diff --git a/stm/stm.mli b/stm/stm.mli index 30e9629c6..699b4f5ff 100644 --- a/stm/stm.mli +++ b/stm/stm.mli @@ -87,7 +87,7 @@ val get_current_state : unit -> Stateid.t val init : unit -> unit (* This returns the node at that position *) -val get_ast : Stateid.t -> (Vernacexpr.vernac_expr * Loc.t) option +val get_ast : Stateid.t -> (Vernacexpr.vernac_expr Loc.located) option (* Filename *) val set_compilation_hints : string -> unit diff --git a/stm/vio_checking.ml b/stm/vio_checking.ml index a6237daa0..b2ea3341b 100644 --- a/stm/vio_checking.ml +++ b/stm/vio_checking.ml @@ -107,7 +107,7 @@ let schedule_vio_compilation j fs = let long_f_dot_v = Loadpath.locate_file (f^".v") in let aux = Aux_file.load_aux_file_for long_f_dot_v in let eta = - try float_of_string (Aux_file.get aux Loc.ghost "vo_compile_time") + try float_of_string (Aux_file.get aux "vo_compile_time") with Not_found -> 0.0 in jobs := (f, eta) :: !jobs) fs; diff --git a/tactics/contradiction.ml b/tactics/contradiction.ml index 63f923dfd..768d6860d 100644 --- a/tactics/contradiction.ml +++ b/tactics/contradiction.ml @@ -31,7 +31,7 @@ let absurd c = let env = Proofview.Goal.env gl in let sigma = Sigma.to_evar_map sigma in let j = Retyping.get_judgment_of env sigma c in - let sigma, j = Coercion.inh_coerce_to_sort Loc.ghost env sigma j in + let sigma, j = Coercion.inh_coerce_to_sort env sigma j in let t = j.Environ.utj_val in let tac = Tacticals.New.tclTHENLIST [ diff --git a/tactics/equality.ml b/tactics/equality.ml index 7ae7446c8..ace7c30e9 100644 --- a/tactics/equality.ml +++ b/tactics/equality.ml @@ -1462,7 +1462,7 @@ let simpleInjClause with_evars = function | Some c -> onInductionArg (fun clear_flag -> onEquality with_evars (injEq ~old:true with_evars clear_flag None)) c let injConcl = injClause None false None -let injHyp clear_flag id = injClause None false (Some (clear_flag,ElimOnIdent (Loc.ghost,id))) +let injHyp clear_flag id = injClause None false (Some (clear_flag,ElimOnIdent (Loc.tag id))) let decompEqThen ntac (lbeq,_,(t,t1,t2) as u) clause = Proofview.Goal.enter { enter = begin fun gl -> diff --git a/tactics/inv.ml b/tactics/inv.ml index 904a17417..f2147db40 100644 --- a/tactics/inv.ml +++ b/tactics/inv.ml @@ -283,7 +283,7 @@ let generalizeRewriteIntros as_mode tac depids id = end } let error_too_many_names pats = - let loc = Loc.join_loc (fst (List.hd pats)) (fst (List.last pats)) in + let loc = Loc.merge (fst (List.hd pats)) (fst (List.last pats)) in Proofview.tclENV >>= fun env -> tclZEROMSG ~loc ( str "Unexpected " ++ diff --git a/tactics/tacticals.ml b/tactics/tacticals.ml index 90b7d6581..074c8b9de 100644 --- a/tactics/tacticals.ml +++ b/tactics/tacticals.ml @@ -171,23 +171,23 @@ let fix_empty_or_and_pattern nv l = | IntroOrPattern [[]] -> IntroOrPattern (List.make nv []) | _ -> l -let check_or_and_pattern_size check_and loc names branchsigns = +let check_or_and_pattern_size ?loc check_and names branchsigns = let n = Array.length branchsigns in let msg p1 p2 = strbrk "a conjunctive pattern made of " ++ int p1 ++ (if p1 == p2 then mt () else str " or " ++ int p2) ++ str " patterns" in let err1 p1 p2 = - user_err ~loc (str "Expects " ++ msg p1 p2 ++ str ".") in + user_err ?loc (str "Expects " ++ msg p1 p2 ++ str ".") in let errn n = - user_err ~loc (str "Expects a disjunctive pattern with " ++ int n + user_err ?loc (str "Expects a disjunctive pattern with " ++ int n ++ str " branches.") in let err1' p1 p2 = - user_err ~loc (strbrk "Expects a disjunctive pattern with 1 branch or " ++ msg p1 p2 ++ str ".") in - let errforthcoming loc = - user_err ~loc (strbrk "Unexpected non atomic pattern.") in + user_err ?loc (strbrk "Expects a disjunctive pattern with 1 branch or " ++ msg p1 p2 ++ str ".") in + let errforthcoming ?loc = + user_err ?loc (strbrk "Unexpected non atomic pattern.") in match names with | 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 (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 @@ -195,7 +195,7 @@ let check_or_and_pattern_size check_and loc names branchsigns = if not (Int.equal p p1 || Int.equal p p2) then err1 p1 p2; if Int.equal p p1 then IntroAndPattern - (List.extend branchsigns.(0) (Loc.ghost,IntroNaming IntroAnonymous) l) + (List.extend branchsigns.(0) (Loc.tag @@ IntroNaming IntroAnonymous) l) else names else @@ -208,20 +208,20 @@ let check_or_and_pattern_size check_and loc names branchsigns = err1' p1 p2 else errn n; names -let get_and_check_or_and_pattern_gen check_and loc names branchsigns = - let names = check_or_and_pattern_size check_and loc names branchsigns in +let get_and_check_or_and_pattern_gen ?loc check_and names branchsigns = + let names = check_or_and_pattern_size ?loc check_and names branchsigns in match names with | IntroAndPattern l -> [|l|] | IntroOrPattern l -> Array.of_list l -let get_and_check_or_and_pattern = get_and_check_or_and_pattern_gen true +let get_and_check_or_and_pattern ?loc = get_and_check_or_and_pattern_gen ?loc true let compute_induction_names_gen check_and branchletsigns = function | None -> 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 diff --git a/tactics/tacticals.mli b/tactics/tacticals.mli index 3b90ec514..d60196213 100644 --- a/tactics/tacticals.mli +++ b/tactics/tacticals.mli @@ -116,7 +116,7 @@ type branch_assumptions = private { error message if |pats| <> |branchsign|; extends them if no pattern is given for let-ins in the case of a conjunctive pattern *) val get_and_check_or_and_pattern : - Loc.t -> delayed_open_constr or_and_intro_pattern_expr -> + ?loc:Loc.t -> delayed_open_constr or_and_intro_pattern_expr -> bool list array -> intro_patterns array (** Tolerate "[]" to mean a disjunctive pattern of any length *) diff --git a/tactics/tactics.ml b/tactics/tactics.ml index e79258582..7dad90242 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -53,8 +53,6 @@ module NamedDecl = Context.Named.Declaration let inj_with_occurrences e = (AllOccurrences,e) -let dloc = Loc.ghost - let typ_of env sigma c = let open Retyping in try get_type_of ~lax:true env (Sigma.to_evar_map sigma) c @@ -427,11 +425,11 @@ let default_id env sigma decl = type name_flag = | NamingAvoid of Id.t list | NamingBasedOn of Id.t * Id.t list - | NamingMustBe of Loc.t * Id.t + | NamingMustBe of Id.t Loc.located let naming_of_name = function | Anonymous -> NamingAvoid [] - | Name id -> NamingMustBe (dloc,id) + | Name id -> NamingMustBe (Loc.tag id) let find_name mayrepl decl naming gl = match naming with | NamingAvoid idl -> @@ -469,7 +467,7 @@ let assert_before_gen b naming t = assert_before_then_gen b naming t (fun _ -> Proofview.tclUNIT ()) let assert_before na = assert_before_gen false (naming_of_name na) -let assert_before_replacing id = assert_before_gen true (NamingMustBe (dloc,id)) +let assert_before_replacing id = assert_before_gen true (NamingMustBe (Loc.tag id)) let assert_after_then_gen b naming t tac = let open Context.Rel.Declaration in @@ -488,7 +486,7 @@ let assert_after_gen b naming t = assert_after_then_gen b naming t (fun _ -> (Proofview.tclUNIT ())) let assert_after na = assert_after_gen false (naming_of_name na) -let assert_after_replacing id = assert_after_gen true (NamingMustBe (dloc,id)) +let assert_after_replacing id = assert_after_gen true (NamingMustBe (Loc.tag id)) (**************************************************************) (* Fixpoints and CoFixpoints *) @@ -969,7 +967,7 @@ let rec intro_then_gen name_flag move_flag force_flag dep_flag tac = end } let intro_gen n m f d = intro_then_gen n m f d (fun _ -> Proofview.tclUNIT ()) -let intro_mustbe_force id = intro_gen (NamingMustBe (dloc,id)) MoveLast true false +let intro_mustbe_force id = intro_gen (NamingMustBe (Loc.tag id)) MoveLast true false let intro_using id = intro_gen (NamingBasedOn (id,[])) MoveLast false false let intro_then = intro_then_gen (NamingAvoid []) MoveLast false false @@ -979,7 +977,7 @@ let intro_avoiding l = intro_gen (NamingAvoid l) MoveLast false false let intro_move_avoid idopt avoid hto = match idopt with | None -> intro_gen (NamingAvoid avoid) hto true false - | Some id -> intro_gen (NamingMustBe (dloc,id)) hto true false + | Some id -> intro_gen (NamingMustBe (Loc.tag id)) hto true false let intro_move idopt hto = intro_move_avoid idopt [] hto @@ -1139,7 +1137,7 @@ let try_intros_until tac = function let rec intros_move = function | [] -> Proofview.tclUNIT () | (hyp,destopt) :: rest -> - Tacticals.New.tclTHEN (intro_gen (NamingMustBe (dloc,hyp)) destopt false false) + Tacticals.New.tclTHEN (intro_gen (NamingMustBe (Loc.tag hyp)) destopt false false) (intros_move rest) let run_delayed env sigma c = @@ -1291,7 +1289,7 @@ let clenv_refine_in ?(sidecond_first=false) with_evars ?(with_classes=true) error_uninstantiated_metas new_hyp_typ clenv; let new_hyp_prf = clenv_value clenv in let exact_tac = Proofview.V82.tactic (Tacmach.refine_no_check new_hyp_prf) in - let naming = NamingMustBe (dloc,targetid) in + let naming = NamingMustBe (Loc.tag targetid) in let with_clear = do_replace (Some id) naming in Tacticals.New.tclTHEN (Proofview.Unsafe.tclEVARS (clear_metas clenv.evd)) @@ -1798,13 +1796,13 @@ let apply_with_delayed_bindings_gen b e l = (one k f) (aux cbl) in aux l -let apply_with_bindings cb = apply_with_bindings_gen false false [None,(dloc,cb)] +let apply_with_bindings cb = apply_with_bindings_gen false false [None,(Loc.tag cb)] -let eapply_with_bindings cb = apply_with_bindings_gen false true [None,(dloc,cb)] +let eapply_with_bindings cb = apply_with_bindings_gen false true [None,(Loc.tag cb)] -let apply c = apply_with_bindings_gen false false [None,(dloc,(c,NoBindings))] +let apply c = apply_with_bindings_gen false false [None,(Loc.tag (c,NoBindings))] -let eapply c = apply_with_bindings_gen false true [None,(dloc,(c,NoBindings))] +let eapply c = apply_with_bindings_gen false true [None,(Loc.tag (c,NoBindings))] let apply_list = function | c::l -> apply_with_bindings (c,ImplicitBindings l) @@ -2216,7 +2214,7 @@ let constructor_tac with_evars expctdnumopt i lbind = (Proofview.Goal.env gl) sigma (fst mind, i) in let cons = mkConstructU (cons, EInstance.make u) in - let apply_tac = general_apply true false with_evars None (dloc,(cons,lbind)) in + let apply_tac = general_apply true false with_evars None (Loc.tag (cons,lbind)) in let tac = (Tacticals.New.tclTHENLIST [ @@ -2301,7 +2299,7 @@ let my_find_eq_data_decompose gl t = -> None | Constr_matching.PatternMatchingFailure -> None -let intro_decomp_eq loc l thin tac id = +let intro_decomp_eq ?loc l thin tac id = Proofview.Goal.enter { enter = begin fun gl -> let c = mkVar id in let t = Tacmach.New.pf_unsafe_type_of gl c in @@ -2309,13 +2307,13 @@ let intro_decomp_eq loc l thin tac id = match my_find_eq_data_decompose gl t with | Some (eq,u,eq_args) -> !intro_decomp_eq_function - (fun n -> tac ((dloc,id)::thin) (Some (true,n)) l) + (fun n -> tac ((Loc.tag id)::thin) (Some (true,n)) l) (eq,t,eq_args) (c, t) | None -> Tacticals.New.tclZEROMSG (str "Not a primitive equality here.") end } -let intro_or_and_pattern loc with_evars bracketed ll thin tac id = +let intro_or_and_pattern ?loc with_evars bracketed ll thin tac id = Proofview.Goal.enter { enter = begin fun gl -> let c = mkVar id in let t = Tacmach.New.pf_unsafe_type_of gl c in @@ -2323,7 +2321,7 @@ let intro_or_and_pattern loc with_evars bracketed ll thin tac id = let branchsigns = compute_constructor_signatures false ind in let nv_with_let = Array.map List.length branchsigns in let ll = fix_empty_or_and_pattern (Array.length branchsigns) ll in - let ll = get_and_check_or_and_pattern loc ll branchsigns in + let ll = get_and_check_or_and_pattern ?loc ll branchsigns in Tacticals.New.tclTHENLASTn (Tacticals.New.tclTHEN (simplest_ecase c) (clear [id])) (Array.map2 (fun n l -> tac thin (Some (bracketed,n)) l) @@ -2372,8 +2370,8 @@ let rewrite_hyp_then assert_style with_evars thin l2r id tac = Tacticals.New.tclTHENFIRST eqtac (tac thin) end } -let prepare_naming loc = function - | IntroIdentifier id -> NamingMustBe (loc,id) +let prepare_naming ?loc = function + | IntroIdentifier id -> NamingMustBe (Loc.tag ?loc id) | IntroAnonymous -> NamingAvoid [] | IntroFresh id -> NamingBasedOn (id,[]) @@ -2415,7 +2413,7 @@ let make_tmp_naming avoid l = function case of IntroFresh, we should use check_thin_clash_then anyway to prevent the case of an IntroFresh precisely using the wild_id *) | IntroWildcard -> NamingBasedOn (wild_id,avoid@explicit_intro_names l) - | pat -> NamingAvoid(avoid@explicit_intro_names ((dloc,IntroAction pat)::l)) + | pat -> NamingAvoid(avoid@explicit_intro_names ((Loc.tag @@ IntroAction pat)::l)) let fit_bound n = function | None -> true @@ -2451,7 +2449,7 @@ let rec intro_patterns_core with_evars b avoid ids thin destopt bound n tac = | [] -> (* Behave as IntroAnonymous *) intro_patterns_core with_evars b avoid ids thin destopt bound n tac - [dloc,IntroNaming IntroAnonymous] + [Loc.tag @@ IntroNaming IntroAnonymous] | (loc,pat) :: l -> if exceed_bound n bound then error_unexpected_extra_pattern loc bound pat else match pat with @@ -2463,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 -> @@ -2488,21 +2486,21 @@ and intro_pattern_naming loc with_evars b avoid ids pat thin destopt bound n tac destopt true false (fun id -> intro_patterns_core with_evars b avoid (id::ids) thin destopt bound n tac l) -and intro_pattern_action loc with_evars b style pat thin destopt tac id = +and intro_pattern_action ?loc with_evars b style pat thin destopt tac id = match pat with | IntroWildcard -> - tac ((loc,id)::thin) None [] + tac ((Loc.tag ?loc id)::thin) None [] | IntroOrAndPattern ll -> - intro_or_and_pattern loc with_evars b ll thin tac id + intro_or_and_pattern ?loc with_evars b ll thin tac id | IntroInjection l' -> - intro_decomp_eq loc l' thin tac id + intro_decomp_eq ?loc l' thin tac id | IntroRewrite l2r -> 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 loc with_evars (IntroIdentifier id) destopt pat in + prepare_intros ~loc with_evars (IntroIdentifier id) destopt pat in let doclear = - if naming = NamingMustBe (loc,id) then + if naming = NamingMustBe (Loc.tag ~loc id) then Proofview.tclUNIT () (* apply_in_once do a replacement *) else clear [id] in @@ -2513,18 +2511,18 @@ and intro_pattern_action loc with_evars b style pat thin destopt tac id = apply_in_delayed_once false true true with_evars naming id (None,(loc',f)) (fun id -> Tacticals.New.tclTHENLIST [doclear; tac_ipat id; tac thin None []]) -and prepare_intros_loc loc with_evars dft destopt = function +and prepare_intros ?loc with_evars dft destopt = function | IntroNaming ipat -> - prepare_naming loc ipat, + prepare_naming ?loc ipat, (fun id -> move_hyp id destopt) | IntroAction ipat -> - prepare_naming loc dft, + prepare_naming ?loc dft, (let tac thin bound = intro_patterns_core with_evars true [] [] thin destopt bound 0 (fun _ l -> clear_wildcards l) in fun id -> - intro_pattern_action loc with_evars true true ipat [] destopt tac id) - | IntroForthcoming _ -> user_err ~loc + intro_pattern_action ?loc with_evars true true ipat [] destopt tac id) + | IntroForthcoming _ -> user_err ?loc (str "Introduction pattern for one hypothesis expected.") let intro_patterns_bound_to with_evars n destopt = @@ -2536,7 +2534,7 @@ let intro_patterns_to with_evars destopt = [] [] [] destopt None 0 (fun _ l -> clear_wildcards l) let intro_pattern_to with_evars destopt pat = - intro_patterns_to with_evars destopt [dloc,pat] + intro_patterns_to with_evars destopt [Loc.tag pat] let intro_patterns with_evars = intro_patterns_to with_evars MoveLast @@ -2549,20 +2547,20 @@ let intros_patterns with_evars = function (* Forward reasoning *) (**************************) -let prepare_intros with_evars dft destopt = function - | None -> prepare_naming dloc dft, (fun _id -> Proofview.tclUNIT ()) - | Some (loc,ipat) -> prepare_intros_loc loc with_evars dft destopt ipat +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 let ipat_of_name = function | Anonymous -> None - | Name id -> Some (dloc, IntroNaming (IntroIdentifier id)) + | Name id -> Some (Loc.tag @@ IntroNaming (IntroIdentifier id)) let head_ident sigma c = let c = fst (decompose_app sigma (snd (decompose_lam_assum sigma c))) in if isVar sigma c then Some (destVar sigma c) else None let assert_as first hd ipat t = - let naming,tac = prepare_intros false IntroAnonymous MoveLast ipat in + let naming,tac = prepare_intros_opt false IntroAnonymous MoveLast ipat in let repl = do_replace hd naming in let tac = if repl then (fun id -> Proofview.tclUNIT ()) else tac in if first then assert_before_then_gen repl naming t tac @@ -2580,10 +2578,10 @@ let general_apply_in sidecond_first with_delta with_destruct with_evars if with_evars then MoveLast (* evars would depend on the whole context *) else get_previous_hyp_position id gl in let naming,ipat_tac = - prepare_intros with_evars (IntroIdentifier id) destopt ipat in + prepare_intros_opt with_evars (IntroIdentifier id) destopt ipat in let lemmas_target, last_lemma_target = let last,first = List.sep_last lemmas in - List.map (fun lem -> (NamingMustBe (dloc,id),lem)) first, (naming,last) + List.map (fun lem -> (NamingMustBe (Loc.tag id),lem)) first, (naming,last) in (* We chain apply_in_once, ending with an intro pattern *) List.fold_right tac lemmas_target (tac last_lemma_target ipat_tac) id @@ -2661,7 +2659,7 @@ let letin_tac_gen with_eq (id,depdecls,lastlhyp,ccl,c) ty = let tac = Tacticals.New.tclTHENLIST [ convert_concl_no_check newcl DEFAULTcast; - intro_gen (NamingMustBe (dloc,id)) (decode_hyp lastlhyp) true false; + intro_gen (NamingMustBe (Loc.tag id)) (decode_hyp lastlhyp) true false; Tacticals.New.tclMAP convert_hyp_no_check depdecls; eq_tac ] in @@ -2961,7 +2959,7 @@ let specialize (c,lbind) ipat = | Var id when Id.List.mem id (Tacmach.New.pf_ids_of_hyps gl) -> (* Like assert (id:=id args) but with the concept of specialization *) let naming,tac = - prepare_intros false (IntroIdentifier id) MoveLast ipat in + prepare_intros_opt false (IntroIdentifier id) MoveLast ipat in let repl = do_replace (Some id) naming in Tacticals.New.tclTHENFIRST (assert_before_then_gen repl naming typ tac) @@ -2975,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 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) @@ -3063,7 +3061,7 @@ let intropattern_of_name gl avoid = function | Name id -> IntroNaming (IntroIdentifier (new_fresh_id avoid id gl)) let rec consume_pattern avoid na isdep gl = function - | [] -> ((dloc, intropattern_of_name gl avoid na), []) + | [] -> ((Loc.tag @@ intropattern_of_name gl avoid na), []) | (loc,IntroForthcoming true)::names when not isdep -> consume_pattern avoid na isdep gl names | (loc,IntroForthcoming _)::names as fullpat -> @@ -3140,7 +3138,7 @@ let induct_discharge with_evars dests avoid' tac (avoid,ra) names = let (recpat,names) = match names with | [loc,IntroNaming (IntroIdentifier id) as pat] -> let id' = next_ident_away (add_prefix "IH" id) avoid in - (pat, [dloc, IntroNaming (IntroIdentifier id')]) + (pat, [Loc.tag @@ IntroNaming (IntroIdentifier id')]) | _ -> consume_pattern avoid (Name recvarname) deprec gl names in let dest = get_recarg_dest dests in dest_intro_patterns with_evars avoid thin dest [recpat] (fun ids thin -> @@ -5026,14 +5024,14 @@ module Simple = struct let intro x = intro_move (Some x) MoveLast let apply c = - apply_with_bindings_gen false false [None,(Loc.ghost,(c,NoBindings))] + apply_with_bindings_gen false false [None,(Loc.tag (c,NoBindings))] let eapply c = - apply_with_bindings_gen false true [None,(Loc.ghost,(c,NoBindings))] + apply_with_bindings_gen false true [None,(Loc.tag (c,NoBindings))] let elim c = elim false None (c,NoBindings) None let case c = general_case_analysis false None (c,NoBindings) let apply_in id c = - apply_in false false id [None,(Loc.ghost, (c, NoBindings))] None + apply_in false false id [None,(Loc.tag (c, NoBindings))] None end diff --git a/toplevel/coqtop.ml b/toplevel/coqtop.ml index f5f43ff66..b7065993f 100644 --- a/toplevel/coqtop.ml +++ b/toplevel/coqtop.ml @@ -185,7 +185,7 @@ let load_vernacular sid = let load_vernacular_obj = ref ([] : string list) let add_vernac_obj s = load_vernacular_obj := s :: !load_vernacular_obj let load_vernac_obj () = - let map dir = Qualid (Loc.ghost, qualid_of_string dir) in + let map dir = Qualid (Loc.tag @@ qualid_of_string dir) in Vernacentries.vernac_require None None (List.rev_map map !load_vernacular_obj) let require_prelude () = @@ -200,7 +200,7 @@ let require_list = ref ([] : string list) let add_require s = require_list := s :: !require_list let require () = let () = if !load_init then silently require_prelude () in - let map dir = Qualid (Loc.ghost, qualid_of_string dir) in + let map dir = Qualid (Loc.tag @@ qualid_of_string dir) in Vernacentries.vernac_require None (Some false) (List.rev_map map !require_list) let add_compat_require v = diff --git a/toplevel/vernac.ml b/toplevel/vernac.ml index 3359a1672..d6bcd2f15 100644 --- a/toplevel/vernac.ml +++ b/toplevel/vernac.ml @@ -146,9 +146,11 @@ let rec interp_vernac sid po (loc,com) = interp com with reraise -> let (reraise, info) = CErrors.push reraise in - let loc' = Option.default Loc.ghost (Loc.get_loc info) in - if Loc.is_ghost loc' then iraise (reraise, Loc.add_loc info loc) - else iraise (reraise, info) + let info = begin + match Loc.get_loc info with + | None -> Loc.add_loc info loc + | Some _ -> info + end in iraise (reraise, info) (* Load a vernac file. CErrors are annotated with file and location *) and load_vernac verbosely sid file = @@ -281,7 +283,7 @@ let compile verbosely f = let wall_clock2 = Unix.gettimeofday () in check_pending_proofs (); Library.save_library_to ldir long_f_dot_vo (Global.opaque_tables ()); - Aux_file.record_in_aux_at Loc.ghost "vo_compile_time" + Aux_file.record_in_aux_at "vo_compile_time" (Printf.sprintf "%.3f" (wall_clock2 -. wall_clock1)); Aux_file.stop_aux_file (); if !Flags.xml_export then Hook.get f_xml_end_library (); diff --git a/vernac/auto_ind_decl.ml b/vernac/auto_ind_decl.ml index b9c4c6cc5..fd454b67e 100644 --- a/vernac/auto_ind_decl.ml +++ b/vernac/auto_ind_decl.ml @@ -60,8 +60,6 @@ exception NonSingletonProp of inductive exception DecidabilityMutualNotSupported exception NoDecidabilityCoInductive -let dl = Loc.ghost - let constr_of_global g = lazy (Universes.constr_of_global g) (* Some pre declaration of constant we are going to use *) @@ -88,12 +86,12 @@ let destruct_on c = destruct false None c None None let destruct_on_using c id = destruct false None c - (Some (dl,IntroOrPattern [[dl,IntroNaming IntroAnonymous]; - [dl,IntroNaming (IntroIdentifier id)]])) + (Some (Loc.tag @@ IntroOrPattern [[Loc.tag @@ IntroNaming IntroAnonymous]; + [Loc.tag @@ IntroNaming (IntroIdentifier id)]])) None let destruct_on_as c l = - destruct false None c (Some (dl,l)) None + destruct false None c (Some (Loc.tag l)) None (* reconstruct the inductive with the correct deBruijn indexes *) let mkFullInd (ind,u) n = @@ -611,8 +609,8 @@ repeat ( apply andb_prop in z;let z1:= fresh "Z" in destruct z as [z1 z]). Proofview.Goal.enter { enter = begin fun gl -> let fresht = fresh_id (Id.of_string "Z") gl in destruct_on_as (EConstr.mkVar freshz) - (IntroOrPattern [[dl,IntroNaming (IntroIdentifier fresht); - dl,IntroNaming (IntroIdentifier freshz)]]) + (IntroOrPattern [[Loc.tag @@ IntroNaming (IntroIdentifier fresht); + Loc.tag @@ IntroNaming (IntroIdentifier freshz)]]) end } ]); (* diff --git a/vernac/classes.ml b/vernac/classes.ml index ffe03bfb7..bdaa3ece6 100644 --- a/vernac/classes.ml +++ b/vernac/classes.ml @@ -408,7 +408,7 @@ let context poly l = let decl = (Discharge, poly, Definitional) in let nstatus = pi3 (Command.declare_assumption false decl (t, !uctx) [] [] impl - Vernacexpr.NoInline (Loc.ghost, id)) + Vernacexpr.NoInline (Loc.tag id)) in let () = uctx := Univ.ContextSet.empty in status && nstatus diff --git a/vernac/command.ml b/vernac/command.ml index 446afb578..fbaa09430 100644 --- a/vernac/command.ml +++ b/vernac/command.ml @@ -267,7 +267,7 @@ match local with (gr,inst,Lib.is_modtype_strict ()) let interp_assumption evdref env impls bl c = - let c = mkCProdN (local_binders_loc bl) bl c in + let c = mkCProdN ?loc:(local_binders_loc bl) bl c in let ty, impls = interp_type_evars_impls env evdref ~impls c in let ty = EConstr.Unsafe.to_constr ty in (ty, impls) @@ -917,7 +917,7 @@ let mkSubset name typ prop = [| typ; mkLambda (name, typ, prop) |]) let sigT = Lazy.from_fun build_sigma_type -let make_qref s = Qualid (Loc.ghost, qualid_of_string s) +let make_qref s = Qualid (Loc.tag @@ qualid_of_string s) let lt_ref = make_qref "Init.Peano.lt" let rec telescope l = @@ -1059,7 +1059,7 @@ let build_wellfounded (recname,pl,n,bl,arityc,body) poly r measure notation = mkApp (EConstr.of_constr (Universes.constr_of_global (delayed_force fix_sub_ref)), [| argtyp ; wf_rel ; Evarutil.e_new_evar env evdref - ~src:(Loc.ghost, Evar_kinds.QuestionMark (Evar_kinds.Define false)) wf_proof; + ~src:(Loc.tag @@ Evar_kinds.QuestionMark (Evar_kinds.Define false)) wf_proof; prop |]) in let def = Typing.e_solve_evars env evdref def in @@ -1212,7 +1212,7 @@ let declare_fixpoint local poly ((fixnames,fixdefs,fixtypes),pl,ctx,fiximps) ind let fixdefs = List.map Option.get fixdefs in let fixdecls = prepare_recursive_declaration fixnames fixtypes fixdefs in let env = Global.env() in - let indexes = search_guard Loc.ghost env indexes fixdecls in + let indexes = search_guard env indexes fixdecls in let fiximps = List.map (fun (n,r,p) -> r) fiximps in let vars = Universes.universes_of_constr (mkFix ((indexes,0),fixdecls)) in let fixdecls = @@ -1327,8 +1327,7 @@ let do_program_recursive local p fixkind fixl ntns = Array.of_list (List.map (subst_vars (List.rev fixnames)) fixdefs) in let indexes = - Pretyping.search_guard - Loc.ghost (Global.env ()) possible_indexes fixdecls in + Pretyping.search_guard (Global.env ()) possible_indexes fixdecls in List.iteri (fun i _ -> Inductive.check_fix env ((indexes,i),fixdecls)) diff --git a/vernac/explainErr.ml b/vernac/explainErr.ml index 04841c922..ed2dd274a 100644 --- a/vernac/explainErr.ml +++ b/vernac/explainErr.ml @@ -119,6 +119,6 @@ let process_vernac_interp_error ?(allow_uncaught=true) (exc, info) = in match e' with | None -> e - | Some (None, loc) -> (fst e, Loc.add_loc (snd e) loc) - | Some (Some msg, loc) -> + | Some (loc, None) -> (fst e, Loc.add_loc (snd e) loc) + | Some (loc, Some msg) -> (EvaluatedError (msg, Some (fst e)), Loc.add_loc (snd e) loc) diff --git a/vernac/explainErr.mli b/vernac/explainErr.mli index 370ad7e3b..9202729ce 100644 --- a/vernac/explainErr.mli +++ b/vernac/explainErr.mli @@ -18,4 +18,4 @@ val process_vernac_interp_error : ?allow_uncaught:bool -> Util.iexn -> Util.iexn val explain_exn_default : exn -> Pp.std_ppcmds -val register_additional_error_info : (Util.iexn -> (Pp.std_ppcmds option * Loc.t) option) -> unit +val register_additional_error_info : (Util.iexn -> (Pp.std_ppcmds option Loc.located) option) -> unit diff --git a/vernac/indschemes.ml b/vernac/indschemes.ml index 9ba4e46e4..2367177e2 100644 --- a/vernac/indschemes.ml +++ b/vernac/indschemes.ml @@ -379,7 +379,7 @@ requested | InType -> recs ^ "t_nodep") ) in let newid = add_suffix (basename_of_global (IndRef ind)) suffix in - let newref = (Loc.ghost,newid) in + let newref = Loc.tag newid in ((newref,isdep,ind,z)::l1),l2 in match t with diff --git a/vernac/lemmas.ml b/vernac/lemmas.ml index 993a2c260..b51792b1e 100644 --- a/vernac/lemmas.ml +++ b/vernac/lemmas.ml @@ -81,7 +81,7 @@ let adjust_guardness_conditions const = function List.fold_left (fun e (_,c,cb,_) -> add c cb e) env l) env (Safe_typing.side_effects_of_private_constants eff) in let indexes = - search_guard Loc.ghost env + search_guard env possible_indexes fixdecls in (mkFix ((indexes,0),fixdecls), ctx), eff | _ -> (body, ctx), eff) } diff --git a/vernac/metasyntax.ml b/vernac/metasyntax.ml index 98b2c3729..a586fe9f2 100644 --- a/vernac/metasyntax.ml +++ b/vernac/metasyntax.ml @@ -1416,7 +1416,7 @@ let add_notation_extra_printing_rule df k v = (* Infix notations *) -let inject_var x = Loc.tag @@ CRef (Ident (Loc.ghost, Id.of_string x),None) +let inject_var x = Loc.tag @@ CRef (Ident (Loc.tag @@ Id.of_string x),None) let add_infix local ((loc,inf),modifiers) pr sc = check_infix_modifiers modifiers; diff --git a/vernac/obligations.ml b/vernac/obligations.ml index ea2401b5c..5e65855ef 100644 --- a/vernac/obligations.ml +++ b/vernac/obligations.ml @@ -558,8 +558,7 @@ let declare_mutual_definition l = List.map3 compute_possible_guardness_evidences wfl fixdefs fixtypes in let indexes = - Pretyping.search_guard - Loc.ghost (Global.env()) + Pretyping.search_guard (Global.env()) possible_indexes fixdecls in Some indexes, List.map_i (fun i _ -> @@ -674,7 +673,7 @@ let init_prog_info ?(opaque = false) sign n pl b t ctx deps fixkind assert(Int.equal (Array.length obls) 0); let n = Nameops.add_suffix n "_obligation" in [| { obl_name = n; obl_body = None; - obl_location = Loc.ghost, Evar_kinds.InternalHole; obl_type = t; + obl_location = Loc.tag Evar_kinds.InternalHole; obl_type = t; obl_status = false, Evar_kinds.Expand; obl_deps = Int.Set.empty; obl_tac = None } |], mkVar n diff --git a/vernac/vernacentries.ml b/vernac/vernacentries.ml index 86406dbe5..75f403e33 100644 --- a/vernac/vernacentries.ml +++ b/vernac/vernacentries.ml @@ -643,7 +643,7 @@ let vernac_declare_module export (loc, id) binders_ast mty_ast = in 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.ghost,id)]) export + Option.iter (fun export -> vernac_import export [Ident (Loc.tag id)]) export let vernac_define_module export (loc, id) binders_ast mty_ast_o mexpr_ast_l = (* We check the state of the system (in section, in module type) @@ -668,7 +668,7 @@ let vernac_define_module export (loc, id) binders_ast mty_ast_o mexpr_ast_l = List.iter (fun (export,id) -> Option.iter - (fun export -> vernac_import export [Ident (Loc.ghost,id)]) export + (fun export -> vernac_import export [Ident (Loc.tag id)]) export ) argsexport | _::_ -> let binders_ast = List.map @@ -683,7 +683,7 @@ let vernac_define_module export (loc, id) binders_ast mty_ast_o mexpr_ast_l = 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.ghost,id)]) + Option.iter (fun export -> vernac_import export [Ident (Loc.tag id)]) export let vernac_end_module export (loc,id as lid) = @@ -715,7 +715,7 @@ let vernac_declare_module_type (loc,id) binders_ast mty_sign mty_ast_l = List.iter (fun (export,id) -> Option.iter - (fun export -> vernac_import export [Ident (Loc.ghost,id)]) export + (fun export -> vernac_import export [Ident (Loc.tag id)]) export ) argsexport | _ :: _ -> @@ -1176,10 +1176,10 @@ let vernac_arguments locality reference args more_implicits nargs_for_red flags end; if scopes_specified || clear_scopes_flag then begin - let scopes = List.map (Option.map (fun (o,k) -> + let scopes = List.map (Option.map (fun (loc,k) -> try ignore (Notation.find_scope k); k with UserError _ -> - Notation.find_delimiters_scope o k)) scopes + Notation.find_delimiters_scope ~loc k)) scopes in vernac_arguments_scope locality reference scopes end; @@ -1708,7 +1708,7 @@ let interp_search_about_item env = | SearchString (s,sc) -> try let ref = - Notation.interp_notation_as_global_reference Loc.ghost + Notation.interp_notation_as_global_reference (fun _ -> true) s sc in GlobSearchSubPattern (Pattern.PRef ref) with UserError _ -> @@ -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 diff --git a/vernac/vernacentries.mli b/vernac/vernacentries.mli index fb2899515..f75f7656d 100644 --- a/vernac/vernacentries.mli +++ b/vernac/vernacentries.mli @@ -18,7 +18,7 @@ val vernac_require : val interp : ?verbosely:bool -> ?proof:Proof_global.closed_proof -> - Loc.t * Vernacexpr.vernac_expr -> unit + Vernacexpr.vernac_expr Loc.located -> unit (** Prepare a "match" template for a given inductive type. For each branch of the match, we list the constructor name -- cgit v1.2.3