aboutsummaryrefslogtreecommitdiff
path: root/interp
diff options
context:
space:
mode:
authorEnrico Tassi2021-03-19 12:27:14 +0100
committerEnrico Tassi2021-04-07 19:59:46 +0200
commitb47931125432df88171c7e8a879294508a603aa9 (patch)
tree64dca669a00da9d64f1ad687e9f18c65b69ec3c0 /interp
parent59d0462f35818c12a0727a560d7b9ecf2ceea994 (diff)
cleanup: less exceptions, removal of try_interp_name_alias
Diffstat (limited to 'interp')
-rw-r--r--interp/constrexpr_ops.ml11
-rw-r--r--interp/constrintern.ml92
-rw-r--r--interp/constrintern.mli11
-rw-r--r--interp/smartlocate.ml8
-rw-r--r--interp/smartlocate.mli5
5 files changed, 66 insertions, 61 deletions
diff --git a/interp/constrexpr_ops.ml b/interp/constrexpr_ops.ml
index f02874253e..e72b73c36e 100644
--- a/interp/constrexpr_ops.ml
+++ b/interp/constrexpr_ops.ml
@@ -283,10 +283,13 @@ let local_binders_loc bll = match bll with
(** Folds and maps *)
let is_constructor id =
- try Globnames.isConstructRef
- (Smartlocate.global_of_extended_global
- (Nametab.locate_extended (qualid_of_ident id)))
- with Not_found -> false
+ match
+ Smartlocate.global_of_extended_global
+ (Nametab.locate_extended (qualid_of_ident id))
+ with
+ | exception Not_found -> false
+ | None -> false
+ | Some gref -> Globnames.isConstructRef gref
let rec cases_pattern_fold_names f h nacc pt = match CAst.(pt.v) with
| CPatRecord l ->
diff --git a/interp/constrintern.ml b/interp/constrintern.ml
index 958e1408f8..3d4b6c40cb 100644
--- a/interp/constrintern.ml
+++ b/interp/constrintern.ml
@@ -1121,23 +1121,54 @@ let dump_extended_global loc = function
let intern_extended_global_of_qualid qid =
let r = Nametab.locate_extended qid in dump_extended_global qid.CAst.loc r; r
-let intern_reference qid =
- let r =
- try intern_extended_global_of_qualid qid
- with Not_found as exn ->
- let _, info = Exninfo.capture exn in
- Nametab.error_global_not_found ~info qid
- in
- Smartlocate.global_of_extended_global r
+let intern_sort_name ~local_univs = function
+ | CSProp -> GSProp
+ | CProp -> GProp
+ | CSet -> GSet
+ | CRawType u -> GRawUniv u
+ | CType qid ->
+ let is_id = qualid_is_ident qid in
+ let local = if not is_id then None
+ else Id.Map.find_opt (qualid_basename qid) local_univs.bound
+ in
+ match local with
+ | Some u -> GUniv u
+ | None ->
+ try GUniv (Univ.Level.make (Nametab.locate_universe qid))
+ with Not_found ->
+ if is_id && local_univs.unb_univs
+ then GLocalUniv (CAst.make ?loc:qid.loc (qualid_basename qid))
+ else
+ CErrors.user_err Pp.(str "Undeclared universe " ++ pr_qualid qid ++ str".")
+
+let intern_sort ~local_univs s =
+ map_glob_sort_gen (List.map (on_fst (intern_sort_name ~local_univs))) s
+
+let intern_instance ~local_univs us =
+ Option.map (List.map (map_glob_sort_gen (intern_sort_name ~local_univs))) us
+
+let intern_name_alias = function
+ | { CAst.v = CRef(qid,u) } ->
+ let r =
+ try Some (intern_extended_global_of_qualid qid)
+ with Not_found -> None
+ in
+ Option.bind r Smartlocate.global_of_extended_global |>
+ Option.map (fun r -> r, intern_instance ~local_univs:empty_local_univs u)
+ | _ -> None
let intern_projection qid =
- try
- match Smartlocate.global_of_extended_global (intern_extended_global_of_qualid qid) with
- | GlobRef.ConstRef c as gr ->
- (gr, Structure.find_from_projection c)
- | _ -> raise Not_found
- with Not_found ->
- Loc.raise ?loc:qid.loc (InternalizationError (NotAProjection qid))
+ match
+ Smartlocate.global_of_extended_global (intern_extended_global_of_qualid qid) |>
+ Option.map (function
+ | GlobRef.ConstRef c as x -> x, Structure.find_from_projection c
+ | _ -> raise Not_found)
+ with
+ | exception Not_found ->
+ Loc.raise ?loc:qid.loc (InternalizationError (NotAProjection qid))
+ | None ->
+ Loc.raise ?loc:qid.loc (InternalizationError (NotAProjection qid))
+ | Some x -> x
(**********************************************************************)
(* Interpreting references *)
@@ -1182,37 +1213,6 @@ let glob_sort_of_level (level: glob_level) : glob_sort =
| UAnonymous {rigid} -> UAnonymous {rigid}
| UNamed id -> UNamed [id,0]
-let intern_sort_name ~local_univs = function
- | CSProp -> GSProp
- | CProp -> GProp
- | CSet -> GSet
- | CRawType u -> GRawUniv u
- | CType qid ->
- let is_id = qualid_is_ident qid in
- let local = if not is_id then None
- else Id.Map.find_opt (qualid_basename qid) local_univs.bound
- in
- match local with
- | Some u -> GUniv u
- | None ->
- try GUniv (Univ.Level.make (Nametab.locate_universe qid))
- with Not_found ->
- if is_id && local_univs.unb_univs
- then GLocalUniv (CAst.make ?loc:qid.loc (qualid_basename qid))
- else
- CErrors.user_err Pp.(str "Undeclared universe " ++ pr_qualid qid ++ str".")
-
-let intern_sort ~local_univs s =
- map_glob_sort_gen (List.map (on_fst (intern_sort_name ~local_univs))) s
-
-let intern_instance ~local_univs us =
- Option.map (List.map (map_glob_sort_gen (intern_sort_name ~local_univs))) us
-
-let try_interp_name_alias = function
- | [], { CAst.v = CRef (ref,u) } ->
- NRef (intern_reference ref,intern_instance ~local_univs:empty_local_univs u)
- | _ -> raise Not_found
-
(* Is it a global reference or a syntactic definition? *)
let intern_qualid ?(no_secvar=false) qid intern env ntnvars us args =
let loc = qid.loc in
diff --git a/interp/constrintern.mli b/interp/constrintern.mli
index 65b63962d0..7c1e658ff1 100644
--- a/interp/constrintern.mli
+++ b/interp/constrintern.mli
@@ -147,12 +147,9 @@ val interp_constr_pattern :
env -> evar_map -> ?expected_type:typing_constraint ->
constr_pattern_expr -> constr_pattern
-(** Raise Not_found if syndef not bound to a name and error if unexisting ref *)
-val intern_reference : qualid -> GlobRef.t
-
-(** For syntactic definitions: check if abbreviation to a name
- and avoid early insertion of maximal implicit arguments *)
-val try_interp_name_alias : 'a list * constr_expr -> notation_constr
+(** Returns None if not a reference or a syndef not bound to a name *)
+val intern_name_alias :
+ constr_expr -> (GlobRef.t * Glob_term.glob_level list option) option
(** Expands abbreviations (syndef); raise an error if not existing *)
val interp_reference : ltac_sign -> qualid -> glob_constr
@@ -174,7 +171,7 @@ val interp_context_evars :
(** Locating references of constructions, possibly via a syntactic definition
(these functions do not modify the glob file) *)
-val locate_reference : Libnames.qualid -> GlobRef.t
+val locate_reference : Libnames.qualid -> GlobRef.t option
val is_global : Id.t -> bool
(** Interprets a term as the left-hand side of a notation. The returned map is
diff --git a/interp/smartlocate.ml b/interp/smartlocate.ml
index 91d05f7317..56b3cd9815 100644
--- a/interp/smartlocate.ml
+++ b/interp/smartlocate.ml
@@ -33,7 +33,7 @@ let global_of_extended_global_head = function
| _ -> raise Not_found in
head_of syn_def
-let global_of_extended_global = function
+let global_of_extended_global_exn = function
| TrueGlobal ref -> ref
| SynDef kn ->
match search_syntactic_definition kn with
@@ -45,11 +45,15 @@ let locate_global_with_alias ?(head=false) qid =
let ref = Nametab.locate_extended qid in
try
if head then global_of_extended_global_head ref
- else global_of_extended_global ref
+ else global_of_extended_global_exn ref
with Not_found ->
user_err ?loc:qid.CAst.loc (pr_qualid qid ++
str " is bound to a notation that does not denote a reference.")
+let global_of_extended_global x =
+ try Some (global_of_extended_global_exn x)
+ with Not_found -> None
+
let global_constant_with_alias qid =
try match locate_global_with_alias qid with
| Names.GlobRef.ConstRef c -> c
diff --git a/interp/smartlocate.mli b/interp/smartlocate.mli
index 26f2a4f36d..abf9839c9e 100644
--- a/interp/smartlocate.mli
+++ b/interp/smartlocate.mli
@@ -19,8 +19,9 @@ open Globnames
val locate_global_with_alias : ?head:bool -> qualid -> GlobRef.t
-(** Extract a global_reference from a reference that can be an "alias" *)
-val global_of_extended_global : extended_global_reference -> GlobRef.t
+(** Extract a global_reference from a reference that can be an "alias".
+ If the reference points to a more complex term, we return None *)
+val global_of_extended_global : extended_global_reference -> GlobRef.t option
(** Locate a reference taking into account possible "alias" notations.
May raise [Nametab.GlobalizationError _] for an unknown reference,