diff --git a/heap_lang/derived.v b/heap_lang/derived.v
index 3d2a02b29138ddcd2a6b1d32180a017f821ef2ea..b834c31faf8d7f78bc3420125b000aa2e1e45312 100644
--- a/heap_lang/derived.v
+++ b/heap_lang/derived.v
@@ -30,7 +30,7 @@ Proof.
   by rewrite -wp_let' //= ?to_of_val ?subst_empty.
 Qed.
 
-Lemma wp_skip E Q : ▷(Q (LitV LitUnit))  ⊑ wp E Skip Q.
+Lemma wp_skip E Q : ▷ (Q (LitV LitUnit)) ⊑ wp E Skip Q.
 Proof. rewrite -wp_seq -wp_value // -wp_value //. Qed.
 
 Lemma wp_le E (n1 n2 : Z) P Q :
diff --git a/heap_lang/lifting.v b/heap_lang/lifting.v
index 918810a61b40346834ff684c7b5f1109793448c3..289a24ae72ca365ecc7bb0ed88e6566c2ea4250c 100644
--- a/heap_lang/lifting.v
+++ b/heap_lang/lifting.v
@@ -76,13 +76,12 @@ Proof.
 Qed.
 
 (** Base axioms for core primitives of the language: Stateless reductions *)
-Lemma wp_fork E e :
-  ▷ wp (Σ:=Σ) coPset_all e (λ _, True) ⊑ wp E (Fork e) (λ v, v = LitV LitUnit).
+Lemma wp_fork E e Q :
+  (▷ Q (LitV LitUnit) ★ ▷ wp (Σ:=Σ) coPset_all e (λ _, True)) ⊑ wp E (Fork e) Q.
 Proof.
   rewrite -(wp_lift_pure_det_step (Fork e) (Lit LitUnit) (Some e)) //=;
     last by intros; inv_step; eauto.
-  apply later_mono, sep_intro_True_l; last done.
-  by rewrite -(wp_value _ _ (Lit _)) //; apply const_intro.
+  rewrite later_sep -(wp_value _ _ (Lit _)) //.
 Qed.
 
 (* For the lemmas involving substitution, we only derive a preliminary version.