diff --git a/coq-actris.opam b/coq-actris.opam
index c14b302ed59e8b58aeb41d18543cdfa8b7140959..1e75850c1245897cc304dd2d3c2e6fec756427ee 100644
--- a/coq-actris.opam
+++ b/coq-actris.opam
@@ -8,7 +8,7 @@ bug-reports: "https://gitlab.mpi-sws.org/iris/actris/issues"
 dev-repo: "git+https://gitlab.mpi-sws.org/iris/actris.git"
 
 depends: [
-  "coq-iris-heap-lang" { (= "dev.2024-08-16.3.8890e30a") | (= "dev") }
+  "coq-iris-heap-lang" { (= "dev.2024-09-04.0.0653ba6c") | (= "dev") }
 ]
 
 build: [make "-j%{jobs}%"]
diff --git a/theories/logrel/session_types.v b/theories/logrel/session_types.v
index 95af463ed698183989145ece118c3fa1ca11c0fc..4d0c7ea65ca492e70668ccb3c80bf272def55d9f 100644
--- a/theories/logrel/session_types.v
+++ b/theories/logrel/session_types.v
@@ -108,7 +108,7 @@ Section session_types.
   Global Instance lty_choice_proper a : Proper ((≡) ==> (≡)) (@lty_choice Σ a).
   Proof. apply ne_proper, _. Qed.
   Global Instance lty_choice_contractive a n :
-    Proper (map_relation (dist_later n) (λ _, False) (λ _, False) ==> dist n)
+    Proper (map_relation (λ _, dist_later n) (λ _ _, False) (λ _ _, False) ==> dist n)
            (@lty_choice Σ a).
   Proof.
     intros Ss Ts Heq. rewrite /lty_choice.