Skip to content
Snippets Groups Projects
auth.v 6.83 KiB
Newer Older
Require Export modures.excl.
Robbert Krebbers's avatar
Robbert Krebbers committed
Local Arguments validN _ _ _ !_ /.
Robbert Krebbers's avatar
Robbert Krebbers committed

Record auth (A : Type) : Type := Auth { authoritative : excl A ; own : A }.
Add Printing Constructor auth.
Robbert Krebbers's avatar
Robbert Krebbers committed
Arguments Auth {_} _ _.
Arguments authoritative {_} _.
Robbert Krebbers's avatar
Robbert Krebbers committed
Arguments own {_} _.
Notation "◯ x" := (Auth ExclUnit x) (at level 20).
Notation "● x" := (Auth (Excl x) ) (at level 20).
Robbert Krebbers's avatar
Robbert Krebbers committed

Robbert Krebbers's avatar
Robbert Krebbers committed
(* COFE *)
Section cofe.
Context {A : cofeT}.

Instance auth_equiv : Equiv (auth A) := λ x y,
  authoritative x  authoritative y  own x  own y.
Instance auth_dist : Dist (auth A) := λ n x y,
Robbert Krebbers's avatar
Robbert Krebbers committed
  authoritative x ={n}= authoritative y  own x ={n}= own y.

Global Instance Auth_ne : Proper (dist n ==> dist n ==> dist n) (@Auth A).
Robbert Krebbers's avatar
Robbert Krebbers committed
Proof. by split. Qed.
Global Instance Auth_proper : Proper (() ==> () ==> ()) (@Auth A).
Proof. by split. Qed.
Global Instance authoritative_ne: Proper (dist n ==> dist n) (@authoritative A).
Robbert Krebbers's avatar
Robbert Krebbers committed
Proof. by destruct 1. Qed.
Global Instance authoritative_proper : Proper (() ==> ()) (@authoritative A).
Proof. by destruct 1. Qed.
Global Instance own_ne : Proper (dist n ==> dist n) (@own A).
Robbert Krebbers's avatar
Robbert Krebbers committed
Proof. by destruct 1. Qed.
Global Instance own_proper : Proper (() ==> ()) (@own A).
Proof. by destruct 1. Qed.
Robbert Krebbers's avatar
Robbert Krebbers committed

Instance auth_compl : Compl (auth A) := λ c,
Robbert Krebbers's avatar
Robbert Krebbers committed
  Auth (compl (chain_map authoritative c)) (compl (chain_map own c)).
Definition auth_cofe_mixin : CofeMixin (auth A).
Robbert Krebbers's avatar
Robbert Krebbers committed
Proof.
  split.
  * intros x y; unfold dist, auth_dist, equiv, auth_equiv.
    rewrite !equiv_dist; naive_solver.
  * intros n; split.
    + by intros ?; split.
    + by intros ?? [??]; split; symmetry.
    + intros ??? [??] [??]; split; etransitivity; eauto.
  * by intros n [x1 y1] [x2 y2] [??]; split; apply dist_S.
  * by split.
  * intros c n; split. apply (conv_compl (chain_map authoritative c) n).
    apply (conv_compl (chain_map own c) n).
Qed.
Canonical Structure authC := CofeT auth_cofe_mixin.
Instance Auth_timeless (x : excl A) (y : A) :
Robbert Krebbers's avatar
Robbert Krebbers committed
  Timeless x  Timeless y  Timeless (Auth x y).
Proof. by intros ?? [??] [??]; split; simpl in *; apply (timeless _). Qed.
Global Instance auth_leibniz : LeibnizEquiv A  LeibnizEquiv (auth A).
Proof. by intros ? [??] [??] [??]; f_equal'; apply leibniz_equiv. Qed.
End cofe.

Arguments authC : clear implicits.
Robbert Krebbers's avatar
Robbert Krebbers committed

(* CMRA *)
Section cmra.
Context {A : cmraT}.

Global Instance auth_empty `{Empty A} : Empty (auth A) := Auth  ∅.
Instance auth_validN : ValidN (auth A) := λ n x,
Robbert Krebbers's avatar
Robbert Krebbers committed
  match authoritative x with
  | Excl a => own x {n} a  {n} a
  | ExclUnit => {n} (own x)
Robbert Krebbers's avatar
Robbert Krebbers committed
  | ExclBot => n = 0
  end.
Global Arguments auth_validN _ !_ /.
Instance auth_unit : Unit (auth A) := λ x,
  Auth (unit (authoritative x)) (unit (own x)).
Instance auth_op : Op (auth A) := λ x y,
  Auth (authoritative x  authoritative y) (own x  own y).
Instance auth_minus : Minus (auth A) := λ x y,
  Auth (authoritative x  authoritative y) (own x  own y).
Lemma auth_included (x y : auth A) :
  x  y  authoritative x  authoritative y  own x  own y.
Proof.
  split; [intros [[z1 z2] Hz]; split; [exists z1|exists z2]; apply Hz|].
  intros [[z1 Hz1] [z2 Hz2]]; exists (Auth z1 z2); split; auto.
Qed.
Lemma auth_includedN n (x y : auth A) :
Robbert Krebbers's avatar
Robbert Krebbers committed
  x {n} y  authoritative x {n} authoritative y  own x {n} own y.
Proof.
  split; [intros [[z1 z2] Hz]; split; [exists z1|exists z2]; apply Hz|].
  intros [[z1 Hz1] [z2 Hz2]]; exists (Auth z1 z2); split; auto.
Qed.
Lemma authoritative_validN n (x : auth A) : {n} x  {n} (authoritative x).
Robbert Krebbers's avatar
Robbert Krebbers committed
Proof. by destruct x as [[]]. Qed.
Lemma own_validN n (x : auth A) : {n} x  {n} (own x).
Proof. destruct x as [[]]; naive_solver eauto using cmra_validN_includedN. Qed.

Definition auth_cmra_mixin : CMRAMixin (auth A).
Robbert Krebbers's avatar
Robbert Krebbers committed
Proof.
  split.
  * by intros n x y1 y2 [Hy Hy']; split; simpl; rewrite ?Hy ?Hy'.
  * by intros n y1 y2 [Hy Hy']; split; simpl; rewrite ?Hy ?Hy'.
Robbert Krebbers's avatar
Robbert Krebbers committed
  * intros n [x a] [y b] [Hx Ha]; simpl in *;
      destruct Hx as [[][]| | |]; intros ?; cofe_subst; auto.
  * by intros n x1 x2 [Hx Hx'] y1 y2 [Hy Hy'];
      split; simpl; rewrite ?Hy ?Hy' ?Hx ?Hx'.
Robbert Krebbers's avatar
Robbert Krebbers committed
  * by intros [[] ?]; simpl.
  * intros n [[] ?] ?; naive_solver eauto using cmra_includedN_S, cmra_validN_S.
  * by split; simpl; rewrite associative.
  * by split; simpl; rewrite commutative.
  * by split; simpl; rewrite ?cmra_unit_l.
  * by split; simpl; rewrite ?cmra_unit_idempotent.
Robbert Krebbers's avatar
Robbert Krebbers committed
  * intros n ??; rewrite! auth_includedN; intros [??].
    by split; simpl; apply cmra_unit_preservingN.
  * assert ( n (a b1 b2 : A), b1  b2 {n} a  b1 {n} a).
    { intros n a b1 b2 <-; apply cmra_includedN_l. }
Robbert Krebbers's avatar
Robbert Krebbers committed
   intros n [[a1| |] b1] [[a2| |] b2];
     naive_solver eauto using cmra_validN_op_l, cmra_validN_includedN.
Robbert Krebbers's avatar
Robbert Krebbers committed
  * by intros n ??; rewrite auth_includedN;
      intros [??]; split; simpl; apply cmra_op_minus.
Qed.
Definition auth_cmra_extend_mixin : CMRAExtendMixin (auth A).
Robbert Krebbers's avatar
Robbert Krebbers committed
Proof.
  intros n x y1 y2 ? [??]; simpl in *.
  destruct (cmra_extend_op n (authoritative x) (authoritative y1)
    (authoritative y2)) as (z1&?&?&?); auto using authoritative_validN.
  destruct (cmra_extend_op n (own x) (own y1) (own y2))
    as (z2&?&?&?); auto using own_validN.
  by exists (Auth (z1.1) (z2.1), Auth (z1.2) (z2.2)).
Qed.
Canonical Structure authRA : cmraT :=
  CMRAT auth_cofe_mixin auth_cmra_mixin auth_cmra_extend_mixin.
Instance auth_cmra_identity `{Empty A} : CMRAIdentity A  CMRAIdentity authRA.
Robbert Krebbers's avatar
Robbert Krebbers committed
Proof.
  split; simpl.
  * by apply (@cmra_empty_valid A _).
  * by intros x; constructor; rewrite /= left_id.
  * apply Auth_timeless; apply _.
Robbert Krebbers's avatar
Robbert Krebbers committed
Qed.
Lemma auth_frag_op (a b : A) :  (a  b)   a   b.
Robbert Krebbers's avatar
Robbert Krebbers committed
Proof. done. Qed.
End cmra.

Arguments authRA : clear implicits.
Robbert Krebbers's avatar
Robbert Krebbers committed

(* Functor *)
Instance auth_fmap : FMap auth := λ A B f x,
  Auth (f <$> authoritative x) (f (own x)).
Arguments auth_fmap _ _ _ !_ /.
Lemma auth_fmap_id {A} (x : auth A) : id <$> x = x.
Proof. by destruct x; rewrite /= excl_fmap_id. Qed.
Lemma excl_fmap_compose {A B C} (f : A  B) (g : B  C) (x : auth A) :
  g  f <$> x = g <$> f <$> x.
Proof. by destruct x; rewrite /= excl_fmap_compose. Qed.
Instance auth_fmap_cmra_ne {A B : cmraT} n :
Robbert Krebbers's avatar
Robbert Krebbers committed
  Proper ((dist n ==> dist n) ==> dist n ==> dist n) (@fmap auth _ A B).
Proof.
  intros f g Hf [??] [??] [??]; split; [by apply excl_fmap_cmra_ne|by apply Hf].
Qed.
Instance auth_fmap_cmra_monotone {A B : cmraT} (f : A  B) :
Robbert Krebbers's avatar
Robbert Krebbers committed
  ( n, Proper (dist n ==> dist n) f)  CMRAMonotone f 
  CMRAMonotone (fmap f : auth A  auth B).
Proof.
  split.
  * by intros n [x a] [y b]; rewrite !auth_includedN /=;
      intros [??]; split; simpl; apply: includedN_preserving.
  * intros n [[a| |] b]; rewrite /= /cmra_validN;
Robbert Krebbers's avatar
Robbert Krebbers committed
      naive_solver eauto using @includedN_preserving, @validN_preserving.
Qed.
Definition authRA_map {A B : cmraT} (f : A -n> B) : authRA A -n> authRA B :=
  CofeMor (fmap f : authRA A  authRA B).
Lemma authRA_map_ne A B n : Proper (dist n ==> dist n) (@authRA_map A B).
Proof. intros f f' Hf [[a| |] b]; repeat constructor; apply Hf. Qed.