aboutsummaryrefslogtreecommitdiff
diff options
context:
space:
mode:
-rw-r--r--interp/constrintern.ml6
-rw-r--r--test-suite/bugs/closed/8126.v13
-rw-r--r--test-suite/output/Notations3.out6
-rw-r--r--test-suite/output/Notations3.v11
4 files changed, 33 insertions, 3 deletions
diff --git a/interp/constrintern.ml b/interp/constrintern.ml
index c87768b190..715823e5d0 100644
--- a/interp/constrintern.ml
+++ b/interp/constrintern.ml
@@ -1891,9 +1891,9 @@ let internalize globalenv env pattern_mode (_, ntnvars as lvar) c =
intern_applied_reference intern env (Environ.named_context globalenv)
lvar us args ref
in
- (* Rem: GApp(_,f,[]) stands for @f *)
- DAst.make ?loc @@
- GApp (f, intern_args env args_scopes (List.map fst args))
+ (* Rem: GApp(_,f,[]) stands for @f *)
+ if args = [] then DAst.make ?loc @@ GApp (f,[]) else
+ smart_gapp f loc (intern_args env args_scopes (List.map fst args))
| CApp ((isproj,f), args) ->
let f,args = match f.CAst.v with
diff --git a/test-suite/bugs/closed/8126.v b/test-suite/bugs/closed/8126.v
new file mode 100644
index 0000000000..f52dfc6b47
--- /dev/null
+++ b/test-suite/bugs/closed/8126.v
@@ -0,0 +1,13 @@
+(* See also output test Notations4.v *)
+
+Inductive foo := tt.
+Bind Scope foo_scope with foo.
+Delimit Scope foo_scope with foo.
+Notation "'HI'" := tt : foo_scope.
+Definition myfoo (x : nat) (y : nat) (z : foo) := y.
+Notation myfoo0 := (@myfoo 0).
+Notation myfoo01 := (@myfoo0 1).
+Check myfoo 0 1 HI. (* prints [myfoo0 1 HI], but should print [myfoo01 HI] *)
+Check myfoo0 1 HI. (* prints [myfoo0 1 HI], but should print [myfoo01 HI] *)
+Check myfoo01 tt. (* prints [myfoo0 1 HI], but should print [myfoo01 HI] *)
+Check myfoo01 HI. (* was failing *)
diff --git a/test-suite/output/Notations3.out b/test-suite/output/Notations3.out
index 5ab616160a..d32cf67e28 100644
--- a/test-suite/output/Notations3.out
+++ b/test-suite/output/Notations3.out
@@ -246,3 +246,9 @@ Notation
============================
##@%
^^^
+myfoo01 tt
+ : nat
+myfoo01 tt
+ : nat
+myfoo01 tt
+ : nat
diff --git a/test-suite/output/Notations3.v b/test-suite/output/Notations3.v
index 876aaa3944..180e8d337e 100644
--- a/test-suite/output/Notations3.v
+++ b/test-suite/output/Notations3.v
@@ -399,3 +399,14 @@ Show.
Abort.
End Issue7731.
+
+Module Issue8126.
+
+Definition myfoo (x : nat) (y : nat) (z : unit) := y.
+Notation myfoo0 := (@myfoo 0).
+Notation myfoo01 := (@myfoo0 1).
+Check myfoo 0 1 tt. (* was printing [myfoo0 1 HI], but should print [myfoo01 HI] *)
+Check myfoo0 1 tt. (* was printing [myfoo0 1 HI], but should print [myfoo01 HI] *)
+Check myfoo01 tt. (* was printing [myfoo0 1 HI], but should print [myfoo01 HI] *)
+
+End Issue8126.