From 143fbc524ec6c9f0215a9b0a9e8f8b1730837360 Mon Sep 17 00:00:00 2001
From: Robbert Krebbers <mail@robbertkrebbers.nl>
Date: Thu, 6 Oct 2016 10:09:24 +0200
Subject: [PATCH] Prove that the dec_agree CMRA is total.

---
 algebra/dec_agree.v | 16 ++++++++--------
 1 file changed, 8 insertions(+), 8 deletions(-)

diff --git a/algebra/dec_agree.v b/algebra/dec_agree.v
index 280a80df4..c7ccce543 100644
--- a/algebra/dec_agree.v
+++ b/algebra/dec_agree.v
@@ -31,21 +31,19 @@ Instance dec_agree_pcore : PCore (dec_agree A) := Some.
 
 Definition dec_agree_ra_mixin : RAMixin (dec_agree A).
 Proof.
-  split.
-  - apply _.
-  - intros x y cx ? [=<-]; eauto.
-  - apply _.
+  apply ra_total_mixin; apply _ || eauto.
   - intros [?|] [?|] [?|]; by repeat (simplify_eq/= || case_match).
   - intros [?|] [?|]; by repeat (simplify_eq/= || case_match).
-  - intros [?|] ? [=<-]; by repeat (simplify_eq/= || case_match).
   - intros [?|]; by repeat (simplify_eq/= || case_match).
-  - intros [?|] [?|] ?? [=<-]; eauto.
   - by intros [?|] [?|] ?.
 Qed.
 
 Canonical Structure dec_agreeR : cmraT :=
   discreteR (dec_agree A) dec_agree_ra_mixin.
 
+Global Instance dec_agree_total : CMRATotal dec_agreeR.
+Proof. intros x. by exists x. Qed.
+
 (* Some properties of this CMRA *)
 Global Instance dec_agree_persistent (x : dec_agreeR) : Persistent x.
 Proof. by constructor. Qed.
@@ -59,8 +57,10 @@ Proof. destruct x; by rewrite /= ?decide_True. Qed.
 Lemma dec_agree_op_inv (x1 x2 : dec_agree A) : ✓ (x1 ⋅ x2) → x1 = x2.
 Proof. destruct x1, x2; by repeat (simplify_eq/= || case_match). Qed.
 
-Lemma DecAgree_included a b : DecAgree a ≼ DecAgree b → a = b.
-Proof. intros [[c|] [=]%leibniz_equiv_iff]. by simplify_option_eq. Qed.
+Lemma DecAgree_included a b : DecAgree a ≼ DecAgree b ↔ a = b.
+Proof.
+  split. intros [[c|] [=]%leibniz_equiv]. by simplify_option_eq. by intros ->.
+Qed.
 End dec_agree.
 
 Arguments dec_agreeC : clear implicits.
-- 
GitLab