From 575f23a96bb6100ddd2c4711f32448005b1762c6 Mon Sep 17 00:00:00 2001
From: Ralf Jung <jung@mpi-sws.org>
Date: Fri, 9 Dec 2016 15:30:32 +0100
Subject: [PATCH] rename _setoid suffix to _equiv; add variant of
 fmap_Some_setoid that can be usefully destructed

---
 theories/fin_maps.v | 2 +-
 theories/list.v     | 2 +-
 theories/option.v   | 8 ++++++--
 3 files changed, 8 insertions(+), 4 deletions(-)

diff --git a/theories/fin_maps.v b/theories/fin_maps.v
index 9ea6d5f8..2c8634da 100644
--- a/theories/fin_maps.v
+++ b/theories/fin_maps.v
@@ -556,7 +556,7 @@ Proof. apply map_eq; intros i; by rewrite lookup_fmap, option_fmap_id. Qed.
 Lemma map_fmap_compose {A B C} (f : A → B) (g : B → C) (m : M A) :
   g ∘ f <$> m = g <$> f <$> m.
 Proof. apply map_eq; intros i; by rewrite !lookup_fmap,option_fmap_compose. Qed.
-Lemma map_fmap_setoid_ext `{Equiv A, Equiv B} (f1 f2 : A → B) m :
+Lemma map_fmap_equiv_ext `{Equiv A, Equiv B} (f1 f2 : A → B) m :
   (∀ i x, m !! i = Some x → f1 x ≡ f2 x) → f1 <$> m ≡ f2 <$> m.
 Proof.
   intros Hi i; rewrite !lookup_fmap.
diff --git a/theories/list.v b/theories/list.v
index 074a9048..b17af16f 100644
--- a/theories/list.v
+++ b/theories/list.v
@@ -2802,7 +2802,7 @@ Section fmap.
   Lemma list_fmap_ext (g : A → B) (l1 l2 : list A) :
     (∀ x, f x = g x) → l1 = l2 → fmap f l1 = fmap g l2.
   Proof. intros ? <-. induction l1; f_equal/=; auto. Qed.
-  Lemma list_fmap_setoid_ext `{Equiv B} (g : A → B) l :
+  Lemma list_fmap_equiv_ext `{Equiv B} (g : A → B) l :
     (∀ x, f x ≡ g x) → f <$> l ≡ g <$> l.
   Proof. induction l; constructor; auto. Qed.
 
diff --git a/theories/option.v b/theories/option.v
index 52dc7223..9e7c8896 100644
--- a/theories/option.v
+++ b/theories/option.v
@@ -180,7 +180,7 @@ Proof. unfold is_Some; destruct mx; naive_solver. Qed.
 Lemma fmap_Some {A B} (f : A → B) mx y :
   f <$> mx = Some y ↔ ∃ x, mx = Some x ∧ y = f x.
 Proof. destruct mx; naive_solver. Qed.
-Lemma fmap_Some_setoid {A B} `{Equiv B} `{!Equivalence ((≡) : relation B)}
+Lemma fmap_Some_equiv {A B} `{Equiv B} `{!Equivalence ((≡) : relation B)}
       (f : A → B) mx y :
   f <$> mx ≡ Some y ↔ ∃ x, mx = Some x ∧ y ≡ f x.
 Proof.
@@ -190,6 +190,10 @@ Proof.
   - intros ?%symmetry%equiv_None. done.
   - intros (? & ? & ?). done.
 Qed.
+Lemma fmap_Some_equiv' {A B} `{Equiv B} `{!Equivalence ((≡) : relation B)}
+      (f : A → B) mx y :
+  f <$> mx ≡ Some y → ∃ x, mx = Some x ∧ y ≡ f x.
+Proof. intros. apply fmap_Some_equiv. done. Qed.
 Lemma fmap_None {A B} (f : A → B) mx : f <$> mx = None ↔ mx = None.
 Proof. by destruct mx. Qed.
 Lemma option_fmap_id {A} (mx : option A) : id <$> mx = mx.
@@ -200,7 +204,7 @@ Proof. by destruct mx. Qed.
 Lemma option_fmap_ext {A B} (f g : A → B) mx :
   (∀ x, f x = g x) → f <$> mx = g <$> mx.
 Proof. intros; destruct mx; f_equal/=; auto. Qed.
-Lemma option_fmap_setoid_ext `{Equiv A, Equiv B} (f g : A → B) mx :
+Lemma option_fmap_equiv_ext `{Equiv A, Equiv B} (f g : A → B) mx :
   (∀ x, f x ≡ g x) → f <$> mx ≡ g <$> mx.
 Proof. destruct mx; constructor; auto. Qed.
 Lemma option_fmap_bind {A B C} (f : A → B) (g : B → option C) mx :
-- 
GitLab