diff --git a/program_logic/adequacy.v b/program_logic/adequacy.v
index ce5af9d66b3729cc14d145eeb369324577690677..795b482ad6e7007ec3f1175f4c1744773ab3df85 100644
--- a/program_logic/adequacy.v
+++ b/program_logic/adequacy.v
@@ -94,7 +94,7 @@ Lemma wptp_result n e1 t1 v2 t2 σ1 σ2 φ :
   world σ1 ★ WP e1 {{ v, ■ φ v }} ★ wptp t1 ⊢
   Nat.iter (S (S n)) (λ P, |=r=> ▷ P) (■ φ v2).
 Proof.
-  intros. rewrite wptp_steps // (Nat_iter_S_r (S n)); apply rvs_iter_mono.
+  intros. rewrite wptp_steps // (Nat_iter_S_r (S n)); apply: rvs_iter_mono.
   iDestruct 1 as (e2 t2') "(% & (Hw & HE & _) & H & _)"; simplify_eq.
   iDestruct (wp_value_inv with "H") as "H". rewrite pvs_eq /pvs_def.
   iVs ("H" with "[Hw HE]") as ">(_ & _ & $)"; iFrame; auto.
@@ -114,7 +114,7 @@ Lemma wptp_safe n e1 e2 t1 t2 σ1 σ2 Φ :
   world σ1 ★ WP e1 {{ Φ }} ★ wptp t1 ⊢
   Nat.iter (S (S n)) (λ P, |=r=> ▷ P) (■ (is_Some (to_val e2) ∨ reducible e2 σ2)).
 Proof.
-  intros ? He2. rewrite wptp_steps // (Nat_iter_S_r (S n)); apply rvs_iter_mono.
+  intros ? He2. rewrite wptp_steps // (Nat_iter_S_r (S n)); apply: rvs_iter_mono.
   iDestruct 1 as (e2' t2') "(% & Hw & H & Htp)"; simplify_eq.
   apply elem_of_cons in He2 as [<-|?]; first (iApply wp_safe; by iFrame "Hw H").
   iApply wp_safe. iFrame "Hw".