From b1d749e59444f86e40f897c41739168bb1b1b9b3 Mon Sep 17 00:00:00 2001 From: Emilio Jesus Gallego Arias Date: Sun, 25 Feb 2018 22:43:42 +0100 Subject: [located] Push inner locations in `reference` to a CAst.t node. The `reference` type contains some ad-hoc locations in its constructors, but there is no reason not to handle them with the standard attribute container provided by `CAst.t`. An orthogonal topic to this commit is whether the `reference` type should contain a location or not at all. It seems that many places would become a bit clearer by splitting `reference` into non-located `reference` and `lreference`, however some other places become messier so we maintain the current status-quo for now. --- plugins/ltac/g_ltac.ml4 | 14 +++++------ plugins/ltac/g_obligations.ml4 | 2 +- plugins/ltac/g_tactic.ml4 | 3 +-- plugins/ltac/pptactic.ml | 4 +-- plugins/ltac/rewrite.ml | 37 ++++++++++++++-------------- plugins/ltac/tacentries.ml | 5 ++-- plugins/ltac/tacintern.ml | 55 +++++++++++++++++++++--------------------- plugins/ltac/tacinterp.ml | 10 ++++---- 8 files changed, 63 insertions(+), 67 deletions(-) (limited to 'plugins/ltac') diff --git a/plugins/ltac/g_ltac.ml4 b/plugins/ltac/g_ltac.ml4 index 1909cd96f..0c42a8bb2 100644 --- a/plugins/ltac/g_ltac.ml4 +++ b/plugins/ltac/g_ltac.ml4 @@ -38,11 +38,11 @@ let genarg_of_ipattern pat = in_gen (rawwit Stdarg.wit_intro_pattern) pat let genarg_of_uconstr c = in_gen (rawwit Stdarg.wit_uconstr) c let in_tac tac = in_gen (rawwit Tacarg.wit_ltac) tac -let reference_to_id = function - | Libnames.Ident (loc, id) -> CAst.make ?loc id - | Libnames.Qualid (loc,_) -> +let reference_to_id = CAst.map_with_loc (fun ?loc -> function + | Libnames.Ident id -> id + | Libnames.Qualid _ -> CErrors.user_err ?loc - (str "This expression should be a simple identifier.") + (str "This expression should be a simple identifier.")) let tactic_mode = Gram.entry_create "vernac:tactic_command" @@ -473,7 +473,7 @@ END VERNAC COMMAND EXTEND VernacPrintLtac CLASSIFIED AS QUERY | [ "Print" "Ltac" reference(r) ] -> - [ Feedback.msg_notice (Tacintern.print_ltac (snd (Libnames.qualid_of_reference r))) ] + [ Feedback.msg_notice (Tacintern.print_ltac (Libnames.qualid_of_reference r).CAst.v) ] END VERNAC COMMAND EXTEND VernacLocateLtac CLASSIFIED AS QUERY @@ -508,8 +508,8 @@ VERNAC COMMAND FUNCTIONAL EXTEND VernacDeclareTacticDefinition | [ "Ltac" ne_ltac_tacdef_body_list_sep(l, "with") ] => [ VtSideff (List.map (function | TacticDefinition ({CAst.v=r},_) -> r - | TacticRedefinition (Ident (_,r),_) -> r - | TacticRedefinition (Qualid (_,q),_) -> snd(repr_qualid q)) l), VtLater + | TacticRedefinition ({CAst.v=Ident r},_) -> r + | TacticRedefinition ({CAst.v=Qualid q},_) -> snd(repr_qualid q)) l), VtLater ] -> [ fun ~atts ~st -> let open Vernacinterp in Tacentries.register_ltac (Locality.make_module_locality atts.locality) l; st diff --git a/plugins/ltac/g_obligations.ml4 b/plugins/ltac/g_obligations.ml4 index 54e2ba960..352e92c2a 100644 --- a/plugins/ltac/g_obligations.ml4 +++ b/plugins/ltac/g_obligations.ml4 @@ -49,7 +49,7 @@ module Tactic = Pltac open Pcoq -let sigref = mkRefC (Qualid (Loc.tag @@ Libnames.qualid_of_string "Coq.Init.Specif.sig")) +let sigref = mkRefC (CAst.make @@ Qualid (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/g_tactic.ml4 b/plugins/ltac/g_tactic.ml4 index 7643cc97d..7534e2799 100644 --- a/plugins/ltac/g_tactic.ml4 +++ b/plugins/ltac/g_tactic.ml4 @@ -154,8 +154,7 @@ let mkTacCase with_evar = function (* Reinterpret ident as notations for variables in the context *) (* because we don't know if they are quantified or not *) | [(clear,ElimOnIdent id),(None,None),None],None -> - let id = CAst.(id.loc, id.v) in - TacCase (with_evar,(clear,(CAst.make @@ CRef (Ident id,None),NoBindings))) + TacCase (with_evar,(clear,(CAst.make @@ CRef (CAst.make ?loc:id.CAst.loc @@ Ident id.CAst.v,None),NoBindings))) | ic -> if List.exists (function ((_, ElimOnAnonHyp _),_,_) -> true | _ -> false) (fst ic) then diff --git a/plugins/ltac/pptactic.ml b/plugins/ltac/pptactic.ml index a42994652..5d262ffcb 100644 --- a/plugins/ltac/pptactic.ml +++ b/plugins/ltac/pptactic.ml @@ -181,9 +181,9 @@ let string_of_genarg_arg (ArgumentType arg) = let pr_and_short_name pr (c,_) = pr c - let pr_or_by_notation f = function + let pr_or_by_notation f = CAst.with_val (function | AN v -> f v - | ByNotation {CAst.v=(s,sc)} -> qs s ++ pr_opt (fun sc -> str "%" ++ str sc) sc + | ByNotation (s,sc) -> qs s ++ pr_opt (fun sc -> str "%" ++ str sc) sc) let pr_located pr (loc,x) = pr x diff --git a/plugins/ltac/rewrite.ml b/plugins/ltac/rewrite.ml index e0368153e..d32a2faef 100644 --- a/plugins/ltac/rewrite.ml +++ b/plugins/ltac/rewrite.ml @@ -1773,12 +1773,11 @@ let rec strategy_of_ast = function (* By default the strategy for "rewrite_db" is top-down *) -let mkappc s l = CAst.make @@ CAppExpl ((None,(Libnames.Ident (Loc.tag @@ Id.of_string s)),None),l) +let mkappc s l = CAst.make @@ CAppExpl ((None,CAst.make @@ Libnames.Ident (Id.of_string s),None),l) let declare_an_instance n s args = (((CAst.make @@ Name n),None), Explicit, - CAst.make @@ CAppExpl ((None, Qualid (Loc.tag @@ qualid_of_string s),None), - args)) + CAst.make @@ CAppExpl ((None, CAst.make @@ Qualid (qualid_of_string s),None), args)) let declare_instance a aeq n s = declare_an_instance n s [a;aeq] @@ -1792,17 +1791,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.tag @@ Id.of_string "reflexivity"),lemma)] + [(CAst.make @@ Ident (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.tag @@ Id.of_string "symmetry"),lemma)] + [(CAst.make @@ Ident (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.tag @@ Id.of_string "transitivity"),lemma)] + [(CAst.make @@ Ident (Id.of_string "transitivity"),lemma)] let declare_relation ?locality ?(binders=[]) a aeq n refl symm trans = init_setoid (); @@ -1826,16 +1825,16 @@ let declare_relation ?locality ?(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.tag @@ Id.of_string "PreOrder_Reflexive"), lemma1); - (Ident (Loc.tag @@ Id.of_string "PreOrder_Transitive"),lemma3)]) + [(CAst.make @@ Ident (Id.of_string "PreOrder_Reflexive"), lemma1); + (CAst.make @@ Ident (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.tag @@ Id.of_string "PER_Symmetric"), lemma2); - (Ident (Loc.tag @@ Id.of_string "PER_Transitive"),lemma3)]) + [(CAst.make @@ Ident (Id.of_string "PER_Symmetric"), lemma2); + (CAst.make @@ Ident (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 @@ -1843,9 +1842,9 @@ let declare_relation ?locality ?(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.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)]) + [(CAst.make @@ Ident (Id.of_string "Equivalence_Reflexive"), lemma1); + (CAst.make @@ Ident (Id.of_string "Equivalence_Symmetric"), lemma2); + (CAst.make @@ Ident (Id.of_string "Equivalence_Transitive"), lemma3)]) let cHole = CAst.make @@ CHole (None, Misctypes.IntroAnonymous, None) @@ -1951,16 +1950,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.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])]) + [(CAst.make @@ Ident (Id.of_string "Equivalence_Reflexive"), mkappc "Seq_refl" [a;aeq;t]); + (CAst.make @@ Ident (Id.of_string "Equivalence_Symmetric"), mkappc "Seq_sym" [a;aeq;t]); + (CAst.make @@ Ident (Id.of_string "Equivalence_Transitive"), mkappc "Seq_trans" [a;aeq;t])]) let make_tactic name = let open Tacexpr in let tacpath = Libnames.qualid_of_string name in - let tacname = Qualid (Loc.tag tacpath) in - TacArg (Loc.tag @@ TacCall (Loc.tag (tacname, []))) + let tacname = CAst.make @@ Qualid tacpath in + TacArg (Loc.tag @@ (TacCall (Loc.tag (tacname, [])))) let warn_add_morphism_deprecated = CWarnings.create ~name:"add-morphism" ~category:"deprecated" (fun () -> @@ -2010,7 +2009,7 @@ let add_morphism glob binders m s n = let instance = (((CAst.make @@ Name instance_id),None), Explicit, CAst.make @@ CAppExpl ( - (None, Qualid (Loc.tag @@ Libnames.qualid_of_string "Coq.Classes.Morphisms.Proper"),None), + (None, CAst.make @@ Qualid (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 566fc2873..e510b9f59 100644 --- a/plugins/ltac/tacentries.ml +++ b/plugins/ltac/tacentries.ml @@ -450,11 +450,10 @@ let register_ltac local tacl = let () = if is_shadowed then warn_unusable_identifier id in NewTac id, body | Tacexpr.TacticRedefinition (ident, body) -> - let loc = loc_of_reference ident in let kn = - try Tacenv.locate_tactic (snd (qualid_of_reference ident)) + try Tacenv.locate_tactic (qualid_of_reference ident).CAst.v with Not_found -> - CErrors.user_err ?loc + CErrors.user_err ?loc:ident.CAst.loc (str "There is no Ltac named " ++ pr_reference ident ++ str ".") in UpdateTac kn, body diff --git a/plugins/ltac/tacintern.ml b/plugins/ltac/tacintern.ml index 45f4a45fc..9ad9e1520 100644 --- a/plugins/ltac/tacintern.ml +++ b/plugins/ltac/tacintern.ml @@ -92,15 +92,15 @@ let intern_int_or_var = intern_or_var (fun (n : int) -> n) let intern_string_or_var = intern_or_var (fun (s : string) -> s) let intern_global_reference ist = function - | Ident (loc,id) when find_var id ist -> + | {CAst.loc;v=Ident id} when find_var id ist -> ArgVar (make ?loc id) | r -> - let loc,_ as lqid = qualid_of_reference r in - try ArgArg (loc,locate_global_with_alias lqid) - with Not_found -> error_global_not_found (snd lqid) + let {CAst.loc} as lqid = qualid_of_reference r in + try ArgArg (loc,locate_global_with_alias lqid) + with Not_found -> error_global_not_found lqid let intern_ltac_variable ist = function - | Ident (loc,id) -> + | {loc;v=Ident id} -> if find_var id ist then (* A local variable of any type *) ArgVar (make ?loc id) @@ -109,19 +109,19 @@ let intern_ltac_variable ist = function raise Not_found let intern_constr_reference strict ist = function - | Ident (_,id) as r when not strict && find_hyp id ist -> + | {v=Ident id} as r when not strict && find_hyp id ist -> (DAst.make @@ GVar id), Some (make @@ CRef (r,None)) - | Ident (_,id) as r when find_var id ist -> + | {v=Ident id} as r when find_var id ist -> (DAst.make @@ GVar id), if strict then None else Some (make @@ CRef (r,None)) | r -> - let loc,_ as lqid = qualid_of_reference r in + let {loc} as lqid = qualid_of_reference r in DAst.make @@ GRef (locate_global_with_alias lqid,None), if strict then None else Some (make @@ CRef (r,None)) (* Internalize an isolated reference in position of tactic *) let intern_isolated_global_tactic_reference r = - let (loc,qid) = qualid_of_reference r in + let {loc;v=qid} = qualid_of_reference r in TacCall (Loc.tag ?loc (ArgArg (loc,Tacenv.locate_tactic qid),[])) let intern_isolated_tactic_reference strict ist r = @@ -135,12 +135,12 @@ let intern_isolated_tactic_reference strict ist r = try ConstrMayEval (ConstrTerm (intern_constr_reference strict ist r)) with Not_found -> (* Reference not found *) - error_global_not_found (snd (qualid_of_reference r)) + error_global_not_found (qualid_of_reference r) (* Internalize an applied tactic reference *) let intern_applied_global_tactic_reference r = - let (loc,qid) = qualid_of_reference r in + let {loc;v=qid} = qualid_of_reference r in ArgArg (loc,Tacenv.locate_tactic qid) let intern_applied_tactic_reference ist r = @@ -151,7 +151,7 @@ let intern_applied_tactic_reference ist r = try intern_applied_global_tactic_reference r with Not_found -> (* Reference not found *) - error_global_not_found (snd (qualid_of_reference r)) + error_global_not_found (qualid_of_reference r) (* Intern a reference parsed in a non-tactic entry *) @@ -167,12 +167,12 @@ let intern_non_tactic_reference strict ist r = with Not_found -> (* By convention, use IntroIdentifier for unbound ident, when not in a def *) match r with - | Ident (loc,id) when not strict -> + | {loc;v=Ident id} when not strict -> let ipat = in_gen (glbwit wit_intro_pattern) (make ?loc @@ IntroNaming (IntroIdentifier id)) in TacGeneric ipat | _ -> (* Reference not found *) - error_global_not_found (snd (qualid_of_reference r)) + error_global_not_found (qualid_of_reference r) let intern_message_token ist = function | (MsgString _ | MsgInt _ as x) -> x @@ -268,7 +268,7 @@ let intern_destruction_arg ist = function | clear,ElimOnIdent {loc;v=id} -> if !strict_check then (* If in a defined tactic, no intros-until *) - let c, p = intern_constr ist (make @@ CRef (Ident (Loc.tag id), None)) in + let c, p = intern_constr ist (make @@ CRef (make @@ Ident id, None)) in match DAst.get c with | GVar id -> clear,ElimOnIdent (make ?loc:c.loc id) | _ -> clear,ElimOnConstr ((c, p), NoBindings) @@ -276,7 +276,7 @@ let intern_destruction_arg ist = function clear,ElimOnIdent (make ?loc id) let short_name = function - | AN (Ident (loc,id)) when not !strict_check -> Some (make ?loc id) + | {v=AN {loc;v=Ident id}} when not !strict_check -> Some (make ?loc id) | _ -> None let intern_evaluable_global_reference ist r = @@ -284,20 +284,20 @@ let intern_evaluable_global_reference ist r = try evaluable_of_global_reference ist.genv (locate_global_with_alias ~head:true lqid) with Not_found -> match r with - | Ident (loc,id) when not !strict_check -> EvalVarRef id - | _ -> error_global_not_found (snd lqid) + | {loc;v=Ident id} when not !strict_check -> EvalVarRef id + | _ -> error_global_not_found lqid let intern_evaluable_reference_or_by_notation ist = function - | AN r -> intern_evaluable_global_reference ist r - | ByNotation {loc;v=(ntn,sc)} -> + | {v=AN r} -> intern_evaluable_global_reference ist r + | {v=ByNotation (ntn,sc);loc} -> evaluable_of_global_reference ist.genv (Notation.interp_notation_as_global_reference ?loc (function ConstRef _ | VarRef _ -> true | _ -> false) ntn sc) (* Globalize a reduction expression *) let intern_evaluable ist = function - | AN (Ident (loc,id)) when find_var id ist -> ArgVar (make ?loc id) - | AN (Ident (loc,id)) when not !strict_check && find_hyp id ist -> + | {loc;v=AN {v=Ident id}} when find_var id ist -> ArgVar (make ?loc id) + | {loc;v=AN {v=Ident id}} when not !strict_check && find_hyp id ist -> ArgArg (EvalVarRef id, Some (make ?loc id)) | r -> let e = intern_evaluable_reference_or_by_notation ist r in @@ -353,10 +353,9 @@ let intern_typed_pattern_or_ref_with_occurrences ist (l,p) = ref or a pattern seems interesting, with "head" reduction in case of an evaluable ref, and "strong" reduction in the subterm matched when a pattern *) - let loc = loc_of_smart_reference r in let r = match r with - | AN r -> r - | _ -> Qualid (loc,qualid_of_path (path_of_global (smart_global r))) in + | {v=AN r} -> r + | {loc} -> make ?loc @@ Qualid (qualid_of_path (path_of_global (smart_global r))) in let sign = { Constrintern.ltac_vars = ist.ltacvars; ltac_bound = Id.Set.empty; @@ -376,7 +375,7 @@ let intern_typed_pattern_or_ref_with_occurrences ist (l,p) = | Inl r -> interp_ref r | Inr { v = CAppExpl((None,r,None),[]) } -> (* We interpret similarly @ref and ref *) - interp_ref (AN r) + interp_ref (make @@ AN r) | Inr c -> Inr (snd (intern_typed_pattern ist ~as_type:false ~ltacvars:ist.ltacvars c))) @@ -385,13 +384,13 @@ let intern_typed_pattern_or_ref_with_occurrences ist (l,p) = let dump_glob_red_expr = function | Unfold occs -> List.iter (fun (_, r) -> try - Dumpglob.add_glob ?loc:(loc_of_or_by_notation Libnames.loc_of_reference r) + Dumpglob.add_glob ?loc:r.loc (Smartlocate.smart_global r) with e when CErrors.noncritical e -> ()) occs | Cbv grf | Lazy grf -> List.iter (fun r -> try - Dumpglob.add_glob ?loc:(loc_of_or_by_notation Libnames.loc_of_reference r) + Dumpglob.add_glob ?loc:r.loc (Smartlocate.smart_global r) with e when CErrors.noncritical e -> ()) grf.rConst | _ -> () diff --git a/plugins/ltac/tacinterp.ml b/plugins/ltac/tacinterp.ml index a287b13b9..6a4bf577b 100644 --- a/plugins/ltac/tacinterp.ml +++ b/plugins/ltac/tacinterp.ml @@ -360,7 +360,7 @@ let interp_reference ist env sigma = function with Not_found -> try VarRef (get_id (Environ.lookup_named id env)) - with Not_found -> error_global_not_found ?loc (qualid_of_ident id) + with Not_found -> error_global_not_found (make ?loc @@ qualid_of_ident id) let try_interp_evaluable env (loc, id) = let v = Environ.lookup_named id env in @@ -376,14 +376,14 @@ let interp_evaluable ist env sigma = function with Not_found -> match r with | EvalConstRef _ -> r - | _ -> error_global_not_found ?loc (qualid_of_ident id) + | _ -> error_global_not_found (make ?loc @@ qualid_of_ident id) end | ArgArg (r,None) -> r | ArgVar {loc;v=id} -> try try_interp_ltac_var (coerce_to_evaluable_ref env sigma) ist (Some (env,sigma)) (make ?loc id) with Not_found -> try try_interp_evaluable env (loc, id) - with Not_found -> error_global_not_found ?loc (qualid_of_ident id) + with Not_found -> error_global_not_found (make ?loc @@ qualid_of_ident id) (* Interprets an hypothesis name *) let interp_occurrences ist occs = @@ -642,7 +642,7 @@ let interp_closed_typed_pattern_with_occurrences ist env sigma (occs, a) = Inr (pattern_of_constr env sigma (EConstr.to_constr sigma c)) in (try try_interp_ltac_var coerce_eval_ref_or_constr ist (Some (env,sigma)) (make ?loc id) with Not_found -> - error_global_not_found ?loc (qualid_of_ident id)) + error_global_not_found (make ?loc @@ qualid_of_ident id)) | Inl (ArgArg _ as b) -> Inl (interp_evaluable ist env sigma b) | Inr c -> Inr (interp_typed_pattern ist env sigma c) in interp_occurrences ist occs, p @@ -926,7 +926,7 @@ let interp_destruction_arg ist gl arg = if Tactics.is_quantified_hypothesis id gl then keep,ElimOnIdent (make ?loc id) else - let c = (DAst.make ?loc @@ GVar id,Some (make @@ CRef (Ident (loc,id),None))) in + let c = (DAst.make ?loc @@ GVar id,Some (make @@ CRef (make ?loc @@ Ident id,None))) in let f env sigma = let (sigma,c) = interp_open_constr ist env sigma c in (sigma, (c,NoBindings)) -- cgit v1.2.3