aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
authorGravatar Hugo Herbelin <Hugo.Herbelin@inria.fr>2017-11-03 11:48:36 +0100
committerGravatar Hugo Herbelin <Hugo.Herbelin@inria.fr>2018-02-20 10:03:08 +0100
commit9fe03a830ae7b10e1c5fb2b9d41213257bd617f0 (patch)
tree677d8b51c334389fa1a6bcbc42e8d022db93490b
parent0c4eea2553d5b3b70d0b5baaf92781a544de83bd (diff)
Adding notations of the form "{'pat|P}", "exists2 'pat, P & Q", etc.
-rw-r--r--test-suite/output/Notations3.out12
-rw-r--r--test-suite/output/Notations3.v27
-rw-r--r--theories/Init/Logic.v7
-rw-r--r--theories/Init/Specif.v26
4 files changed, 68 insertions, 4 deletions
diff --git a/test-suite/output/Notations3.out b/test-suite/output/Notations3.out
index 436e97e9f..e6a6e0288 100644
--- a/test-suite/output/Notations3.out
+++ b/test-suite/output/Notations3.out
@@ -194,9 +194,11 @@ pair
(prod nat (prod nat nat))) (prod (prod nat nat) nat)
fun x : nat => if x is n .+ 1 then n else 1
: nat -> nat
+{'{{x, y}} : nat * nat | x + y = 0}
+ : Set
exists2' {{x, y}}, x = 0 & y = 0
: Prop
-exists2 x : nat * nat,
+myexists2 x : nat * nat,
let '{{y, z}} := x in y > z & let '{{y, z}} := x in z > y
: Prop
fun '({{x, y}} as z) => x + y = 0 /\ z = z
@@ -225,3 +227,11 @@ fun N : nat => [N | N + 0]
: nat -> nat * (nat -> nat)
fun S : nat => [[S | S + S]]
: nat -> nat * (nat -> nat)
+{I : nat | I = I}
+ : Set
+{'I : True | I = I}
+ : Prop
+{'{{x, y}} : nat * nat | x + y = 0}
+ : Set
+exists2 '{{y, z}} : nat * nat, y > z & z > y
+ : Prop
diff --git a/test-suite/output/Notations3.v b/test-suite/output/Notations3.v
index df3d86793..f41ddac4d 100644
--- a/test-suite/output/Notations3.v
+++ b/test-suite/output/Notations3.v
@@ -341,6 +341,10 @@ 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.
Notation "'exists2'' x , p & q" := (ex2 (fun x => p) (fun x => q))
(at level 200, x pattern, p at level 200, right associativity,
@@ -353,7 +357,15 @@ End D.
(* Ensuring for reparsability that printer of notations does not use a
pattern where only an ident could be reparsed *)
-Check ex2 (fun x => let '(y,z) := x in y>z) (fun x => let '(y,z) := x in z>y).
+Module E.
+Inductive myex2 {A:Type} (P Q:A -> Prop) : Prop :=
+ myex_intro2 : forall x:A, P x -> Q x -> myex2 P Q.
+Notation "'myexists2' x : A , p & q" := (myex2 (A:=A) (fun x => p) (fun x => q))
+ (at level 200, x ident, A at level 200, p at level 200, right associativity,
+ format "'[' 'myexists2' '/ ' x : A , '/ ' '[' p & '/' q ']' ']'")
+ : type_scope.
+Check myex2 (fun x => let '(y,z) := x in y>z) (fun x => let '(y,z) := x in z>y).
+End E.
(* A canonical example of a notation with a non-recursive binder *)
@@ -391,7 +403,7 @@ Check fun p => if p is S n then n else 0.
Check fun p => if p is Lt then 1 else 0.
(* Check that mixed binders and terms defaults to ident and not pattern *)
-Module E.
+Module F.
(* First without an indirection *)
Notation "[ n | t ]" := (n, (fun n : nat => t)).
Check fun S : nat => [ S | S+S ].
@@ -400,4 +412,13 @@ Check fun N : nat => (N, (fun n => n+0)). (* another test in passing *)
Notation "[[ n | p | t ]]" := (n, (fun p : nat => t)).
Notation "[[ n | t ]]" := [[ n | n | t ]].
Check fun S : nat => [[ S | S+S ]].
-End E.
+End F.
+
+(* Check parsability/printability of {x|P} and variants *)
+
+Check {I:nat|I=I}.
+Check {'I:True|I=I}.
+Check {'(x,y)|x+y=0}.
+
+(* Check exists2 with a pattern *)
+Check ex2 (fun x => let '(y,z) := x in y>z) (fun x => let '(y,z) := x in z>y).
diff --git a/theories/Init/Logic.v b/theories/Init/Logic.v
index 037d37daf..053ed601f 100644
--- a/theories/Init/Logic.v
+++ b/theories/Init/Logic.v
@@ -267,6 +267,13 @@ Notation "'exists2' x : A , p & q" := (ex2 (A:=A) (fun x => p) (fun x => q))
format "'[' 'exists2' '/ ' x : A , '/ ' '[' p & '/' q ']' ']'")
: type_scope.
+Notation "'exists2' ' x , p & q" := (ex2 (fun x => p) (fun x => q))
+ (at level 200, x strict pattern, p at level 200, right associativity) : type_scope.
+Notation "'exists2' ' x : A , p & q" := (ex2 (A:=A) (fun x => p) (fun x => q))
+ (at level 200, x strict pattern, A at level 200, p at level 200, right associativity,
+ format "'[' 'exists2' '/ ' ' x : A , '/ ' '[' p & '/' q ']' ']'")
+ : type_scope.
+
(** Derived rules for universal quantification *)
Section universal_quantification.
diff --git a/theories/Init/Specif.v b/theories/Init/Specif.v
index 3b4f833a3..ef95aef50 100644
--- a/theories/Init/Specif.v
+++ b/theories/Init/Specif.v
@@ -53,6 +53,32 @@ 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.
+Notation "{ ' pat : A | P & Q }" := (sig2 (A:=A) (fun pat => P) (fun pat => Q)) :
+ type_scope.
+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.