aboutsummaryrefslogtreecommitdiff
diff options
context:
space:
mode:
-rw-r--r--doc/changelog/03-notations/12946-master+fix12908-part1-collision-lonely-notation-printing.rst6
-rw-r--r--interp/notation.ml21
-rw-r--r--test-suite/output/bug_12908.out2
-rw-r--r--test-suite/output/bug_12908.v6
4 files changed, 32 insertions, 3 deletions
diff --git a/doc/changelog/03-notations/12946-master+fix12908-part1-collision-lonely-notation-printing.rst b/doc/changelog/03-notations/12946-master+fix12908-part1-collision-lonely-notation-printing.rst
new file mode 100644
index 0000000000..95a9093272
--- /dev/null
+++ b/doc/changelog/03-notations/12946-master+fix12908-part1-collision-lonely-notation-printing.rst
@@ -0,0 +1,6 @@
+- **Fixed:**
+ Undetected collision between a lonely notation and a notation in
+ scope at printing time
+ (`#12946 <https://github.com/coq/coq/pull/12946>`_,
+ fixes the first part of `#12908 <https://github.com/coq/coq/issues/12908>`_,
+ by Hugo Herbelin).
diff --git a/interp/notation.ml b/interp/notation.ml
index 17ae045187..7e90e15b72 100644
--- a/interp/notation.ml
+++ b/interp/notation.ml
@@ -1198,10 +1198,25 @@ let rec find_without_delimiters find (ntn_scope,ntn) = function
find_without_delimiters find (ntn_scope,ntn) scopes
end
| LonelyNotationItem ntn' :: scopes ->
- begin match ntn_scope, ntn with
- | LastLonelyNotation, Some ntn when notation_eq ntn ntn' ->
- Some (None, None)
+ begin match ntn with
+ | Some ntn'' when notation_eq ntn' ntn'' ->
+ begin match ntn_scope with
+ | LastLonelyNotation ->
+ (* If the first notation with same string in the visibility stack
+ is the one we want to print, then it can be used without
+ risking a collision *)
+ Some (None, None)
+ | NotationInScope _ ->
+ (* A lonely notation is liable to hide the scoped notation
+ to print, we check if the lonely notation is active to
+ know if the delimiter of the scoped notationis needed *)
+ if find default_scope then
+ find_with_delimiters ntn_scope
+ else
+ find_without_delimiters find (ntn_scope,ntn) scopes
+ end
| _ ->
+ (* A lonely notation which does not interfere with the notation to use *)
find_without_delimiters find (ntn_scope,ntn) scopes
end
| [] ->
diff --git a/test-suite/output/bug_12908.out b/test-suite/output/bug_12908.out
new file mode 100644
index 0000000000..fca6dde704
--- /dev/null
+++ b/test-suite/output/bug_12908.out
@@ -0,0 +1,2 @@
+forall m n : nat, m * n = (2 * m * n)%nat
+ : Prop
diff --git a/test-suite/output/bug_12908.v b/test-suite/output/bug_12908.v
new file mode 100644
index 0000000000..558c9f9f6a
--- /dev/null
+++ b/test-suite/output/bug_12908.v
@@ -0,0 +1,6 @@
+Definition mult' m n := 2 * m * n.
+Module A.
+(* Test hiding of a scoped notation by a lonely notation *)
+Infix "*" := mult'.
+Check forall m n, mult' m n = Nat.mul (Nat.mul 2 m) n.
+End A.