aboutsummaryrefslogtreecommitdiff
diff options
context:
space:
mode:
authorEmilio Jesus Gallego Arias2018-10-08 22:27:45 +0200
committerEmilio Jesus Gallego Arias2018-10-08 22:27:45 +0200
commitf3c123ea05660ce393260283d64cc3ff0eba2f8c (patch)
treef7901e17eb87f1225df970c184773952c9085231
parent754d9662b9f7619542e77ad85f575d227fb42b2f (diff)
parentab237af7e952281e3fe695d888242cdf2abcaa90 (diff)
Merge PR #8676: Fixes #8672: ill-formed pattern substitution in notation with "let"
-rw-r--r--interp/notation_ops.ml8
-rw-r--r--test-suite/bugs/closed/8672.v5
2 files changed, 10 insertions, 3 deletions
diff --git a/interp/notation_ops.ml b/interp/notation_ops.ml
index ff5e2bb5f3..ab57176643 100644
--- a/interp/notation_ops.ml
+++ b/interp/notation_ops.ml
@@ -153,11 +153,13 @@ let protect g e na =
if disjpat <> None then user_err (Pp.str "Unsupported substitution of an arbitrary pattern.");
e',na
-let apply_cases_pattern ?loc ((ids,disjpat),id) c =
- let tm = DAst.make ?loc (GVar id) in
+let apply_cases_pattern_term ?loc (ids,disjpat) tm c =
let eqns = List.map (fun pat -> (CAst.make ?loc (ids,[pat],c))) disjpat in
DAst.make ?loc @@ GCases (Constr.LetPatternStyle, None, [tm,(Anonymous,None)], eqns)
+let apply_cases_pattern ?loc (ids_disjpat,id) c =
+ apply_cases_pattern_term ?loc ids_disjpat (DAst.make ?loc (GVar id)) c
+
let glob_constr_of_notation_constr_with_binders ?loc g f e nc =
let lt x = DAst.make ?loc x in lt @@ match nc with
| NVar id -> GVar id
@@ -182,7 +184,7 @@ let glob_constr_of_notation_constr_with_binders ?loc g f e nc =
let e',disjpat,na = g e na in
(match disjpat with
| None -> GLetIn (na,f e b,Option.map (f e) t,f e' c)
- | Some disjpat -> DAst.get (apply_cases_pattern ?loc disjpat (f e' c)))
+ | Some (disjpat,_id) -> DAst.get (apply_cases_pattern_term ?loc disjpat (f e b) (f e' c)))
| NCases (sty,rtntypopt,tml,eqnl) ->
let e',tml' = List.fold_right (fun (tm,(na,t)) (e',tml') ->
let e',t' = match t with
diff --git a/test-suite/bugs/closed/8672.v b/test-suite/bugs/closed/8672.v
new file mode 100644
index 0000000000..66cd6dfa8c
--- /dev/null
+++ b/test-suite/bugs/closed/8672.v
@@ -0,0 +1,5 @@
+(* Was generating a dangling "pat" variable at some time *)
+
+Notation "'plet' x := e 'in' t" :=
+ ((fun H => let x := id H in t) e) (at level 0, x pattern).
+Definition bla := plet (pair x y) := pair 1 2 in x.