aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
-rw-r--r--theories/Init/Logic.v20
-rw-r--r--theories/Lists/List.v4
-rw-r--r--theories/Lists/SetoidList.v3
-rw-r--r--theories/Logic/JMeq.v3
-rw-r--r--theories/MSets/MSetAVL.v2
-rw-r--r--theories/MSets/MSetList.v3
-rw-r--r--theories/Program/Syntax.v9
-rw-r--r--theories/Unicode/Utf8.v37
8 files changed, 22 insertions, 59 deletions
diff --git a/theories/Init/Logic.v b/theories/Init/Logic.v
index 14832e0b7..97f3b126c 100644
--- a/theories/Init/Logic.v
+++ b/theories/Init/Logic.v
@@ -216,11 +216,9 @@ Definition all (A:Type) (P:A -> Prop) := forall x:A, P x.
(* Rule order is important to give printing priority to fully typed exists *)
-Notation "'exists' x , p" := (ex (fun x => p))
- (at level 200, x ident, right associativity) : type_scope.
-Notation "'exists' x : t , p" := (ex (fun x:t => p))
- (at level 200, x ident, right associativity,
- format "'[' 'exists' '/ ' x : t , '/ ' p ']'")
+Notation "'exists' x .. y , p" := (ex (fun x => .. (ex (fun y => p)) ..))
+ (at level 200, x binder, right associativity,
+ format "'[' 'exists' '/ ' x .. y , '/ ' p ']'")
: type_scope.
Notation "'exists2' x , p & q" := (ex2 (fun x => p) (fun x => q))
@@ -390,13 +388,11 @@ Definition uniqueness (A:Type) (P:A->Prop) := forall x y, P x -> P y -> x = y.
(** Unique existence *)
-Notation "'exists' ! x , P" := (ex (unique (fun x => P)))
- (at level 200, x ident, right associativity,
- format "'[' 'exists' ! '/ ' x , '/ ' P ']'") : type_scope.
-Notation "'exists' ! x : A , P" :=
- (ex (unique (fun x:A => P)))
- (at level 200, x ident, right associativity,
- format "'[' 'exists' ! '/ ' x : A , '/ ' P ']'") : type_scope.
+Notation "'exists' ! x .. y , p" :=
+ (ex (unique (fun x => .. (ex (unique (fun y => p))) ..)))
+ (at level 200, x binder, right associativity,
+ format "'[' 'exists' ! '/ ' x .. y , '/ ' p ']'")
+ : type_scope.
Lemma unique_existence : forall (A:Type) (P:A->Prop),
((exists x, P x) /\ uniqueness P) <-> (exists! x, P x).
diff --git a/theories/Lists/List.v b/theories/Lists/List.v
index d1e3f90b3..25f96292d 100644
--- a/theories/Lists/List.v
+++ b/theories/Lists/List.v
@@ -117,7 +117,7 @@ Section Facts.
unfold not; intros a H; inversion_clear H.
Qed.
- Theorem in_split : forall x (l:list A), In x l -> exists l1, exists l2, l = l1++x::l2.
+ Theorem in_split : forall x (l:list A), In x l -> exists l1 l2, l = l1++x::l2.
Proof.
induction l; simpl; destruct 1.
subst a; auto.
@@ -1641,7 +1641,7 @@ Proof. exact Forall2_nil. Qed.
Theorem Forall2_app_inv_l : forall A B (R:A->B->Prop) l1 l2 l',
Forall2 R (l1 ++ l2) l' ->
- exists l1', exists l2', Forall2 R l1 l1' /\ Forall2 R l2 l2' /\ l' = l1' ++ l2'.
+ exists l1' l2', Forall2 R l1 l1' /\ Forall2 R l2 l2' /\ l' = l1' ++ l2'.
Proof.
induction l1; intros.
exists [], l'; auto.
diff --git a/theories/Lists/SetoidList.v b/theories/Lists/SetoidList.v
index 88ccdb126..43a965cf1 100644
--- a/theories/Lists/SetoidList.v
+++ b/theories/Lists/SetoidList.v
@@ -157,8 +157,7 @@ Qed.
Hint Resolve In_InA.
Lemma InA_split : forall l x, InA x l ->
- exists l1, exists y, exists l2,
- eqA x y /\ l = l1++y::l2.
+ exists l1 y l2, eqA x y /\ l = l1++y::l2.
Proof.
induction l; intros; inv.
exists (@nil A); exists a; exists l; auto.
diff --git a/theories/Logic/JMeq.v b/theories/Logic/JMeq.v
index f17e403ed..330ee79ef 100644
--- a/theories/Logic/JMeq.v
+++ b/theories/Logic/JMeq.v
@@ -111,8 +111,7 @@ apply JMeq_refl.
Qed.
Lemma eq_dep_strictly_stronger_JMeq :
- exists U, exists P, exists p, exists q, exists x, exists y,
- JMeq x y /\ ~ eq_dep U P p x q y.
+ exists U P p q x y, JMeq x y /\ ~ eq_dep U P p x q y.
Proof.
exists bool. exists (fun _ => True). exists true. exists false.
exists I. exists I.
diff --git a/theories/MSets/MSetAVL.v b/theories/MSets/MSetAVL.v
index 6825e6881..349cdedf7 100644
--- a/theories/MSets/MSetAVL.v
+++ b/theories/MSets/MSetAVL.v
@@ -1689,7 +1689,7 @@ Proof.
Qed.
Definition lt (s1 s2 : t) : Prop :=
- exists s1', exists s2', Ok s1' /\ Ok s2' /\ eq s1 s1' /\ eq s2 s2'
+ exists s1' s2', Ok s1' /\ Ok s2' /\ eq s1 s1' /\ eq s2 s2'
/\ L.lt (elements s1') (elements s2').
Instance lt_strorder : StrictOrder lt.
diff --git a/theories/MSets/MSetList.v b/theories/MSets/MSetList.v
index 45278eaf6..bcf68f1d4 100644
--- a/theories/MSets/MSetList.v
+++ b/theories/MSets/MSetList.v
@@ -786,8 +786,7 @@ Module MakeRaw (X: OrderedType) <: RawSets X.
Definition eq := L.eq.
Definition eq_equiv := L.eq_equiv.
Definition lt l1 l2 :=
- exists l1', exists l2', Ok l1' /\ Ok l2' /\
- eq l1 l1' /\ eq l2 l2' /\ L.lt l1' l2'.
+ exists l1' l2', Ok l1' /\ Ok l2' /\ eq l1 l1' /\ eq l2 l2' /\ L.lt l1' l2'.
Instance lt_strorder : StrictOrder lt.
Proof.
diff --git a/theories/Program/Syntax.v b/theories/Program/Syntax.v
index 2c5fb79c5..bb52a0103 100644
--- a/theories/Program/Syntax.v
+++ b/theories/Program/Syntax.v
@@ -51,15 +51,6 @@ Implicit Arguments Vcons [[A] [n]].
(** Treating n-ary exists *)
-Notation " 'exists' x y , p" := (ex (fun x => (ex (fun y => p))))
- (at level 200, x ident, y ident, right associativity) : type_scope.
-
-Notation " 'exists' x y z , p" := (ex (fun x => (ex (fun y => (ex (fun z => p))))))
- (at level 200, x ident, y ident, z ident, right associativity) : type_scope.
-
-Notation " 'exists' x y z w , p" := (ex (fun x => (ex (fun y => (ex (fun z => (ex (fun w => p))))))))
- (at level 200, x ident, y ident, z ident, w ident, right associativity) : type_scope.
-
Tactic Notation "exists" constr(x) := exists x.
Tactic Notation "exists" constr(x) constr(y) := exists x ; exists y.
Tactic Notation "exists" constr(x) constr(y) constr(z) := exists x ; exists y ; exists z.
diff --git a/theories/Unicode/Utf8.v b/theories/Unicode/Utf8.v
index 3a11c9e5e..f806336ba 100644
--- a/theories/Unicode/Utf8.v
+++ b/theories/Unicode/Utf8.v
@@ -8,29 +8,10 @@
(************************************************************************)
(* Logic *)
-Notation "∀ x , P" := (forall x , P)
- (at level 200, x ident, right associativity) : type_scope.
-Notation "∀ x y , P" := (forall x y , P)
- (at level 200, x ident, y ident, right associativity) : type_scope.
-Notation "∀ x y z , P" := (forall x y z , P)
- (at level 200, x ident, y ident, z ident, right associativity) : type_scope.
-Notation "∀ x y z u , P" := (forall x y z u , P)
- (at level 200, x ident, y ident, z ident, u ident, right associativity)
- : type_scope.
-Notation "∀ x : t , P" := (forall x : t , P)
- (at level 200, x ident, right associativity) : type_scope.
-Notation "∀ x y : t , P" := (forall x y : t , P)
- (at level 200, x ident, y ident, right associativity) : type_scope.
-Notation "∀ x y z : t , P" := (forall x y z : t , P)
- (at level 200, x ident, y ident, z ident, right associativity) : type_scope.
-Notation "∀ x y z u : t , P" := (forall x y z u : t , P)
- (at level 200, x ident, y ident, z ident, u ident, right associativity)
- : type_scope.
-
-Notation "∃ x , P" := (exists x , P)
- (at level 200, x ident, right associativity) : type_scope.
-Notation "∃ x : t , P" := (exists x : t, P)
- (at level 200, x ident, right associativity) : type_scope.
+Notation "∀ x .. y , P" := (forall x, .. (forall y, P) ..)
+ (at level 200, x binder, y binder, right associativity) : type_scope.
+Notation "∃ x .. y , P" := (exists x, .. (exists y, P) ..)
+ (at level 200, x binder, y binder, right associativity) : type_scope.
Notation "x ∨ y" := (x \/ y) (at level 85, right associativity) : type_scope.
Notation "x ∧ y" := (x /\ y) (at level 80, right associativity) : type_scope.
@@ -40,10 +21,8 @@ Notation "¬ x" := (~x) (at level 75, right associativity) : type_scope.
Notation "x ≠ y" := (x <> y) (at level 70) : type_scope.
(* Abstraction *)
-(* Not nice
-Notation "'λ' x : T , y" := ([x:T] y) (at level 1, x,T,y at level 10).
-Notation "'λ' x := T , y" := ([x:=T] y) (at level 1, x,T,y at level 10).
-*)
+Notation "'λ' x .. y , t" := (fun x => .. (fun y => t) ..)
+ (at level 200, x binder, y binder, right associativity).
(* Arithmetic *)
Notation "x ≤ y" := (le x y) (at level 70, no associativity).
@@ -51,10 +30,10 @@ Notation "x ≥ y" := (ge x y) (at level 70, no associativity).
(* test *)
(*
-Goal ∀ x, True -> (∃ y , x ≥ y + 1) ∨ x ≤ 0.
+Check ∀ x z, True -> (∃ y v, x + v ≥ y + z) ∨ x ≤ 0.
*)
(* Integer Arithmetic *)
(* TODO: this should come after ZArith
-Notation "x ≤ y" := (Zle x y) (at level 1, y at level 10).
+Notation "x ≤ y" := (Zle x y) (at level 70, no associativity).
*)