diff --git a/opam b/opam
index 09d87855b6d5d5ab0bc26f3037dabeddb34bf8e9..4c340587b0d956c13128b427a352b9f0e5daf9e1 100644
--- a/opam
+++ b/opam
@@ -11,5 +11,5 @@ install: [make "install"]
 remove: ["rm" "-rf" "%{lib}%/coq/user-contrib/iris"]
 depends: [
   "coq" { (>= "8.7.1" & < "8.10~") | (= "dev") }
-  "coq-stdpp" { (= "dev.2019-04-24.1.809e0d1d") | (= "dev") }
+  "coq-stdpp" { (= "dev.2019-04-25.0.0f2d2c8a") | (= "dev") }
 ]
diff --git a/theories/heap_lang/lang.v b/theories/heap_lang/lang.v
index 78a2aa862b60ea610f3b736f2e032129978aafbd..a8b975d481a95df4b24f2f1f2ae63e9294aa9618 100644
--- a/theories/heap_lang/lang.v
+++ b/theories/heap_lang/lang.v
@@ -1,6 +1,6 @@
 From iris.program_logic Require Export language ectx_language ectxi_language.
 From iris.algebra Require Export ofe.
-From stdpp Require Export strings.
+From stdpp Require Export binders strings.
 From stdpp Require Import gmap.
 Set Default Proof Using "Type".
 
@@ -42,22 +42,6 @@ Inductive bin_op : Set :=
   | ShiftLOp | ShiftROp (* Shifts *)
   | LeOp | LtOp | EqOp. (* Relations *)
 
-Inductive binder := BAnon | BNamed : string → binder.
-Delimit Scope binder_scope with bind.
-Bind Scope binder_scope with binder.
-Definition cons_binder (mx : binder) (X : list string) : list string :=
-  match mx with BAnon => X | BNamed x => x :: X end.
-Infix ":b:" := cons_binder (at level 60, right associativity).
-Instance binder_eq_dec_eq : EqDecision binder.
-Proof. solve_decision. Defined.
-
-Instance set_unfold_cons_binder x mx X P :
-  SetUnfoldElemOf x X P → SetUnfoldElemOf x (mx :b: X) (BNamed x = mx ∨ P).
-Proof.
-  constructor. rewrite -(set_unfold_elem_of x X P).
-  destruct mx; rewrite /= ?elem_of_cons; naive_solver.
-Qed.
-
 Inductive expr :=
   (* Values *)
   | Val (v : val)
@@ -244,11 +228,6 @@ Proof.
   | 10 => LeOp | 11 => LtOp | _ => EqOp
   end) _); by intros [].
 Qed.
-Instance binder_countable : Countable binder.
-Proof.
- refine (inj_countable' (λ b, match b with BNamed s => Some s | BAnon => None end)
-  (λ b, match b with Some s => BNamed s | None => BAnon end) _); by intros [].
-Qed.
 Instance expr_countable : Countable expr.
 Proof.
  set (enc :=
diff --git a/theories/heap_lang/notation.v b/theories/heap_lang/notation.v
index 81e3fc3e54f5ffb8d3f8b83bd8177d67148096d5..891059ebaae8ceeb99030b4492f888b5d7252bcc 100644
--- a/theories/heap_lang/notation.v
+++ b/theories/heap_lang/notation.v
@@ -15,9 +15,6 @@ Coercion Val : val >-> expr.
 
 Coercion Var : string >-> expr.
 
-Coercion BNamed : string >-> binder.
-Notation "<>" := BAnon : binder_scope.
-
 (* No scope for the values, does not conflict and scope is often not inferred
 properly. *)
 Notation "# l" := (LitV l%Z%V) (at level 8, format "# l").
@@ -50,11 +47,11 @@ Moreover, if the branches do not fit on a single line, it will be printed as:
   end
 *)
 Notation "'match:' e0 'with' 'InjL' x1 => e1 | 'InjR' x2 => e2 'end'" :=
-  (Match e0 x1%bind e1 x2%bind e2)
+  (Match e0 x1%binder e1 x2%binder e2)
   (e0, x1, e1, x2, e2 at level 200,
    format "'[hv' 'match:'  e0  'with'  '/  ' '[' 'InjL'  x1  =>  '/  ' e1 ']'  '/' '[' |  'InjR'  x2  =>  '/  ' e2 ']'  '/' 'end' ']'") : expr_scope.
 Notation "'match:' e0 'with' 'InjR' x1 => e1 | 'InjL' x2 => e2 'end'" :=
-  (Match e0 x2%bind e2 x1%bind e1)
+  (Match e0 x2%binder e2 x1%binder e1)
   (e0, x1, e1, x2, e2 at level 200, only parsing) : expr_scope.
 
 Notation "()" := LitUnit : val_scope.
@@ -81,10 +78,10 @@ Notation "e1 <- e2" := (Store e1%E e2%E) (at level 80) : expr_scope.
 
 (* The breaking point '/  ' makes sure that the body of the rec is indented
 by two spaces in case the whole rec does not fit on a single line. *)
-Notation "'rec:' f x := e" := (Rec f%bind x%bind e%E)
+Notation "'rec:' f x := e" := (Rec f%binder x%binder e%E)
   (at level 200, f at level 1, x at level 1, e at level 200,
    format "'[' 'rec:'  f  x  :=  '/  ' e ']'") : expr_scope.
-Notation "'rec:' f x := e" := (RecV f%bind x%bind e%E)
+Notation "'rec:' f x := e" := (RecV f%binder x%binder e%E)
   (at level 200, f at level 1, x at level 1, e at level 200,
    format "'[' 'rec:'  f  x  :=  '/  ' e ']'") : val_scope.
 Notation "'if:' e1 'then' e2 'else' e3" := (If e1%E e2%E e3%E)
@@ -94,30 +91,30 @@ Notation "'if:' e1 'then' e2 'else' e3" := (If e1%E e2%E e3%E)
 are stated explicitly instead of relying on the Notations Let and Seq as
 defined above. This is needed because App is now a coercion, and these
 notations are otherwise not pretty printed back accordingly. *)
-Notation "'rec:' f x y .. z := e" := (Rec f%bind x%bind (Lam y%bind .. (Lam z%bind e%E) ..))
+Notation "'rec:' f x y .. z := e" := (Rec f%binder x%binder (Lam y%binder .. (Lam z%binder e%E) ..))
   (at level 200, f, x, y, z at level 1, e at level 200,
    format "'[' 'rec:'  f  x  y  ..  z  :=  '/  ' e ']'") : expr_scope.
-Notation "'rec:' f x y .. z := e" := (RecV f%bind x%bind (Lam y%bind .. (Lam z%bind e%E) ..))
+Notation "'rec:' f x y .. z := e" := (RecV f%binder x%binder (Lam y%binder .. (Lam z%binder e%E) ..))
   (at level 200, f, x, y, z at level 1, e at level 200,
    format "'[' 'rec:'  f  x  y  ..  z  :=  '/  ' e ']'") : val_scope.
 
 (* The breaking point '/  ' makes sure that the body of the λ: is indented
 by two spaces in case the whole λ: does not fit on a single line. *)
-Notation "λ: x , e" := (Lam x%bind e%E)
+Notation "λ: x , e" := (Lam x%binder e%E)
   (at level 200, x at level 1, e at level 200,
    format "'[' 'λ:'  x ,  '/  ' e ']'") : expr_scope.
-Notation "λ: x y .. z , e" := (Lam x%bind (Lam y%bind .. (Lam z%bind e%E) ..))
+Notation "λ: x y .. z , e" := (Lam x%binder (Lam y%binder .. (Lam z%binder e%E) ..))
   (at level 200, x, y, z at level 1, e at level 200,
    format "'[' 'λ:'  x  y  ..  z ,  '/  ' e ']'") : expr_scope.
 
-Notation "λ: x , e" := (LamV x%bind e%E)
+Notation "λ: x , e" := (LamV x%binder e%E)
   (at level 200, x at level 1, e at level 200,
    format "'[' 'λ:'  x ,  '/  ' e ']'") : val_scope.
-Notation "λ: x y .. z , e" := (LamV x%bind (Lam y%bind .. (Lam z%bind e%E) .. ))
+Notation "λ: x y .. z , e" := (LamV x%binder (Lam y%binder .. (Lam z%binder e%E) .. ))
   (at level 200, x, y, z at level 1, e at level 200,
    format "'[' 'λ:'  x  y  ..  z ,  '/  ' e ']'") : val_scope.
 
-Notation "'let:' x := e1 'in' e2" := (Lam x%bind e2%E e1%E)
+Notation "'let:' x := e1 'in' e2" := (Lam x%binder e2%E e1%E)
   (at level 200, x at level 1, e1, e2 at level 200,
    format "'[' 'let:'  x  :=  '[' e1 ']'  'in'  '/' e2 ']'") : expr_scope.
 Notation "e1 ;; e2" := (Lam BAnon e2%E e1%E)
@@ -137,10 +134,10 @@ Notation SOME x := (InjR x) (only parsing).
 Notation SOMEV x := (InjRV x) (only parsing).
 
 Notation "'match:' e0 'with' 'NONE' => e1 | 'SOME' x => e2 'end'" :=
-  (Match e0 BAnon e1 x%bind e2)
+  (Match e0 BAnon e1 x%binder e2)
   (e0, e1, x, e2 at level 200, only parsing) : expr_scope.
 Notation "'match:' e0 'with' 'SOME' x => e2 | 'NONE' => e1 'end'" :=
-  (Match e0 BAnon e1 x%bind e2)
+  (Match e0 BAnon e1 x%binder e2)
   (e0, e1, x, e2 at level 200, only parsing) : expr_scope.
 
 Notation "'resolve_proph:' p 'to:' v" := (ResolveProph p v) (at level 100) : expr_scope.