diff --git a/heap_lang/notation.v b/heap_lang/notation.v
index 156222098423562f2c1f3386498819c6acebeac8..7ac1a85e78e85e615904f2ee6cfdcfeaef352a23 100644
--- a/heap_lang/notation.v
+++ b/heap_lang/notation.v
@@ -32,6 +32,7 @@ Notation "^ v" := (of_val' v%V) (at level 8, format "^ v") : expr_scope.
 (** Syntax inspired by Coq/Ocaml. Constructions with higher precedence come
     first. *)
 Notation "( e1 , e2 , .. , en )" := (Pair .. (Pair e1 e2) .. en) : expr_scope.
+Notation "( e1 , e2 , .. , en )" := (PairV .. (PairV e1 e2) .. en) : val_scope.
 Notation "'match:' e0 'with' 'InjL' x1 => e1 | 'InjR' x2 => e2 'end'" :=
   (Match e0 x1 e1 x2 e2)
   (e0, x1, e1, x2, e2 at level 200) : expr_scope.
diff --git a/heap_lang/par.v b/heap_lang/par.v
index 9a707763d81f0bae95f6135edba1c196bb56b331..4e5da32f5c03dffe7af3216dee6bbebabc3bd61b 100644
--- a/heap_lang/par.v
+++ b/heap_lang/par.v
@@ -20,21 +20,21 @@ Context (heapN N : namespace).
 Local Notation iProp := (iPropG heap_lang Σ).
 
 Lemma par_spec (Ψ1 Ψ2 : val → iProp) e (f1 f2 : val) (Φ : val → iProp) :
-  heapN ⊥ N → to_val e = Some (PairV f1 f2) →
+  heapN ⊥ N → to_val e = Some (f1,f2)%V →
   (heap_ctx heapN ★ #> f1 #() {{ Ψ1 }} ★ #> f2 #() {{ Ψ2 }} ★
-   ∀ v1 v2, Ψ1 v1 ★ Ψ2 v2 -★ Φ (PairV v1 v2))
+   ∀ v1 v2, Ψ1 v1 ★ Ψ2 v2 -★ Φ (v1,v2)%V)
   ⊑ #> par e {{ Φ }}.
 Proof.
   intros. rewrite /par.
   wp_focus e. etransitivity; last by eapply wp_value. wp_let.
   (* FIXME: wp_proj should not spawn these goals. *)
   wp_proj; eauto using to_of_val.
-  (ewp eapply spawn_spec); eauto using to_of_val.
+  ewp eapply spawn_spec; eauto using to_of_val.
   apply sep_mono_r. apply sep_mono_r.
   apply forall_intro=>h. apply wand_intro_l. wp_let.
   wp_proj; eauto using to_of_val.
   wp_focus (f2 _). rewrite wp_frame_r wp_frame_l. apply wp_mono=>v2. wp_let.
-  (ewp eapply join_spec); eauto using to_of_val. apply sep_mono_r.
+  ewp eapply join_spec; eauto using to_of_val. apply sep_mono_r.
   apply forall_intro=>v1. apply wand_intro_l. wp_let.
   etransitivity; last by (eapply wp_value, to_val_Pair; eapply to_of_val).
   rewrite (forall_elim v1) (forall_elim v2). rewrite assoc.
@@ -44,7 +44,7 @@ Qed.
 Lemma wp_par (Ψ1 Ψ2 : val → iProp) (e1 e2 : expr []) (Φ : val → iProp) :
   heapN ⊥ N →
   (heap_ctx heapN ★ #> e1 {{ Ψ1 }} ★ #> e2 {{ Ψ2 }} ★
-   ∀ v1 v2, Ψ1 v1 ★ Ψ2 v2 -★ Φ (PairV v1 v2))
+   ∀ v1 v2, Ψ1 v1 ★ Ψ2 v2 -★ Φ (v1,v2)%V)
   ⊑ #> ParV e1 e2 {{ Φ }}.
 Proof.
   intros. rewrite -par_spec //. apply sep_mono_r.