aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
authorGravatar Hugo Herbelin <Hugo.Herbelin@inria.fr>2017-08-12 13:32:05 +0200
committerGravatar Hugo Herbelin <Hugo.Herbelin@inria.fr>2017-09-12 20:27:14 +0200
commita267b06c4a237ce3e4ce257e47f6429543804baa (patch)
treef8e84c33cbf8935d19c3e167bb4dbf9cd52e6a16
parentcc94172036789cfef28007f59510b7f17df5d45d (diff)
Fixing little inaccuracy in coercions to ident or name.
For instance, we don't want "id@{u}" to be coerced to id, or "?[n]" to "_".
-rw-r--r--interp/constrexpr_ops.ml6
-rw-r--r--test-suite/success/Notations.v5
2 files changed, 8 insertions, 3 deletions
diff --git a/interp/constrexpr_ops.ml b/interp/constrexpr_ops.ml
index 2d0a19b9a..771c13734 100644
--- a/interp/constrexpr_ops.ml
+++ b/interp/constrexpr_ops.ml
@@ -320,13 +320,13 @@ let coerce_reference_to_id = function
(str "This expression should be a simple identifier.")
let coerce_to_id = function
- | { CAst.v = CRef (Ident (loc,id),_); _ } -> (loc,id)
+ | { CAst.v = CRef (Ident (loc,id),None) } -> (loc,id)
| { CAst.loc; _ } -> CErrors.user_err ?loc
~hdr:"coerce_to_id"
(str "This expression should be a simple identifier.")
let coerce_to_name = function
- | { CAst.v = CRef (Ident (loc,id),_) } -> (loc,Name id)
- | { CAst.loc; CAst.v = CHole (_,_,_) } -> (loc,Anonymous)
+ | { CAst.v = CRef (Ident (loc,id),None) } -> (loc,Name id)
+ | { CAst.loc; CAst.v = CHole (None,Misctypes.IntroAnonymous,None) } -> (loc,Anonymous)
| { CAst.loc; _ } -> CErrors.user_err ?loc ~hdr:"coerce_to_name"
(str "This expression should be a name.")
diff --git a/test-suite/success/Notations.v b/test-suite/success/Notations.v
index 837f2efd0..4d04f2cf9 100644
--- a/test-suite/success/Notations.v
+++ b/test-suite/success/Notations.v
@@ -142,3 +142,8 @@ Fail Notation "'foobarkeyword'" := (@nil) (only parsing, only printing).
Reserved Notation "x === y" (at level 50).
Inductive EQ {A} (x:A) : A -> Prop := REFL : x === x
where "x === y" := (EQ x y).
+
+(* Check that strictly ident or _ are coerced to a name *)
+
+Fail Check {x@{u},y|x=x}.
+Fail Check {?[n],y|0=0}.