Skip to content
Snippets Groups Projects
Commit 4f1ed7c9 authored by Robbert Krebbers's avatar Robbert Krebbers
Browse files

Clean up dec_agree.

Most notably, there is no need to internalize stuff into the logic
as it follows from generic lemmas for discrete COFEs/CMRAs.
parent eb8dd726
No related branches found
No related tags found
No related merge requests found
From algebra Require Export cmra. From algebra Require Export cmra.
From algebra Require Import functor upred.
Local Arguments validN _ _ _ !_ /. Local Arguments validN _ _ _ !_ /.
Local Arguments valid _ _ !_ /. Local Arguments valid _ _ !_ /.
Local Arguments op _ _ _ !_ /. Local Arguments op _ _ _ !_ /.
Local Arguments unit _ _ !_ /. Local Arguments unit _ _ !_ /.
(* This is isomorphic to optiob, but has a very different RA structure. *) (* This is isomorphic to option, but has a very different RA structure. *)
Inductive dec_agree (A : Type) : Type := Inductive dec_agree (A : Type) : Type :=
| DecAgree : A dec_agree A | DecAgree : A dec_agree A
| DecAgreeBot : dec_agree A. | DecAgreeBot : dec_agree A.
...@@ -35,33 +34,23 @@ Proof. ...@@ -35,33 +34,23 @@ Proof.
- apply _. - apply _.
- apply _. - apply _.
- apply _. - apply _.
- intros [?|] [?|] [?|]; simpl; repeat (case_match; simpl); subst; congruence. - intros [?|] [?|] [?|]; by repeat (simplify_eq/= || case_match).
- intros [?|] [?|]; simpl; repeat (case_match; simpl); try subst; congruence. - intros [?|] [?|]; by repeat (simplify_eq/= || case_match).
- intros [?|]; simpl; repeat (case_match; simpl); try subst; congruence. - intros [?|]; by repeat (simplify_eq/= || case_match).
- intros [?|]; simpl; repeat (case_match; simpl); try subst; congruence. - intros [?|]; by repeat (simplify_eq/= || case_match).
- intros [?|] [?|] ?; simpl; done. - by intros [?|] [?|] ?.
- intros [?|] [?|] ?; simpl; done. - by intros [?|] [?|] ?.
- intros [?|] [?|] [[?|]]; simpl; repeat (case_match; simpl); subst; - intros [?|] [?|] [[?|]]; fold_leibniz;
try congruence; []. intros; by repeat (simplify_eq/= || case_match).
case=>EQ. destruct EQ. done.
Qed. Qed.
Canonical Structure dec_agreeRA : cmraT := discreteRA dec_agree_ra. Canonical Structure dec_agreeRA : cmraT := discreteRA dec_agree_ra.
(* Some properties of this CMRA *) (* Some properties of this CMRA *)
Lemma dec_agree_idemp (x : dec_agree A) : x x x. Lemma dec_agree_idemp (x : dec_agree A) : x x x.
Proof. Proof. destruct x; by repeat (simplify_eq/= || case_match). Qed.
destruct x as [x|]; simpl; repeat (case_match; simpl); try subst; congruence.
Qed.
Lemma dec_agree_op_inv (x1 x2 : dec_agree A) : (x1 x2) x1 x2. Lemma dec_agree_op_inv (x1 x2 : dec_agree A) : (x1 x2) x1 x2.
Proof. Proof. destruct x1, x2; by repeat (simplify_eq/= || case_match). Qed.
destruct x1 as [x1|], x2 as [x2|]; simpl;repeat (case_match; simpl); by subst.
Qed.
Lemma dec_agree_equivI {M} a b : (DecAgree a DecAgree b)%I (a = b : uPred M)%I.
Proof. do 2 split. by case. by destruct 1. Qed.
Lemma dec_agree_validI {M} (x y : dec_agreeRA) : (x y) (x = y : uPred M).
Proof. split=> r n _ ?. by apply: dec_agree_op_inv. Qed.
End dec_agree. End dec_agree.
0% Loading or .
You are about to add 0 people to the discussion. Proceed with caution.
Please register or to comment