diff --git a/CHANGELOG.md b/CHANGELOG.md
index 2a5a0cee5c915d483bb75e5f33207a1bb1695998..9500ee9d0e6d6c991c365d4cff2e850588b0a0c6 100644
--- a/CHANGELOG.md
+++ b/CHANGELOG.md
@@ -43,6 +43,8 @@ s/\bmap_disjoint_subset\b/kmap_subset/g
 s/\blookup_union_l\b/lookup_union_l'/g
 EOF
 ```
+- Add the bind operation `set_bind` for finite sets. (by Dan Frumin and Paolo G.
+  Giarrusso)
 
 ## std++ 1.7.0 (2022-01-22)
 
diff --git a/theories/fin_sets.v b/theories/fin_sets.v
index 41f8311c9f2e5884d6cc3dbf328e57a5cf668909..54bae6793cc3240d018a16d2a74ea8b76e5c76e2 100644
--- a/theories/fin_sets.v
+++ b/theories/fin_sets.v
@@ -27,6 +27,12 @@ Definition set_map `{Elements A C, Singleton B D, Empty D, Union D}
 Typeclasses Opaque set_map.
 Global Instance: Params (@set_map) 8 := {}.
 
+Definition set_bind `{Elements A SA, Empty SB, Union SB}
+    (f : A → SB) (X : SA) : SB :=
+  ⋃ (f <$> elements X).
+Typeclasses Opaque set_bind.
+Global Instance: Params (@set_bind) 6 := {}.
+
 Global Instance set_fresh `{Elements A C, Fresh A (list A)} : Fresh A C :=
   fresh ∘ elements.
 Typeclasses Opaque set_fresh.
@@ -436,6 +442,53 @@ Section map.
   Proof. unfold_leibniz. apply set_map_singleton. Qed.
 End map.
 
+(** * Bind *)
+Section set_bind.
+  Context `{FinSet B SB}.
+
+  Local Notation set_bind := (set_bind (A:=A) (SA:=C) (SB:=SB)).
+
+  Lemma elem_of_set_bind (f : A → SB) (X : C) y :
+    y ∈ set_bind f X ↔ ∃ x, x ∈ X ∧ y ∈ f x.
+  Proof.
+    unfold set_bind. rewrite !elem_of_union_list. set_solver.
+  Qed.
+
+  Global Instance set_unfold_set_bind (f : A → SB) (X : C)
+       (y : B) (P : A → B → Prop) (Q : A → Prop) :
+    (∀ x y, SetUnfoldElemOf y (f x) (P x y)) →
+    (∀ x, SetUnfoldElemOf x X (Q x)) →
+    SetUnfoldElemOf y (set_bind f X) (∃ x, Q x ∧ P x y).
+  Proof.
+    intros HSU1 HSU2. constructor.
+    rewrite elem_of_set_bind. set_solver.
+  Qed.
+
+  Global Instance set_bind_proper :
+    Proper (pointwise_relation _ (≡) ==> (≡) ==> (≡)) set_bind.
+  Proof. unfold pointwise_relation; intros f1 f2 Hf X1 X2 HX. set_solver. Qed.
+
+  Global Instance set_bind_mono :
+    Proper (pointwise_relation _ (⊆) ==> (⊆) ==> (⊆)) set_bind.
+  Proof. unfold pointwise_relation; intros f1 f2 Hf X1 X2 HX. set_solver. Qed.
+
+  Lemma set_bind_ext (f g : A → SB) (X Y : C) :
+    (∀ x, x ∈ X → x ∈ Y → f x ≡ g x) → X ≡ Y → set_bind f X ≡ set_bind g Y.
+  Proof. set_solver. Qed.
+
+  Lemma set_bind_singleton f x : set_bind f {[x]} ≡ f x.
+  Proof. set_solver. Qed.
+  Lemma set_bind_singleton_L `{!LeibnizEquiv SB} f x : set_bind f {[x]} = f x.
+  Proof. unfold_leibniz. apply set_bind_singleton. Qed.
+
+  Lemma set_bind_disj_union f (X Y : C) :
+    X ## Y → set_bind f (X ∪ Y) ≡ set_bind f X ∪ set_bind f Y.
+  Proof. set_solver. Qed.
+  Lemma set_bind_disj_union_L `{!LeibnizEquiv SB} f (X Y : C) :
+    X ## Y → set_bind f (X ∪ Y) = set_bind f X ∪ set_bind f Y.
+  Proof. unfold_leibniz. apply set_bind_disj_union. Qed.
+End set_bind.
+
 (** * Decision procedures *)
 Lemma set_Forall_elements P X : set_Forall P X ↔ Forall P (elements X).
 Proof. rewrite Forall_forall. by setoid_rewrite elem_of_elements. Qed.