diff --git a/barrier/client.v b/barrier/client.v
index fa7b5133fedd2bf5e67d817c947a8026f33704e5..54cae55da3bb32d951ee60cd6adf138d6e212ea3 100644
--- a/barrier/client.v
+++ b/barrier/client.v
@@ -53,7 +53,8 @@ Section client.
     wp_let. (ewp eapply (wp_par heapN N (λ _, True%I) (λ _, True%I))); eauto.
     rewrite 2!{1}[heap_ctx _]always_sep_dup !assoc [(_ ★ heap_ctx _)%I]comm.
     ecancel [heap_ctx _]. sep_split right: []; last first.
-    { do 2 apply forall_intro=>_. apply wand_intro_l. eauto with I. }
+    { do 2 apply forall_intro=>_. apply wand_intro_l.
+      eauto using later_intro with I. }
     sep_split left: [send heapN _ _ _; heap_ctx _; y ↦ _]%I.
     - (* The original thread, the sender. *)
       (ewp eapply wp_store); eauto with I. strip_later.
@@ -68,7 +69,8 @@ Section client.
       (ewp eapply (wp_par heapN N (λ _, True%I) (λ _, True%I))); eauto.
       do 2 rewrite {1}[heap_ctx _]always_sep_dup.
       ecancel [heap_ctx _]. rewrite !assoc. sep_split right: []; last first.
-      { do 2 apply forall_intro=>_. apply wand_intro_l. eauto with I. }
+      { do 2 apply forall_intro=>_. apply wand_intro_l.
+        eauto using later_intro with I. }
       sep_split left: [recv heapN _ _ _; heap_ctx _]%I; by rewrite -worker_safe comm.
 Qed.
 End client.
diff --git a/heap_lang/par.v b/heap_lang/par.v
index 4e5da32f5c03dffe7af3216dee6bbebabc3bd61b..a70c29b4211fd28efba7056f944d7aa766f844bb 100644
--- a/heap_lang/par.v
+++ b/heap_lang/par.v
@@ -22,7 +22,7 @@ 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 (f1,f2)%V →
   (heap_ctx heapN ★ #> f1 #() {{ Ψ1 }} ★ #> f2 #() {{ Ψ2 }} ★
-   ∀ v1 v2, Ψ1 v1 ★ Ψ2 v2 -★ Φ (v1,v2)%V)
+   ∀ v1 v2, Ψ1 v1 ★ Ψ2 v2 -★ ▷ Φ (v1,v2)%V)
   ⊑ #> par e {{ Φ }}.
 Proof.
   intros. rewrite /par.
@@ -35,16 +35,15 @@ Proof.
   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.
-  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.
-  eapply wand_apply_r'; done.
+  apply forall_intro=>v1. apply wand_intro_l. 
+  rewrite (forall_elim v1) (forall_elim v2). rewrite assoc wand_elim_r.
+  wp_let. eapply wp_value, to_val_Pair; eapply to_of_val.
 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 -★ Φ (v1,v2)%V)
+   ∀ v1 v2, Ψ1 v1 ★ Ψ2 v2 -★ ▷ Φ (v1,v2)%V)
   ⊑ #> ParV e1 e2 {{ Φ }}.
 Proof.
   intros. rewrite -par_spec //. apply sep_mono_r.