From feb655ce464ed1be44183fa07b7fc3f6b1c86e55 Mon Sep 17 00:00:00 2001
From: "Paolo G. Giarrusso" <p.giarrusso@gmail.com>
Date: Sun, 24 Jul 2022 01:49:28 +0200
Subject: [PATCH] set_bind theory: revise setoid rewriting

- strengthen set_bind_subset to have the same arity
- add missing Params instance to prevent slow failures in setoid rewriting
- `Proper` instances can be proven before `set_bind_ext` via `set_solver`.
- `set_bind_ext` is then provable by set_solver directly.
---
 theories/fin_sets.v | 17 ++++++++---------
 1 file changed, 8 insertions(+), 9 deletions(-)

diff --git a/theories/fin_sets.v b/theories/fin_sets.v
index f5494f3e..3ffb9ed6 100644
--- a/theories/fin_sets.v
+++ b/theories/fin_sets.v
@@ -31,6 +31,7 @@ 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.
@@ -463,17 +464,15 @@ Section set_bind.
     rewrite elem_of_set_bind. 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.
-    intros Hfg HXY b. rewrite !elem_of_set_bind. set_solver.
-  Qed.
-
   Global Instance set_bind_proper : Proper (pointwise_relation _ (≡) ==> (≡) ==> (≡)) set_bind.
-  Proof. intros f1 f2 Hf X1 X2 HX. by apply set_bind_ext. Qed.
+  Proof. unfold pointwise_relation; intros f1 f2 Hf X1 X2 HX. set_solver. Qed.
 
-  Global Instance set_bind_subset f : Proper ((⊆) ==> (⊆)) (set_bind f).
-  Proof. intros X Y HXY. set_solver. Qed.
+  Global Instance set_bind_subset : 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.
-- 
GitLab