aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
-rw-r--r--parsing/g_constr.ml42
-rw-r--r--test-suite/output/Notations3.v2
-rw-r--r--theories/Init/Notations.v15
-rw-r--r--theories/Init/Specif.v17
4 files changed, 17 insertions, 19 deletions
diff --git a/parsing/g_constr.ml4 b/parsing/g_constr.ml4
index c2ce003f1..9f12db649 100644
--- a/parsing/g_constr.ml4
+++ b/parsing/g_constr.ml4
@@ -219,6 +219,8 @@ GEXTEND Gram
CAst.make ~loc:(!@loc) @@ CNotation("( _ )",([c],[],[],[]))
| _ -> c)
| "{|"; c = record_declaration; "|}" -> c
+ | "{"; c = binder_constr ; "}" ->
+ CAst.make ~loc:(!@loc) @@ CNotation(("{ _ }"),([c],[],[],[]))
| "`{"; c = operconstr LEVEL "200"; "}" ->
CAst.make ~loc:(!@loc) @@ CGeneralization (Implicit, None, c)
| "`("; c = operconstr LEVEL "200"; ")" ->
diff --git a/test-suite/output/Notations3.v b/test-suite/output/Notations3.v
index f41ddac4d..c98bfff41 100644
--- a/test-suite/output/Notations3.v
+++ b/test-suite/output/Notations3.v
@@ -341,8 +341,6 @@ Check fun x => if x is n.+1 then n else 1.
(* Examples with binding patterns *)
-Import SpecifPatternNotations.
-
Check {'(x,y)|x+y=0}.
Module D.
diff --git a/theories/Init/Notations.v b/theories/Init/Notations.v
index 0c212b025..a9051e761 100644
--- a/theories/Init/Notations.v
+++ b/theories/Init/Notations.v
@@ -78,6 +78,21 @@ Reserved Notation "{ x : A | P & Q }" (at level 0, x at level 99).
Reserved Notation "{ x : A & P }" (at level 0, x at level 99).
Reserved Notation "{ x : A & P & Q }" (at level 0, x at level 99).
+Reserved Notation "{ ' pat | P }"
+ (at level 0, pat strict pattern, format "{ ' pat | P }").
+Reserved Notation "{ ' pat | P & Q }"
+ (at level 0, pat strict pattern, format "{ ' pat | P & Q }").
+
+Reserved Notation "{ ' pat : A | P }"
+ (at level 0, pat strict pattern, format "{ ' pat : A | P }").
+Reserved Notation "{ ' pat : A | P & Q }"
+ (at level 0, pat strict pattern, format "{ ' pat : A | P & Q }").
+
+Reserved Notation "{ ' pat : A & P }"
+ (at level 0, pat strict pattern, format "{ ' pat : A & P }").
+Reserved Notation "{ ' pat : A & P & Q }"
+ (at level 0, pat strict pattern, format "{ ' pat : A & P & Q }").
+
(** Support for Gonthier-Ssreflect's "if c is pat then u else v" *)
Module IfNotations.
diff --git a/theories/Init/Specif.v b/theories/Init/Specif.v
index ef95aef50..47e8a7558 100644
--- a/theories/Init/Specif.v
+++ b/theories/Init/Specif.v
@@ -53,21 +53,6 @@ Notation "{ x : A & P }" := (sigT (A:=A) (fun x => P)) : type_scope.
Notation "{ x : A & P & Q }" := (sigT2 (A:=A) (fun x => P) (fun x => Q)) :
type_scope.
-Module SpecifPatternNotations.
-
-Reserved Notation "{ ' pat | P }"
- (at level 0, pat strict pattern, format "{ ' pat | P }").
-Reserved Notation "{ ' pat | P & Q }"
- (at level 0, pat strict pattern, format "{ ' pat | P & Q }").
-Reserved Notation "{ ' pat : A | P }"
- (at level 0, pat strict pattern, format "{ ' pat : A | P }").
-Reserved Notation "{ ' pat : A | P & Q }"
- (at level 0, pat strict pattern, format "{ ' pat : A | P & Q }").
-Reserved Notation "{ ' pat : A & P }"
- (at level 0, pat strict pattern, format "{ ' pat : A & P }").
-Reserved Notation "{ ' pat : A & P & Q }"
- (at level 0, pat strict pattern, format "{ ' pat : A & P & Q }").
-
Notation "{ ' pat | P }" := (sig (fun pat => P)) : type_scope.
Notation "{ ' pat | P & Q }" := (sig2 (fun pat => P) (fun pat => Q)) : type_scope.
Notation "{ ' pat : A | P }" := (sig (A:=A) (fun pat => P)) : type_scope.
@@ -77,8 +62,6 @@ Notation "{ ' pat : A & P }" := (sigT (A:=A) (fun pat => P)) : type_scope.
Notation "{ ' pat : A & P & Q }" := (sigT2 (A:=A) (fun pat => P) (fun pat => Q)) :
type_scope.
-End SpecifPatternNotations.
-
Add Printing Let sig.
Add Printing Let sig2.
Add Printing Let sigT.