Skip to content
Snippets Groups Projects

Compare revisions

Changes are shown as if the source revision was being merged into the target revision. Learn more about comparing revisions.

Source

Select target project
No results found

Target

Select target project
  • iris/iris
  • jeehoon.kang/iris-coq
  • amintimany/iris-coq
  • dfrumin/iris-coq
  • Villetaneuse/iris
  • gares/iris
  • shiatsumat/iris
  • Blaisorblade/iris
  • jihgfee/iris-coq
  • mrhaandi/iris
  • tlsomers/iris
  • Quarkbeast/iris-coq
  • janno/iris
  • amaurremi/iris-coq
  • proux/iris
  • tchajed/iris
  • herbelin/iris-coq
  • msammler/iris-coq
  • maximedenes/iris-coq
  • bpeters/iris
  • haidang/iris
  • lepigre/iris
  • lczch/iris
  • simonspies/iris
  • gpirlea/iris
  • dkhalanskiyjb/iris
  • gmalecha/iris
  • germanD/iris
  • aa755/iris
  • jules/iris
  • abeln/iris
  • simonfv/iris
  • atrieu/iris
  • arthuraa/iris
  • simonh/iris
  • jung/iris
  • mattam82/iris
  • Armael/iris
  • adamAndMath/iris
  • gmevel/iris
  • snyke7/iris
  • johannes/iris
  • NiklasM/iris
  • simonspies/iris-parametric-index
  • svancollem/iris
  • proux1/iris
  • wmansky/iris
  • LukeXuan/iris
  • ivanbakel/iris
  • SkySkimmer/iris
  • tjhance/iris
  • yiyunliu/iris
  • Lee-Janggun/iris
  • thomas-lamiaux/iris
  • dongjae/iris
  • dnezam/iris
  • Tragicus/iris
  • clef-men/iris
  • ffengyu/iris
59 results
Show changes
Showing
with 7529 additions and 170 deletions
From iris.algebra Require Export cmra.
From iris.prelude Require Import options.
Local Arguments validN _ _ _ _ !_ /.
Local Arguments valid _ _ !_ /.
Local Arguments op _ _ _ !_ /.
Local Arguments pcore _ _ !_ /.
(** Define an agreement construction such that Agree A is discrete when A is discrete.
Notice that this construction is NOT complete. The following is due to Aleš:
Proposition: Ag(T) is not necessarily complete.
Proof.
Let T be the set of binary streams (infinite sequences) with the usual
ultrametric, measuring how far they agree.
Let Aₙ be the set of all binary strings of length n. Thus for Aₙ to be a
subset of T we have them continue as a stream of zeroes.
Now Aₙ is a finite non-empty subset of T. Moreover {Aₙ} is a Cauchy sequence
in the defined (Hausdorff) metric.
However the limit (if it were to exist as an element of Ag(T)) would have to
be the set of all binary streams, which is not exactly finite.
Thus Ag(T) is not necessarily complete.
*)
(** Note that the projection [agree_car] is not non-expansive, so it cannot be
used in the logic. If you need to get a witness out, you should use the
lemma [to_agree_uninjN] instead. In general, [agree_car] should ONLY be used
internally in this file. *)
Record agree (A : Type) : Type := {
agree_car : list A;
agree_not_nil : bool_decide (agree_car = []) = false
}.
Global Arguments agree_car {_} _.
Global Arguments agree_not_nil {_} _.
Local Coercion agree_car : agree >-> list.
Definition to_agree {A} (a : A) : agree A :=
{| agree_car := [a]; agree_not_nil := eq_refl |}.
Lemma elem_of_agree {A} (x : agree A) : a, a agree_car x.
Proof. destruct x as [[|a ?] ?]; set_solver+. Qed.
Lemma agree_eq {A} (x y : agree A) : agree_car x = agree_car y x = y.
Proof.
destruct x as [a ?], y as [b ?]; simpl.
intros ->; f_equal. apply (proof_irrel _).
Qed.
Section agree.
Context {SI : sidx} {A : ofe}.
Implicit Types a b : A.
Implicit Types x y : agree A.
(* OFE *)
Local Instance agree_dist : Dist (agree A) := λ n x y,
( a, a agree_car x b, b agree_car y a {n} b)
( b, b agree_car y a, a agree_car x a {n} b).
Local Instance agree_equiv : Equiv (agree A) := λ x y, n, x {n} y.
Definition agree_ofe_mixin : OfeMixin (agree A).
Proof.
split.
- done.
- intros n; split; rewrite /dist /agree_dist.
+ intros x; split; eauto.
+ intros x y [??]. naive_solver eauto.
+ intros x y z [H1 H1'] [H2 H2']; split.
* intros a ?. destruct (H1 a) as (b&?&?); auto.
destruct (H2 b) as (c&?&?); eauto. by exists c; split; last etrans.
* intros a ?. destruct (H2' a) as (b&?&?); auto.
destruct (H1' b) as (c&?&?); eauto. by exists c; split; last etrans.
- intros n m x y [??]; split; naive_solver eauto using dist_le.
Qed.
Canonical Structure agreeO := Ofe (agree A) agree_ofe_mixin.
(* CMRA *)
(* agree_validN is carefully written such that, when applied to a singleton, it
is convertible to True. This makes working with agreement much more pleasant. *)
Local Instance agree_validN_instance : ValidN (agree A) := λ n x,
match agree_car x with
| [a] => True
| _ => a b, a agree_car x b agree_car x a {n} b
end.
Local Instance agree_valid_instance : Valid (agree A) := λ x, n, {n} x.
Local Program Instance agree_op_instance : Op (agree A) := λ x y,
{| agree_car := agree_car x ++ agree_car y |}.
Next Obligation. by intros [[|??]] y. Qed.
Local Instance agree_pcore_instance : PCore (agree A) := Some.
Lemma agree_validN_def n x :
{n} x a b, a agree_car x b agree_car x a {n} b.
Proof.
rewrite /validN /agree_validN_instance. destruct (agree_car _) as [|? [|??]]; auto.
setoid_rewrite elem_of_list_singleton; naive_solver.
Qed.
Local Instance agree_comm : Comm () (@op (agree A) _).
Proof. intros x y n; split=> a /=; setoid_rewrite elem_of_app; naive_solver. Qed.
Local Instance agree_assoc : Assoc () (@op (agree A) _).
Proof.
intros x y z n; split=> a /=; repeat setoid_rewrite elem_of_app; naive_solver.
Qed.
Lemma agree_idemp x : x x x.
Proof. intros n; split=> a /=; setoid_rewrite elem_of_app; naive_solver. Qed.
Local Instance agree_validN_ne n :
Proper (dist n ==> impl) (@validN SI (agree A) _ n).
Proof.
intros x y [H H']; rewrite /impl !agree_validN_def; intros Hv a b Ha Hb.
destruct (H' a) as (a'&?&<-); auto. destruct (H' b) as (b'&?&<-); auto.
Qed.
Local Instance agree_validN_proper n :
Proper (equiv ==> iff) (@validN SI (agree A) _ n).
Proof. move=> x y /equiv_dist H. by split; rewrite (H n). Qed.
Local Instance agree_op_ne' x : NonExpansive (op x).
Proof.
intros n y1 y2 [H H']; split=> a /=; setoid_rewrite elem_of_app; naive_solver.
Qed.
Local Instance agree_op_ne : NonExpansive2 (@op (agree A) _).
Proof. by intros n x1 x2 Hx y1 y2 Hy; rewrite Hy !(comm _ _ y2) Hx. Qed.
Local Instance agree_op_proper : Proper (() ==> () ==> ()) (op (A := agree A)) :=
ne_proper_2 _.
Lemma agree_included x y : x y y x y.
Proof.
split; [|by intros ?; exists y].
by intros [z Hz]; rewrite Hz assoc agree_idemp.
Qed.
Lemma agree_includedN n x y : x {n} y y {n} x y.
Proof.
split; [|by intros ?; exists y].
by intros [z Hz]; rewrite Hz assoc agree_idemp.
Qed.
Lemma agree_op_invN n x1 x2 : {n} (x1 x2) x1 {n} x2.
Proof.
rewrite agree_validN_def /=. setoid_rewrite elem_of_app=> Hv; split=> a Ha.
- destruct (elem_of_agree x2); naive_solver.
- destruct (elem_of_agree x1); naive_solver.
Qed.
Definition agree_cmra_mixin : CmraMixin (agree A).
Proof.
apply cmra_total_mixin; try apply _ || by eauto.
- intros n m x; rewrite !agree_validN_def; eauto using dist_le.
- intros x. apply agree_idemp.
- intros n x y; rewrite !agree_validN_def /=.
setoid_rewrite elem_of_app; naive_solver.
- intros n x y1 y2 Hval Hx; exists x, x; simpl; split.
+ by rewrite agree_idemp.
+ by move: Hval; rewrite Hx; move=> /agree_op_invN->; rewrite agree_idemp.
Qed.
Canonical Structure agreeR : cmra := Cmra (agree A) agree_cmra_mixin.
Global Instance agree_cmra_total : CmraTotal agreeR.
Proof. rewrite /CmraTotal; eauto. Qed.
Global Instance agree_core_id x : CoreId x.
Proof. by constructor. Qed.
Global Instance agree_cmra_discrete : OfeDiscrete A CmraDiscrete agreeR.
Proof.
intros HD. split.
- intros x y [H H'] n; split=> a; setoid_rewrite <-(discrete_iff_0 _ _); auto.
- intros x; rewrite agree_validN_def=> Hv n. apply agree_validN_def=> a b ??.
apply discrete_iff_0; auto.
Qed.
Global Instance to_agree_ne : NonExpansive to_agree.
Proof.
intros n a1 a2 Hx; split=> b /=;
setoid_rewrite elem_of_list_singleton; naive_solver.
Qed.
Global Instance to_agree_proper : Proper (() ==> ()) to_agree := ne_proper _.
Global Instance to_agree_discrete a : Discrete a Discrete (to_agree a).
Proof.
intros ? y [H H'] n; split.
- intros a' ->%elem_of_list_singleton. destruct (H a) as [b ?]; first by left.
exists b. by rewrite -discrete_iff_0.
- intros b Hb. destruct (H' b) as (b'&->%elem_of_list_singleton&?); auto.
exists a. by rewrite elem_of_list_singleton -discrete_iff_0.
Qed.
Lemma agree_op_inv x y : (x y) x y.
Proof.
intros ?. apply equiv_dist=> n. by apply agree_op_invN, cmra_valid_validN.
Qed.
Global Instance to_agree_injN n : Inj (dist n) (dist n) (to_agree).
Proof.
move=> a b [_] /=. setoid_rewrite elem_of_list_singleton. naive_solver.
Qed.
Global Instance to_agree_inj : Inj () () (to_agree).
Proof. intros a b ?. apply equiv_dist=>n. by apply (inj to_agree), equiv_dist. Qed.
Lemma to_agree_uninjN n x : {n} x a, to_agree a {n} x.
Proof.
rewrite agree_validN_def=> Hv.
destruct (elem_of_agree x) as [a ?].
exists a; split=> b /=; setoid_rewrite elem_of_list_singleton; naive_solver.
Qed.
Lemma to_agree_uninj x : x a, to_agree a x.
Proof.
rewrite /valid /agree_valid_instance; setoid_rewrite agree_validN_def.
destruct (elem_of_agree x) as [a ?].
exists a; split=> b /=; setoid_rewrite elem_of_list_singleton; naive_solver.
Qed.
Lemma agree_valid_includedN n x y : {n} y x {n} y x {n} y.
Proof.
move=> Hval [z Hy]; move: Hval; rewrite Hy.
by move=> /agree_op_invN->; rewrite agree_idemp.
Qed.
Lemma agree_valid_included x y : y x y x y.
Proof.
move=> Hval [z Hy]; move: Hval; rewrite Hy.
by move=> /agree_op_inv->; rewrite agree_idemp.
Qed.
Lemma to_agree_includedN n a b : to_agree a {n} to_agree b a {n} b.
Proof.
split; last by intros ->.
intros. by apply (inj to_agree), agree_valid_includedN.
Qed.
Lemma to_agree_included a b : to_agree a to_agree b a b.
Proof.
split; last by intros ->.
intros. by apply (inj to_agree), agree_valid_included.
Qed.
Lemma to_agree_included_L `{!LeibnizEquiv A} a b : to_agree a to_agree b a = b.
Proof. unfold_leibniz. apply to_agree_included. Qed.
Global Instance agree_cancelable x : Cancelable x.
Proof.
intros n y z Hv Heq.
destruct (to_agree_uninjN n x) as [x' EQx]; first by eapply cmra_validN_op_l.
destruct (to_agree_uninjN n y) as [y' EQy]; first by eapply cmra_validN_op_r.
destruct (to_agree_uninjN n z) as [z' EQz].
{ eapply (cmra_validN_op_r n x z). by rewrite -Heq. }
assert (Hx'y' : x' {n} y').
{ apply (inj to_agree), agree_op_invN. by rewrite EQx EQy. }
assert (Hx'z' : x' {n} z').
{ apply (inj to_agree), agree_op_invN. by rewrite EQx EQz -Heq. }
by rewrite -EQy -EQz -Hx'y' -Hx'z'.
Qed.
Lemma to_agree_op_invN a b n : {n} (to_agree a to_agree b) a {n} b.
Proof. by intros ?%agree_op_invN%(inj to_agree). Qed.
Lemma to_agree_op_inv a b : (to_agree a to_agree b) a b.
Proof. by intros ?%agree_op_inv%(inj to_agree). Qed.
Lemma to_agree_op_inv_L `{!LeibnizEquiv A} a b : (to_agree a to_agree b) a = b.
Proof. by intros ?%to_agree_op_inv%leibniz_equiv. Qed.
Lemma to_agree_op_validN a b n : {n} (to_agree a to_agree b) a {n} b.
Proof.
split; first by apply to_agree_op_invN.
intros ->. rewrite agree_idemp //.
Qed.
Lemma to_agree_op_valid a b : (to_agree a to_agree b) a b.
Proof.
split; first by apply to_agree_op_inv.
intros ->. rewrite agree_idemp //.
Qed.
Lemma to_agree_op_valid_L `{!LeibnizEquiv A} a b : (to_agree a to_agree b) a = b.
Proof. rewrite to_agree_op_valid. by fold_leibniz. Qed.
End agree.
Global Instance: Params (@to_agree) 1 := {}.
Global Arguments agreeO {_} _.
Global Arguments agreeR {_} _.
Program Definition agree_map {A B} (f : A B) (x : agree A) : agree B :=
{| agree_car := f <$> agree_car x |}.
Next Obligation. by intros A B f [[|??] ?]. Qed.
Lemma agree_map_id {A} (x : agree A) : agree_map id x = x.
Proof. apply agree_eq. by rewrite /= list_fmap_id. Qed.
Lemma agree_map_compose {A B C} (f : A B) (g : B C) (x : agree A) :
agree_map (g f) x = agree_map g (agree_map f x).
Proof. apply agree_eq. by rewrite /= list_fmap_compose. Qed.
Lemma agree_map_to_agree {A B} (f : A B) (x : A) :
agree_map f (to_agree x) = to_agree (f x).
Proof. by apply agree_eq. Qed.
Section agree_map.
Context {SI : sidx} {A B : ofe} (f : A B) {Hf: NonExpansive f}.
Local Instance agree_map_ne : NonExpansive (agree_map f).
Proof using Type*.
intros n x y [H H']; split=> b /=; setoid_rewrite elem_of_list_fmap.
- intros (a&->&?). destruct (H a) as (a'&?&?); auto. naive_solver.
- intros (a&->&?). destruct (H' a) as (a'&?&?); auto. naive_solver.
Qed.
Local Instance agree_map_proper : Proper (() ==> ()) (agree_map f) := ne_proper _.
Lemma agree_map_ext (g : A B) x :
( a, f a g a) agree_map f x agree_map g x.
Proof using Hf.
intros Hfg n; split=> b /=; setoid_rewrite elem_of_list_fmap.
- intros (a&->&?). exists (g a). rewrite Hfg; eauto.
- intros (a&->&?). exists (f a). rewrite -Hfg; eauto.
Qed.
Global Instance agree_map_morphism : CmraMorphism (agree_map f).
Proof using Hf.
split; first apply _.
- intros n x. rewrite !agree_validN_def=> Hv b b' /=.
intros (a&->&?)%elem_of_list_fmap (a'&->&?)%elem_of_list_fmap.
apply Hf; eauto.
- done.
- intros x y n; split=> b /=;
rewrite !fmap_app; setoid_rewrite elem_of_app; eauto.
Qed.
End agree_map.
Definition agreeO_map {SI : sidx} {A B : ofe}
(f : A -n> B) : agreeO A -n> agreeO B :=
OfeMor (agree_map f : agreeO A agreeO B).
Global Instance agreeO_map_ne {SI : sidx} A B : NonExpansive (@agreeO_map SI A B).
Proof.
intros n f g Hfg x; split=> b /=;
setoid_rewrite elem_of_list_fmap; naive_solver.
Qed.
Program Definition agreeRF {SI : sidx} (F : oFunctor) : rFunctor := {|
rFunctor_car A _ B _ := agreeR (oFunctor_car F A B);
rFunctor_map A1 _ A2 _ B1 _ B2 _ fg := agreeO_map (oFunctor_map F fg)
|}.
Next Obligation.
intros ? ? A1 ? A2 ? B1 ? B2 ? n ???; simpl.
by apply agreeO_map_ne, oFunctor_map_ne.
Qed.
Next Obligation.
intros ? F A ? B ? x; simpl. rewrite -{2}(agree_map_id x).
apply (agree_map_ext _)=>y. by rewrite oFunctor_map_id.
Qed.
Next Obligation.
intros ? F A1 ? A2 ? A3 ? B1 ? B2 ? B3 ? f g f' g' x; simpl.
rewrite -agree_map_compose.
apply (agree_map_ext _)=>y; apply oFunctor_map_compose.
Qed.
Global Instance agreeRF_contractive {SI : sidx} F :
oFunctorContractive F rFunctorContractive (agreeRF F).
Proof.
intros ? A1 ? A2 ? B1 ? B2 ? n ???; simpl.
by apply agreeO_map_ne, oFunctor_map_contractive.
Qed.
From iris.algebra Require Export view frac.
From iris.algebra Require Import proofmode_classes big_op.
From iris.prelude Require Import options.
(** The authoritative camera with fractional authoritative elements *)
(** The authoritative camera has 2 types of elements: the authoritative element
[●{dq} a] and the fragment [◯ b] (of which there can be several). To enable
sharing of the authoritative element [●{dq} a], it is equipped with a
discardable fraction [dq]. Updates are only possible with the full
authoritative element [● a] (syntax for [●{#1} a]]), while fractional
authoritative elements have agreement, i.e., [✓ (●{dq1} a1 ⋅ ●{dq2} a2) → a1 ≡
a2]. *)
(** * Definition of the view relation *)
(** The authoritative camera is obtained by instantiating the view camera. *)
Definition auth_view_rel_raw {SI : sidx} {A : ucmra} (n : SI) (a b : A) : Prop :=
b {n} a {n} a.
Lemma auth_view_rel_raw_mono {SI : sidx} (A : ucmra) n1 n2 (a1 a2 b1 b2 : A) :
auth_view_rel_raw n1 a1 b1
a1 {n2} a2
b2 {n2} b1
(n2 n1)%sidx
auth_view_rel_raw n2 a2 b2.
Proof.
intros [??] Ha12 ??. split.
- trans b1; [done|]. rewrite -Ha12. by apply cmra_includedN_le with n1.
- rewrite -Ha12. by apply cmra_validN_le with n1.
Qed.
Lemma auth_view_rel_raw_valid {SI : sidx} (A : ucmra) n (a b : A) :
auth_view_rel_raw n a b {n} b.
Proof. intros [??]; eauto using cmra_validN_includedN. Qed.
Lemma auth_view_rel_raw_unit {SI : sidx} (A : ucmra) n :
a : A, auth_view_rel_raw n a ε.
Proof. exists ε. split; [done|]. apply ucmra_unit_validN. Qed.
Canonical Structure auth_view_rel {SI : sidx} {A : ucmra} : view_rel A A :=
ViewRel auth_view_rel_raw (auth_view_rel_raw_mono A)
(auth_view_rel_raw_valid A) (auth_view_rel_raw_unit A).
Lemma auth_view_rel_unit {SI : sidx} {A : ucmra} n (a : A) :
auth_view_rel n a ε {n} a.
Proof. split; [by intros [??]|]. split; auto using ucmra_unit_leastN. Qed.
Lemma auth_view_rel_exists {SI : sidx} {A : ucmra} n (b : A) :
( a, auth_view_rel n a b) {n} b.
Proof.
split; [|intros; exists b; by split].
intros [a Hrel]. eapply auth_view_rel_raw_valid, Hrel.
Qed.
Global Instance auth_view_rel_discrete {SI : sidx} {A : ucmra} :
CmraDiscrete A ViewRelDiscrete (auth_view_rel (A:=A)).
Proof.
intros ? n a b [??]; split.
- by apply cmra_discrete_included_iff_0.
- by apply cmra_discrete_valid_iff_0.
Qed.
(** * Definition and operations on the authoritative camera *)
(** The type [auth] is not defined as a [Definition], but as a [Notation].
This way, one can use [auth A] with [A : Type] instead of [A : ucmra], and let
canonical structure search determine the corresponding camera instance. *)
Notation auth A := (view (A:=A) (B:=A) auth_view_rel_raw).
Definition authO {SI : sidx} (A : ucmra) : ofe :=
viewO (A:=A) (B:=A) auth_view_rel.
Definition authR {SI : sidx} (A : ucmra) : cmra :=
viewR (A:=A) (B:=A) auth_view_rel.
Definition authUR {SI : sidx} (A : ucmra) : ucmra :=
viewUR (A:=A) (B:=A) auth_view_rel.
Definition auth_auth {SI : sidx} {A: ucmra} : dfrac A auth A := view_auth.
Definition auth_frag {SI : sidx} {A: ucmra} : A auth A := view_frag.
Global Typeclasses Opaque auth_auth auth_frag.
Global Instance: Params (@auth_auth) 3 := {}.
Global Instance: Params (@auth_frag) 2 := {}.
Notation "● dq a" := (auth_auth dq a)
(at level 20, dq custom dfrac at level 1, format "● dq a").
Notation "◯ a" := (auth_frag a) (at level 20).
(** * Laws of the authoritative camera *)
(** We omit the usual [equivI] lemma because it is hard to state a suitably
general version in terms of [●] and [◯], and because such a lemma has never
been needed in practice. *)
Section auth.
Context {SI : sidx} {A : ucmra}.
Implicit Types a b : A.
Implicit Types x y : auth A.
Implicit Types q : frac.
Implicit Types dq : dfrac.
Global Instance auth_auth_ne dq : NonExpansive (@auth_auth SI A dq).
Proof. rewrite /auth_auth. apply _. Qed.
Global Instance auth_auth_proper dq : Proper (() ==> ()) (@auth_auth SI A dq).
Proof. rewrite /auth_auth. apply _. Qed.
Global Instance auth_frag_ne : NonExpansive (@auth_frag SI A).
Proof. rewrite /auth_frag. apply _. Qed.
Global Instance auth_frag_proper : Proper (() ==> ()) (@auth_frag SI A).
Proof. rewrite /auth_frag. apply _. Qed.
Global Instance auth_auth_dist_inj n :
Inj2 (=) (dist n) (dist n) (@auth_auth SI A).
Proof. rewrite /auth_auth. apply _. Qed.
Global Instance auth_auth_inj : Inj2 (=) () () (@auth_auth SI A).
Proof. rewrite /auth_auth. apply _. Qed.
Global Instance auth_frag_dist_inj n : Inj (dist n) (dist n) (@auth_frag SI A).
Proof. rewrite /auth_frag. apply _. Qed.
Global Instance auth_frag_inj : Inj () () (@auth_frag SI A).
Proof. rewrite /auth_frag. apply _. Qed.
Global Instance auth_ofe_discrete : OfeDiscrete A OfeDiscrete (authO A).
Proof. apply _. Qed.
Global Instance auth_auth_discrete dq a :
Discrete a Discrete (ε : A) Discrete ({dq} a).
Proof. rewrite /auth_auth. apply _. Qed.
Global Instance auth_frag_discrete a : Discrete a Discrete ( a).
Proof. rewrite /auth_frag. apply _. Qed.
Global Instance auth_cmra_discrete : CmraDiscrete A CmraDiscrete (authR A).
Proof. apply _. Qed.
(** Operation *)
Lemma auth_auth_dfrac_op dq1 dq2 a : {dq1 dq2} a {dq1} a {dq2} a.
Proof. apply view_auth_dfrac_op. Qed.
Global Instance auth_auth_dfrac_is_op dq dq1 dq2 a :
IsOp dq dq1 dq2 IsOp' ({dq} a) ({dq1} a) ({dq2} a).
Proof. rewrite /auth_auth. apply _. Qed.
Lemma auth_frag_op a b : (a b) = a b.
Proof. apply view_frag_op. Qed.
Lemma auth_frag_mono a b : a b a b.
Proof. apply view_frag_mono. Qed.
Lemma auth_frag_core a : core ( a) = (core a).
Proof. apply view_frag_core. Qed.
Lemma auth_both_core_discarded a b :
core (●□ a b) ●□ a (core b).
Proof. apply view_both_core_discarded. Qed.
Lemma auth_both_core_frac q a b :
core ({#q} a b) (core b).
Proof. apply view_both_core_frac. Qed.
Global Instance auth_auth_core_id a : CoreId (●□ a).
Proof. rewrite /auth_auth. apply _. Qed.
Global Instance auth_frag_core_id a : CoreId a CoreId ( a).
Proof. rewrite /auth_frag. apply _. Qed.
Global Instance auth_both_core_id a1 a2 : CoreId a2 CoreId (●□ a1 a2).
Proof. rewrite /auth_auth /auth_frag. apply _. Qed.
Global Instance auth_frag_is_op a b1 b2 :
IsOp a b1 b2 IsOp' ( a) ( b1) ( b2).
Proof. rewrite /auth_frag. apply _. Qed.
Global Instance auth_frag_sep_homomorphism :
MonoidHomomorphism op op () (@auth_frag SI A).
Proof. rewrite /auth_frag. apply _. Qed.
Lemma big_opL_auth_frag {B} (g : nat B A) (l : list B) :
( [^op list] kx l, g k x) [^op list] kx l, (g k x).
Proof. apply (big_opL_commute _). Qed.
Lemma big_opM_auth_frag `{Countable K} {B} (g : K B A) (m : gmap K B) :
( [^op map] kx m, g k x) [^op map] kx m, (g k x).
Proof. apply (big_opM_commute _). Qed.
Lemma big_opS_auth_frag `{Countable B} (g : B A) (X : gset B) :
( [^op set] x X, g x) [^op set] x X, (g x).
Proof. apply (big_opS_commute _). Qed.
Lemma big_opMS_auth_frag `{Countable B} (g : B A) (X : gmultiset B) :
( [^op mset] x X, g x) [^op mset] x X, (g x).
Proof. apply (big_opMS_commute _). Qed.
(** Validity *)
Lemma auth_auth_dfrac_op_invN n dq1 a dq2 b : {n} ({dq1} a {dq2} b) a {n} b.
Proof. apply view_auth_dfrac_op_invN. Qed.
Lemma auth_auth_dfrac_op_inv dq1 a dq2 b : ({dq1} a {dq2} b) a b.
Proof. apply view_auth_dfrac_op_inv. Qed.
Lemma auth_auth_dfrac_op_inv_L `{!LeibnizEquiv A} dq1 a dq2 b :
({dq1} a {dq2} b) a = b.
Proof. by apply view_auth_dfrac_op_inv_L. Qed.
Lemma auth_auth_dfrac_validN n dq a : {n} ({dq} a) dq {n} a.
Proof. by rewrite view_auth_dfrac_validN auth_view_rel_unit. Qed.
Lemma auth_auth_validN n a : {n} ( a) {n} a.
Proof. by rewrite view_auth_validN auth_view_rel_unit. Qed.
Lemma auth_auth_dfrac_op_validN n dq1 dq2 a1 a2 :
{n} ({dq1} a1 {dq2} a2) (dq1 dq2) a1 {n} a2 {n} a1.
Proof. by rewrite view_auth_dfrac_op_validN auth_view_rel_unit. Qed.
Lemma auth_auth_op_validN n a1 a2 : {n} ( a1 a2) False.
Proof. apply view_auth_op_validN. Qed.
(** The following lemmas are also stated as implications, which can be used
to force [apply] to use the lemma in the right direction. *)
Lemma auth_frag_validN n b : {n} ( b) {n} b.
Proof. by rewrite view_frag_validN auth_view_rel_exists. Qed.
Lemma auth_frag_validN_1 n b : {n} ( b) {n} b.
Proof. apply auth_frag_validN. Qed.
Lemma auth_frag_validN_2 n b : {n} b {n} ( b).
Proof. apply auth_frag_validN. Qed.
Lemma auth_frag_op_validN n b1 b2 : {n} ( b1 b2) {n} (b1 b2).
Proof. apply auth_frag_validN. Qed.
Lemma auth_frag_op_validN_1 n b1 b2 : {n} ( b1 b2) {n} (b1 b2).
Proof. apply auth_frag_op_validN. Qed.
Lemma auth_frag_op_validN_2 n b1 b2 : {n} (b1 b2) {n} ( b1 b2).
Proof. apply auth_frag_op_validN. Qed.
Lemma auth_both_dfrac_validN n dq a b :
{n} ({dq} a b) dq b {n} a {n} a.
Proof. apply view_both_dfrac_validN. Qed.
Lemma auth_both_validN n a b : {n} ( a b) b {n} a {n} a.
Proof. apply view_both_validN. Qed.
Lemma auth_auth_dfrac_valid dq a : ({dq} a) dq a.
Proof.
rewrite view_auth_dfrac_valid !cmra_valid_validN.
by setoid_rewrite auth_view_rel_unit.
Qed.
Lemma auth_auth_valid a : ( a) a.
Proof.
rewrite view_auth_valid !cmra_valid_validN.
by setoid_rewrite auth_view_rel_unit.
Qed.
Lemma auth_auth_dfrac_op_valid dq1 dq2 a1 a2 :
({dq1} a1 {dq2} a2) (dq1 dq2) a1 a2 a1.
Proof.
rewrite view_auth_dfrac_op_valid !cmra_valid_validN.
setoid_rewrite auth_view_rel_unit. done.
Qed.
Lemma auth_auth_op_valid a1 a2 : ( a1 a2) False.
Proof. apply view_auth_op_valid. Qed.
(** The following lemmas are also stated as implications, which can be used
to force [apply] to use the lemma in the right direction. *)
Lemma auth_frag_valid b : ( b) b.
Proof.
rewrite view_frag_valid cmra_valid_validN.
by setoid_rewrite auth_view_rel_exists.
Qed.
Lemma auth_frag_valid_1 b : ( b) b.
Proof. apply auth_frag_valid. Qed.
Lemma auth_frag_valid_2 b : b ( b).
Proof. apply auth_frag_valid. Qed.
Lemma auth_frag_op_valid b1 b2 : ( b1 b2) (b1 b2).
Proof. apply auth_frag_valid. Qed.
Lemma auth_frag_op_valid_1 b1 b2 : ( b1 b2) (b1 b2).
Proof. apply auth_frag_op_valid. Qed.
Lemma auth_frag_op_valid_2 b1 b2 : (b1 b2) ( b1 b2).
Proof. apply auth_frag_op_valid. Qed.
(** These lemma statements are a bit awkward as we cannot possibly extract a
single witness for [b ≼ a] from validity, we have to make do with one witness
per step-index, i.e., [∀ n, b ≼{n} a]. *)
Lemma auth_both_dfrac_valid dq a b :
({dq} a b) dq ( n, b {n} a) a.
Proof.
rewrite view_both_dfrac_valid. apply and_iff_compat_l. split.
- intros Hrel. split.
+ intros n. by destruct (Hrel n).
+ apply cmra_valid_validN=> n. by destruct (Hrel n).
- intros [Hincl Hval] n. split; [done|by apply cmra_valid_validN].
Qed.
Lemma auth_both_valid a b :
( a b) ( n, b {n} a) a.
Proof. rewrite auth_both_dfrac_valid. split; [naive_solver|done]. Qed.
(* The reverse direction of the two lemmas below only holds if the camera is
discrete. *)
Lemma auth_both_dfrac_valid_2 dq a b : dq a b a ({dq} a b).
Proof.
intros. apply auth_both_dfrac_valid.
naive_solver eauto using cmra_included_includedN.
Qed.
Lemma auth_both_valid_2 a b : a b a ( a b).
Proof. intros ??. by apply auth_both_dfrac_valid_2. Qed.
Lemma auth_both_dfrac_valid_discrete `{!CmraDiscrete A} dq a b :
({dq} a b) dq b a a.
Proof.
rewrite auth_both_dfrac_valid. setoid_rewrite <-cmra_discrete_included_iff.
pose 0. naive_solver.
Qed.
Lemma auth_both_valid_discrete `{!CmraDiscrete A} a b :
( a b) b a a.
Proof. rewrite auth_both_dfrac_valid_discrete. split; [naive_solver|done]. Qed.
(** Inclusion *)
Lemma auth_auth_dfrac_includedN n dq1 dq2 a1 a2 b :
{dq1} a1 {n} {dq2} a2 b (dq1 dq2 dq1 = dq2) a1 {n} a2.
Proof. apply view_auth_dfrac_includedN. Qed.
Lemma auth_auth_dfrac_included dq1 dq2 a1 a2 b :
{dq1} a1 {dq2} a2 b (dq1 dq2 dq1 = dq2) a1 a2.
Proof. apply view_auth_dfrac_included. Qed.
Lemma auth_auth_includedN n a1 a2 b :
a1 {n} a2 b a1 {n} a2.
Proof. apply view_auth_includedN. Qed.
Lemma auth_auth_included a1 a2 b :
a1 a2 b a1 a2.
Proof. apply view_auth_included. Qed.
Lemma auth_frag_includedN n dq a b1 b2 :
b1 {n} {dq} a b2 b1 {n} b2.
Proof. apply view_frag_includedN. Qed.
Lemma auth_frag_included dq a b1 b2 :
b1 {dq} a b2 b1 b2.
Proof. apply view_frag_included. Qed.
(** The weaker [auth_both_included] lemmas below are a consequence of the
[auth_auth_included] and [auth_frag_included] lemmas above. *)
Lemma auth_both_dfrac_includedN n dq1 dq2 a1 a2 b1 b2 :
{dq1} a1 b1 {n} {dq2} a2 b2
(dq1 dq2 dq1 = dq2) a1 {n} a2 b1 {n} b2.
Proof. apply view_both_dfrac_includedN. Qed.
Lemma auth_both_dfrac_included dq1 dq2 a1 a2 b1 b2 :
{dq1} a1 b1 {dq2} a2 b2
(dq1 dq2 dq1 = dq2) a1 a2 b1 b2.
Proof. apply view_both_dfrac_included. Qed.
Lemma auth_both_includedN n a1 a2 b1 b2 :
a1 b1 {n} a2 b2 a1 {n} a2 b1 {n} b2.
Proof. apply view_both_includedN. Qed.
Lemma auth_both_included a1 a2 b1 b2 :
a1 b1 a2 b2 a1 a2 b1 b2.
Proof. apply view_both_included. Qed.
(** Updates *)
Lemma auth_update a b a' b' :
(a,b) ~l~> (a',b') a b ~~> a' b'.
Proof.
intros Hup. apply view_update=> n bf [[bf' Haeq] Hav].
destruct (Hup n (Some (bf bf'))); simpl in *; [done|by rewrite assoc|].
split; [|done]. exists bf'. by rewrite -assoc.
Qed.
Lemma auth_update_alloc a a' b' : (a,ε) ~l~> (a',b') a ~~> a' b'.
Proof. intros. rewrite -(right_id _ _ ( a)). by apply auth_update. Qed.
Lemma auth_update_dealloc a b a' : (a,b) ~l~> (a',ε) a b ~~> a'.
Proof. intros. rewrite -(right_id _ _ ( a')). by apply auth_update. Qed.
Lemma auth_update_auth a a' b' : (a,ε) ~l~> (a',b') a ~~> a'.
Proof.
intros. etrans; first exact: auth_update_alloc.
exact: cmra_update_op_l.
Qed.
Lemma auth_update_auth_persist dq a : {dq} a ~~> ●□ a.
Proof. apply view_update_auth_persist. Qed.
Lemma auth_updateP_auth_unpersist a : ●□ a ~~>: λ k, q, k = {#q} a.
Proof. apply view_updateP_auth_unpersist. Qed.
Lemma auth_updateP_both_unpersist a b : ●□ a b ~~>: λ k, q, k = {#q} a b.
Proof. apply view_updateP_both_unpersist. Qed.
Lemma auth_update_dfrac_alloc dq a b `{!CoreId b} :
b a {dq} a ~~> {dq} a b.
Proof.
intros Ha%(core_id_extract _ _). apply view_update_dfrac_alloc=> n bf [??].
split; [|done]. rewrite Ha (comm _ a). by apply cmra_monoN_l.
Qed.
Lemma auth_local_update a b0 b1 a' b0' b1' :
(b0, b1) ~l~> (b0', b1') b0' a' a'
( a b0, a b1) ~l~> ( a' b0', a' b1').
Proof.
intros. apply view_local_update; [done|]=> n [??]. split.
- by apply cmra_included_includedN.
- by apply cmra_valid_validN.
Qed.
End auth.
(** * Functor *)
Program Definition authURF {SI : sidx} (F : urFunctor) : urFunctor := {|
urFunctor_car A _ B _ := authUR (urFunctor_car F A B);
urFunctor_map A1 _ A2 _ B1 _ B2 _ fg :=
viewO_map (urFunctor_map F fg) (urFunctor_map F fg)
|}.
Next Obligation.
intros ? F A1 ? A2 ? B1 ? B2 ? n f g Hfg.
apply viewO_map_ne; by apply urFunctor_map_ne.
Qed.
Next Obligation.
intros ? F A ? B ? x; simpl in *. rewrite -{2}(view_map_id x).
apply (view_map_ext _ _ _ _)=> y; apply urFunctor_map_id.
Qed.
Next Obligation.
intros ? F A1 ? A2 ? A3 ? B1 ? B2 ? B3 ? f g f' g' x; simpl in *.
rewrite -view_map_compose.
apply (view_map_ext _ _ _ _)=> y; apply urFunctor_map_compose.
Qed.
Next Obligation.
intros ? F A1 ? A2 ? B1 ? B2 ? fg; simpl.
apply view_map_cmra_morphism; [apply _..|]=> n a b [??]; split.
- by apply (cmra_morphism_monotoneN _).
- by apply (cmra_morphism_validN _).
Qed.
Global Instance authURF_contractive {SI : sidx} F :
urFunctorContractive F urFunctorContractive (authURF F).
Proof.
intros ? A1 ? A2 ? B1 ? B2 ? n f g Hfg.
apply viewO_map_ne; by apply urFunctor_map_contractive.
Qed.
Program Definition authRF {SI : sidx} (F : urFunctor) : rFunctor := {|
rFunctor_car A _ B _ := authR (urFunctor_car F A B);
rFunctor_map A1 _ A2 _ B1 _ B2 _ fg :=
viewO_map (urFunctor_map F fg) (urFunctor_map F fg)
|}.
Solve Obligations with apply @authURF.
Global Instance authRF_contractive {SI : sidx} F :
urFunctorContractive F rFunctorContractive (authRF F).
Proof. apply authURF_contractive. Qed.
From stdpp Require Export functions gmap gmultiset.
From iris.algebra Require Export monoid.
From iris.prelude Require Import options.
Local Existing Instances monoid_ne monoid_assoc monoid_comm
monoid_left_id monoid_right_id monoid_proper
monoid_homomorphism_rel_po monoid_homomorphism_rel_proper
monoid_homomorphism_op_proper
monoid_homomorphism_ne weak_monoid_homomorphism_proper.
(** We define the following big operators with binders build in:
- The operator [ [^o list] k ↦ x ∈ l, P ] folds over a list [l]. The binder [x]
refers to each element at index [k].
- The operator [ [^o map] k ↦ x ∈ m, P ] folds over a map [m]. The binder [x]
refers to each element at index [k].
- The operator [ [^o set] x ∈ X, P ] folds over a set [X]. The binder [x] refers
to each element.
Since these big operators are like quantifiers, they have the same precedence as
[∀] and [∃]. *)
(** * Big ops over lists *)
Fixpoint big_opL {SI : sidx} {M : ofe}
{o : M M M} `{!Monoid o} {A} (f : nat A M) (xs : list A) : M :=
match xs with
| [] => monoid_unit
| x :: xs => o (f 0 x) (big_opL (λ n, f (S n)) xs)
end.
Global Instance: Params (@big_opL) 5 := {}.
Global Arguments big_opL {SI} {M} o {_ A} _ !_ /.
Global Typeclasses Opaque big_opL.
Notation "'[^' o 'list]' k ↦ x ∈ l , P" := (big_opL o (λ k x, P) l)
(at level 200, o at level 1, l at level 10, k, x at level 1, right associativity,
format "[^ o list] k ↦ x ∈ l , P") : stdpp_scope.
Notation "'[^' o 'list]' x ∈ l , P" := (big_opL o (λ _ x, P) l)
(at level 200, o at level 1, l at level 10, x at level 1, right associativity,
format "[^ o list] x ∈ l , P") : stdpp_scope.
Local Definition big_opM_def {SI : sidx} {M : ofe}
{o : M M M} `{!Monoid o} `{Countable K} {A} (f : K A M)
(m : gmap K A) : M := big_opL o (λ _, uncurry f) (map_to_list m).
Local Definition big_opM_aux : seal (@big_opM_def). Proof. by eexists. Qed.
Definition big_opM := big_opM_aux.(unseal).
Global Arguments big_opM {SI} {M} o {_ K _ _ A} _ _.
Local Definition big_opM_unseal :
@big_opM = @big_opM_def := big_opM_aux.(seal_eq).
Global Instance: Params (@big_opM) 8 := {}.
Notation "'[^' o 'map]' k ↦ x ∈ m , P" := (big_opM o (λ k x, P) m)
(at level 200, o at level 1, m at level 10, k, x at level 1, right associativity,
format "[^ o map] k ↦ x ∈ m , P") : stdpp_scope.
Notation "'[^' o 'map]' x ∈ m , P" := (big_opM o (λ _ x, P) m)
(at level 200, o at level 1, m at level 10, x at level 1, right associativity,
format "[^ o map] x ∈ m , P") : stdpp_scope.
Local Definition big_opS_def {SI : sidx} {M : ofe}
{o : M M M} `{!Monoid o} `{Countable A} (f : A M)
(X : gset A) : M := big_opL o (λ _, f) (elements X).
Local Definition big_opS_aux : seal (@big_opS_def). Proof. by eexists. Qed.
Definition big_opS := big_opS_aux.(unseal).
Global Arguments big_opS {SI} {M} o {_ A _ _} _ _.
Local Definition big_opS_unseal :
@big_opS = @big_opS_def := big_opS_aux.(seal_eq).
Global Instance: Params (@big_opS) 7 := {}.
Notation "'[^' o 'set]' x ∈ X , P" := (big_opS o (λ x, P) X)
(at level 200, o at level 1, X at level 10, x at level 1, right associativity,
format "[^ o set] x ∈ X , P") : stdpp_scope.
Local Definition big_opMS_def {SI : sidx} {M : ofe}
{o : M M M} `{!Monoid o} `{Countable A} (f : A M)
(X : gmultiset A) : M := big_opL o (λ _, f) (elements X).
Local Definition big_opMS_aux : seal (@big_opMS_def). Proof. by eexists. Qed.
Definition big_opMS := big_opMS_aux.(unseal).
Global Arguments big_opMS {SI} {M} o {_ A _ _} _ _.
Local Definition big_opMS_unseal :
@big_opMS = @big_opMS_def := big_opMS_aux.(seal_eq).
Global Instance: Params (@big_opMS) 8 := {}.
Notation "'[^' o 'mset]' x ∈ X , P" := (big_opMS o (λ x, P) X)
(at level 200, o at level 1, X at level 10, x at level 1, right associativity,
format "[^ o mset] x ∈ X , P") : stdpp_scope.
(** * Properties about big ops *)
Section big_op.
Context {SI : sidx} {M : ofe} {o : M M M} `{!Monoid o}.
Implicit Types xs : list M.
Infix "`o`" := o (at level 50, left associativity).
(** ** Big ops over lists *)
Section list.
Context {A : Type}.
Implicit Types l : list A.
Implicit Types f g : nat A M.
Lemma big_opL_nil f : ([^o list] ky [], f k y) = monoid_unit.
Proof. done. Qed.
Lemma big_opL_cons f x l :
([^o list] ky x :: l, f k y) = f 0 x `o` [^o list] ky l, f (S k) y.
Proof. done. Qed.
Lemma big_opL_singleton f x : ([^o list] ky [x], f k y) f 0 x.
Proof. by rewrite /= right_id. Qed.
Lemma big_opL_app f l1 l2 :
([^o list] ky l1 ++ l2, f k y)
([^o list] ky l1, f k y) `o` ([^o list] ky l2, f (length l1 + k) y).
Proof.
revert f. induction l1 as [|x l1 IH]=> f /=; first by rewrite left_id.
by rewrite IH assoc.
Qed.
Lemma big_opL_snoc f l x :
([^o list] ky l ++ [x], f k y) ([^o list] ky l, f k y) `o` f (length l) x.
Proof. rewrite big_opL_app big_opL_singleton Nat.add_0_r //. Qed.
Lemma big_opL_unit l : ([^o list] ky l, monoid_unit) (monoid_unit : M).
Proof. induction l; rewrite /= ?left_id //. Qed.
Lemma big_opL_take_drop Φ l n :
([^o list] k x l, Φ k x)
([^o list] k x take n l, Φ k x) `o` ([^o list] k x drop n l, Φ (n + k) x).
Proof.
rewrite -{1}(take_drop n l) big_opL_app length_take.
destruct (decide (length l n)).
- rewrite drop_ge //=.
- rewrite Nat.min_l //=; lia.
Qed.
Lemma big_opL_gen_proper_2 {B} (R : relation M) f (g : nat B M)
l1 (l2 : list B) :
R monoid_unit monoid_unit
Proper (R ==> R ==> R) o
( k,
match l1 !! k, l2 !! k with
| Some y1, Some y2 => R (f k y1) (g k y2)
| None, None => True
| _, _ => False
end)
R ([^o list] k y l1, f k y) ([^o list] k y l2, g k y).
Proof.
intros ??. revert l2 f g. induction l1 as [|x1 l1 IH]=> -[|x2 l2] //= f g Hfg.
- by specialize (Hfg 0).
- by specialize (Hfg 0).
- f_equiv; [apply (Hfg 0)|]. apply IH. intros k. apply (Hfg (S k)).
Qed.
Lemma big_opL_gen_proper R f g l :
Reflexive R
Proper (R ==> R ==> R) o
( k y, l !! k = Some y R (f k y) (g k y))
R ([^o list] k y l, f k y) ([^o list] k y l, g k y).
Proof.
intros. apply big_opL_gen_proper_2; [done..|].
intros k. destruct (l !! k) eqn:?; auto.
Qed.
Lemma big_opL_ext f g l :
( k y, l !! k = Some y f k y = g k y)
([^o list] k y l, f k y) = [^o list] k y l, g k y.
Proof. apply big_opL_gen_proper; apply _. Qed.
Lemma big_opL_permutation (f : A M) l1 l2 :
l1 l2 ([^o list] x l1, f x) ([^o list] x l2, f x).
Proof.
induction 1 as [|x xs1 xs2 ? IH|x y xs|xs1 xs2 xs3]; simpl; auto.
- by rewrite IH.
- by rewrite !assoc (comm _ (f x)).
- by etrans.
Qed.
Global Instance big_opL_permutation' (f : A M) :
Proper (() ==> ()) (big_opL o (λ _, f)).
Proof. intros xs1 xs2. apply big_opL_permutation. Qed.
(** The lemmas [big_opL_ne] and [big_opL_proper] are more generic than the
instances as they also give [l !! k = Some y] in the premise. *)
Lemma big_opL_ne f g l n :
( k y, l !! k = Some y f k y {n} g k y)
([^o list] k y l, f k y) {n} ([^o list] k y l, g k y).
Proof. apply big_opL_gen_proper; apply _. Qed.
Lemma big_opL_proper f g l :
( k y, l !! k = Some y f k y g k y)
([^o list] k y l, f k y) ([^o list] k y l, g k y).
Proof. apply big_opL_gen_proper; apply _. Qed.
(** The version [big_opL_proper_2] with [≡] for the list arguments can only be
used if there is a setoid on [A]. The version for [dist n] can be found in
[algebra.list]. We do not define this lemma as a [Proper] instance, since
[f_equiv] will then use sometimes use this one, and other times
[big_opL_proper'], depending on whether a setoid on [A] exists. *)
Lemma big_opL_proper_2 `{!Equiv A} f g l1 l2 :
l1 l2
( k y1 y2,
l1 !! k = Some y1 l2 !! k = Some y2 y1 y2 f k y1 g k y2)
([^o list] k y l1, f k y) ([^o list] k y l2, g k y).
Proof.
intros Hl Hf. apply big_opL_gen_proper_2; try (apply _ || done).
(* FIXME (Coq #14441) unnecessary type annotation *)
intros k. assert (l1 !! k ≡@{option A} l2 !! k) as Hlk by (by f_equiv).
destruct (l1 !! k) eqn:?, (l2 !! k) eqn:?; inversion Hlk; naive_solver.
Qed.
Global Instance big_opL_ne' n :
Proper (pointwise_relation _ (pointwise_relation _ (dist n)) ==> (=) ==> dist n)
(big_opL o (A:=A)).
Proof. intros f f' Hf l ? <-. apply big_opL_ne; intros; apply Hf. Qed.
Global Instance big_opL_proper' :
Proper (pointwise_relation _ (pointwise_relation _ ()) ==> (=) ==> ())
(big_opL o (A:=A)).
Proof. intros f f' Hf l ? <-. apply big_opL_proper; intros; apply Hf. Qed.
Lemma big_opL_consZ_l (f : Z A M) x l :
([^o list] ky x :: l, f k y) = f 0 x `o` [^o list] ky l, f (1 + k)%Z y.
Proof. rewrite big_opL_cons. auto using big_opL_ext with f_equal lia. Qed.
Lemma big_opL_consZ_r (f : Z A M) x l :
([^o list] ky x :: l, f k y) = f 0 x `o` [^o list] ky l, f (k + 1)%Z y.
Proof. rewrite big_opL_cons. auto using big_opL_ext with f_equal lia. Qed.
Lemma big_opL_fmap {B} (h : A B) (f : nat B M) l :
([^o list] ky h <$> l, f k y) ([^o list] ky l, f k (h y)).
Proof. revert f. induction l as [|x l IH]=> f; csimpl=> //. by rewrite IH. Qed.
Lemma big_opL_omap {B} (h : A option B) (f : B M) l :
([^o list] y omap h l, f y) ([^o list] y l, from_option f monoid_unit (h y)).
Proof.
revert f. induction l as [|x l IH]=> f //; csimpl.
case_match; csimpl; by rewrite IH // left_id.
Qed.
Lemma big_opL_op f g l :
([^o list] kx l, f k x `o` g k x)
([^o list] kx l, f k x) `o` ([^o list] kx l, g k x).
Proof.
revert f g; induction l as [|x l IH]=> f g /=; first by rewrite left_id.
by rewrite IH -!assoc (assoc _ (g _ _)) [(g _ _ `o` _)]comm -!assoc.
Qed.
(** Shows that some property [P] is closed under [big_opL]. Examples of [P]
are [Persistent], [Affine], [Timeless]. *)
Lemma big_opL_closed (P : M Prop) f l :
P monoid_unit
( x y, P x P y P (x `o` y))
( k x, l !! k = Some x P (f k x))
P ([^o list] kx l, f k x).
Proof.
intros Hunit Hop. revert f. induction l as [|x l IH]=> f Hf /=; [done|].
apply Hop; first by auto. apply IH=> k. apply (Hf (S k)).
Qed.
End list.
Lemma big_opL_bind {A B} (h : A list B) (f : B M) l :
([^o list] y l ≫= h, f y) ([^o list] x l, [^o list] y h x, f y).
Proof.
revert f. induction l as [|x l IH]=> f; csimpl=> //. by rewrite big_opL_app IH.
Qed.
Lemma big_opL_sep_zip_with {A B C} (f : A B C) (g1 : C A) (g2 : C B)
(h1 : nat A M) (h2 : nat B M) l1 l2 :
( x y, g1 (f x y) = x)
( x y, g2 (f x y) = y)
length l1 = length l2
([^o list] kxy zip_with f l1 l2, h1 k (g1 xy) `o` h2 k (g2 xy))
([^o list] kx l1, h1 k x) `o` ([^o list] ky l2, h2 k y).
Proof.
intros Hlen Hg1 Hg2. rewrite big_opL_op.
rewrite -(big_opL_fmap g1) -(big_opL_fmap g2).
rewrite fmap_zip_with_r; [|auto with lia..].
by rewrite fmap_zip_with_l; [|auto with lia..].
Qed.
Lemma big_opL_sep_zip {A B} (h1 : nat A M) (h2 : nat B M) l1 l2 :
length l1 = length l2
([^o list] kxy zip l1 l2, h1 k xy.1 `o` h2 k xy.2)
([^o list] kx l1, h1 k x) `o` ([^o list] ky l2, h2 k y).
Proof. by apply big_opL_sep_zip_with. Qed.
(** ** Big ops over finite maps *)
Lemma big_opM_empty `{Countable K} {B} (f : K B M) :
([^o map] kx , f k x) = monoid_unit.
Proof. by rewrite big_opM_unseal /big_opM_def map_to_list_empty. Qed.
Lemma big_opM_insert `{Countable K} {B} (f : K B M) (m : gmap K B) i x :
m !! i = None
([^o map] ky <[i:=x]> m, f k y) f i x `o` [^o map] ky m, f k y.
Proof. intros ?. by rewrite big_opM_unseal /big_opM_def map_to_list_insert. Qed.
Lemma big_opM_delete `{Countable K} {B} (f : K B M) (m : gmap K B) i x :
m !! i = Some x
([^o map] ky m, f k y) f i x `o` [^o map] ky delete i m, f k y.
Proof.
intros. rewrite -big_opM_insert ?lookup_delete //.
by rewrite insert_delete.
Qed.
Section gmap.
Context `{Countable K} {A : Type}.
Implicit Types m : gmap K A.
Implicit Types f g : K A M.
Lemma big_opM_gen_proper_2 {B} (R : relation M) f (g : K B M)
m1 (m2 : gmap K B) :
subrelation () R Equivalence R
Proper (R ==> R ==> R) o
( k,
match m1 !! k, m2 !! k with
| Some y1, Some y2 => R (f k y1) (g k y2)
| None, None => True
| _, _ => False
end)
R ([^o map] k x m1, f k x) ([^o map] k x m2, g k x).
Proof.
intros HR ??. revert m2 f g.
induction m1 as [|k x1 m1 Hm1k IH] using map_ind=> m2 f g Hfg.
{ destruct m2 as [|k x2 m2 _ _] using map_ind.
{ rewrite !big_opM_empty. by apply HR. }
generalize (Hfg k). by rewrite lookup_empty lookup_insert. }
generalize (Hfg k). rewrite lookup_insert.
destruct (m2 !! k) as [x2|] eqn:Hm2k; [intros Hk|done].
etrans; [by apply HR, big_opM_insert|].
etrans; [|by symmetry; apply HR, big_opM_delete].
f_equiv; [done|]. apply IH=> k'. destruct (decide (k = k')) as [->|?].
- by rewrite lookup_delete Hm1k.
- generalize (Hfg k'). rewrite lookup_insert_ne // lookup_delete_ne //.
Qed.
Lemma big_opM_gen_proper R f g m :
Reflexive R
Proper (R ==> R ==> R) o
( k x, m !! k = Some x R (f k x) (g k x))
R ([^o map] k x m, f k x) ([^o map] k x m, g k x).
Proof.
intros ?? Hf. rewrite big_opM_unseal. apply (big_opL_gen_proper R); auto.
intros k [i x] ?%elem_of_list_lookup_2. by apply Hf, elem_of_map_to_list.
Qed.
Lemma big_opM_ext f g m :
( k x, m !! k = Some x f k x = g k x)
([^o map] k x m, f k x) = ([^o map] k x m, g k x).
Proof. apply big_opM_gen_proper; apply _. Qed.
(** The lemmas [big_opM_ne] and [big_opM_proper] are more generic than the
instances as they also give [m !! k = Some y] in the premise. *)
Lemma big_opM_ne f g m n :
( k x, m !! k = Some x f k x {n} g k x)
([^o map] k x m, f k x) {n} ([^o map] k x m, g k x).
Proof. apply big_opM_gen_proper; apply _. Qed.
Lemma big_opM_proper f g m :
( k x, m !! k = Some x f k x g k x)
([^o map] k x m, f k x) ([^o map] k x m, g k x).
Proof. apply big_opM_gen_proper; apply _. Qed.
(** The version [big_opM_proper_2] with [≡] for the map arguments can only be
used if there is a setoid on [A]. The version for [dist n] can be found in
[algebra.gmap]. We do not define this lemma as a [Proper] instance, since
[f_equiv] will then use sometimes use this one, and other times
[big_opM_proper'], depending on whether a setoid on [A] exists. *)
Lemma big_opM_proper_2 `{!Equiv A} f g m1 m2 :
m1 m2
( k y1 y2,
m1 !! k = Some y1 m2 !! k = Some y2 y1 y2 f k y1 g k y2)
([^o map] k y m1, f k y) ([^o map] k y m2, g k y).
Proof.
intros Hl Hf. apply big_opM_gen_proper_2; try (apply _ || done).
(* FIXME (Coq #14441) unnecessary type annotation *)
intros k. assert (m1 !! k ≡@{option A} m2 !! k) as Hlk by (by f_equiv).
destruct (m1 !! k) eqn:?, (m2 !! k) eqn:?; inversion Hlk; naive_solver.
Qed.
Global Instance big_opM_ne' n :
Proper (pointwise_relation _ (pointwise_relation _ (dist n)) ==> (=) ==> dist n)
(big_opM o (K:=K) (A:=A)).
Proof. intros f g Hf m ? <-. apply big_opM_ne; intros; apply Hf. Qed.
Global Instance big_opM_proper' :
Proper (pointwise_relation _ (pointwise_relation _ ()) ==> (=) ==> ())
(big_opM o (K:=K) (A:=A)).
Proof. intros f g Hf m ? <-. apply big_opM_proper; intros; apply Hf. Qed.
(* FIXME: This lemma could be generalized from [≡] to [=], but that breaks
[setoid_rewrite] in the proof of [big_sepS_sepS]. See Coq issue #14349. *)
Lemma big_opM_map_to_list f m :
([^o map] kx m, f k x) [^o list] xk map_to_list m, f (xk.1) (xk.2).
Proof. rewrite big_opM_unseal. apply big_opL_proper'; [|done]. by intros ? [??]. Qed.
Lemma big_opM_list_to_map f l :
NoDup l.*1
([^o map] kx list_to_map l, f k x) [^o list] xk l, f (xk.1) (xk.2).
Proof.
intros. rewrite big_opM_map_to_list.
by apply big_opL_permutation, map_to_list_to_map.
Qed.
Lemma big_opM_singleton f i x : ([^o map] ky {[i:=x]}, f k y) f i x.
Proof.
rewrite -insert_empty big_opM_insert/=; last eauto using lookup_empty.
by rewrite big_opM_empty right_id.
Qed.
Lemma big_opM_unit m : ([^o map] ky m, monoid_unit) (monoid_unit : M).
Proof.
by induction m using map_ind; rewrite /= ?big_opM_insert ?left_id // big_opM_unseal.
Qed.
Lemma big_opM_fmap {B} (h : A B) (f : K B M) m :
([^o map] ky h <$> m, f k y) ([^o map] ky m, f k (h y)).
Proof.
rewrite big_opM_unseal /big_opM_def map_to_list_fmap big_opL_fmap.
by apply big_opL_proper=> ? [??].
Qed.
Lemma big_opM_omap {B} (h : A option B) (f : K B M) m :
([^o map] ky omap h m, f k y)
[^o map] ky m, from_option (f k) monoid_unit (h y).
Proof.
revert f. induction m as [|i x m Hmi IH] using map_ind=> f.
{ by rewrite omap_empty !big_opM_empty. }
assert (omap h m !! i = None) by (by rewrite lookup_omap Hmi).
destruct (h x) as [y|] eqn:Hhx.
- by rewrite omap_insert Hhx //= !big_opM_insert // IH Hhx.
- rewrite omap_insert_None // delete_notin // big_opM_insert //.
by rewrite Hhx /= left_id.
Qed.
Lemma big_opM_insert_delete `{Countable K} {B} (f : K B M) (m : gmap K B) i x :
([^o map] ky <[i:=x]> m, f k y) f i x `o` [^o map] ky delete i m, f k y.
Proof.
rewrite -insert_delete_insert big_opM_insert; first done. by rewrite lookup_delete.
Qed.
Lemma big_opM_insert_override (f : K A M) m i x x' :
m !! i = Some x f i x f i x'
([^o map] ky <[i:=x']> m, f k y) ([^o map] ky m, f k y).
Proof.
intros ? Hx. rewrite -insert_delete_insert big_opM_insert ?lookup_delete //.
by rewrite -Hx -big_opM_delete.
Qed.
Lemma big_opM_fn_insert {B} (g : K A B M) (f : K B) m i (x : A) b :
m !! i = None
([^o map] ky <[i:=x]> m, g k y (<[i:=b]> f k))
g i x b `o` [^o map] ky m, g k y (f k).
Proof.
intros. rewrite big_opM_insert // fn_lookup_insert.
f_equiv; apply big_opM_proper; auto=> k y ?.
by rewrite fn_lookup_insert_ne; last set_solver.
Qed.
Lemma big_opM_fn_insert' (f : K M) m i x P :
m !! i = None
([^o map] ky <[i:=x]> m, <[i:=P]> f k) (P `o` [^o map] ky m, f k).
Proof. apply (big_opM_fn_insert (λ _ _, id)). Qed.
Lemma big_opM_filter' (φ : K * A Prop) `{ kx, Decision (φ kx)} f m :
([^o map] k x filter φ m, f k x)
([^o map] k x m, if decide (φ (k, x)) then f k x else monoid_unit).
Proof.
induction m as [|k v m ? IH] using map_ind.
{ by rewrite map_filter_empty !big_opM_empty. }
destruct (decide (φ (k, v))).
- rewrite map_filter_insert_True //.
assert (filter φ m !! k = None) by (apply map_lookup_filter_None; eauto).
by rewrite !big_opM_insert // decide_True // IH.
- rewrite map_filter_insert_not' //; last by congruence.
rewrite !big_opM_insert // decide_False // IH. by rewrite left_id.
Qed.
Lemma big_opM_union f m1 m2 :
m1 ## m2
([^o map] ky m1 m2, f k y)
([^o map] ky m1, f k y) `o` ([^o map] ky m2, f k y).
Proof.
intros. induction m1 as [|i x m ? IH] using map_ind.
{ by rewrite big_opM_empty !left_id. }
decompose_map_disjoint.
rewrite -insert_union_l !big_opM_insert //;
last by apply lookup_union_None.
rewrite -assoc IH //.
Qed.
Lemma big_opM_op f g m :
([^o map] kx m, f k x `o` g k x)
([^o map] kx m, f k x) `o` ([^o map] kx m, g k x).
Proof.
rewrite big_opM_unseal /big_opM_def -big_opL_op. by apply big_opL_proper=> ? [??].
Qed.
(** Shows that some property [P] is closed under [big_opM]. Examples of [P]
are [Persistent], [Affine], [Timeless]. *)
Lemma big_opM_closed (P : M Prop) f m :
Proper (() ==> iff) P
P monoid_unit
( x y, P x P y P (x `o` y))
( k x, m !! k = Some x P (f k x))
P ([^o map] kx m, f k x).
Proof.
intros ?? Hop Hf. induction m as [|k x ?? IH] using map_ind.
{ by rewrite big_opM_empty. }
rewrite big_opM_insert //. apply Hop.
{ apply Hf. by rewrite lookup_insert. }
apply IH=> k' x' ?. apply Hf. rewrite lookup_insert_ne; naive_solver.
Qed.
End gmap.
Lemma big_opM_sep_zip_with `{Countable K} {A B C}
(f : A B C) (g1 : C A) (g2 : C B)
(h1 : K A M) (h2 : K B M) m1 m2 :
( x y, g1 (f x y) = x)
( x y, g2 (f x y) = y)
( k, is_Some (m1 !! k) is_Some (m2 !! k))
([^o map] kxy map_zip_with f m1 m2, h1 k (g1 xy) `o` h2 k (g2 xy))
([^o map] kx m1, h1 k x) `o` ([^o map] ky m2, h2 k y).
Proof.
intros Hdom Hg1 Hg2. rewrite big_opM_op.
rewrite -(big_opM_fmap g1) -(big_opM_fmap g2).
rewrite map_fmap_zip_with_r; [|naive_solver..].
by rewrite map_fmap_zip_with_l; [|naive_solver..].
Qed.
Lemma big_opM_sep_zip `{Countable K} {A B}
(h1 : K A M) (h2 : K B M) m1 m2 :
( k, is_Some (m1 !! k) is_Some (m2 !! k))
([^o map] kxy map_zip m1 m2, h1 k xy.1 `o` h2 k xy.2)
([^o map] kx m1, h1 k x) `o` ([^o map] ky m2, h2 k y).
Proof. intros. by apply big_opM_sep_zip_with. Qed.
(** ** Big ops over finite sets *)
Section gset.
Context `{Countable A}.
Implicit Types X : gset A.
Implicit Types f : A M.
Lemma big_opS_gen_proper R f g X :
Reflexive R Proper (R ==> R ==> R) o
( x, x X R (f x) (g x))
R ([^o set] x X, f x) ([^o set] x X, g x).
Proof.
rewrite big_opS_unseal. intros ?? Hf. apply (big_opL_gen_proper R); auto.
intros k x ?%elem_of_list_lookup_2. by apply Hf, elem_of_elements.
Qed.
Lemma big_opS_ext f g X :
( x, x X f x = g x)
([^o set] x X, f x) = ([^o set] x X, g x).
Proof. apply big_opS_gen_proper; apply _. Qed.
(** The lemmas [big_opS_ne] and [big_opS_proper] are more generic than the
instances as they also give [x ∈ X] in the premise. *)
Lemma big_opS_ne f g X n :
( x, x X f x {n} g x)
([^o set] x X, f x) {n} ([^o set] x X, g x).
Proof. apply big_opS_gen_proper; apply _. Qed.
Lemma big_opS_proper f g X :
( x, x X f x g x)
([^o set] x X, f x) ([^o set] x X, g x).
Proof. apply big_opS_gen_proper; apply _. Qed.
Global Instance big_opS_ne' n :
Proper (pointwise_relation _ (dist n) ==> (=) ==> dist n) (big_opS o (A:=A)).
Proof. intros f g Hf m ? <-. apply big_opS_ne; intros; apply Hf. Qed.
Global Instance big_opS_proper' :
Proper (pointwise_relation _ () ==> (=) ==> ()) (big_opS o (A:=A)).
Proof. intros f g Hf m ? <-. apply big_opS_proper; intros; apply Hf. Qed.
(* FIXME: This lemma could be generalized from [≡] to [=], but that breaks
[setoid_rewrite] in the proof of [big_sepS_sepS]. See Coq issue #14349. *)
Lemma big_opS_elements f X :
([^o set] x X, f x) [^o list] x elements X, f x.
Proof. by rewrite big_opS_unseal. Qed.
Lemma big_opS_empty f : ([^o set] x , f x) = monoid_unit.
Proof. by rewrite big_opS_unseal /big_opS_def elements_empty. Qed.
Lemma big_opS_insert f X x :
x X ([^o set] y {[ x ]} X, f y) (f x `o` [^o set] y X, f y).
Proof. intros. by rewrite !big_opS_elements elements_union_singleton. Qed.
Lemma big_opS_fn_insert {B} (f : A B M) h X x b :
x X
([^o set] y {[ x ]} X, f y (<[x:=b]> h y))
f x b `o` [^o set] y X, f y (h y).
Proof.
intros. rewrite big_opS_insert // fn_lookup_insert.
f_equiv; apply big_opS_proper; auto=> y ?.
by rewrite fn_lookup_insert_ne; last set_solver.
Qed.
Lemma big_opS_fn_insert' f X x P :
x X ([^o set] y {[ x ]} X, <[x:=P]> f y) (P `o` [^o set] y X, f y).
Proof. apply (big_opS_fn_insert (λ y, id)). Qed.
Lemma big_opS_union f X Y :
X ## Y
([^o set] y X Y, f y) ([^o set] y X, f y) `o` ([^o set] y Y, f y).
Proof.
intros. induction X as [|x X ? IH] using set_ind_L.
{ by rewrite left_id_L big_opS_empty left_id. }
rewrite -assoc_L !big_opS_insert; [|set_solver..].
by rewrite -assoc IH; last set_solver.
Qed.
Lemma big_opS_delete f X x :
x X ([^o set] y X, f y) f x `o` [^o set] y X {[ x ]}, f y.
Proof.
intros. rewrite -big_opS_insert; last set_solver.
by rewrite -union_difference_L; last set_solver.
Qed.
Lemma big_opS_singleton f x : ([^o set] y {[ x ]}, f y) f x.
Proof. intros. by rewrite big_opS_elements elements_singleton /= right_id. Qed.
Lemma big_opS_unit X : ([^o set] y X, monoid_unit) (monoid_unit : M).
Proof.
by induction X using set_ind_L; rewrite /= ?big_opS_insert ?left_id // big_opS_unseal.
Qed.
Lemma big_opS_filter' (φ : A Prop) `{ x, Decision (φ x)} f X :
([^o set] y filter φ X, f y)
([^o set] y X, if decide (φ y) then f y else monoid_unit).
Proof.
induction X as [|x X ? IH] using set_ind_L.
{ by rewrite filter_empty_L !big_opS_empty. }
destruct (decide (φ x)).
- rewrite filter_union_L filter_singleton_L //.
rewrite !big_opS_insert //; last set_solver.
by rewrite decide_True // IH.
- rewrite filter_union_L filter_singleton_not_L // left_id_L.
by rewrite !big_opS_insert // decide_False // IH left_id.
Qed.
Lemma big_opS_op f g X :
([^o set] y X, f y `o` g y) ([^o set] y X, f y) `o` ([^o set] y X, g y).
Proof. by rewrite !big_opS_elements -big_opL_op. Qed.
Lemma big_opS_list_to_set f (l : list A) :
NoDup l
([^o set] x list_to_set l, f x) [^o list] x l, f x.
Proof.
induction 1 as [|x l ?? IHl].
- rewrite big_opS_empty //.
- rewrite /= big_opS_union; last set_solver.
by rewrite big_opS_singleton IHl.
Qed.
(** Shows that some property [P] is closed under [big_opS]. Examples of [P]
are [Persistent], [Affine], [Timeless]. *)
Lemma big_opS_closed (P : M Prop) f X :
Proper (() ==> iff) P
P monoid_unit
( x y, P x P y P (x `o` y))
( x, x X P (f x))
P ([^o set] x X, f x).
Proof.
intros ?? Hop Hf. induction X as [|x X ? IH] using set_ind_L.
{ by rewrite big_opS_empty. }
rewrite big_opS_insert //. apply Hop.
{ apply Hf. set_solver. }
apply IH=> x' ?. apply Hf. set_solver.
Qed.
End gset.
Lemma big_opS_set_map `{Countable A, Countable B} (h : A B) (X : gset A) (f : B M) :
Inj (=) (=) h
([^o set] x set_map h X, f x) ([^o set] x X, f (h x)).
Proof.
intros Hinj.
induction X as [|x X ? IH] using set_ind_L.
{ by rewrite set_map_empty !big_opS_empty. }
rewrite set_map_union_L set_map_singleton_L.
rewrite !big_opS_union; [|set_solver..].
rewrite !big_opS_singleton IH //.
Qed.
Lemma big_opM_dom `{Countable K} {A} (f : K M) (m : gmap K A) :
([^o map] k↦_ m, f k) ([^o set] k dom m, f k).
Proof.
induction m as [|i x ?? IH] using map_ind.
{ by rewrite big_opM_unseal big_opS_unseal dom_empty_L. }
by rewrite dom_insert_L big_opM_insert // IH big_opS_insert ?not_elem_of_dom.
Qed.
Lemma big_opM_gset_to_gmap `{Countable K} {A} (f : K A M) (X : gset K) c :
([^o map] ka gset_to_gmap c X, f k a) ([^o set] k X, f k c).
Proof.
rewrite -{2}(dom_gset_to_gmap X c) -big_opM_dom.
apply big_opM_proper. by intros k ? [_ ->]%lookup_gset_to_gmap_Some.
Qed.
(** ** Big ops over finite msets *)
Section gmultiset.
Context `{Countable A}.
Implicit Types X : gmultiset A.
Implicit Types f : A M.
Lemma big_opMS_gen_proper R f g X :
Reflexive R Proper (R ==> R ==> R) o
( x, x X R (f x) (g x))
R ([^o mset] x X, f x) ([^o mset] x X, g x).
Proof.
rewrite big_opMS_unseal. intros ?? Hf. apply (big_opL_gen_proper R); auto.
intros k x ?%elem_of_list_lookup_2. by apply Hf, gmultiset_elem_of_elements.
Qed.
Lemma big_opMS_ext f g X :
( x, x X f x = g x)
([^o mset] x X, f x) = ([^o mset] x X, g x).
Proof. apply big_opMS_gen_proper; apply _. Qed.
(** The lemmas [big_opMS_ne] and [big_opMS_proper] are more generic than the
instances as they also give [x ∈ X] in the premise. *)
Lemma big_opMS_ne f g X n :
( x, x X f x {n} g x)
([^o mset] x X, f x) {n} ([^o mset] x X, g x).
Proof. apply big_opMS_gen_proper; apply _. Qed.
Lemma big_opMS_proper f g X :
( x, x X f x g x)
([^o mset] x X, f x) ([^o mset] x X, g x).
Proof. apply big_opMS_gen_proper; apply _. Qed.
Global Instance big_opMS_ne' n :
Proper (pointwise_relation _ (dist n) ==> (=) ==> dist n) (big_opMS o (A:=A)).
Proof. intros f g Hf m ? <-. apply big_opMS_ne; intros; apply Hf. Qed.
Global Instance big_opMS_proper' :
Proper (pointwise_relation _ () ==> (=) ==> ()) (big_opMS o (A:=A)).
Proof. intros f g Hf m ? <-. apply big_opMS_proper; intros; apply Hf. Qed.
(* FIXME: This lemma could be generalized from [≡] to [=], but that breaks
[setoid_rewrite] in the proof of [big_sepS_sepS]. See Coq issue #14349. *)
Lemma big_opMS_elements f X :
([^o mset] x X, f x) [^o list] x elements X, f x.
Proof. by rewrite big_opMS_unseal. Qed.
Lemma big_opMS_empty f : ([^o mset] x , f x) = monoid_unit.
Proof. by rewrite big_opMS_unseal /big_opMS_def gmultiset_elements_empty. Qed.
Lemma big_opMS_disj_union f X Y :
([^o mset] y X Y, f y) ([^o mset] y X, f y) `o` [^o mset] y Y, f y.
Proof. by rewrite big_opMS_unseal /big_opMS_def gmultiset_elements_disj_union big_opL_app. Qed.
Lemma big_opMS_singleton f x : ([^o mset] y {[+ x +]}, f y) f x.
Proof.
intros. by rewrite big_opMS_unseal /big_opMS_def gmultiset_elements_singleton /= right_id.
Qed.
Lemma big_opMS_insert f X x :
([^o mset] y {[+ x +]} X, f y) (f x `o` [^o mset] y X, f y).
Proof. intros. rewrite big_opMS_disj_union big_opMS_singleton //. Qed.
Lemma big_opMS_delete f X x :
x X ([^o mset] y X, f y) f x `o` [^o mset] y X {[+ x +]}, f y.
Proof.
intros. rewrite -big_opMS_singleton -big_opMS_disj_union.
by rewrite -gmultiset_disj_union_difference'.
Qed.
Lemma big_opMS_unit X : ([^o mset] y X, monoid_unit) (monoid_unit : M).
Proof.
by induction X using gmultiset_ind;
rewrite /= ?big_opMS_disj_union ?big_opMS_singleton ?left_id // big_opMS_unseal.
Qed.
Lemma big_opMS_op f g X :
([^o mset] y X, f y `o` g y) ([^o mset] y X, f y) `o` ([^o mset] y X, g y).
Proof. by rewrite big_opMS_unseal /big_opMS_def -big_opL_op. Qed.
(** Shows that some property [P] is closed under [big_opMS]. Examples of [P]
are [Persistent], [Affine], [Timeless]. *)
Lemma big_opMS_closed (P : M Prop) f X :
Proper (() ==> iff) P
P monoid_unit
( x y, P x P y P (x `o` y))
( x, x X P (f x))
P ([^o mset] x X, f x).
Proof.
intros ?? Hop Hf. induction X as [|x X IH] using gmultiset_ind.
{ by rewrite big_opMS_empty. }
rewrite big_opMS_insert //. apply Hop.
{ apply Hf. set_solver. }
apply IH=> x' ?. apply Hf. set_solver.
Qed.
End gmultiset.
(** Commuting lemmas *)
Lemma big_opL_opL {A B} (f : nat A nat B M) (l1 : list A) (l2 : list B) :
([^o list] k1x1 l1, [^o list] k2x2 l2, f k1 x1 k2 x2)
([^o list] k2x2 l2, [^o list] k1x1 l1, f k1 x1 k2 x2).
Proof.
revert f l2. induction l1 as [|x1 l1 IH]; simpl; intros Φ l2.
{ by rewrite big_opL_unit. }
by rewrite IH big_opL_op.
Qed.
Lemma big_opL_opM {A} `{Countable K} {B}
(f : nat A K B M) (l1 : list A) (m2 : gmap K B) :
([^o list] k1x1 l1, [^o map] k2x2 m2, f k1 x1 k2 x2)
([^o map] k2x2 m2, [^o list] k1x1 l1, f k1 x1 k2 x2).
Proof. repeat setoid_rewrite big_opM_map_to_list. by rewrite big_opL_opL. Qed.
Lemma big_opL_opS {A} `{Countable B}
(f : nat A B M) (l1 : list A) (X2 : gset B) :
([^o list] k1x1 l1, [^o set] x2 X2, f k1 x1 x2)
([^o set] x2 X2, [^o list] k1x1 l1, f k1 x1 x2).
Proof. repeat setoid_rewrite big_opS_elements. by rewrite big_opL_opL. Qed.
Lemma big_opL_opMS {A} `{Countable B}
(f : nat A B M) (l1 : list A) (X2 : gmultiset B) :
([^o list] k1x1 l1, [^o mset] x2 X2, f k1 x1 x2)
([^o mset] x2 X2, [^o list] k1x1 l1, f k1 x1 x2).
Proof. repeat setoid_rewrite big_opMS_elements. by rewrite big_opL_opL. Qed.
Lemma big_opM_opL {A} `{Countable K} {B}
(f : K A nat B M) (m1 : gmap K A) (l2 : list B) :
([^o map] k1x1 m1, [^o list] k2x2 l2, f k1 x1 k2 x2)
([^o list] k2x2 l2, [^o map] k1x1 m1, f k1 x1 k2 x2).
Proof. symmetry. apply big_opL_opM. Qed.
Lemma big_opM_opM `{Countable K1} {A} `{Countable K2} {B}
(f : K1 A K2 B M) (m1 : gmap K1 A) (m2 : gmap K2 B) :
([^o map] k1x1 m1, [^o map] k2x2 m2, f k1 x1 k2 x2)
([^o map] k2x2 m2, [^o map] k1x1 m1, f k1 x1 k2 x2).
Proof. repeat setoid_rewrite big_opM_map_to_list. by rewrite big_opL_opL. Qed.
Lemma big_opM_opS `{Countable K} {A} `{Countable B}
(f : K A B M) (m1 : gmap K A) (X2 : gset B) :
([^o map] k1x1 m1, [^o set] x2 X2, f k1 x1 x2)
([^o set] x2 X2, [^o map] k1x1 m1, f k1 x1 x2).
Proof.
repeat setoid_rewrite big_opM_map_to_list.
repeat setoid_rewrite big_opS_elements. by rewrite big_opL_opL.
Qed.
Lemma big_opM_opMS `{Countable K} {A} `{Countable B} (f : K A B M)
(m1 : gmap K A) (X2 : gmultiset B) :
([^o map] k1x1 m1, [^o mset] x2 X2, f k1 x1 x2)
([^o mset] x2 X2, [^o map] k1x1 m1, f k1 x1 x2).
Proof.
repeat setoid_rewrite big_opM_map_to_list.
repeat setoid_rewrite big_opMS_elements. by rewrite big_opL_opL.
Qed.
Lemma big_opS_opL `{Countable A} {B}
(f : A nat B M) (X1 : gset A) (l2 : list B) :
([^o set] x1 X1, [^o list] k2x2 l2, f x1 k2 x2)
([^o list] k2x2 l2, [^o set] x1 X1, f x1 k2 x2).
Proof. symmetry. apply big_opL_opS. Qed.
Lemma big_opS_opM `{Countable A} `{Countable K} {B}
(f : A K B M) (X1 : gset A) (m2 : gmap K B) :
([^o set] x1 X1, [^o map] k2x2 m2, f x1 k2 x2)
([^o map] k2x2 m2, [^o set] x1 X1, f x1 k2 x2).
Proof. symmetry. apply big_opM_opS. Qed.
Lemma big_opS_opS `{Countable A, Countable B}
(X : gset A) (Y : gset B) (f : A B M) :
([^o set] x X, [^o set] y Y, f x y) ([^o set] y Y, [^o set] x X, f x y).
Proof. repeat setoid_rewrite big_opS_elements. by rewrite big_opL_opL. Qed.
Lemma big_opS_opMS `{Countable A, Countable B}
(X : gset A) (Y : gmultiset B) (f : A B M) :
([^o set] x X, [^o mset] y Y, f x y) ([^o mset] y Y, [^o set] x X, f x y).
Proof.
repeat setoid_rewrite big_opS_elements.
repeat setoid_rewrite big_opMS_elements. by rewrite big_opL_opL.
Qed.
Lemma big_opMS_opL `{Countable A} {B}
(f : A nat B M) (X1 : gmultiset A) (l2 : list B) :
([^o mset] x1 X1, [^o list] k2x2 l2, f x1 k2 x2)
([^o list] k2x2 l2, [^o mset] x1 X1, f x1 k2 x2).
Proof. symmetry. apply big_opL_opMS. Qed.
Lemma big_opMS_opM `{Countable A} `{Countable K} {B} (f : A K B M)
(X1 : gmultiset A) (m2 : gmap K B) :
([^o mset] x1 X1, [^o map] k2x2 m2, f x1 k2 x2)
([^o map] k2x2 m2, [^o mset] x1 X1, f x1 k2 x2).
Proof. symmetry. apply big_opM_opMS. Qed.
Lemma big_opMS_opS `{Countable A, Countable B}
(X : gmultiset A) (Y : gset B) (f : A B M) :
([^o mset] x X, [^o set] y Y, f x y) ([^o set] y Y, [^o mset] x X, f x y).
Proof. symmetry. apply big_opS_opMS. Qed.
Lemma big_opMS_opMS `{Countable A, Countable B}
(X : gmultiset A) (Y : gmultiset B) (f : A B M) :
([^o mset] x X, [^o mset] y Y, f x y) ([^o mset] y Y, [^o mset] x X, f x y).
Proof. repeat setoid_rewrite big_opMS_elements. by rewrite big_opL_opL. Qed.
End big_op.
Section homomorphisms.
Context {SI : sidx} {M1 M2 : ofe}.
Context {o1 : M1 M1 M1} {o2 : M2 M2 M2} `{!Monoid o1, !Monoid o2}.
Infix "`o1`" := o1 (at level 50, left associativity).
Infix "`o2`" := o2 (at level 50, left associativity).
(** The ssreflect rewrite tactic only works for relations that have a
[RewriteRelation] instance. For the purpose of this section, we want to
rewrite with arbitrary relations, so we declare any relation to be a
[RewriteRelation]. *)
Local Instance: {A} (R : relation A), RewriteRelation R := {}.
Lemma big_opL_commute {A} (h : M1 M2) `{!MonoidHomomorphism o1 o2 R h}
(f : nat A M1) l :
R (h ([^o1 list] kx l, f k x)) ([^o2 list] kx l, h (f k x)).
Proof.
revert f. induction l as [|x l IH]=> f /=.
- apply monoid_homomorphism_unit.
- by rewrite monoid_homomorphism IH.
Qed.
Lemma big_opL_commute1 {A} (h : M1 M2) `{!WeakMonoidHomomorphism o1 o2 R h}
(f : nat A M1) l :
l [] R (h ([^o1 list] kx l, f k x)) ([^o2 list] kx l, h (f k x)).
Proof.
intros ?. revert f. induction l as [|x [|x' l'] IH]=> f //.
- by rewrite !big_opL_singleton.
- by rewrite !(big_opL_cons _ x) monoid_homomorphism IH.
Qed.
Lemma big_opM_commute `{Countable K} {A} (h : M1 M2)
`{!MonoidHomomorphism o1 o2 R h} (f : K A M1) m :
R (h ([^o1 map] kx m, f k x)) ([^o2 map] kx m, h (f k x)).
Proof.
intros. induction m as [|i x m ? IH] using map_ind.
- by rewrite !big_opM_empty monoid_homomorphism_unit.
- by rewrite !big_opM_insert // monoid_homomorphism -IH.
Qed.
Lemma big_opM_commute1 `{Countable K} {A} (h : M1 M2)
`{!WeakMonoidHomomorphism o1 o2 R h} (f : K A M1) m :
m R (h ([^o1 map] kx m, f k x)) ([^o2 map] kx m, h (f k x)).
Proof.
intros. induction m as [|i x m ? IH] using map_ind; [done|].
destruct (decide (m = )) as [->|].
- by rewrite !big_opM_insert // !big_opM_empty !right_id.
- by rewrite !big_opM_insert // monoid_homomorphism -IH //.
Qed.
Lemma big_opS_commute `{Countable A} (h : M1 M2)
`{!MonoidHomomorphism o1 o2 R h} (f : A M1) X :
R (h ([^o1 set] x X, f x)) ([^o2 set] x X, h (f x)).
Proof.
intros. induction X as [|x X ? IH] using set_ind_L.
- by rewrite !big_opS_empty monoid_homomorphism_unit.
- by rewrite !big_opS_insert // monoid_homomorphism -IH.
Qed.
Lemma big_opS_commute1 `{Countable A} (h : M1 M2)
`{!WeakMonoidHomomorphism o1 o2 R h} (f : A M1) X :
X R (h ([^o1 set] x X, f x)) ([^o2 set] x X, h (f x)).
Proof.
intros. induction X as [|x X ? IH] using set_ind_L; [done|].
destruct (decide (X = )) as [->|].
- by rewrite !big_opS_insert // !big_opS_empty !right_id.
- by rewrite !big_opS_insert // monoid_homomorphism -IH //.
Qed.
Lemma big_opMS_commute `{Countable A} (h : M1 M2)
`{!MonoidHomomorphism o1 o2 R h} (f : A M1) X :
R (h ([^o1 mset] x X, f x)) ([^o2 mset] x X, h (f x)).
Proof.
intros. induction X as [|x X IH] using gmultiset_ind.
- by rewrite !big_opMS_empty monoid_homomorphism_unit.
- by rewrite !big_opMS_disj_union !big_opMS_singleton monoid_homomorphism -IH.
Qed.
Lemma big_opMS_commute1 `{Countable A} (h : M1 M2)
`{!WeakMonoidHomomorphism o1 o2 R h} (f : A M1) X :
X R (h ([^o1 mset] x X, f x)) ([^o2 mset] x X, h (f x)).
Proof.
intros. induction X as [|x X IH] using gmultiset_ind; [done|].
destruct (decide (X = )) as [->|].
- by rewrite !big_opMS_disj_union !big_opMS_singleton !big_opMS_empty !right_id.
- by rewrite !big_opMS_disj_union !big_opMS_singleton monoid_homomorphism -IH //.
Qed.
Context `{!LeibnizEquiv M2}.
Lemma big_opL_commute_L {A} (h : M1 M2)
`{!MonoidHomomorphism o1 o2 () h} (f : nat A M1) l :
h ([^o1 list] kx l, f k x) = ([^o2 list] kx l, h (f k x)).
Proof using Type*. unfold_leibniz. by apply big_opL_commute. Qed.
Lemma big_opL_commute1_L {A} (h : M1 M2)
`{!WeakMonoidHomomorphism o1 o2 () h} (f : nat A M1) l :
l [] h ([^o1 list] kx l, f k x) = ([^o2 list] kx l, h (f k x)).
Proof using Type*. unfold_leibniz. by apply big_opL_commute1. Qed.
Lemma big_opM_commute_L `{Countable K} {A} (h : M1 M2)
`{!MonoidHomomorphism o1 o2 () h} (f : K A M1) m :
h ([^o1 map] kx m, f k x) = ([^o2 map] kx m, h (f k x)).
Proof using Type*. unfold_leibniz. by apply big_opM_commute. Qed.
Lemma big_opM_commute1_L `{Countable K} {A} (h : M1 M2)
`{!WeakMonoidHomomorphism o1 o2 () h} (f : K A M1) m :
m h ([^o1 map] kx m, f k x) = ([^o2 map] kx m, h (f k x)).
Proof using Type*. unfold_leibniz. by apply big_opM_commute1. Qed.
Lemma big_opS_commute_L `{Countable A} (h : M1 M2)
`{!MonoidHomomorphism o1 o2 () h} (f : A M1) X :
h ([^o1 set] x X, f x) = ([^o2 set] x X, h (f x)).
Proof using Type*. unfold_leibniz. by apply big_opS_commute. Qed.
Lemma big_opS_commute1_L `{ Countable A} (h : M1 M2)
`{!WeakMonoidHomomorphism o1 o2 () h} (f : A M1) X :
X h ([^o1 set] x X, f x) = ([^o2 set] x X, h (f x)).
Proof using Type*. intros. rewrite <-leibniz_equiv_iff. by apply big_opS_commute1. Qed.
Lemma big_opMS_commute_L `{Countable A} (h : M1 M2)
`{!MonoidHomomorphism o1 o2 () h} (f : A M1) X :
h ([^o1 mset] x X, f x) = ([^o2 mset] x X, h (f x)).
Proof using Type*. unfold_leibniz. by apply big_opMS_commute. Qed.
Lemma big_opMS_commute1_L `{Countable A} (h : M1 M2)
`{!WeakMonoidHomomorphism o1 o2 () h} (f : A M1) X :
X h ([^o1 mset] x X, f x) = ([^o2 mset] x X, h (f x)).
Proof using Type*. intros. rewrite <-leibniz_equiv_iff. by apply big_opMS_commute1. Qed.
End homomorphisms.
From stdpp Require Import finite.
From iris.algebra Require Export ofe monoid.
From iris.prelude Require Import options.
Local Set Primitive Projections.
Local Open Scope sidx_scope.
Class PCore (A : Type) := pcore : A option A.
Global Hint Mode PCore ! : typeclass_instances.
Global Instance: Params (@pcore) 2 := {}.
Class Op (A : Type) := op : A A A.
Global Hint Mode Op ! : typeclass_instances.
Global Instance: Params (@op) 2 := {}.
Infix "⋅" := op (at level 50, left associativity) : stdpp_scope.
Notation "(⋅)" := op (only parsing) : stdpp_scope.
(* The inclusion quantifies over [A], not [option A]. This means we do not get
reflexivity. However, if we used [option A], the following would no longer
hold:
x ≼ y ↔ x.1 ≼ y.1 ∧ x.2 ≼ y.2
If you need the reflexive closure of the inclusion relation, you can use
[Some a ≼ Some b]. There are various [Some_included] lemmas that help
deal with propositions of this shape.
*)
Definition included {A} `{!Equiv A, !Op A} (x y : A) := z, y x z.
Infix "≼" := included (at level 70) : stdpp_scope.
Notation "(≼)" := included (only parsing) : stdpp_scope.
Global Hint Extern 0 (_ _) => reflexivity : core.
Global Instance: Params (@included) 3 := {}.
(** [opM] is used in some lemma statements where [A] has not yet been shown to
be a CMRA, so we define it directly in terms of [Op]. *)
Definition opM `{!Op A} (x : A) (my : option A) :=
match my with Some y => x y | None => x end.
Infix "⋅?" := opM (at level 50, left associativity) : stdpp_scope.
Class ValidN {SI : sidx} (A : Type) := validN : SI A Prop.
Global Hint Mode ValidN - ! : typeclass_instances.
Global Instance: Params (@validN) 4 := {}.
Notation "✓{ n } x" := (validN n x)
(at level 20, n at next level, format "✓{ n } x").
Class Valid (A : Type) := valid : A Prop.
Global Hint Mode Valid ! : typeclass_instances.
Global Instance: Params (@valid) 2 := {}.
Notation "✓ x" := (valid x) (at level 20) : stdpp_scope.
Definition includedN {SI : sidx} `{!Dist A, Op A} (n : SI) (x y : A) :=
z, y {n} x z.
Notation "x ≼{ n } y" := (includedN n x y)
(at level 70, n at next level, format "x ≼{ n } y") : stdpp_scope.
Global Instance: Params (@includedN) 5 := {}.
Global Hint Extern 0 (_ {_} _) => reflexivity : core.
Section mixin.
Record CmraMixin {SI : sidx} A
`{!Dist A, !Equiv A, !PCore A, !Op A, !Valid A, !ValidN A} := {
(* setoids *)
mixin_cmra_op_ne (x : A) : NonExpansive (op x);
mixin_cmra_pcore_ne n (x y : A) cx :
x {n} y pcore x = Some cx cy, pcore y = Some cy cx {n} cy;
mixin_cmra_validN_ne n : Proper (dist (A := A) n ==> impl) (validN n);
(* valid *)
mixin_cmra_valid_validN (x : A) : x n, {n} x;
mixin_cmra_validN_le n n' (x : A) : {n} x n' n {n'} x;
(* monoid *)
mixin_cmra_assoc : Assoc (≡@{A}) ();
mixin_cmra_comm : Comm (≡@{A}) ();
mixin_cmra_pcore_l (x : A) cx : pcore x = Some cx cx x x;
mixin_cmra_pcore_idemp (x : A) cx : pcore x = Some cx pcore cx Some cx;
mixin_cmra_pcore_mono (x y : A) cx :
x y pcore x = Some cx cy, pcore y = Some cy cx cy;
mixin_cmra_validN_op_l n (x y : A) : {n} (x y) {n} x;
mixin_cmra_extend n (x y1 y2 : A) :
{n} x x {n} y1 y2
{ z1 : A & { z2 | x z1 z2 z1 {n} y1 z2 {n} y2 } }
}.
End mixin.
(** Bundled version *)
#[projections(primitive=no)] (* FIXME: making this primitive leads to strange
TC resolution failures in view.v *)
Structure cmra {SI : sidx} := Cmra' {
cmra_car :> Type;
cmra_equiv : Equiv cmra_car;
cmra_dist : Dist cmra_car;
cmra_pcore : PCore cmra_car;
cmra_op : Op cmra_car;
cmra_valid : Valid cmra_car;
cmra_validN : ValidN cmra_car;
cmra_ofe_mixin : OfeMixin cmra_car;
cmra_mixin : CmraMixin cmra_car;
}.
Global Arguments Cmra' {_} _ {_ _ _ _ _ _} _ _.
(* Given [m : CmraMixin A], the notation [Cmra A m] provides a smart
constructor, which uses [ofe_mixin_of A] to infer the canonical OFE mixin of
the type [A], so that it does not have to be given manually. *)
Notation Cmra A m := (Cmra' A (ofe_mixin_of A%type) m) (only parsing).
Global Arguments cmra_car : simpl never.
Global Arguments cmra_equiv : simpl never.
Global Arguments cmra_dist : simpl never.
Global Arguments cmra_pcore : simpl never.
Global Arguments cmra_op : simpl never.
Global Arguments cmra_valid : simpl never.
Global Arguments cmra_validN : simpl never.
Global Arguments cmra_ofe_mixin : simpl never.
Global Arguments cmra_mixin : simpl never.
Add Printing Constructor cmra.
(* FIXME(Coq #6294) : we need the new unification algorithm here. *)
Global Hint Extern 0 (PCore _) => refine (cmra_pcore _); shelve : typeclass_instances.
Global Hint Extern 0 (Op _) => refine (cmra_op _); shelve : typeclass_instances.
Global Hint Extern 0 (Valid _) => refine (cmra_valid _); shelve : typeclass_instances.
Global Hint Extern 0 (ValidN _) => refine (cmra_validN _); shelve : typeclass_instances.
Coercion cmra_ofeO {SI : sidx} (A : cmra) : ofe := Ofe A (cmra_ofe_mixin A).
Canonical Structure cmra_ofeO.
(** As explained more thoroughly in iris#539, Coq can run into trouble when
[cmra] combinators (such as [optionUR]) are stacked and combined with
coercions like [cmra_ofeO]. To partially address this, we give Coq's
type-checker some directions for unfolding, with the Strategy command.
For these structures, we instruct Coq to eagerly _expand_ all projections,
except for the coercion to type (in this case, [cmra_car]), since that causes
problem with canonical structure inference. Additionally, we make Coq
eagerly expand the coercions that go from one structure to another, like
[cmra_ofeO] in this case. *)
Global Strategy expand [cmra_ofeO cmra_equiv cmra_dist cmra_pcore cmra_op
cmra_valid cmra_validN cmra_ofe_mixin cmra_mixin].
Definition cmra_mixin_of' {SI : sidx} A {Ac : cmra}
(f : Ac A) : CmraMixin Ac := cmra_mixin Ac.
Notation cmra_mixin_of A :=
ltac:(let H := eval hnf in (cmra_mixin_of' A id) in exact H) (only parsing).
(** Lifting properties from the mixin *)
Section cmra_mixin.
Context {SI : sidx} {A : cmra}.
Implicit Types x y : A.
Global Instance cmra_op_ne (x : A) : NonExpansive (op x).
Proof. apply (mixin_cmra_op_ne _ (cmra_mixin A)). Qed.
Lemma cmra_pcore_ne n x y cx :
x {n} y pcore x = Some cx cy, pcore y = Some cy cx {n} cy.
Proof. apply (mixin_cmra_pcore_ne _ (cmra_mixin A)). Qed.
Global Instance cmra_validN_ne n : Proper (dist n ==> impl) (@validN _ A _ n).
Proof. apply (mixin_cmra_validN_ne _ (cmra_mixin A)). Qed.
Lemma cmra_valid_validN x : x n, {n} x.
Proof. apply (mixin_cmra_valid_validN _ (cmra_mixin A)). Qed.
Lemma cmra_validN_le n n' x : {n} x n' n {n'} x.
Proof. apply (mixin_cmra_validN_le _ (cmra_mixin A)). Qed.
Global Instance cmra_assoc : Assoc () (@op A _).
Proof. apply (mixin_cmra_assoc _ (cmra_mixin A)). Qed.
Global Instance cmra_comm : Comm () (@op A _).
Proof. apply (mixin_cmra_comm _ (cmra_mixin A)). Qed.
Lemma cmra_pcore_l x cx : pcore x = Some cx cx x x.
Proof. apply (mixin_cmra_pcore_l _ (cmra_mixin A)). Qed.
Lemma cmra_pcore_idemp x cx : pcore x = Some cx pcore cx Some cx.
Proof. apply (mixin_cmra_pcore_idemp _ (cmra_mixin A)). Qed.
Lemma cmra_pcore_mono x y cx :
x y pcore x = Some cx cy, pcore y = Some cy cx cy.
Proof. apply (mixin_cmra_pcore_mono _ (cmra_mixin A)). Qed.
Lemma cmra_validN_op_l n x y : {n} (x y) {n} x.
Proof. apply (mixin_cmra_validN_op_l _ (cmra_mixin A)). Qed.
Lemma cmra_extend n x y1 y2 :
{n} x x {n} y1 y2
{ z1 : A & { z2 | x z1 z2 z1 {n} y1 z2 {n} y2 } }.
Proof. apply (mixin_cmra_extend _ (cmra_mixin A)). Qed.
End cmra_mixin.
(** * CoreId elements *)
Class CoreId {SI : sidx} {A : cmra} (x : A) := core_id : pcore x Some x.
Global Arguments core_id {_ _} _ {_}.
Global Hint Mode CoreId - + ! : typeclass_instances.
Global Instance: Params (@CoreId) 2 := {}.
(** * Exclusive elements (i.e., elements that cannot have a frame). *)
Class Exclusive {SI : sidx} {A : cmra} (x : A) :=
exclusive0_l y : {0} (x y) False.
Global Arguments exclusive0_l {_ _} _ {_} _ _.
Global Hint Mode Exclusive - + ! : typeclass_instances.
Global Instance: Params (@Exclusive) 2 := {}.
(** * Cancelable elements. *)
Class Cancelable {SI : sidx} {A : cmra} (x : A) :=
cancelableN n y z : {n} (x y) x y {n} x z y {n} z.
Global Arguments cancelableN {_ _} _ {_} _ _ _ _.
Global Hint Mode Cancelable - + ! : typeclass_instances.
Global Instance: Params (@Cancelable) 2 := {}.
(** * Identity-free elements. *)
Class IdFree {SI : sidx} {A : cmra} (x : A) :=
id_free0_r y : {0} x x y {0} x False.
Global Arguments id_free0_r {_ _} _ {_} _ _.
Global Hint Mode IdFree - + ! : typeclass_instances.
Global Instance: Params (@IdFree) 2 := {}.
(** * CMRAs whose core is total *)
Class CmraTotal {SI : sidx} (A : cmra) :=
cmra_total (x : A) : is_Some (pcore x).
Global Hint Mode CmraTotal - ! : typeclass_instances.
(** The function [core] returns a dummy when used on CMRAs without total
core. We only ever use this for [CmraTotal] CMRAs, but it is more convenient
to not require that proof to be able to call this function. *)
Definition core {A} `{!PCore A} (x : A) : A := default x (pcore x).
Global Instance: Params (@core) 2 := {}.
(** * CMRAs with a unit element *)
Class Unit (A : Type) := ε : A.
Global Hint Mode Unit ! : typeclass_instances.
Global Arguments ε {_ _}.
Record UcmraMixin {SI : sidx} A
`{!Dist A, !Equiv A, !PCore A, !Op A, !Valid A, !Unit A} := {
mixin_ucmra_unit_valid : (ε : A);
mixin_ucmra_unit_left_id : LeftId (≡@{A}) ε ();
mixin_ucmra_pcore_unit : pcore ε ≡@{option A} Some ε
}.
#[projections(primitive=no)] (* FIXME: making this primitive leads to strange
TC resolution failures in view.v *)
Structure ucmra {SI : sidx} := Ucmra' {
ucmra_car :> Type;
ucmra_equiv : Equiv ucmra_car;
ucmra_dist : Dist ucmra_car;
ucmra_pcore : PCore ucmra_car;
ucmra_op : Op ucmra_car;
ucmra_valid : Valid ucmra_car;
ucmra_validN : ValidN ucmra_car;
ucmra_unit : Unit ucmra_car;
ucmra_ofe_mixin : OfeMixin ucmra_car;
ucmra_cmra_mixin : CmraMixin ucmra_car;
ucmra_mixin : UcmraMixin ucmra_car;
}.
Global Arguments Ucmra' {_} _ {_ _ _ _ _ _ _} _ _ _.
Notation Ucmra A m :=
(Ucmra' A (ofe_mixin_of A%type) (cmra_mixin_of A%type) m) (only parsing).
Global Arguments ucmra_car : simpl never.
Global Arguments ucmra_equiv : simpl never.
Global Arguments ucmra_dist : simpl never.
Global Arguments ucmra_pcore : simpl never.
Global Arguments ucmra_op : simpl never.
Global Arguments ucmra_valid : simpl never.
Global Arguments ucmra_validN : simpl never.
Global Arguments ucmra_ofe_mixin : simpl never.
Global Arguments ucmra_cmra_mixin : simpl never.
Global Arguments ucmra_mixin : simpl never.
Add Printing Constructor ucmra.
(* FIXME(Coq #6294) : we need the new unification algorithm here. *)
Global Hint Extern 0 (Unit _) => refine (ucmra_unit _); shelve : typeclass_instances.
Coercion ucmra_ofeO {SI : sidx} (A : ucmra) : ofe := Ofe A (ucmra_ofe_mixin A).
Canonical Structure ucmra_ofeO.
Coercion ucmra_cmraR {SI : sidx} (A : ucmra) : cmra :=
Cmra' A (ucmra_ofe_mixin A) (ucmra_cmra_mixin A).
Canonical Structure ucmra_cmraR.
(** As for CMRAs above, we instruct Coq to eagerly _expand_ all projections,
except for the coercion to type (in this case, [ucmra_car]), since that causes
problem with canonical structure inference. Additionally, we make Coq
eagerly expand the coercions that go from one structure to another, like
[ucmra_cmraR] and [ucmra_ofeO] in this case. *)
Global Strategy expand [ucmra_cmraR ucmra_ofeO ucmra_equiv ucmra_dist ucmra_pcore
ucmra_op ucmra_valid ucmra_validN ucmra_unit
ucmra_ofe_mixin ucmra_cmra_mixin].
(** Lifting properties from the mixin *)
Section ucmra_mixin.
Context {SI : sidx} {A : ucmra}.
Implicit Types x y : A.
Lemma ucmra_unit_valid : (ε : A).
Proof. apply (mixin_ucmra_unit_valid _ (ucmra_mixin A)). Qed.
Global Instance ucmra_unit_left_id : LeftId () ε (@op A _).
Proof. apply (mixin_ucmra_unit_left_id _ (ucmra_mixin A)). Qed.
Lemma ucmra_pcore_unit : pcore (ε:A) Some ε.
Proof. apply (mixin_ucmra_pcore_unit _ (ucmra_mixin A)). Qed.
End ucmra_mixin.
(** * Discrete CMRAs *)
#[projections(primitive=no)] (* FIXME: making this primitive means we cannot use
the projections with eauto any more (see https://github.com/coq/coq/issues/17561) *)
Class CmraDiscrete {SI : sidx} (A : cmra) := {
#[global] cmra_discrete_ofe_discrete :: OfeDiscrete A;
cmra_discrete_valid (x : A) : {0} x x
}.
Global Hint Mode CmraDiscrete - ! : typeclass_instances.
(** * Morphisms *)
Class CmraMorphism {SI : sidx} {A B : cmra} (f : A B) := {
#[global] cmra_morphism_ne :: NonExpansive f;
cmra_morphism_validN n x : {n} x {n} f x;
cmra_morphism_pcore x : f <$> pcore x pcore (f x);
cmra_morphism_op x y : f (x y) f x f y
}.
Global Hint Mode CmraMorphism - - - ! : typeclass_instances.
Global Arguments cmra_morphism_validN {_ _ _} _ {_} _ _ _.
Global Arguments cmra_morphism_pcore {_ _ _} _ {_} _.
Global Arguments cmra_morphism_op {_ _ _} _ {_} _ _.
(** * Properties **)
Section cmra.
Context {SI : sidx} {A : cmra}.
Implicit Types x y z : A.
Implicit Types xs ys zs : list A.
(** ** Setoids *)
Global Instance cmra_pcore_ne' : NonExpansive (@pcore A _).
Proof.
intros n x y Hxy. destruct (pcore x) as [cx|] eqn:?.
{ destruct (cmra_pcore_ne n x y cx) as (cy&->&->); auto. }
destruct (pcore y) as [cy|] eqn:?; auto.
destruct (cmra_pcore_ne n y x cy) as (cx&?&->); simplify_eq/=; auto.
Qed.
Lemma cmra_pcore_proper x y cx :
x y pcore x = Some cx cy, pcore y = Some cy cx cy.
Proof.
intros. destruct (cmra_pcore_ne 0 x y cx) as (cy&?&?); auto.
exists cy; split; [done|apply equiv_dist=> n].
destruct (cmra_pcore_ne n x y cx) as (cy'&?&?); naive_solver.
Qed.
Global Instance cmra_pcore_proper' : Proper (() ==> ()) (@pcore A _).
Proof. apply (ne_proper _). Qed.
Global Instance cmra_op_ne' : NonExpansive2 (@op A _).
Proof. intros n x1 x2 Hx y1 y2 Hy. by rewrite Hy (comm _ x1) Hx (comm _ y2). Qed.
Global Instance cmra_op_proper' : Proper (() ==> () ==> ()) (@op A _).
Proof. apply (ne_proper_2 _). Qed.
Global Instance cmra_validN_ne' n : Proper (dist n ==> iff) (@validN SI A _ n) | 1.
Proof. by split; apply cmra_validN_ne. Qed.
Global Instance cmra_validN_proper n : Proper (() ==> iff) (@validN SI A _ n) | 1.
Proof. by intros x1 x2 Hx; apply cmra_validN_ne', equiv_dist. Qed.
Global Instance cmra_valid_proper : Proper (() ==> iff) (@valid A _).
Proof.
intros x y Hxy; rewrite !cmra_valid_validN.
by split=> ? n; [rewrite -Hxy|rewrite Hxy].
Qed.
Global Instance cmra_includedN_ne n :
Proper (dist n ==> dist n ==> iff) (@includedN SI A _ _ n) | 1.
Proof.
intros x x' Hx y y' Hy.
by split; intros [z ?]; exists z; [rewrite -Hx -Hy|rewrite Hx Hy].
Qed.
Global Instance cmra_includedN_proper n :
Proper (() ==> () ==> iff) (@includedN SI A _ _ n) | 1.
Proof.
intros x x' Hx y y' Hy; revert Hx Hy; rewrite !equiv_dist=> Hx Hy.
by rewrite (Hx n) (Hy n).
Qed.
Global Instance cmra_included_proper :
Proper (() ==> () ==> iff) (@included A _ _) | 1.
Proof.
intros x x' Hx y y' Hy.
by split; intros [z ?]; exists z; [rewrite -Hx -Hy|rewrite Hx Hy].
Qed.
Global Instance cmra_opM_ne : NonExpansive2 (@opM A _).
Proof. destruct 2; by ofe_subst. Qed.
Global Instance cmra_opM_proper : Proper (() ==> () ==> ()) (@opM A _).
Proof. destruct 2; by setoid_subst. Qed.
Global Instance CoreId_proper : Proper (() ==> iff) (@CoreId SI A).
Proof. solve_proper. Qed.
Global Instance Exclusive_proper : Proper (() ==> iff) (@Exclusive SI A).
Proof. intros x y Hxy. rewrite /Exclusive. by setoid_rewrite Hxy. Qed.
Global Instance Cancelable_proper : Proper (() ==> iff) (@Cancelable SI A).
Proof. intros x y Hxy. rewrite /Cancelable. by setoid_rewrite Hxy. Qed.
Global Instance IdFree_proper : Proper (() ==> iff) (@IdFree SI A).
Proof. intros x y Hxy. rewrite /IdFree. by setoid_rewrite Hxy. Qed.
(** ** Op *)
Lemma cmra_op_opM_assoc x y mz : (x y) ? mz x (y ? mz).
Proof. destruct mz; by rewrite /= -?assoc. Qed.
(** ** Validity *)
Lemma cmra_validN_lt n n' x : {n} x n' < n {n'} x.
Proof. eauto using cmra_validN_le, SIdx.lt_le_incl. Qed.
Lemma cmra_valid_op_l x y : (x y) x.
Proof. rewrite !cmra_valid_validN; eauto using cmra_validN_op_l. Qed.
Lemma cmra_validN_op_r n x y : {n} (x y) {n} y.
Proof. rewrite (comm _ x); apply cmra_validN_op_l. Qed.
Lemma cmra_valid_op_r x y : (x y) y.
Proof. rewrite !cmra_valid_validN; eauto using cmra_validN_op_r. Qed.
(** ** Core *)
Lemma cmra_pcore_l' x cx : pcore x Some cx cx x x.
Proof. intros (cx'&?&<-)%Some_equiv_eq. by apply cmra_pcore_l. Qed.
Lemma cmra_pcore_r x cx : pcore x = Some cx x cx x.
Proof. intros. rewrite comm. by apply cmra_pcore_l. Qed.
Lemma cmra_pcore_r' x cx : pcore x Some cx x cx x.
Proof. intros (cx'&?&<-)%Some_equiv_eq. by apply cmra_pcore_r. Qed.
Lemma cmra_pcore_idemp' x cx : pcore x Some cx pcore cx Some cx.
Proof. intros (cx'&?&<-)%Some_equiv_eq. eauto using cmra_pcore_idemp. Qed.
Lemma cmra_pcore_dup x cx : pcore x = Some cx cx cx cx.
Proof. intros; symmetry; eauto using cmra_pcore_r', cmra_pcore_idemp. Qed.
Lemma cmra_pcore_dup' x cx : pcore x Some cx cx cx cx.
Proof. intros; symmetry; eauto using cmra_pcore_r', cmra_pcore_idemp'. Qed.
Lemma cmra_pcore_validN n x cx : {n} x pcore x = Some cx {n} cx.
Proof.
intros Hvx Hx%cmra_pcore_l. move: Hvx; rewrite -Hx. apply cmra_validN_op_l.
Qed.
Lemma cmra_pcore_valid x cx : x pcore x = Some cx cx.
Proof.
intros Hv Hx%cmra_pcore_l. move: Hv; rewrite -Hx. apply cmra_valid_op_l.
Qed.
(** ** Exclusive elements *)
Lemma exclusiveN_l n x `{!Exclusive x} y : {n} (x y) False.
Proof. intros. by eapply (exclusive0_l x y), cmra_validN_le, SIdx.le_0_l. Qed.
Lemma exclusiveN_r n x `{!Exclusive x} y : {n} (y x) False.
Proof. rewrite comm. by apply exclusiveN_l. Qed.
Lemma exclusive_l x `{!Exclusive x} y : (x y) False.
Proof. by move /cmra_valid_validN /(_ 0) /exclusive0_l. Qed.
Lemma exclusive_r x `{!Exclusive x} y : (y x) False.
Proof. rewrite comm. by apply exclusive_l. Qed.
Lemma exclusiveN_opM n x `{!Exclusive x} my : {n} (x ? my) my = None.
Proof. destruct my as [y|]; last done. move=> /(exclusiveN_l _ x) []. Qed.
Lemma exclusive_includedN n x `{!Exclusive x} y : x {n} y {n} y False.
Proof. intros [? ->]. by apply exclusiveN_l. Qed.
Lemma exclusive_included x `{!Exclusive x} y : x y y False.
Proof. intros [? ->]. by apply exclusive_l. Qed.
(** ** Order *)
Lemma cmra_included_includedN n x y : x y x {n} y.
Proof. intros [z ->]. by exists z. Qed.
Global Instance cmra_includedN_trans n : Transitive (@includedN SI A _ _ n).
Proof.
intros x y z [z1 Hy] [z2 Hz]; exists (z1 z2). by rewrite assoc -Hy -Hz.
Qed.
Global Instance cmra_included_trans: Transitive (@included A _ _).
Proof.
intros x y z [z1 Hy] [z2 Hz]; exists (z1 z2). by rewrite assoc -Hy -Hz.
Qed.
Lemma cmra_valid_included x y : y x y x.
Proof. intros Hyv [z ?]; setoid_subst; eauto using cmra_valid_op_l. Qed.
Lemma cmra_validN_includedN n x y : {n} y x {n} y {n} x.
Proof. intros Hyv [z ?]; ofe_subst y; eauto using cmra_validN_op_l. Qed.
Lemma cmra_validN_included n x y : {n} y x y {n} x.
Proof. intros Hyv [z ?]; setoid_subst; eauto using cmra_validN_op_l. Qed.
Lemma cmra_includedN_le n m x y : x {n} y m n x {m} y.
Proof. by intros [z Hz] H; exists z; eapply dist_le. Qed.
Lemma cmra_includedN_S n x y : x {Sᵢ n} y x {n} y.
Proof. intros ?. by eapply cmra_includedN_le, SIdx.le_succ_diag_r. Qed.
Lemma cmra_includedN_l n x y : x {n} x y.
Proof. by exists y. Qed.
Lemma cmra_included_l x y : x x y.
Proof. by exists y. Qed.
Lemma cmra_includedN_r n x y : y {n} x y.
Proof. rewrite (comm op); apply cmra_includedN_l. Qed.
Lemma cmra_included_r x y : y x y.
Proof. rewrite (comm op); apply cmra_included_l. Qed.
Lemma cmra_pcore_mono' x y cx :
x y pcore x Some cx cy, pcore y = Some cy cx cy.
Proof.
intros ? (cx'&?&Hcx)%Some_equiv_eq.
destruct (cmra_pcore_mono x y cx') as (cy&->&?); auto.
exists cy; by rewrite -Hcx.
Qed.
Lemma cmra_pcore_monoN' n x y cx :
x {n} y pcore x {n} Some cx cy, pcore y = Some cy cx {n} cy.
Proof.
intros [z Hy] (cx'&?&Hcx)%dist_Some_inv_r'.
destruct (cmra_pcore_mono x (x z) cx')
as (cy&Hxy&?); auto using cmra_included_l.
assert (pcore y {n} Some cy) as (cy'&?&Hcy')%dist_Some_inv_r'.
{ by rewrite Hy Hxy. }
exists cy'; split; first done.
rewrite Hcx -Hcy'; auto using cmra_included_includedN.
Qed.
Lemma cmra_included_pcore x cx : pcore x = Some cx cx x.
Proof. exists x. by rewrite cmra_pcore_l. Qed.
Lemma cmra_monoN_l n x y z : x {n} y z x {n} z y.
Proof. by intros [z1 Hz1]; exists z1; rewrite Hz1 (assoc op). Qed.
Lemma cmra_mono_l x y z : x y z x z y.
Proof. by intros [z1 Hz1]; exists z1; rewrite Hz1 (assoc op). Qed.
Lemma cmra_monoN_r n x y z : x {n} y x z {n} y z.
Proof. by intros; rewrite -!(comm _ z); apply cmra_monoN_l. Qed.
Lemma cmra_mono_r x y z : x y x z y z.
Proof. by intros; rewrite -!(comm _ z); apply cmra_mono_l. Qed.
Lemma cmra_monoN n x1 x2 y1 y2 : x1 {n} y1 x2 {n} y2 x1 x2 {n} y1 y2.
Proof. intros; etrans; eauto using cmra_monoN_l, cmra_monoN_r. Qed.
Lemma cmra_mono x1 x2 y1 y2 : x1 y1 x2 y2 x1 x2 y1 y2.
Proof. intros; etrans; eauto using cmra_mono_l, cmra_mono_r. Qed.
Global Instance cmra_monoN' n :
Proper (includedN n ==> includedN n ==> includedN n) (@op A _).
Proof. intros x1 x2 Hx y1 y2 Hy. by apply cmra_monoN. Qed.
Global Instance cmra_mono' :
Proper (included ==> included ==> included) (@op A _).
Proof. intros x1 x2 Hx y1 y2 Hy. by apply cmra_mono. Qed.
Lemma cmra_included_dist_l n x1 x2 x1' :
x1 x2 x1' {n} x1 x2', x1' x2' x2' {n} x2.
Proof.
intros [z Hx2] Hx1; exists (x1' z); split; auto using cmra_included_l.
by rewrite Hx1 Hx2.
Qed.
(** ** CoreId elements *)
Lemma core_id_dup x `{!CoreId x} : x x x.
Proof. by apply cmra_pcore_dup' with x. Qed.
Lemma core_id_extract x y `{!CoreId x} :
x y y y x.
Proof.
intros ?.
destruct (cmra_pcore_mono' x y x) as (cy & Hcy & [x' Hx']); [done|exact: core_id|].
rewrite -(cmra_pcore_r y) //. rewrite Hx' -!assoc. f_equiv.
rewrite [x' x]comm assoc -core_id_dup. done.
Qed.
(** ** Total core *)
Section total_core.
Local Set Default Proof Using "Type*".
Context `{!CmraTotal A}.
Lemma cmra_pcore_core x : pcore x = Some (core x).
Proof.
rewrite /core. destruct (cmra_total x) as [cx ->]. done.
Qed.
Lemma cmra_core_l x : core x x x.
Proof.
destruct (cmra_total x) as [cx Hcx]. by rewrite /core /= Hcx cmra_pcore_l.
Qed.
Lemma cmra_core_idemp x : core (core x) core x.
Proof.
destruct (cmra_total x) as [cx Hcx]. by rewrite /core /= Hcx cmra_pcore_idemp.
Qed.
Lemma cmra_core_mono x y : x y core x core y.
Proof.
intros; destruct (cmra_total x) as [cx Hcx].
destruct (cmra_pcore_mono x y cx) as (cy&Hcy&?); auto.
by rewrite /core /= Hcx Hcy.
Qed.
Global Instance cmra_core_ne : NonExpansive (@core A _).
Proof.
intros n x y Hxy. destruct (cmra_total x) as [cx Hcx].
by rewrite /core /= -Hxy Hcx.
Qed.
Global Instance cmra_core_proper : Proper (() ==> ()) (@core A _).
Proof. apply (ne_proper _). Qed.
Lemma cmra_core_r x : x core x x.
Proof. by rewrite (comm _ x) cmra_core_l. Qed.
Lemma cmra_core_dup x : core x core x core x.
Proof. by rewrite -{3}(cmra_core_idemp x) cmra_core_r. Qed.
Lemma cmra_core_validN n x : {n} x {n} core x.
Proof. rewrite -{1}(cmra_core_l x); apply cmra_validN_op_l. Qed.
Lemma cmra_core_valid x : x core x.
Proof. rewrite -{1}(cmra_core_l x); apply cmra_valid_op_l. Qed.
Lemma core_id_total x : CoreId x core x x.
Proof.
split; [intros; by rewrite /core /= (core_id x)|].
rewrite /CoreId /core /=.
destruct (cmra_total x) as [? ->]. by constructor.
Qed.
Lemma core_id_core x `{!CoreId x} : core x x.
Proof. by apply core_id_total. Qed.
(** Not an instance since TC search cannot solve the premise. *)
Lemma cmra_pcore_core_id x y : pcore x = Some y CoreId y.
Proof. rewrite /CoreId. eauto using cmra_pcore_idemp. Qed.
Global Instance cmra_core_core_id x : CoreId (core x).
Proof. eapply cmra_pcore_core_id. rewrite cmra_pcore_core. done. Qed.
Lemma cmra_included_core x : core x x.
Proof. by exists x; rewrite cmra_core_l. Qed.
Global Instance cmra_includedN_preorder n : PreOrder (@includedN SI A _ _ n).
Proof.
split; [|apply _]. by intros x; exists (core x); rewrite cmra_core_r.
Qed.
Global Instance cmra_included_preorder : PreOrder (@included A _ _).
Proof.
split; [|apply _]. by intros x; exists (core x); rewrite cmra_core_r.
Qed.
Lemma cmra_core_monoN n x y : x {n} y core x {n} core y.
Proof.
intros [z ->].
apply cmra_included_includedN, cmra_core_mono, cmra_included_l.
Qed.
End total_core.
(** ** Discrete *)
Lemma cmra_discrete_included_l x y : Discrete x {0} y x {0} y x y.
Proof.
intros ?? [x' ?].
destruct (cmra_extend 0 y x x') as (z&z'&Hy&Hz&Hz'); auto; simpl in *.
by exists z'; rewrite Hy (discrete_0 x z).
Qed.
Lemma cmra_discrete_included_r x y : Discrete y x {0} y x y.
Proof. intros ? [x' ?]. exists x'. by apply (discrete_0 y). Qed.
Lemma cmra_op_discrete x1 x2 :
{0} (x1 x2) Discrete x1 Discrete x2 Discrete (x1 x2).
Proof.
intros ??? z Hz.
destruct (cmra_extend 0 z x1 x2) as (y1&y2&Hz'&?&?); auto; simpl in *.
{ rewrite -?Hz. done. }
by rewrite Hz' (discrete_0 x1 y1) // (discrete_0 x2 y2).
Qed.
(** ** Discrete *)
Lemma cmra_discrete_valid_iff `{!CmraDiscrete A} n x : x {n} x.
Proof.
split; first by rewrite cmra_valid_validN.
eauto using cmra_discrete_valid, cmra_validN_le, SIdx.le_0_l.
Qed.
Lemma cmra_discrete_valid_iff_0 `{!CmraDiscrete A} n x : {0} x {n} x.
Proof. by rewrite -!cmra_discrete_valid_iff. Qed.
Lemma cmra_discrete_included_iff `{!OfeDiscrete A} n x y : x y x {n} y.
Proof.
split; first by apply cmra_included_includedN.
intros [z ->%(discrete_iff _ _)]; eauto using cmra_included_l.
Qed.
Lemma cmra_discrete_included_iff_0 `{!OfeDiscrete A} n x y : x {0} y x {n} y.
Proof. by rewrite -!cmra_discrete_included_iff. Qed.
(** Cancelable elements *)
Global Instance cancelable_proper : Proper (equiv ==> iff) (@Cancelable SI A).
Proof. unfold Cancelable. intros x x' EQ. by setoid_rewrite EQ. Qed.
Lemma cancelable x `{!Cancelable x} y z : (x y) x y x z y z.
Proof. rewrite !equiv_dist cmra_valid_validN. intros. by apply (cancelableN x). Qed.
Lemma discrete_cancelable x `{!CmraDiscrete A}:
( y z, (x y) x y x z y z) Cancelable x.
Proof. intros ????. rewrite -!discrete_iff -cmra_discrete_valid_iff. auto. Qed.
Global Instance cancelable_op x y :
Cancelable x Cancelable y Cancelable (x y).
Proof.
intros ?? n z z' ??. apply (cancelableN y), (cancelableN x).
- eapply cmra_validN_op_r. by rewrite assoc.
- by rewrite assoc.
- by rewrite !assoc.
Qed.
Global Instance exclusive_cancelable (x : A) : Exclusive x Cancelable x.
Proof. intros ? n z z' []%(exclusiveN_l _ x). Qed.
(** Id-free elements *)
Global Instance id_free_ne n : Proper (dist n ==> iff) (@IdFree SI A).
Proof.
intros x x' EQ%(dist_le _ 0); [|apply SIdx.le_0_l]. rewrite /IdFree.
split=> y ?; (rewrite -EQ || rewrite EQ); eauto.
Qed.
Global Instance id_free_proper : Proper (equiv ==> iff) (@IdFree SI A).
Proof. by move=> P Q /equiv_dist /(_ 0)=> ->. Qed.
Lemma id_freeN_r n n' x `{!IdFree x} y : {n}x x y {n'} x False.
Proof. eauto using cmra_validN_le, dist_le, SIdx.le_0_l. Qed.
Lemma id_freeN_l n n' x `{!IdFree x} y : {n}x y x {n'} x False.
Proof. rewrite comm. eauto using id_freeN_r. Qed.
Lemma id_free_r x `{!IdFree x} y : x x y x False.
Proof. move=> /cmra_valid_validN ? /equiv_dist. eauto. Qed.
Lemma id_free_l x `{!IdFree x} y : x y x x False.
Proof. rewrite comm. eauto using id_free_r. Qed.
Lemma discrete_id_free x `{!CmraDiscrete A}:
( y, x x y x False) IdFree x.
Proof.
intros Hx y ??. apply (Hx y), (discrete_0 _); eauto using cmra_discrete_valid.
Qed.
Global Instance id_free_op_r x y : IdFree y Cancelable x IdFree (x y).
Proof.
intros ?? z ? Hid%symmetry. revert Hid. rewrite -assoc=>/(cancelableN x) ?.
eapply (id_free0_r y); [by eapply cmra_validN_op_r |symmetry; eauto].
Qed.
Global Instance id_free_op_l x y : IdFree x Cancelable y IdFree (x y).
Proof. intros. rewrite comm. apply _. Qed.
Global Instance exclusive_id_free x : Exclusive x IdFree x.
Proof. intros ? z ? Hid. apply (exclusiveN_l 0 x z). by rewrite Hid. Qed.
End cmra.
(* We use a [Hint Extern] with [apply:], instead of [Hint Immediate], to invoke
the new unification algorithm. The old unification algorithm sometimes gets
confused by going from [ucmra]'s to [cmra]'s and back. *)
Global Hint Extern 0 (?a ?a _) => apply: cmra_included_l : core.
Global Hint Extern 0 (?a _ ?a) => apply: cmra_included_r : core.
(** * Properties about CMRAs with a unit element **)
Section ucmra.
Context {SI : sidx} {A : ucmra}.
Implicit Types x y z : A.
Lemma ucmra_unit_validN n : {n} (ε:A).
Proof. apply cmra_valid_validN, ucmra_unit_valid. Qed.
Lemma ucmra_unit_leastN n x : ε {n} x.
Proof. by exists x; rewrite left_id. Qed.
Lemma ucmra_unit_least x : ε x.
Proof. by exists x; rewrite left_id. Qed.
Global Instance ucmra_unit_right_id : RightId () ε (@op A _).
Proof. by intros x; rewrite (comm op) left_id. Qed.
Global Instance ucmra_unit_core_id : CoreId (ε:A).
Proof. apply ucmra_pcore_unit. Qed.
Global Instance cmra_unit_cmra_total : CmraTotal A.
Proof.
intros x. destruct (cmra_pcore_mono' ε x ε) as (cx&->&?); [..|by eauto].
- apply ucmra_unit_least.
- apply (core_id _).
Qed.
Global Instance empty_cancelable : Cancelable (ε:A).
Proof. intros ???. by rewrite !left_id. Qed.
(* For big ops *)
Global Instance cmra_monoid : Monoid (@op A _) := {| monoid_unit := ε |}.
End ucmra.
Global Hint Immediate cmra_unit_cmra_total : core.
Global Hint Extern 0 (ε _) => apply: ucmra_unit_least : core.
(** * Properties about CMRAs with Leibniz equality *)
Section cmra_leibniz.
Local Set Default Proof Using "Type*".
Context {SI : sidx} {A : cmra} `{!LeibnizEquiv A}.
Implicit Types x y : A.
Global Instance cmra_assoc_L : Assoc (=) (@op A _).
Proof. intros x y z. unfold_leibniz. by rewrite assoc. Qed.
Global Instance cmra_comm_L : Comm (=) (@op A _).
Proof. intros x y. unfold_leibniz. by rewrite comm. Qed.
Lemma cmra_pcore_l_L x cx : pcore x = Some cx cx x = x.
Proof. unfold_leibniz. apply cmra_pcore_l'. Qed.
Lemma cmra_pcore_idemp_L x cx : pcore x = Some cx pcore cx = Some cx.
Proof. unfold_leibniz. apply cmra_pcore_idemp'. Qed.
Lemma cmra_op_opM_assoc_L x y mz : (x y) ? mz = x (y ? mz).
Proof. unfold_leibniz. apply cmra_op_opM_assoc. Qed.
(** ** Core *)
Lemma cmra_pcore_r_L x cx : pcore x = Some cx x cx = x.
Proof. unfold_leibniz. apply cmra_pcore_r'. Qed.
Lemma cmra_pcore_dup_L x cx : pcore x = Some cx cx = cx cx.
Proof. unfold_leibniz. apply cmra_pcore_dup'. Qed.
(** ** CoreId elements *)
Lemma core_id_dup_L x `{!CoreId x} : x = x x.
Proof. unfold_leibniz. by apply core_id_dup. Qed.
(** ** Total core *)
Section total_core.
Context `{!CmraTotal A}.
Lemma cmra_core_r_L x : x core x = x.
Proof. unfold_leibniz. apply cmra_core_r. Qed.
Lemma cmra_core_l_L x : core x x = x.
Proof. unfold_leibniz. apply cmra_core_l. Qed.
Lemma cmra_core_idemp_L x : core (core x) = core x.
Proof. unfold_leibniz. apply cmra_core_idemp. Qed.
Lemma cmra_core_dup_L x : core x = core x core x.
Proof. unfold_leibniz. apply cmra_core_dup. Qed.
Lemma core_id_total_L x : CoreId x core x = x.
Proof. unfold_leibniz. apply core_id_total. Qed.
Lemma core_id_core_L x `{!CoreId x} : core x = x.
Proof. by apply core_id_total_L. Qed.
End total_core.
End cmra_leibniz.
Section ucmra_leibniz.
Local Set Default Proof Using "Type*".
Context {SI : sidx} {A : ucmra} `{!LeibnizEquiv A}.
Implicit Types x y z : A.
Global Instance ucmra_unit_left_id_L : LeftId (=) ε (@op A _).
Proof. intros x. unfold_leibniz. by rewrite left_id. Qed.
Global Instance ucmra_unit_right_id_L : RightId (=) ε (@op A _).
Proof. intros x. unfold_leibniz. by rewrite right_id. Qed.
End ucmra_leibniz.
(** * Constructing a CMRA with total core *)
Section cmra_total.
Context {SI : sidx} A `{!Dist A, !Equiv A, !PCore A, !Op A, !Valid A, !ValidN A}.
Context (total : x : A, is_Some (pcore x)).
Context (op_ne : x : A, NonExpansive (op x)).
Context (core_ne : NonExpansive (@core A _)).
Context (validN_ne : n, Proper (dist n ==> impl) (@validN SI A _ n)).
Context (valid_validN : (x : A), x n, {n} x).
Context (validN_le : n n' (x : A), {n} x n' n {n'} x).
Context (op_assoc : Assoc () (@op A _)).
Context (op_comm : Comm () (@op A _)).
Context (core_l : x : A, core x x x).
Context (core_idemp : x : A, core (core x) core x).
Context (core_mono : x y : A, x y core x core y).
Context (validN_op_l : n (x y : A), {n} (x y) {n} x).
Context (extend : n (x y1 y2 : A),
{n} x x {n} y1 y2
{ z1 : A & { z2 | x z1 z2 z1 {n} y1 z2 {n} y2 } }).
Lemma cmra_total_mixin : CmraMixin A.
Proof using Type*.
split; auto.
- intros n x y ? Hcx%core_ne Hx; move: Hcx. rewrite /core /= Hx /=.
case (total y)=> [cy ->]; eauto.
- intros x cx Hcx. move: (core_l x). by rewrite /core /= Hcx.
- intros x cx Hcx. move: (core_idemp x). rewrite /core /= Hcx /=.
case (total cx)=>[ccx ->]; by constructor.
- intros x y cx Hxy%core_mono Hx. move: Hxy.
rewrite /core /= Hx /=. case (total y)=> [cy ->]; eauto.
Qed.
End cmra_total.
(** * Properties about morphisms *)
Global Instance cmra_morphism_id {SI : sidx} {A : cmra} : CmraMorphism (@id A).
Proof.
split => /=.
- apply _.
- done.
- intros. by rewrite option_fmap_id.
- done.
Qed.
Global Instance cmra_morphism_proper {SI : sidx}
{A B : cmra} (f : A B) `{!CmraMorphism f} :
Proper (() ==> ()) f := ne_proper _.
Global Instance cmra_morphism_compose {SI : sidx}
{A B C : cmra} (f : A B) (g : B C) :
CmraMorphism f CmraMorphism g CmraMorphism (g f).
Proof.
split.
- apply _.
- move=> n x Hx /=. by apply cmra_morphism_validN, cmra_morphism_validN.
- move=> x /=. by rewrite option_fmap_compose !cmra_morphism_pcore.
- move=> x y /=. by rewrite !cmra_morphism_op.
Qed.
Section cmra_morphism.
Local Set Default Proof Using "Type*".
Context {SI : sidx} {A B : cmra} (f : A B) `{!CmraMorphism f}.
Lemma cmra_morphism_core x : f (core x) core (f x).
Proof. unfold core. rewrite -cmra_morphism_pcore. by destruct (pcore x). Qed.
Lemma cmra_morphism_monotone x y : x y f x f y.
Proof. intros [z ->]. exists (f z). by rewrite cmra_morphism_op. Qed.
Lemma cmra_morphism_monotoneN n x y : x {n} y f x {n} f y.
Proof. intros [z ->]. exists (f z). by rewrite cmra_morphism_op. Qed.
Lemma cmra_morphism_valid x : x f x.
Proof. rewrite !cmra_valid_validN; eauto using cmra_morphism_validN. Qed.
End cmra_morphism.
(** COFE → CMRA Functors *)
Record rFunctor {SI : sidx} := RFunctor {
rFunctor_car : A `{!Cofe A} B `{!Cofe B}, cmra;
rFunctor_map `{!Cofe A1, !Cofe A2, !Cofe B1, !Cofe B2} :
((A2 -n> A1) * (B1 -n> B2)) rFunctor_car A1 B1 -n> rFunctor_car A2 B2;
rFunctor_map_ne `{!Cofe A1, !Cofe A2, !Cofe B1, !Cofe B2} :
NonExpansive (@rFunctor_map A1 _ A2 _ B1 _ B2 _);
rFunctor_map_id `{!Cofe A, !Cofe B} (x : rFunctor_car A B) :
rFunctor_map (cid,cid) x x;
rFunctor_map_compose `{!Cofe A1, !Cofe A2, !Cofe A3, !Cofe B1, !Cofe B2, !Cofe B3}
(f : A2 -n> A1) (g : A3 -n> A2) (f' : B1 -n> B2) (g' : B2 -n> B3) x :
rFunctor_map (fg, g'f') x rFunctor_map (g,g') (rFunctor_map (f,f') x);
rFunctor_mor `{!Cofe A1, !Cofe A2, !Cofe B1, !Cofe B2}
(fg : (A2 -n> A1) * (B1 -n> B2)) :
CmraMorphism (rFunctor_map fg)
}.
Global Existing Instances rFunctor_map_ne rFunctor_mor.
Global Instance: Params (@rFunctor_map) 10 := {}.
Declare Scope rFunctor_scope.
Delimit Scope rFunctor_scope with RF.
Bind Scope rFunctor_scope with rFunctor.
Class rFunctorContractive {SI : sidx} (F : rFunctor) :=
#[global] rFunctor_map_contractive `{!Cofe A1, !Cofe A2, !Cofe B1, !Cofe B2} ::
Contractive (@rFunctor_map SI F A1 _ A2 _ B1 _ B2 _).
Global Hint Mode rFunctorContractive - ! : typeclass_instances.
Definition rFunctor_apply {SI : sidx} (F: rFunctor) (A: ofe) `{!Cofe A} : cmra :=
rFunctor_car F A A.
Program Definition rFunctor_to_oFunctor {SI : sidx} (F: rFunctor) : oFunctor := {|
oFunctor_car A _ B _ := rFunctor_car F A B;
oFunctor_map A1 _ A2 _ B1 _ B2 _ fg := rFunctor_map F fg
|}.
Next Obligation.
intros ? F A ? B ? x. simpl in *. apply rFunctor_map_id.
Qed.
Next Obligation.
intros ? F A1 ? A2 ? A3 ? B1 ? B2 ? B3 ? f g f' g' x. simpl in *.
apply rFunctor_map_compose.
Qed.
Global Instance rFunctor_to_oFunctor_contractive {SI : sidx} F :
rFunctorContractive F oFunctorContractive (rFunctor_to_oFunctor F).
Proof.
intros ? A1 ? A2 ? B1 ? B2 ? n f g Hfg. apply rFunctor_map_contractive. done.
Qed.
Program Definition rFunctor_oFunctor_compose
{SI : sidx} (F1 : rFunctor) (F2 : oFunctor)
`{!∀ `{!Cofe A, !Cofe B}, Cofe (oFunctor_car F2 A B)} : rFunctor := {|
rFunctor_car A _ B _ := rFunctor_car F1 (oFunctor_car F2 B A) (oFunctor_car F2 A B);
rFunctor_map A1 _ A2 _ B1 _ B2 _ 'fg :=
rFunctor_map F1 (oFunctor_map F2 (fg.2,fg.1),oFunctor_map F2 fg)
|}.
Next Obligation.
intros ? F1 F2 ? A1 ? A2 ? B1 ? B2 ? n [f1 g1] [f2 g2] [??]; simpl in *.
apply rFunctor_map_ne; split; apply oFunctor_map_ne; by split.
Qed.
Next Obligation.
intros ? F1 F2 ? A ? B ? x; simpl in *. rewrite -{2}(rFunctor_map_id F1 x).
apply equiv_dist=> n. apply rFunctor_map_ne.
split=> y /=; by rewrite !oFunctor_map_id.
Qed.
Next Obligation.
intros ? F1 F2 ? A1 ? A2 ? A3 ? B1 ? B2 ? B3 ? f g f' g' x; simpl in *.
rewrite -rFunctor_map_compose. apply equiv_dist=> n. apply rFunctor_map_ne.
split=> y /=; by rewrite !oFunctor_map_compose.
Qed.
Global Instance rFunctor_oFunctor_compose_contractive_1
{SI : sidx} (F1 : rFunctor) (F2 : oFunctor)
`{!∀ `{!Cofe A, !Cofe B}, Cofe (oFunctor_car F2 A B)} :
rFunctorContractive F1 rFunctorContractive (rFunctor_oFunctor_compose F1 F2).
Proof.
intros ? A1 ? A2 ? B1 ? B2 ? n [f1 g1] [f2 g2] Hfg; simpl in *.
f_contractive; destruct Hfg; split; simpl in *; apply oFunctor_map_ne; by split.
Qed.
Global Instance rFunctor_oFunctor_compose_contractive_2
{SI : sidx} (F1 : rFunctor) (F2 : oFunctor)
`{!∀ `{!Cofe A, !Cofe B}, Cofe (oFunctor_car F2 A B)} :
oFunctorContractive F2 rFunctorContractive (rFunctor_oFunctor_compose F1 F2).
Proof.
intros ? A1 ? A2 ? B1 ? B2 ? n [f1 g1] [f2 g2] Hfg; simpl in *.
f_equiv; split; simpl in *; f_contractive; destruct Hfg; by split.
Qed.
Program Definition constRF {SI : sidx} (B : cmra) : rFunctor :=
{| rFunctor_car A1 _ A2 _ := B; rFunctor_map A1 _ A2 _ B1 _ B2 _ f := cid |}.
Solve Obligations with done.
Coercion constRF : cmra >-> rFunctor.
Global Instance constRF_contractive {SI : sidx} B : rFunctorContractive (constRF B).
Proof. rewrite /rFunctorContractive; apply _. Qed.
(** COFE → UCMRA Functors *)
Record urFunctor {SI : sidx} := URFunctor {
urFunctor_car : A `{!Cofe A} B `{!Cofe B}, ucmra;
urFunctor_map `{!Cofe A1, !Cofe A2, !Cofe B1, !Cofe B2} :
((A2 -n> A1) * (B1 -n> B2)) urFunctor_car A1 B1 -n> urFunctor_car A2 B2;
urFunctor_map_ne `{!Cofe A1, !Cofe A2, !Cofe B1, !Cofe B2} :
NonExpansive (@urFunctor_map A1 _ A2 _ B1 _ B2 _);
urFunctor_map_id `{!Cofe A, !Cofe B} (x : urFunctor_car A B) :
urFunctor_map (cid,cid) x x;
urFunctor_map_compose `{!Cofe A1, !Cofe A2, !Cofe A3, !Cofe B1, !Cofe B2, !Cofe B3}
(f : A2 -n> A1) (g : A3 -n> A2) (f' : B1 -n> B2) (g' : B2 -n> B3) x :
urFunctor_map (fg, g'f') x urFunctor_map (g,g') (urFunctor_map (f,f') x);
urFunctor_mor `{!Cofe A1, !Cofe A2, !Cofe B1, !Cofe B2}
(fg : (A2 -n> A1) * (B1 -n> B2)) :
CmraMorphism (urFunctor_map fg)
}.
Global Existing Instances urFunctor_map_ne urFunctor_mor.
Global Instance: Params (@urFunctor_map) 10 := {}.
Declare Scope urFunctor_scope.
Delimit Scope urFunctor_scope with URF.
Bind Scope urFunctor_scope with urFunctor.
Class urFunctorContractive {SI : sidx} (F : urFunctor) :=
#[global] urFunctor_map_contractive `{!Cofe A1, !Cofe A2, !Cofe B1, !Cofe B2} ::
Contractive (@urFunctor_map SI F A1 _ A2 _ B1 _ B2 _).
Global Hint Mode urFunctorContractive - ! : typeclass_instances.
Definition urFunctor_apply {SI : sidx} (F: urFunctor) (A: ofe) `{!Cofe A} : ucmra :=
urFunctor_car F A A.
Program Definition urFunctor_to_rFunctor {SI : sidx} (F: urFunctor) : rFunctor := {|
rFunctor_car A _ B _ := urFunctor_car F A B;
rFunctor_map A1 _ A2 _ B1 _ B2 _ fg := urFunctor_map F fg
|}.
Next Obligation.
intros ? F A ? B ? x. simpl in *. apply urFunctor_map_id.
Qed.
Next Obligation.
intros ? F A1 ? A2 ? A3 ? B1 ? B2 ? B3 ? f g f' g' x. simpl in *.
apply urFunctor_map_compose.
Qed.
Global Instance urFunctor_to_rFunctor_contractive {SI : sidx} F :
urFunctorContractive F rFunctorContractive (urFunctor_to_rFunctor F).
Proof.
intros ? A1 ? A2 ? B1 ? B2 ? n f g Hfg. apply urFunctor_map_contractive. done.
Qed.
Program Definition urFunctor_oFunctor_compose
{SI : sidx} (F1 : urFunctor) (F2 : oFunctor)
`{!∀ `{!Cofe A, !Cofe B}, Cofe (oFunctor_car F2 A B)} : urFunctor := {|
urFunctor_car A _ B _ := urFunctor_car F1 (oFunctor_car F2 B A) (oFunctor_car F2 A B);
urFunctor_map A1 _ A2 _ B1 _ B2 _ 'fg :=
urFunctor_map F1 (oFunctor_map F2 (fg.2,fg.1),oFunctor_map F2 fg)
|}.
Next Obligation.
intros ? F1 F2 ? A1 ? A2 ? B1 ? B2 ? n [f1 g1] [f2 g2] [??]; simpl in *.
apply urFunctor_map_ne; split; apply oFunctor_map_ne; by split.
Qed.
Next Obligation.
intros ? F1 F2 ? A ? B ? x; simpl in *. rewrite -{2}(urFunctor_map_id F1 x).
apply equiv_dist=> n. apply urFunctor_map_ne.
split=> y /=; by rewrite !oFunctor_map_id.
Qed.
Next Obligation.
intros ? F1 F2 ? A1 ? A2 ? A3 ? B1 ? B2 ? B3 ? f g f' g' x; simpl in *.
rewrite -urFunctor_map_compose. apply equiv_dist=> n. apply urFunctor_map_ne.
split=> y /=; by rewrite !oFunctor_map_compose.
Qed.
Global Instance urFunctor_oFunctor_compose_contractive_1
{SI : sidx} (F1 : urFunctor) (F2 : oFunctor)
`{!∀ `{!Cofe A, !Cofe B}, Cofe (oFunctor_car F2 A B)} :
urFunctorContractive F1 urFunctorContractive (urFunctor_oFunctor_compose F1 F2).
Proof.
intros ? A1 ? A2 ? B1 ? B2 ? n [f1 g1] [f2 g2] Hfg; simpl in *.
f_contractive; destruct Hfg; split; simpl in *; apply oFunctor_map_ne; by split.
Qed.
Global Instance urFunctor_oFunctor_compose_contractive_2
{SI : sidx} (F1 : urFunctor) (F2 : oFunctor)
`{!∀ `{!Cofe A, !Cofe B}, Cofe (oFunctor_car F2 A B)} :
oFunctorContractive F2 urFunctorContractive (urFunctor_oFunctor_compose F1 F2).
Proof.
intros ? A1 ? A2 ? B1 ? B2 ? n [f1 g1] [f2 g2] Hfg; simpl in *.
f_equiv; split; simpl in *; f_contractive; destruct Hfg; by split.
Qed.
Program Definition constURF {SI : sidx} (B : ucmra) : urFunctor :=
{| urFunctor_car A1 _ A2 _ := B; urFunctor_map A1 _ A2 _ B1 _ B2 _ f := cid |}.
Solve Obligations with done.
Coercion constURF : ucmra >-> urFunctor.
Global Instance constURF_contractive {SI : sidx} B :
urFunctorContractive (constURF B).
Proof. rewrite /urFunctorContractive; apply _. Qed.
(** * Transporting a CMRA equality *)
Definition cmra_transport {SI : sidx} {A B : cmra} (H : A = B) (x : A) : B :=
eq_rect A id x _ H.
Lemma cmra_transport_trans {SI : sidx} {A B C : cmra} (H1 : A = B) (H2 : B = C) x :
cmra_transport H2 (cmra_transport H1 x) = cmra_transport (eq_trans H1 H2) x.
Proof. by destruct H2. Qed.
Section cmra_transport.
Context {SI : sidx} {A B : cmra} (H : A = B).
Notation T := (cmra_transport H).
Global Instance cmra_transport_ne : NonExpansive T.
Proof. by intros ???; destruct H. Qed.
Global Instance cmra_transport_proper : Proper (() ==> ()) T.
Proof. by intros ???; destruct H. Qed.
Lemma cmra_transport_op x y : T (x y) = T x T y.
Proof. by destruct H. Qed.
Lemma cmra_transport_core x : T (core x) = core (T x).
Proof. by destruct H. Qed.
Lemma cmra_transport_validN n x : {n} T x {n} x.
Proof. by destruct H. Qed.
Lemma cmra_transport_valid x : T x x.
Proof. by destruct H. Qed.
Global Instance cmra_transport_discrete x : Discrete x Discrete (T x).
Proof. by destruct H. Qed.
Global Instance cmra_transport_core_id x : CoreId x CoreId (T x).
Proof. by destruct H. Qed.
End cmra_transport.
(** * Instances *)
(** ** Discrete CMRA *)
Record RAMixin A `{Equiv A, PCore A, Op A, Valid A} := {
(* setoids *)
ra_op_proper (x : A) : Proper (() ==> ()) (op x);
ra_core_proper (x y : A) cx :
x y pcore x = Some cx cy, pcore y = Some cy cx cy;
ra_validN_proper : Proper ((≡@{A}) ==> impl) valid;
(* monoid *)
ra_assoc : Assoc (≡@{A}) ();
ra_comm : Comm (≡@{A}) ();
ra_pcore_l (x : A) cx : pcore x = Some cx cx x x;
ra_pcore_idemp (x : A) cx : pcore x = Some cx pcore cx Some cx;
ra_pcore_mono (x y : A) cx :
x y pcore x = Some cx cy, pcore y = Some cy cx cy;
ra_valid_op_l (x y : A) : (x y) x
}.
Section discrete.
Local Set Default Proof Using "Type*".
Context {SI : sidx} `{!Equiv A, !PCore A, !Op A, !Valid A}.
Context (Heq : @Equivalence A ()).
Context (ra_mix : RAMixin A).
Existing Instances discrete_dist.
Local Instance discrete_validN_instance : ValidN A := λ n x, x.
Definition discrete_cmra_mixin : CmraMixin A.
Proof.
destruct ra_mix; split; try done.
- intros x; split; first done. by move=> /(_ 0).
- intros n x y1 y2 ??; by exists y1, y2.
Qed.
Local Instance discrete_cmra_discrete :
CmraDiscrete (Cmra' A (discrete_ofe_mixin Heq) discrete_cmra_mixin).
Proof. split; first apply _. done. Qed.
End discrete.
(** A smart constructor for the discrete RA over a carrier [A]. It uses
[ofe_discrete_equivalence_of A] to make sure the same [Equivalence] proof is
used as when constructing the OFE. *)
Notation discreteR A ra_mix :=
(Cmra A (discrete_cmra_mixin (discrete_ofe_equivalence_of A%type) ra_mix))
(only parsing).
Section ra_total.
Local Set Default Proof Using "Type*".
Context A `{Equiv A, PCore A, Op A, Valid A}.
Context (total : x : A, is_Some (pcore x)).
Context (op_proper : x : A, Proper (() ==> ()) (op x)).
Context (core_proper: Proper (() ==> ()) (@core A _)).
Context (valid_proper : Proper (() ==> impl) (@valid A _)).
Context (op_assoc : Assoc () (@op A _)).
Context (op_comm : Comm () (@op A _)).
Context (core_l : x : A, core x x x).
Context (core_idemp : x : A, core (core x) core x).
Context (core_mono : x y : A, x y core x core y).
Context (valid_op_l : x y : A, (x y) x).
Lemma ra_total_mixin : RAMixin A.
Proof.
split; auto.
- intros x y ? Hcx%core_proper Hx; move: Hcx. rewrite /core /= Hx /=.
case (total y)=> [cy ->]; eauto.
- intros x cx Hcx. move: (core_l x). by rewrite /core /= Hcx.
- intros x cx Hcx. move: (core_idemp x). rewrite /core /= Hcx /=.
case (total cx)=>[ccx ->]; by constructor.
- intros x y cx Hxy%core_mono Hx. move: Hxy.
rewrite /core /= Hx /=. case (total y)=> [cy ->]; eauto.
Qed.
End ra_total.
(** ** CMRA for the unit type *)
Section unit.
Context {SI : sidx}.
Local Instance unit_valid_instance : Valid () := λ x, True.
Local Instance unit_validN_instance : ValidN () := λ n x, True.
Local Instance unit_pcore_instance : PCore () := λ x, Some x.
Local Instance unit_op_instance : Op () := λ x y, ().
Lemma unit_cmra_mixin : CmraMixin ().
Proof. apply discrete_cmra_mixin, ra_total_mixin; by eauto. Qed.
Canonical Structure unitR : cmra := Cmra unit unit_cmra_mixin.
Local Instance unit_unit_instance : Unit () := ().
Lemma unit_ucmra_mixin : UcmraMixin ().
Proof. done. Qed.
Canonical Structure unitUR : ucmra := Ucmra unit unit_ucmra_mixin.
Global Instance unit_cmra_discrete : CmraDiscrete unitR.
Proof. done. Qed.
Global Instance unit_core_id (x : ()) : CoreId x.
Proof. by constructor. Qed.
Global Instance unit_cancelable (x : ()) : Cancelable x.
Proof. by constructor. Qed.
End unit.
(** ** CMRA for the empty type *)
Section empty.
Context {SI : sidx}.
Local Instance Empty_set_valid_instance : Valid Empty_set := λ x, False.
Local Instance Empty_set_validN_instance : ValidN Empty_set := λ n x, False.
Local Instance Empty_set_pcore_instance : PCore Empty_set := λ x, Some x.
Local Instance Empty_set_op_instance : Op Empty_set := λ x y, x.
Lemma Empty_set_cmra_mixin : CmraMixin Empty_set.
Proof. apply discrete_cmra_mixin, ra_total_mixin; by (intros [] || done). Qed.
Canonical Structure Empty_setR : cmra := Cmra Empty_set Empty_set_cmra_mixin.
Global Instance Empty_set_cmra_discrete : CmraDiscrete Empty_setR.
Proof. done. Qed.
Global Instance Empty_set_core_id (x : Empty_set) : CoreId x.
Proof. by constructor. Qed.
Global Instance Empty_set_cancelable (x : Empty_set) : Cancelable x.
Proof. by constructor. Qed.
End empty.
(** ** Product *)
Section prod.
Context {SI : sidx} {A B : cmra}.
Local Arguments pcore _ _ !_ /.
Local Arguments cmra_pcore _ !_/.
Local Instance prod_op_instance : Op (A * B) := λ x y, (x.1 y.1, x.2 y.2).
Local Instance prod_pcore_instance : PCore (A * B) := λ x,
c1 pcore (x.1); c2 pcore (x.2); Some (c1, c2).
Local Arguments prod_pcore_instance !_ /.
Local Instance prod_valid_instance : Valid (A * B) := λ x, x.1 x.2.
Local Instance prod_validN_instance : ValidN (A * B) := λ n x, {n} x.1 {n} x.2.
Lemma prod_pcore_Some (x cx : A * B) :
pcore x = Some cx pcore (x.1) = Some (cx.1) pcore (x.2) = Some (cx.2).
Proof. destruct x, cx; by intuition simplify_option_eq. Qed.
Lemma prod_pcore_Some' (x cx : A * B) :
pcore x Some cx pcore (x.1) Some (cx.1) pcore (x.2) Some (cx.2).
Proof.
split; [by intros (cx'&[-> ->]%prod_pcore_Some&<-)%Some_equiv_eq|].
rewrite {3}/pcore /prod_pcore_instance. (* TODO: use setoid rewrite *)
intros [Hx1 Hx2]; inversion_clear Hx1; simpl; inversion_clear Hx2.
by constructor.
Qed.
Lemma prod_included (x y : A * B) : x y x.1 y.1 x.2 y.2.
Proof.
split; [intros [z Hz]; split; [exists (z.1)|exists (z.2)]; apply Hz|].
intros [[z1 Hz1] [z2 Hz2]]; exists (z1,z2); split; auto.
Qed.
Lemma prod_includedN (x y : A * B) n : x {n} y x.1 {n} y.1 x.2 {n} y.2.
Proof.
split; [intros [z Hz]; split; [exists (z.1)|exists (z.2)]; apply Hz|].
intros [[z1 Hz1] [z2 Hz2]]; exists (z1,z2); split; auto.
Qed.
Definition prod_cmra_mixin : CmraMixin (A * B).
Proof.
split; try apply _.
- by intros n x y1 y2 [Hy1 Hy2]; split; rewrite /= ?Hy1 ?Hy2.
- intros n x y cx; setoid_rewrite prod_pcore_Some=> -[??] [??].
destruct (cmra_pcore_ne n (x.1) (y.1) (cx.1)) as (z1&->&?); auto.
destruct (cmra_pcore_ne n (x.2) (y.2) (cx.2)) as (z2&->&?); auto.
exists (z1,z2); repeat constructor; auto.
- by intros n y1 y2 [Hy1 Hy2] [??]; split; rewrite /= -?Hy1 -?Hy2.
- intros x; split.
+ intros [??] n; split; by apply cmra_valid_validN.
+ intros Hxy; split; apply cmra_valid_validN=> n; apply Hxy.
- intros n m x [??]; split; by eapply cmra_validN_le.
- by split; rewrite /= assoc.
- by split; rewrite /= comm.
- intros x y [??]%prod_pcore_Some;
constructor; simpl; eauto using cmra_pcore_l.
- intros x y; rewrite prod_pcore_Some prod_pcore_Some'.
naive_solver eauto using cmra_pcore_idemp.
- intros x y cx; rewrite prod_included prod_pcore_Some=> -[??] [??].
destruct (cmra_pcore_mono (x.1) (y.1) (cx.1)) as (z1&?&?); auto.
destruct (cmra_pcore_mono (x.2) (y.2) (cx.2)) as (z2&?&?); auto.
exists (z1,z2). by rewrite prod_included prod_pcore_Some.
- intros n x y [??]; split; simpl in *; eauto using cmra_validN_op_l.
- intros n x y1 y2 [??] [??]; simpl in *.
destruct (cmra_extend n (x.1) (y1.1) (y2.1)) as (z11&z12&?&?&?); auto.
destruct (cmra_extend n (x.2) (y1.2) (y2.2)) as (z21&z22&?&?&?); auto.
by exists (z11,z21), (z12,z22).
Qed.
Canonical Structure prodR := Cmra (prod A B) prod_cmra_mixin.
Lemma pair_op (a a' : A) (b b' : B) : (a a', b b') = (a, b) (a', b').
Proof. done. Qed.
Lemma pair_valid (a : A) (b : B) : (a, b) a b.
Proof. done. Qed.
Lemma pair_validN (a : A) (b : B) n : {n} (a, b) {n} a {n} b.
Proof. done. Qed.
Lemma pair_included (a a' : A) (b b' : B) :
(a, b) (a', b') a a' b b'.
Proof. apply prod_included. Qed.
Lemma pair_includedN (a a' : A) (b b' : B) n :
(a, b) {n} (a', b') a {n} a' b {n} b'.
Proof. apply prod_includedN. Qed.
Lemma pair_pcore (a : A) (b : B) :
pcore (a, b) = c1 pcore a; c2 pcore b; Some (c1, c2).
Proof. done. Qed.
Lemma pair_core `{!CmraTotal A, !CmraTotal B} (a : A) (b : B) :
core (a, b) = (core a, core b).
Proof.
rewrite /core {1}/pcore /=.
rewrite (cmra_pcore_core a) /= (cmra_pcore_core b). done.
Qed.
Global Instance prod_cmra_total : CmraTotal A CmraTotal B CmraTotal prodR.
Proof.
intros H1 H2 [a b]. destruct (H1 a) as [ca ?], (H2 b) as [cb ?].
exists (ca,cb); by simplify_option_eq.
Qed.
Global Instance prod_cmra_discrete :
CmraDiscrete A CmraDiscrete B CmraDiscrete prodR.
Proof. split; [apply _|]. by intros ? []; split; apply cmra_discrete_valid. Qed.
(* FIXME(Coq #6294): This is not an instance because we need it to use the new
unification. *)
Lemma pair_core_id x y :
CoreId x CoreId y CoreId (x,y).
Proof. by rewrite /CoreId prod_pcore_Some'. Qed.
Global Instance pair_exclusive_l x y : Exclusive x Exclusive (x,y).
Proof. by intros ?[][?%exclusive0_l]. Qed.
Global Instance pair_exclusive_r x y : Exclusive y Exclusive (x,y).
Proof. by intros ?[][??%exclusive0_l]. Qed.
Global Instance pair_cancelable x y :
Cancelable x Cancelable y Cancelable (x,y).
Proof. intros ???[][][][]. constructor; simpl in *; by eapply cancelableN. Qed.
Global Instance pair_id_free_l x y : IdFree x IdFree (x,y).
Proof. move=> Hx [a b] [? _] [/=? _]. apply (Hx a); eauto. Qed.
Global Instance pair_id_free_r x y : IdFree y IdFree (x,y).
Proof. move=> Hy [a b] [_ ?] [_ /=?]. apply (Hy b); eauto. Qed.
End prod.
(* Registering as [Hint Extern] with new unification. *)
Global Hint Extern 4 (CoreId _) =>
notypeclasses refine (pair_core_id _ _ _ _) : typeclass_instances.
Global Arguments prodR {_} _ _.
Section prod_unit.
Context {SI : sidx} {A B : ucmra}.
Local Instance prod_unit_instance `{Unit A, Unit B} : Unit (A * B) := (ε, ε).
Lemma prod_ucmra_mixin : UcmraMixin (A * B).
Proof.
split.
- split; apply ucmra_unit_valid.
- by split; rewrite /=left_id.
- rewrite prod_pcore_Some'; split; apply (core_id _).
Qed.
Canonical Structure prodUR := Ucmra (prod A B) prod_ucmra_mixin.
Lemma pair_split (a : A) (b : B) : (a, b) (a, ε) (ε, b).
Proof. by rewrite -pair_op left_id right_id. Qed.
Lemma pair_split_L `{!LeibnizEquiv A, !LeibnizEquiv B} (x : A) (y : B) :
(x, y) = (x, ε) (ε, y).
Proof. unfold_leibniz. apply pair_split. Qed.
Lemma pair_op_1 (a a': A) : (a a', ε) ≡@{A*B} (a, ε) (a', ε).
Proof. by rewrite -pair_op ucmra_unit_left_id. Qed.
Lemma pair_op_1_L `{!LeibnizEquiv A, !LeibnizEquiv B} (a a': A) :
(a a', ε) =@{A*B} (a, ε) (a', ε).
Proof. unfold_leibniz. apply pair_op_1. Qed.
Lemma pair_op_2 (b b': B) : (ε, b b') ≡@{A*B} (ε, b) (ε, b').
Proof. by rewrite -pair_op ucmra_unit_left_id. Qed.
Lemma pair_op_2_L `{!LeibnizEquiv A, !LeibnizEquiv B} (b b': B) :
(ε, b b') =@{A*B} (ε, b) (ε, b').
Proof. unfold_leibniz. apply pair_op_2. Qed.
End prod_unit.
Global Arguments prodUR {_} _ _.
Global Instance prod_map_cmra_morphism
{SI : sidx} {A A' B B' : cmra} (f : A A') (g : B B') :
CmraMorphism f CmraMorphism g CmraMorphism (prod_map f g).
Proof.
split; first apply _.
- by intros n x [??]; split; simpl; apply cmra_morphism_validN.
- intros [x1 x2]. rewrite /= !pair_pcore /=.
pose proof (Hf := cmra_morphism_pcore f (x1)).
destruct (pcore (f (x1))), (pcore (x1)); inv Hf=>//=.
pose proof (Hg := cmra_morphism_pcore g (x2)).
destruct (pcore (g (x2))), (pcore (x2)); inv Hg=>//=.
by setoid_subst.
- intros. by rewrite /prod_map /= !cmra_morphism_op.
Qed.
Program Definition prodRF {SI : sidx} (F1 F2 : rFunctor) : rFunctor := {|
rFunctor_car A _ B _ := prodR (rFunctor_car F1 A B) (rFunctor_car F2 A B);
rFunctor_map A1 _ A2 _ B1 _ B2 _ fg :=
prodO_map (rFunctor_map F1 fg) (rFunctor_map F2 fg)
|}.
Next Obligation.
intros ? F1 F2 A1 ? A2 ? B1 ? B2 ? n ???.
by apply prodO_map_ne; apply rFunctor_map_ne.
Qed.
Next Obligation. by intros ? F1 F2 A ? B ? [??]; rewrite /= !rFunctor_map_id. Qed.
Next Obligation.
intros ? F1 F2 A1 ? A2 ? A3 ? B1 ? B2 ? B3 ? f g f' g' [??]; simpl.
by rewrite !rFunctor_map_compose.
Qed.
Notation "F1 * F2" := (prodRF F1%RF F2%RF) : rFunctor_scope.
Global Instance prodRF_contractive {SI : sidx} F1 F2 :
rFunctorContractive F1 rFunctorContractive F2
rFunctorContractive (prodRF F1 F2).
Proof.
intros ?? A1 ? A2 ? B1 ? B2 ? n ???;
by apply prodO_map_ne; apply rFunctor_map_contractive.
Qed.
Program Definition prodURF {SI : sidx} (F1 F2 : urFunctor) : urFunctor := {|
urFunctor_car A _ B _ := prodUR (urFunctor_car F1 A B) (urFunctor_car F2 A B);
urFunctor_map A1 _ A2 _ B1 _ B2 _ fg :=
prodO_map (urFunctor_map F1 fg) (urFunctor_map F2 fg)
|}.
Next Obligation.
intros ? F1 F2 A1 ? A2 ? B1 ? B2 ? n ???.
by apply prodO_map_ne; apply urFunctor_map_ne.
Qed.
Next Obligation. by intros ? F1 F2 A ? B ? [??]; rewrite /= !urFunctor_map_id. Qed.
Next Obligation.
intros ? F1 F2 A1 ? A2 ? A3 ? B1 ? B2 ? B3 ? f g f' g' [??]; simpl.
by rewrite !urFunctor_map_compose.
Qed.
Notation "F1 * F2" := (prodURF F1%URF F2%URF) : urFunctor_scope.
Global Instance prodURF_contractive {SI : sidx} F1 F2 :
urFunctorContractive F1 urFunctorContractive F2
urFunctorContractive (prodURF F1 F2).
Proof.
intros ?? A1 ? A2 ? B1 ? B2 ? n ???;
by apply prodO_map_ne; apply urFunctor_map_contractive.
Qed.
(** ** CMRA for the option type *)
Section option.
Context {SI : sidx} {A : cmra}.
Implicit Types a b : A.
Implicit Types ma mb : option A.
Local Arguments core _ _ !_ /.
Local Arguments pcore _ _ !_ /.
Local Instance option_valid_instance : Valid (option A) := λ ma,
match ma with Some a => a | None => True end.
Local Instance option_validN_instance : ValidN (option A) := λ n ma,
match ma with Some a => {n} a | None => True end.
Local Instance option_pcore_instance : PCore (option A) := λ ma,
Some (ma ≫= pcore).
Local Arguments option_pcore_instance !_ /.
Local Instance option_op_instance : Op (option A) :=
union_with (λ a b, Some (a b)).
Definition Some_valid a : Some a a := reflexivity _.
Definition Some_validN a n : {n} Some a {n} a := reflexivity _.
Definition Some_op a b : Some (a b) = Some a Some b := eq_refl.
Lemma Some_core `{!CmraTotal A} a : Some (core a) = core (Some a).
Proof. rewrite /core /=. by destruct (cmra_total a) as [? ->]. Qed.
Lemma pcore_Some a : pcore (Some a) = Some (pcore a).
Proof. done. Qed.
Lemma Some_op_opM a ma : Some a ma = Some (a ? ma).
Proof. by destruct ma. Qed.
Lemma option_included ma mb :
ma mb ma = None a b, ma = Some a mb = Some b (a b a b).
Proof.
split.
- intros [mc Hmc].
destruct ma as [a|]; [right|by left].
destruct mb as [b|]; [exists a, b|destruct mc; inversion_clear Hmc].
destruct mc as [c|]; inversion_clear Hmc; split_and?; auto;
setoid_subst; eauto.
- intros [->|(a&b&->&->&[Hc|[c Hc]])].
+ exists mb. by destruct mb.
+ exists None; by constructor.
+ exists (Some c); by constructor.
Qed.
Lemma option_included_total `{!CmraTotal A} ma mb :
ma mb ma = None a b, ma = Some a mb = Some b a b.
Proof.
rewrite option_included. split; last naive_solver.
intros [->|(a&b&->&->&[Hab|?])]; [by eauto| |by eauto 10].
right. exists a, b. by rewrite {3}Hab.
Qed.
Lemma option_includedN n ma mb :
ma {n} mb ma = None
x y, ma = Some x mb = Some y (x {n} y x {n} y).
Proof.
split.
- intros [mc Hmc].
destruct ma as [a|]; [right|by left].
destruct mb as [b|]; [exists a, b|destruct mc; inversion_clear Hmc].
destruct mc as [c|]; inversion_clear Hmc; split_and?; auto;
ofe_subst; eauto using cmra_includedN_l.
- intros [->|(a&y&->&->&[Hc|[c Hc]])].
+ exists mb. by destruct mb.
+ exists None; by constructor.
+ exists (Some c); by constructor.
Qed.
Lemma option_includedN_total `{!CmraTotal A} n ma mb :
ma {n} mb ma = None a b, ma = Some a mb = Some b a {n} b.
Proof.
rewrite option_includedN. split; last naive_solver.
intros [->|(a&b&->&->&[Hab|?])]; [by eauto| |by eauto 10].
right. exists a, b. by rewrite {3}Hab.
Qed.
(* See below for more [included] lemmas. *)
Lemma option_cmra_mixin : CmraMixin (option A).
Proof.
apply cmra_total_mixin.
- eauto.
- by intros [a|] n; destruct 1; constructor; ofe_subst.
- destruct 1; by ofe_subst.
- by destruct 1; rewrite /validN /option_validN_instance //=; ofe_subst.
- intros [a|]; [apply cmra_valid_validN|done].
- intros n m [a|];
unfold validN, option_validN_instance; eauto using cmra_validN_le.
- intros [a|] [b|] [c|]; constructor; rewrite ?assoc; auto.
- intros [a|] [b|]; constructor; rewrite 1?comm; auto.
- intros [a|]; simpl; auto.
destruct (pcore a) as [ca|] eqn:?; constructor; eauto using cmra_pcore_l.
- intros [a|]; simpl; auto.
destruct (pcore a) as [ca|] eqn:?; simpl; eauto using cmra_pcore_idemp.
- intros ma mb; setoid_rewrite option_included.
intros [->|(a&b&->&->&[?|?])]; simpl; eauto.
+ destruct (pcore a) as [ca|] eqn:?; eauto.
destruct (cmra_pcore_proper a b ca) as (?&?&?); eauto 10.
+ destruct (pcore a) as [ca|] eqn:?; eauto.
destruct (cmra_pcore_mono a b ca) as (?&?&?); eauto 10.
- intros n [a|] [b|]; rewrite /validN /option_validN_instance /=;
eauto using cmra_validN_op_l.
- intros n ma mb1 mb2.
destruct ma as [a|], mb1 as [b1|], mb2 as [b2|]; intros Hx Hx';
(try by exfalso; inversion Hx'); (try (apply (inj Some) in Hx')).
+ destruct (cmra_extend n a b1 b2) as (c1&c2&?&?&?); auto.
by exists (Some c1), (Some c2); repeat constructor.
+ by exists (Some a), None; repeat constructor.
+ by exists None, (Some a); repeat constructor.
+ exists None, None; repeat constructor.
Qed.
Canonical Structure optionR := Cmra (option A) option_cmra_mixin.
Global Instance option_cmra_discrete : CmraDiscrete A CmraDiscrete optionR.
Proof. split; [apply _|]. by intros [a|]; [apply (cmra_discrete_valid a)|]. Qed.
Local Instance option_unit_instance : Unit (option A) := None.
Lemma option_ucmra_mixin : UcmraMixin optionR.
Proof. split; [done| |done]. by intros []. Qed.
Canonical Structure optionUR := Ucmra (option A) option_ucmra_mixin.
(** Misc *)
Lemma op_None ma mb : ma mb = None ma = None mb = None.
Proof. destruct ma, mb; naive_solver. Qed.
Lemma op_is_Some ma mb : is_Some (ma mb) is_Some ma is_Some mb.
Proof. rewrite -!not_eq_None_Some op_None. destruct ma, mb; naive_solver. Qed.
(* When the goal is already of the form [None ⋅ x], the [LeftId] instance for
[ε] does not fire. *)
Global Instance op_None_left_id : LeftId (=) None (@op (option A) _).
Proof. intros [a|]; done. Qed.
Global Instance op_None_right_id : RightId (=) None (@op (option A) _).
Proof. intros [a|]; done. Qed.
Lemma cmra_opM_opM_assoc a mb mc : a ? mb ? mc a ? (mb mc).
Proof. destruct mb, mc; by rewrite /= -?assoc. Qed.
Lemma cmra_opM_opM_assoc_L `{!LeibnizEquiv A} a mb mc :
a ? mb ? mc = a ? (mb mc).
Proof. unfold_leibniz. apply cmra_opM_opM_assoc. Qed.
Lemma cmra_opM_opM_swap a mb mc : a ? mb ? mc a ? mc ? mb.
Proof. by rewrite !cmra_opM_opM_assoc (comm _ mb). Qed.
Lemma cmra_opM_opM_swap_L `{!LeibnizEquiv A} a mb mc :
a ? mb ? mc = a ? mc ? mb.
Proof. by rewrite !cmra_opM_opM_assoc_L (comm_L _ mb). Qed.
Lemma cmra_opM_fmap_Some ma1 ma2 : ma1 ? (Some <$> ma2) = ma1 ma2.
Proof. by destruct ma1, ma2. Qed.
Global Instance Some_core_id a : CoreId a CoreId (Some a).
Proof. by constructor. Qed.
Global Instance option_core_id ma : ( x : A, CoreId x) CoreId ma.
Proof. intros. destruct ma; apply _. Qed.
Lemma exclusiveN_Some_l n a `{!Exclusive a} mb :
{n} (Some a mb) mb = None.
Proof. destruct mb; last done. move=> /(exclusiveN_l _ a) []. Qed.
Lemma exclusiveN_Some_r n a `{!Exclusive a} mb :
{n} (mb Some a) mb = None.
Proof. rewrite comm. by apply exclusiveN_Some_l. Qed.
Lemma exclusive_Some_l a `{!Exclusive a} mb : (Some a mb) mb = None.
Proof. destruct mb; last done. move=> /(exclusive_l a) []. Qed.
Lemma exclusive_Some_r a `{!Exclusive a} mb : (mb Some a) mb = None.
Proof. rewrite comm. by apply exclusive_Some_l. Qed.
Lemma Some_includedN n a b : Some a {n} Some b a {n} b a {n} b.
Proof. rewrite option_includedN; naive_solver. Qed.
Lemma Some_includedN_1 n a b : Some a {n} Some b a {n} b a {n} b.
Proof. rewrite Some_includedN. auto. Qed.
Lemma Some_includedN_2 n a b : a {n} b a {n} b Some a {n} Some b.
Proof. rewrite Some_includedN. auto. Qed.
Lemma Some_includedN_mono n a b : a {n} b Some a {n} Some b.
Proof. rewrite Some_includedN. auto. Qed.
Lemma Some_includedN_refl n a b : a {n} b Some a {n} Some b.
Proof. rewrite Some_includedN. auto. Qed.
Lemma Some_includedN_is_Some n x mb : Some x {n} mb is_Some mb.
Proof. rewrite option_includedN. naive_solver. Qed.
Lemma Some_included a b : Some a Some b a b a b.
Proof. rewrite option_included; naive_solver. Qed.
Lemma Some_included_1 a b : Some a Some b a b a b.
Proof. rewrite Some_included. auto. Qed.
Lemma Some_included_2 a b : a b a b Some a Some b.
Proof. rewrite Some_included. auto. Qed.
Lemma Some_included_mono a b : a b Some a Some b.
Proof. rewrite Some_included. auto. Qed.
Lemma Some_included_refl a b : a b Some a Some b.
Proof. rewrite Some_included. auto. Qed.
Lemma Some_included_is_Some x mb : Some x mb is_Some mb.
Proof. rewrite option_included. naive_solver. Qed.
Lemma Some_includedN_opM n a b : Some a {n} Some b mc, b {n} a ? mc.
Proof.
rewrite /includedN. f_equiv=> mc. by rewrite -(inj_iff Some b) Some_op_opM.
Qed.
Lemma Some_included_opM a b : Some a Some b mc, b a ? mc.
Proof.
rewrite /included. f_equiv=> mc. by rewrite -(inj_iff Some b) Some_op_opM.
Qed.
Lemma cmra_validN_Some_includedN n a b : {n} a Some b {n} Some a {n} b.
Proof. apply: (cmra_validN_includedN _ (Some _) (Some _)). Qed.
Lemma cmra_valid_Some_included a b : a Some b Some a b.
Proof. apply: (cmra_valid_included (Some _) (Some _)). Qed.
Lemma Some_includedN_total `{!CmraTotal A} n a b : Some a {n} Some b a {n} b.
Proof. rewrite Some_includedN. split; [|by eauto]. by intros [->|?]. Qed.
Lemma Some_included_total `{!CmraTotal A} a b : Some a Some b a b.
Proof. rewrite Some_included. split; [|by eauto]. by intros [->|?]. Qed.
Lemma Some_includedN_exclusive n a `{!Exclusive a} b :
Some a {n} Some b {n} b a {n} b.
Proof. move=> /Some_includedN [//|/exclusive_includedN]; tauto. Qed.
Lemma Some_included_exclusive a `{!Exclusive a} b :
Some a Some b b a b.
Proof. move=> /Some_included [//|/exclusive_included]; tauto. Qed.
Lemma is_Some_includedN n ma mb : ma {n} mb is_Some ma is_Some mb.
Proof. rewrite -!not_eq_None_Some option_includedN. naive_solver. Qed.
Lemma is_Some_included ma mb : ma mb is_Some ma is_Some mb.
Proof. rewrite -!not_eq_None_Some option_included. naive_solver. Qed.
Global Instance cancelable_Some a :
IdFree a Cancelable a Cancelable (Some a).
Proof.
intros Hirr ? n [b|] [c|] ? EQ; inversion_clear EQ.
- constructor. by apply (cancelableN a).
- destruct (Hirr b); [|eauto using dist_le, SIdx.le_0_l].
by eapply (cmra_validN_op_l 0 a b), (cmra_validN_le n), SIdx.le_0_l.
- destruct (Hirr c); [|symmetry; eauto using dist_le, SIdx.le_0_l].
by eapply (cmra_validN_le n), SIdx.le_0_l.
- done.
Qed.
Global Instance option_cancelable (ma : option A) :
( a : A, IdFree a) ( a : A, Cancelable a) Cancelable ma.
Proof. destruct ma; apply _. Qed.
End option.
Global Arguments optionR {_} _.
Global Arguments optionUR {_} _.
Section option_prod.
Context {SI : sidx} {A B : cmra}.
Implicit Types a : A.
Implicit Types b : B.
Lemma Some_pair_includedN n a1 a2 b1 b2 :
Some (a1,b1) {n} Some (a2,b2) Some a1 {n} Some a2 Some b1 {n} Some b2.
Proof. rewrite !Some_includedN. intros [[??]|[??]%prod_includedN]; eauto. Qed.
Lemma Some_pair_includedN_l n a1 a2 b1 b2 :
Some (a1,b1) {n} Some (a2,b2) Some a1 {n} Some a2.
Proof. intros. eapply Some_pair_includedN. done. Qed.
Lemma Some_pair_includedN_r n a1 a2 b1 b2 :
Some (a1,b1) {n} Some (a2,b2) Some b1 {n} Some b2.
Proof. intros. eapply Some_pair_includedN. done. Qed.
Lemma Some_pair_includedN_total_1 `{!CmraTotal A} n a1 a2 b1 b2 :
Some (a1,b1) {n} Some (a2,b2) a1 {n} a2 Some b1 {n} Some b2.
Proof. intros ?%Some_pair_includedN. by rewrite -(Some_includedN_total _ a1). Qed.
Lemma Some_pair_includedN_total_2 `{!CmraTotal B} n a1 a2 b1 b2 :
Some (a1,b1) {n} Some (a2,b2) Some a1 {n} Some a2 b1 {n} b2.
Proof. intros ?%Some_pair_includedN. by rewrite -(Some_includedN_total _ b1). Qed.
Lemma Some_pair_included a1 a2 b1 b2 :
Some (a1,b1) Some (a2,b2) Some a1 Some a2 Some b1 Some b2.
Proof. rewrite !Some_included. intros [[??]|[??]%prod_included]; eauto. Qed.
Lemma Some_pair_included_l a1 a2 b1 b2 :
Some (a1,b1) Some (a2,b2) Some a1 Some a2.
Proof. intros. eapply Some_pair_included. done. Qed.
Lemma Some_pair_included_r a1 a2 b1 b2 :
Some (a1,b1) Some (a2,b2) Some b1 Some b2.
Proof. intros. eapply Some_pair_included. done. Qed.
Lemma Some_pair_included_total_1 `{!CmraTotal A} a1 a2 b1 b2 :
Some (a1,b1) Some (a2,b2) a1 a2 Some b1 Some b2.
Proof. intros ?%Some_pair_included. by rewrite -(Some_included_total a1). Qed.
Lemma Some_pair_included_total_2 `{!CmraTotal B} a1 a2 b1 b2 :
Some (a1,b1) Some (a2,b2) Some a1 Some a2 b1 b2.
Proof. intros ?%Some_pair_included. by rewrite -(Some_included_total b1). Qed.
End option_prod.
Lemma option_fmap_mono {SI : sidx} {A B : cmra} (f : A B) (ma mb : option A) :
Proper (() ==> ()) f
( a b, a b f a f b)
ma mb f <$> ma f <$> mb.
Proof.
intros ??. rewrite !option_included; intros [->|(a&b&->&->&?)]; naive_solver.
Qed.
Global Instance option_fmap_cmra_morphism {SI : sidx}
{A B : cmra} (f: A B) `{!CmraMorphism f} :
CmraMorphism (fmap f : option A option B).
Proof.
split; first apply _.
- intros n [a|] ?; rewrite /cmra_validN //=. by apply (cmra_morphism_validN f).
- move=> [a|] //. by apply Some_proper, cmra_morphism_pcore.
- move=> [a|] [b|] //=. by rewrite (cmra_morphism_op f).
Qed.
Program Definition optionURF {SI : sidx} (F : rFunctor) : urFunctor := {|
urFunctor_car A _ B _ := optionUR (rFunctor_car F A B);
urFunctor_map A1 _ A2 _ B1 _ B2 _ fg := optionO_map (rFunctor_map F fg)
|}.
Next Obligation.
intros ? F A1 ? A2 ? B1 ? B2 ? n f g Hfg.
by apply optionO_map_ne, rFunctor_map_ne.
Qed.
Next Obligation.
intros ? F A ? B ? x. rewrite /= -{2}(option_fmap_id x).
apply option_fmap_equiv_ext=>y; apply rFunctor_map_id.
Qed.
Next Obligation.
intros ? F A1 ? A2 ? A3 ? B1 ? B2 ? B3 ? f g f' g' x.
rewrite /= -option_fmap_compose.
apply option_fmap_equiv_ext=>y; apply rFunctor_map_compose.
Qed.
Global Instance optionURF_contractive {SI : sidx} F :
rFunctorContractive F urFunctorContractive (optionURF F).
Proof.
intros ? A1 ? A2 ? B1 ? B2 ? n f g Hfg.
by apply optionO_map_ne, rFunctor_map_contractive.
Qed.
Program Definition optionRF {SI : sidx} (F : rFunctor) : rFunctor := {|
rFunctor_car A _ B _ := optionR (rFunctor_car F A B);
rFunctor_map A1 _ A2 _ B1 _ B2 _ fg := optionO_map (rFunctor_map F fg)
|}.
Solve Obligations with apply @optionURF.
Global Instance optionRF_contractive {SI : sidx} F :
rFunctorContractive F rFunctorContractive (optionRF F).
Proof. apply optionURF_contractive. Qed.
(* Dependently-typed functions over a discrete domain *)
Section discrete_fun_cmra.
Context {SI : sidx} {A: Type} {B : A ucmra}.
Implicit Types f g : discrete_fun B.
Local Instance discrete_fun_op_instance : Op (discrete_fun B) := λ f g x,
f x g x.
Local Instance discrete_fun_pcore_instance : PCore (discrete_fun B) := λ f,
Some (λ x, core (f x)).
Local Instance discrete_fun_valid_instance : Valid (discrete_fun B) := λ f,
x, f x.
Local Instance discrete_fun_validN_instance : ValidN (discrete_fun B) := λ n f,
x, {n} f x.
Definition discrete_fun_lookup_op f g x : (f g) x = f x g x := eq_refl.
Definition discrete_fun_lookup_core f x : (core f) x = core (f x) := eq_refl.
Lemma discrete_fun_included_spec_1 (f g : discrete_fun B) x : f g f x g x.
Proof.
by intros [h Hh]; exists (h x); rewrite /op /discrete_fun_op_instance (Hh x).
Qed.
Lemma discrete_fun_included_spec `{Finite A} (f g : discrete_fun B) :
f g x, f x g x.
Proof.
split; [by intros; apply discrete_fun_included_spec_1|].
intros [h ?]%finite_choice; by exists h.
Qed.
Lemma discrete_fun_cmra_mixin : CmraMixin (discrete_fun B).
Proof.
apply cmra_total_mixin.
- eauto.
- intros n f1 f2 f3 Hf x. by rewrite discrete_fun_lookup_op (Hf x).
- intros n f1 f2 Hf x. by rewrite discrete_fun_lookup_core (Hf x).
- intros n f1 f2 Hf ? x. by rewrite -(Hf x).
- intros g; split.
+ intros Hg n i; apply cmra_valid_validN, Hg.
+ intros Hg i; apply cmra_valid_validN=> n; apply Hg.
- intros n n' f Hf ? x. eauto using cmra_validN_le.
- intros f1 f2 f3 x. by rewrite discrete_fun_lookup_op assoc.
- intros f1 f2 x. by rewrite discrete_fun_lookup_op comm.
- intros f x.
by rewrite discrete_fun_lookup_op discrete_fun_lookup_core cmra_core_l.
- intros f x. by rewrite discrete_fun_lookup_core cmra_core_idemp.
- intros f1 f2 Hf12. exists (core f2)=>x. rewrite discrete_fun_lookup_op.
apply (discrete_fun_included_spec_1 _ _ x), (cmra_core_mono (f1 x)) in Hf12.
rewrite !discrete_fun_lookup_core. destruct Hf12 as [? ->].
rewrite assoc -cmra_core_dup //.
- intros n f1 f2 Hf x. apply cmra_validN_op_l with (f2 x), Hf.
- intros n f f1 f2 Hf Hf12.
assert (FUN := λ x, cmra_extend n (f x) (f1 x) (f2 x) (Hf x) (Hf12 x)).
exists (λ x, projT1 (FUN x)), (λ x, proj1_sig (projT2 (FUN x))).
split; [|split]=>x; [rewrite discrete_fun_lookup_op| |];
by destruct (FUN x) as (?&?&?&?&?).
Qed.
Canonical Structure discrete_funR :=
Cmra (discrete_fun B) discrete_fun_cmra_mixin.
Local Instance discrete_fun_unit_instance : Unit (discrete_fun B) := λ x, ε.
Definition discrete_fun_lookup_empty x : ε x = ε := eq_refl.
Lemma discrete_fun_ucmra_mixin : UcmraMixin (discrete_fun B).
Proof.
split.
- intros x. apply ucmra_unit_valid.
- intros f x. by rewrite discrete_fun_lookup_op left_id.
- constructor=> x. apply core_id_core, _.
Qed.
Canonical Structure discrete_funUR :=
Ucmra (discrete_fun B) discrete_fun_ucmra_mixin.
Global Instance discrete_fun_unit_discrete :
( i, Discrete (ε : B i)) Discrete (ε : discrete_fun B).
Proof. intros ? f Hf x. by apply: discrete. Qed.
End discrete_fun_cmra.
Global Arguments discrete_funR {_ _} _.
Global Arguments discrete_funUR {_ _} _.
Global Instance discrete_fun_map_cmra_morphism
{SI : sidx} {A} {B1 B2 : A ucmra} (f : x, B1 x B2 x) :
( x, CmraMorphism (f x)) CmraMorphism (discrete_fun_map f).
Proof.
split; first apply _.
- intros n g Hg x. rewrite /discrete_fun_map.
apply (cmra_morphism_validN (f _)), Hg.
- intros. apply Some_proper=>i. apply (cmra_morphism_core (f i)).
- intros g1 g2 i.
by rewrite /discrete_fun_map discrete_fun_lookup_op cmra_morphism_op.
Qed.
Program Definition discrete_funURF
{SI : sidx} {C} (F : C urFunctor) : urFunctor := {|
urFunctor_car A _ B _ := discrete_funUR (λ c, urFunctor_car (F c) A B);
urFunctor_map A1 _ A2 _ B1 _ B2 _ fg :=
discrete_funO_map (λ c, urFunctor_map (F c) fg)
|}.
Next Obligation.
intros ? C F A1 ? A2 ? B1 ? B2 ? n ?? g.
by apply discrete_funO_map_ne=>?; apply urFunctor_map_ne.
Qed.
Next Obligation.
intros ? C F A ? B ? g; simpl. rewrite -{2}(discrete_fun_map_id g).
apply discrete_fun_map_ext=> y; apply urFunctor_map_id.
Qed.
Next Obligation.
intros ? C F A1 ? A2 ? A3 ? B1 ? B2 ? B3 ? f1 f2 f1' f2' g.
rewrite /=-discrete_fun_map_compose.
apply discrete_fun_map_ext=>y; apply urFunctor_map_compose.
Qed.
Global Instance discrete_funURF_contractive {SI : sidx} {C} (F : C urFunctor) :
( c, urFunctorContractive (F c)) urFunctorContractive (discrete_funURF F).
Proof.
intros ? A1 ? A2 ? B1 ? B2 ? n ?? g.
by apply discrete_funO_map_ne=> c; apply urFunctor_map_contractive.
Qed.
(** * Constructing a camera [B] through a mapping into [A]
The mapping may restrict the domain (i.e., we have an injection from [B] to [A],
not a bijection) and validity. These two restrictions work on opposite "ends" of
[A] according to [≼]: domain restriction must prove that when an element is in
the domain, so is its composition with other elements; validity restriction must
prove that if the composition of two elements is valid, then so are both of the
elements. The "domain" is the image of [g] in [A], or equivalently the part of
[A] where [f] returns [Some]. *)
Lemma inj_cmra_mixin_restrict_validity {SI : sidx} {A : cmra} {B : ofe}
`{!PCore B, !Op B, !Valid B, !ValidN B}
(f : A option B) (g : B A)
(* [g] is proper/non-expansive and injective w.r.t. OFE equality *)
(g_dist : n y1 y2, y1 {n} y2 g y1 {n} g y2)
(* [g] is surjective into the part of [A] where [is_Some ∘ f] holds
(and [f] its inverse) *)
(gf_dist : (x : A) (y : B) n, f x {n} Some y g y {n} x)
(* [g] commutes with [pcore] (on the part where it is defined) and [op] *)
(g_pcore_dist : (y cy : B) n,
pcore y {n} Some cy pcore (g y) {n} Some (g cy))
(g_op : (y1 y2 : B), g (y1 y2) g y1 g y2)
(* [g] also commutes with [opM] when the right-hand side is produced by [f],
and cancels the [f]. In particular this axiom implies that when taking an
element in the domain ([g y]), its composition with *any* [x : A] is still in
the domain, and [f] computes the preimage properly.
Note that just requiring "the composition of two elements from the domain
is in the domain" is insufficient for this lemma to hold. [g_op] already shows
that this is the case, but the issue is that in [pcore_mono] we obtain a
[g y1 ≼ g y2], and the existentially quantified "remainder" in the [≼] has no
reason to be in the domain, so [g_op] is too weak to turn this into some
relation between [y1] and [y2] in [B]. At the same time, [g_opM_f] does not
impl [g_op] since we need [g_op] to prove that [⋅] in [B] respects [≡].
Therefore both [g_op] and [g_opM_f] are required for this lemma to work. *)
(g_opM_f : (x : A) (y : B), g (y ? f x) g y x)
(* The validity predicate on [B] restricts the one on [A] *)
(g_validN : n (y : B), {n} y {n} (g y))
(* The validity predicate on [B] satisfies the laws of validity *)
(valid_validN_ne : n, Proper (dist n ==> impl) (validN (A:=B) n))
(valid_rvalidN : y : B, y n, {n} y)
(validN_le : n n' (y : B), {n} y n' n {n'} y)
(validN_op_l : n (y1 y2 : B), {n} (y1 y2) {n} y1) :
CmraMixin B.
Proof.
(* Some general derived facts that will be useful later. *)
assert (g_equiv : y1 y2, y1 y2 g y1 g y2).
{ intros ??. rewrite !equiv_dist. naive_solver. }
assert (g_pcore : (y cy : B),
pcore y Some cy pcore (g y) Some (g cy)).
{ intros. rewrite !equiv_dist. naive_solver. }
assert (gf : x y, f x Some y g y x).
{ intros. rewrite !equiv_dist. naive_solver. }
assert (fg : y, f (g y) Some y).
{ intros. apply gf. done. }
assert (g_ne : NonExpansive g).
{ intros n ???. apply g_dist. done. }
(* Some of the CMRA properties are useful in proving the others. *)
assert (b_pcore_l' : y cy : B, pcore y Some cy cy y y).
{ intros y cy Hy. apply g_equiv. rewrite g_op. apply cmra_pcore_l'.
apply g_pcore. done. }
assert (b_pcore_idemp : y cy : B, pcore y Some cy pcore cy Some cy).
{ intros y cy Hy. eapply g_pcore, cmra_pcore_idemp', g_pcore. done. }
(* Now prove all the mixin laws. *)
split.
- intros y n z1 z2 Hz%g_dist. apply g_dist. by rewrite !g_op Hz.
- intros n y1 y2 cy Hy%g_dist Hy1.
assert (g <$> pcore y2 {n} Some (g cy))
as (cx & (cy'&->&->)%fmap_Some_1 & ?%g_dist)%dist_Some_inv_r'; [|by eauto].
assert (Hgcy : pcore (g y1) Some (g cy)).
{ apply g_pcore. rewrite Hy1. done. }
rewrite equiv_dist in Hgcy. specialize (Hgcy n).
rewrite Hy in Hgcy. apply g_pcore_dist in Hgcy. rewrite Hgcy. done.
- done.
- done.
- done.
- intros y1 y2 y3. apply g_equiv. by rewrite !g_op assoc.
- intros y1 y2. apply g_equiv. by rewrite !g_op comm.
- intros y cy Hcy. apply b_pcore_l'. by rewrite Hcy.
- intros y cy Hcy. eapply b_pcore_idemp. by rewrite -Hcy.
- intros y1 y2 cy [z Hy2] Hy1.
destruct (cmra_pcore_mono' (g y1) (g y2) (g cy)) as (cx&Hcgy2&[x Hcx]).
{ exists (g z). rewrite -g_op. by apply g_equiv. }
{ apply g_pcore. rewrite Hy1 //. }
apply (reflexive_eq (R:=equiv)) in Hcgy2.
rewrite -g_opM_f in Hcx. rewrite Hcx in Hcgy2.
apply g_pcore in Hcgy2.
apply Some_equiv_eq in Hcgy2 as [cy' [-> Hcy']].
eexists. split; first done.
destruct (f x) as [y|].
+ exists y. done.
+ exists cy. apply (reflexive_eq (R:=equiv)), b_pcore_idemp, b_pcore_l' in Hy1.
rewrite Hy1 //.
- done.
- intros n y z1 z2 ?%g_validN ?.
destruct (cmra_extend n (g y) (g z1) (g z2)) as (x1&x2&Hgy&Hx1&Hx2).
{ done. }
{ rewrite -g_op. by apply g_dist. }
symmetry in Hx1, Hx2.
apply gf_dist in Hx1, Hx2.
destruct (f x1) as [y1|] eqn:Hy1; last first.
{ exfalso. inversion Hx1. }
destruct (f x2) as [y2|] eqn:Hy2; last first.
{ exfalso. inversion Hx2. }
exists y1, y2. split_and!.
+ apply g_equiv. rewrite Hgy g_op.
f_equiv; symmetry; apply gf; rewrite ?Hy1 ?Hy2 //.
+ apply g_dist. apply (inj Some) in Hx1. rewrite Hx1 //.
+ apply g_dist. apply (inj Some) in Hx2. rewrite Hx2 //.
Qed.
(** Constructing a CMRA through an isomorphism that may restrict validity. *)
Lemma iso_cmra_mixin_restrict_validity {SI : sidx} {A : cmra} {B : ofe}
`{!PCore B, !Op B, !Valid B, !ValidN B}
(f : A B) (g : B A)
(* [g] is proper/non-expansive and injective w.r.t. setoid and OFE equality *)
(g_dist : n y1 y2, y1 {n} y2 g y1 {n} g y2)
(* [g] is surjective (and [f] its inverse) *)
(gf : x : A, g (f x) x)
(* [g] commutes with [pcore] and [op] *)
(g_pcore : y : B, pcore (g y) g <$> pcore y)
(g_op : y1 y2, g (y1 y2) g y1 g y2)
(* The validity predicate on [B] restricts the one on [A] *)
(g_validN : n y, {n} y {n} (g y))
(* The validity predicate on [B] satisfies the laws of validity *)
(valid_validN_ne : n, Proper (dist n ==> impl) (validN (A:=B) n))
(valid_rvalidN : y : B, y n, {n} y)
(validN_le: n m (y : B), {n} y m n {m} y)
(validN_op_l : n (y1 y2 : B), {n} (y1 y2) {n} y1) :
CmraMixin B.
Proof.
assert (g_ne : NonExpansive g).
{ intros n ???. apply g_dist. done. }
assert (g_equiv : y1 y2, y1 y2 g y1 g y2).
{ intros ??.
split; intros ?; apply equiv_dist; intros; apply g_dist, equiv_dist; done. }
apply (inj_cmra_mixin_restrict_validity (λ x, Some (f x)) g); try done.
- intros. split.
+ intros Hy%(inj Some). rewrite -Hy gf //.
+ intros ?. f_equiv. apply g_dist. rewrite gf. done.
- intros. rewrite g_pcore. split.
+ intros ->. done.
+ intros (? & -> & ->%g_dist)%fmap_Some_dist. done.
- intros ??. simpl. rewrite g_op gf //.
Qed.
(** * Constructing a camera through an isomorphism *)
Lemma iso_cmra_mixin {SI : sidx} {A : cmra} {B : ofe}
`{!PCore B, !Op B, !Valid B, !ValidN B}
(f : A B) (g : B A)
(* [g] is proper/non-expansive and injective w.r.t. OFE equality *)
(g_dist : n y1 y2, y1 {n} y2 g y1 {n} g y2)
(* [g] is surjective (and [f] its inverse) *)
(gf : x : A, g (f x) x)
(* [g] commutes with [pcore], [op], [valid], and [validN] *)
(g_pcore : y : B, pcore (g y) g <$> pcore y)
(g_op : y1 y2, g (y1 y2) g y1 g y2)
(g_valid : y, (g y) y)
(g_validN : n y, {n} (g y) {n} y) :
CmraMixin B.
Proof.
apply (iso_cmra_mixin_restrict_validity f g); auto.
- by intros n y ?%g_validN.
- intros n y1 y2 Hy%g_dist Hy1. by rewrite -g_validN -Hy g_validN.
- intros y. rewrite -g_valid cmra_valid_validN. naive_solver.
- intros n m y. rewrite -!g_validN. apply cmra_validN_le.
- intros n y1 y2. rewrite -!g_validN g_op. apply cmra_validN_op_l.
Qed.
From stdpp Require Import gmap gmultiset.
From iris.algebra Require Export big_op cmra.
From iris.prelude Require Import options.
(** Option *)
Lemma big_opL_None {SI : sidx} {M : cmra} {A} (f : nat A option M) l :
([^op list] kx l, f k x) = None k x, l !! k = Some x f k x = None.
Proof.
revert f. induction l as [|x l IH]=> f //=. rewrite op_None IH. split.
- intros [??] [|k] y ?; naive_solver.
- intros Hl. split; [by apply (Hl 0)|]. intros k. apply (Hl (S k)).
Qed.
Lemma big_opM_None {SI : sidx}
{M : cmra} `{Countable K} {A} (f : K A option M) m :
([^op map] kx m, f k x) = None k x, m !! k = Some x f k x = None.
Proof.
induction m as [|i x m ? IH] using map_ind=> /=.
{ by rewrite big_opM_empty. }
rewrite -None_equiv_eq big_opM_insert // None_equiv_eq op_None IH. split.
{ intros [??] k y. rewrite lookup_insert_Some; naive_solver. }
intros Hm; split.
- apply (Hm i). by simplify_map_eq.
- intros k y ?. apply (Hm k). by simplify_map_eq.
Qed.
Lemma big_opS_None {SI : sidx} {M : cmra} `{Countable A} (f : A option M) X :
([^op set] x X, f x) = None x, x X f x = None.
Proof.
induction X as [|x X ? IH] using set_ind_L.
{ by rewrite big_opS_empty. }
rewrite -None_equiv_eq big_opS_insert // None_equiv_eq op_None IH. set_solver.
Qed.
Lemma big_opMS_None {SI : sidx} {M : cmra} `{Countable A} (f : A option M) X :
([^op mset] x X, f x) = None x, x X f x = None.
Proof.
induction X as [|x X IH] using gmultiset_ind.
{ rewrite big_opMS_empty. multiset_solver. }
rewrite -None_equiv_eq big_opMS_disj_union big_opMS_singleton None_equiv_eq op_None IH.
multiset_solver.
Qed.
From stdpp Require Export sets coPset.
From iris.algebra Require Export cmra. From iris.algebra Require Export cmra.
From iris.algebra Require Import updates local_updates. From iris.algebra Require Import updates local_updates.
From iris.prelude Require Export collections coPset. From iris.prelude Require Import options.
(** This is pretty much the same as algebra/gset, but I was not able to (** This is pretty much the same as algebra/gset, but I was not able to
generalize the construction without breaking canonical structures. *) generalize the construction without breaking canonical structures. *)
(* The union CMRA *) (* The union CMRA *)
Section coPset. Section coPset.
Context {SI : sidx}.
Implicit Types X Y : coPset. Implicit Types X Y : coPset.
Canonical Structure coPsetC := leibnizC coPset. Canonical Structure coPsetO := discreteO coPset.
Instance coPset_valid : Valid coPset := λ _, True. Local Instance coPset_valid_instance : Valid coPset := λ _, True.
Instance coPset_op : Op coPset := union. Local Instance coPset_unit_instance : Unit coPset := ( : coPset).
Instance coPset_pcore : PCore coPset := Some. Local Instance coPset_op_instance : Op coPset := union.
Local Instance coPset_pcore_instance : PCore coPset := Some.
Lemma coPset_op_union X Y : X Y = X Y. Lemma coPset_op X Y : X Y = X Y.
Proof. done. Qed. Proof. done. Qed.
Lemma coPset_core_self X : core X = X. Lemma coPset_core X : core X = X.
Proof. done. Qed. Proof. done. Qed.
Lemma coPset_included X Y : X Y X Y. Lemma coPset_included X Y : X Y X Y.
Proof. Proof.
split. split.
- intros [Z ->]. rewrite coPset_op_union. set_solver. - intros [Z ->]. rewrite coPset_op. set_solver.
- intros (Z&->&?)%subseteq_disjoint_union_L. by exists Z. - intros (Z&->&?)%subseteq_disjoint_union_L. by exists Z.
Qed. Qed.
...@@ -31,19 +34,21 @@ Section coPset. ...@@ -31,19 +34,21 @@ Section coPset.
- solve_proper. - solve_proper.
- solve_proper. - solve_proper.
- solve_proper. - solve_proper.
- intros X1 X2 X3. by rewrite !coPset_op_union assoc_L. - intros X1 X2 X3. by rewrite !coPset_op assoc_L.
- intros X1 X2. by rewrite !coPset_op_union comm_L. - intros X1 X2. by rewrite !coPset_op comm_L.
- intros X. by rewrite coPset_core_self idemp_L. - intros X. by rewrite coPset_core idemp_L.
Qed. Qed.
Canonical Structure coPsetR := discreteR coPset coPset_ra_mixin. Canonical Structure coPsetR := discreteR coPset coPset_ra_mixin.
Lemma coPset_ucmra_mixin : UCMRAMixin coPset. Global Instance coPset_cmra_discrete : CmraDiscrete coPsetR.
Proof. split. done. intros X. by rewrite coPset_op_union left_id_L. done. Qed. Proof. apply discrete_cmra_discrete. Qed.
Canonical Structure coPsetUR :=
discreteUR coPset coPset_ra_mixin coPset_ucmra_mixin.
Lemma coPset_opM X mY : X ? mY = X from_option id mY. Lemma coPset_ucmra_mixin : UcmraMixin coPset.
Proof. destruct mY; by rewrite /= ?right_id_L. Qed. Proof. split; [done | | done]. intros X. by rewrite coPset_op left_id_L. Qed.
Canonical Structure coPsetUR := Ucmra coPset coPset_ucmra_mixin.
Lemma coPset_opM X mY : X ? mY = X default mY.
Proof using SI. destruct mY; by rewrite /= ?right_id_L. Qed.
Lemma coPset_update X Y : X ~~> Y. Lemma coPset_update X Y : X ~~> Y.
Proof. done. Qed. Proof. done. Qed.
...@@ -52,28 +57,29 @@ Section coPset. ...@@ -52,28 +57,29 @@ Section coPset.
Proof. Proof.
intros (Z&->&?)%subseteq_disjoint_union_L. intros (Z&->&?)%subseteq_disjoint_union_L.
rewrite local_update_unital_discrete=> Z' _ /leibniz_equiv_iff->. rewrite local_update_unital_discrete=> Z' _ /leibniz_equiv_iff->.
split. done. rewrite coPset_op_union. set_solver. split; first done. rewrite coPset_op. set_solver.
Qed. Qed.
End coPset. End coPset.
(* The disjoiny union CMRA *) (* The disjoint union CMRA *)
Inductive coPset_disj := Inductive coPset_disj :=
| CoPset : coPset coPset_disj | CoPset : coPset coPset_disj
| CoPsetBot : coPset_disj. | CoPsetInvalid : coPset_disj.
Section coPset_disj. Section coPset_disj.
Arguments op _ _ !_ !_ /. Context {SI : sidx}.
Canonical Structure coPset_disjC := leibnizC coPset_disj. Local Arguments op _ _ !_ !_ /.
Canonical Structure coPset_disjO := leibnizO coPset_disj.
Instance coPset_disj_valid : Valid coPset_disj := λ X,
match X with CoPset _ => True | CoPsetBot => False end. Local Instance coPset_disj_valid_instance : Valid coPset_disj := λ X,
Instance coPset_disj_empty : Empty coPset_disj := CoPset ∅. match X with CoPset _ => True | CoPsetInvalid => False end.
Instance coPset_disj_op : Op coPset_disj := λ X Y, Local Instance coPset_disj_unit_instance : Unit coPset_disj := CoPset ∅.
Local Instance coPset_disj_op_instance : Op coPset_disj := λ X Y,
match X, Y with match X, Y with
| CoPset X, CoPset Y => if decide (X Y) then CoPset (X Y) else CoPsetBot | CoPset X, CoPset Y => if decide (X ## Y) then CoPset (X Y) else CoPsetInvalid
| _, _ => CoPsetBot | _, _ => CoPsetInvalid
end. end.
Instance coPset_disj_pcore : PCore coPset_disj := λ _, Some . Local Instance coPset_disj_pcore_instance : PCore coPset_disj := λ _, Some ε.
Ltac coPset_disj_solve := Ltac coPset_disj_solve :=
repeat (simpl || case_decide); repeat (simpl || case_decide);
...@@ -87,11 +93,11 @@ Section coPset_disj. ...@@ -87,11 +93,11 @@ Section coPset_disj.
exists (CoPset Z). coPset_disj_solve. exists (CoPset Z). coPset_disj_solve.
Qed. Qed.
Lemma coPset_disj_valid_inv_l X Y : Lemma coPset_disj_valid_inv_l X Y :
(CoPset X Y) Y', Y = CoPset Y' X Y'. (CoPset X Y) Y', Y = CoPset Y' X ## Y'.
Proof. destruct Y; repeat (simpl || case_decide); by eauto. Qed. Proof. destruct Y; repeat (simpl || case_decide); by eauto. Qed.
Lemma coPset_disj_union X Y : X Y CoPset X CoPset Y = CoPset (X Y). Lemma coPset_disj_union X Y : X ## Y CoPset X CoPset Y = CoPset (X Y).
Proof. intros. by rewrite /= decide_True. Qed. Proof. intros. by rewrite /= decide_True. Qed.
Lemma coPset_disj_valid_op X Y : (CoPset X CoPset Y) X Y. Lemma coPset_disj_valid_op X Y : (CoPset X CoPset Y) X ## Y.
Proof. simpl. case_decide; by split. Qed. Proof. simpl. case_decide; by split. Qed.
Lemma coPset_disj_ra_mixin : RAMixin coPset_disj. Lemma coPset_disj_ra_mixin : RAMixin coPset_disj.
...@@ -108,8 +114,10 @@ Section coPset_disj. ...@@ -108,8 +114,10 @@ Section coPset_disj.
Qed. Qed.
Canonical Structure coPset_disjR := discreteR coPset_disj coPset_disj_ra_mixin. Canonical Structure coPset_disjR := discreteR coPset_disj coPset_disj_ra_mixin.
Lemma coPset_disj_ucmra_mixin : UCMRAMixin coPset_disj. Global Instance coPset_disj_cmra_discrete : CmraDiscrete coPset_disjR.
Proof. apply discrete_cmra_discrete. Qed.
Lemma coPset_disj_ucmra_mixin : UcmraMixin coPset_disj.
Proof. split; try apply _ || done. intros [X|]; coPset_disj_solve. Qed. Proof. split; try apply _ || done. intros [X|]; coPset_disj_solve. Qed.
Canonical Structure coPset_disjUR := Canonical Structure coPset_disjUR := Ucmra coPset_disj coPset_disj_ucmra_mixin.
discreteUR coPset_disj coPset_disj_ra_mixin coPset_disj_ucmra_mixin.
End coPset_disj. End coPset_disj.
From iris.algebra Require Export cofe. From iris.algebra Require Export ofe.
From iris.algebra Require Export stepindex_finite.
Record solution (F : cFunctor) := Solution { From iris.prelude Require Import options.
solution_car :> cofeT;
solution_unfold : solution_car -n> F solution_car; (** * Solver for recursive domain equations over Cofes for FINITE step-indices *)
solution_fold : F solution_car -n> solution_car; (** This file implements a solver for recursive equations of the form [F X ≃ X],
solution_fold_unfold X : solution_fold (solution_unfold X) X; where [F] is a locally contractive functor of Cofes. As such, it is an
solution_unfold_fold X : solution_unfold (solution_fold X) X implementation of America and Rutten's theorem. More details can be found in the
Iris Reference.
This implementation only works for the [nat] index type. Importing this file
will globally fix the index type to [nat]. *)
(* Note that [Inhabited] is not derivable. Take [F X := ▶ X], then a possible
solution is [Empty_set]. *)
Record solution (F : oFunctor) := Solution {
solution_car :> ofe;
solution_inhabited : Inhabited solution_car;
solution_cofe : Cofe solution_car;
solution_iso :> ofe_iso (oFunctor_apply F solution_car) solution_car;
}. }.
Arguments solution_unfold {_} _. Global Existing Instances solution_inhabited solution_cofe.
Arguments solution_fold {_} _.
Module solver. Section solver. Module solver. Section solver.
Context (F : cFunctor) `{Fcontr : cFunctorContractive F} Context (F : oFunctor) `{Fcontr : !oFunctorContractive F}.
`{Finhab : Inhabited (F unitC)}. Context `{Fcofe : (T : ofe) `{!Cofe T}, Cofe (oFunctor_apply F T)}.
Notation map := (cFunctor_map F). Context `{Finh : Inhabited (oFunctor_apply F unitO)}.
Notation map := (oFunctor_map F).
Fixpoint A' (k : nat) : { C : ofe & Cofe C } :=
match k with
| 0 => existT (P:=Cofe) unitO _
| S k => existT (P:=Cofe) (@oFunctor_apply _ F (projT1 (A' k)) (projT2 (A' k))) _
end.
Notation A k := (projT1 (A' k)).
Local Instance A_cofe k : Cofe (A k) := projT2 (A' k).
Fixpoint A (k : nat) : cofeT :=
match k with 0 => unitC | S k => F (A k) end.
Fixpoint f (k : nat) : A k -n> A (S k) := Fixpoint f (k : nat) : A k -n> A (S k) :=
match k with 0 => CofeMor (λ _, inhabitant) | S k => map (g k,f k) end match k with 0 => OfeMor (λ _, inhabitant) | S k => map (g k,f k) end
with g (k : nat) : A (S k) -n> A k := with g (k : nat) : A (S k) -n> A k :=
match k with 0 => CofeMor (λ _, ()) | S k => map (f k,g k) end. match k with 0 => OfeMor (λ _, ()) | S k => map (f k,g k) end.
Definition f_S k (x : A (S k)) : f (S k) x = map (g k,f k) x := eq_refl. Definition f_S k (x : A (S k)) : f (S k) x = map (g k,f k) x := eq_refl.
Definition g_S k (x : A (S (S k))) : g (S k) x = map (f k,g k) x := eq_refl. Definition g_S k (x : A (S (S k))) : g (S k) x = map (f k,g k) x := eq_refl.
Arguments A : simpl never. Global Arguments f : simpl never.
Arguments f : simpl never. Global Arguments g : simpl never.
Arguments g : simpl never.
Lemma gf {k} (x : A k) : g k (f k x) x. Lemma gf {k} (x : A k) : g k (f k x) x.
Proof. Proof using Fcontr.
induction k as [|k IH]; simpl in *; [by destruct x|]. induction k as [|k IH]; simpl in *; [by destruct x|].
rewrite -cFunctor_compose -{2}[x]cFunctor_id. by apply (contractive_proper map). rewrite -oFunctor_map_compose -{2}[x]oFunctor_map_id.
by apply (contractive_proper map).
Qed. Qed.
Lemma fg {k} (x : A (S (S k))) : f (S k) (g (S k) x) {k} x. Lemma fg {k} (x : A (S (S k))) : f (S k) (g (S k) x) {k} x.
Proof. Proof using Fcontr.
induction k as [|k IH]; simpl. induction k as [|k IH]; simpl.
- rewrite f_S g_S -{2}[x]cFunctor_id -cFunctor_compose. - rewrite f_S g_S -{2}[x]oFunctor_map_id -oFunctor_map_compose.
apply (contractive_0 map). apply (contractive_0 map).
- rewrite f_S g_S -{2}[x]cFunctor_id -cFunctor_compose. - rewrite f_S g_S -{2}[x]oFunctor_map_id -oFunctor_map_compose.
by apply (contractive_S map). by apply (contractive_S map).
Qed. Qed.
...@@ -45,53 +63,57 @@ Record tower := { ...@@ -45,53 +63,57 @@ Record tower := {
tower_car k :> A k; tower_car k :> A k;
g_tower k : g k (tower_car (S k)) tower_car k g_tower k : g k (tower_car (S k)) tower_car k
}. }.
Instance tower_equiv : Equiv tower := λ X Y, k, X k Y k. Global Instance tower_equiv : Equiv tower := λ X Y, k, X k Y k.
Instance tower_dist : Dist tower := λ n X Y, k, X k {n} Y k. Global Instance tower_dist : Dist tower := λ n X Y, k, X k {n} Y k.
Program Definition tower_chain (c : chain tower) (k : nat) : chain (A k) := Definition tower_ofe_mixin : OfeMixin tower.
{| chain_car i := c i k |}.
Next Obligation. intros c k n i ?; apply (chain_cauchy c n); lia. Qed.
Program Instance tower_compl : Compl tower := λ c,
{| tower_car n := compl (tower_chain c n) |}.
Next Obligation.
intros c k; apply equiv_dist=> n.
by rewrite (conv_compl n (tower_chain c k))
(conv_compl n (tower_chain c (S k))) /= (g_tower (c _) k).
Qed.
Definition tower_cofe_mixin : CofeMixin tower.
Proof. Proof.
split. apply ofe_mixin_finite.
- intros X Y; split; [by intros HXY n k; apply equiv_dist|]. - intros X Y; split; [by intros HXY n k; apply equiv_dist|].
intros HXY k; apply equiv_dist; intros n; apply HXY. intros HXY k; apply equiv_dist; intros n; apply HXY.
- intros k; split. - intros k; split.
+ by intros X n. + by intros X n.
+ by intros X Y ? n. + by intros X Y ? n.
+ by intros X Y Z ?? n; trans (Y n). + by intros X Y Z ?? n; trans (Y n).
- intros k X Y HXY n; apply dist_S. - intros k X Y HXY n. specialize (HXY (S n)).
by rewrite -(g_tower X) (HXY (S n)) g_tower. apply (dist_le _ k) in HXY; [|apply SIdx.le_succ_diag_r].
- intros n c k; rewrite /= (conv_compl n (tower_chain c k)). by rewrite -(g_tower X) HXY g_tower.
apply (chain_cauchy c); lia. Qed.
Definition T : ofe := Ofe tower tower_ofe_mixin.
Program Definition tower_chain (c : chain T) (k : nat) : chain (A k) :=
{| chain_car i := c i k |}.
Next Obligation. intros c k n i ?; by apply (chain_cauchy c n). Qed.
Program Definition tower_compl : Compl T := λ c,
{| tower_car n := compl (tower_chain c n) |}.
Next Obligation.
intros c k; apply equiv_dist=> n.
by rewrite (conv_compl n (tower_chain c k))
(conv_compl n (tower_chain c (S k))) /= (g_tower (c _) k).
Qed.
Global Program Instance tower_cofe : Cofe T := cofe_finite tower_compl _.
Next Obligation.
intros n c k; rewrite /= (conv_compl n (tower_chain c k)). done.
Qed. Qed.
Definition T : cofeT := CofeT tower tower_cofe_mixin.
Fixpoint ff {k} (i : nat) : A k -n> A (i + k) := Fixpoint ff {k} (i : nat) : A k -n> A (i + k) :=
match i with 0 => cid | S i => f (i + k) ff i end. match i with 0 => cid | S i => f (i + k) ff i end.
Fixpoint gg {k} (i : nat) : A (i + k) -n> A k := Fixpoint gg {k} (i : nat) : A (i + k) -n> A k :=
match i with 0 => cid | S i => gg i g (i + k) end. match i with 0 => cid | S i => gg i g (i + k) end.
Lemma ggff {k i} (x : A k) : gg i (ff i x) x. Lemma ggff {k i} (x : A k) : gg i (ff i x) x.
Proof. induction i as [|i IH]; simpl; [done|by rewrite (gf (ff i x)) IH]. Qed. Proof using Fcontr. induction i as [|i IH]; simpl; [done|by rewrite (gf (ff i x)) IH]. Qed.
Lemma f_tower k (X : tower) : f (S k) (X (S k)) {k} X (S (S k)). Lemma f_tower k (X : tower) : f (S k) (X (S k)) {k} X (S (S k)).
Proof. intros. by rewrite -(fg (X (S (S k)))) -(g_tower X). Qed. Proof using Fcontr. intros. by rewrite -(fg (X (S (S k)))) -(g_tower X). Qed.
Lemma ff_tower k i (X : tower) : ff i (X (S k)) {k} X (i + S k). Lemma ff_tower k i (X : tower) : ff i (X (S k)) {k} X (i + S k).
Proof. Proof using Fcontr.
intros; induction i as [|i IH]; simpl; [done|]. intros; induction i as [|i IH]; simpl; [done|].
by rewrite IH Nat.add_succ_r (dist_le _ _ _ _ (f_tower _ X)); last omega. by rewrite IH Nat.add_succ_r (dist_le _ _ _ _ (f_tower _ X)); last lia.
Qed. Qed.
Lemma gg_tower k i (X : tower) : gg i (X (i + k)) X k. Lemma gg_tower k i (X : tower) : gg i (X (i + k)) X k.
Proof. by induction i as [|i IH]; simpl; [done|rewrite g_tower IH]. Qed. Proof. by induction i as [|i IH]; simpl; [done|rewrite g_tower IH]. Qed.
Instance tower_car_ne n k : Proper (dist n ==> dist n) (λ X, tower_car X k). Global Instance tower_car_ne k : NonExpansive (λ X, tower_car X k).
Proof. by intros X Y HX. Qed. Proof. by intros X Y HX. Qed.
Definition project (k : nat) : T -n> A k := CofeMor (λ X : T, tower_car X k). Definition project (k : nat) : T -n> A k := OfeMor (λ X : T, tower_car X k).
Definition coerce {i j} (H : i = j) : A i -n> A j := Definition coerce {i j} (H : i = j) : A i -n> A j :=
eq_rect _ (λ i', A i -n> A i') cid _ H. eq_rect _ (λ i', A i -n> A i') cid _ H.
...@@ -109,8 +131,8 @@ Proof. by assert (k = j) by lia; subst; rewrite !coerce_id. Qed. ...@@ -109,8 +131,8 @@ Proof. by assert (k = j) by lia; subst; rewrite !coerce_id. Qed.
Lemma gg_gg {k i i1 i2 j} : (H1: k = i + j) (H2: k = i2 + (i1 + j)) (x: A k), Lemma gg_gg {k i i1 i2 j} : (H1: k = i + j) (H2: k = i2 + (i1 + j)) (x: A k),
gg i (coerce H1 x) = gg i1 (gg i2 (coerce H2 x)). gg i (coerce H1 x) = gg i1 (gg i2 (coerce H2 x)).
Proof. Proof.
intros ? -> x. assert (i = i2 + i1) as -> by lia. revert j x H1. intros Hij -> x. assert (i = i2 + i1) as -> by lia. revert j x Hij.
induction i2 as [|i2 IH]; intros j X H1; simplify_eq/=; induction i2 as [|i2 IH]; intros j X Hij; simplify_eq/=;
[by rewrite coerce_id|by rewrite g_coerce IH]. [by rewrite coerce_id|by rewrite g_coerce IH].
Qed. Qed.
Lemma ff_ff {k i i1 i2 j} : (H1: i + k = j) (H2: i1 + (i2 + k) = j) (x: A k), Lemma ff_ff {k i i1 i2 j} : (H1: i + k = j) (H2: i1 + (i2 + k) = j) (x: A k),
...@@ -128,7 +150,7 @@ Definition embed_coerce {k} (i : nat) : A k -n> A i := ...@@ -128,7 +150,7 @@ Definition embed_coerce {k} (i : nat) : A k -n> A i :=
end. end.
Lemma g_embed_coerce {k i} (x : A k) : Lemma g_embed_coerce {k i} (x : A k) :
g i (embed_coerce (S i) x) embed_coerce i x. g i (embed_coerce (S i) x) embed_coerce i x.
Proof. Proof using Fcontr.
unfold embed_coerce; destruct (le_lt_dec (S i) k), (le_lt_dec i k); simpl. unfold embed_coerce; destruct (le_lt_dec (S i) k), (le_lt_dec i k); simpl.
- symmetry; by erewrite (@gg_gg _ _ 1 (k - S i)); simpl. - symmetry; by erewrite (@gg_gg _ _ 1 (k - S i)); simpl.
- exfalso; lia. - exfalso; lia.
...@@ -142,10 +164,10 @@ Qed. ...@@ -142,10 +164,10 @@ Qed.
Program Definition embed (k : nat) (x : A k) : T := Program Definition embed (k : nat) (x : A k) : T :=
{| tower_car n := embed_coerce n x |}. {| tower_car n := embed_coerce n x |}.
Next Obligation. intros k x i. apply g_embed_coerce. Qed. Next Obligation. intros k x i. apply g_embed_coerce. Qed.
Instance: Params (@embed) 1. Global Instance: Params (@embed) 1 := {}.
Instance embed_ne k n : Proper (dist n ==> dist n) (embed k). Global Instance embed_ne k : NonExpansive (embed k).
Proof. by intros x y Hxy i; rewrite /= Hxy. Qed. Proof. by intros n x y Hxy i; rewrite /= Hxy. Qed.
Definition embed' (k : nat) : A k -n> T := CofeMor (embed k). Definition embed' (k : nat) : A k -n> T := OfeMor (embed k).
Lemma embed_f k (x : A k) : embed (S k) (f k x) embed k x. Lemma embed_f k (x : A k) : embed (S k) (f k x) embed k x.
Proof. Proof.
rewrite equiv_dist=> n i; rewrite /embed /= /embed_coerce. rewrite equiv_dist=> n i; rewrite /embed /= /embed_coerce.
...@@ -168,43 +190,45 @@ Proof. ...@@ -168,43 +190,45 @@ Proof.
- rewrite (ff_tower k (i - S k) X). by destruct (Nat.sub_add _ _ _). - rewrite (ff_tower k (i - S k) X). by destruct (Nat.sub_add _ _ _).
Qed. Qed.
Program Definition unfold_chain (X : T) : chain (F T) := Global Instance tower_inhabited : Inhabited tower := populate (embed 0 ()).
Program Definition unfold_chain (X : T) : chain (oFunctor_apply F T) :=
{| chain_car n := map (project n,embed' n) (X (S n)) |}. {| chain_car n := map (project n,embed' n) (X (S n)) |}.
Next Obligation. Next Obligation.
intros X n i Hi. simpl; intros X n i Hi.
assert ( k, i = k + n) as [k ?] by (exists (i - n); lia); subst; clear Hi. assert ( k, i = k + n) as [k ?] by (exists (i - n); lia); subst; clear Hi.
induction k as [|k IH]; simpl; first done. induction k as [|k IH]; simpl; first done.
rewrite -IH -(dist_le _ _ _ _ (f_tower (k + n) _)); last lia. rewrite -IH -(dist_le _ _ _ _ (f_tower (k + n) _)); last lia.
rewrite f_S -cFunctor_compose. rewrite f_S -oFunctor_map_compose.
by apply (contractive_ne map); split=> Y /=; rewrite ?g_tower ?embed_f. by apply (contractive_ne map); split=> Y /=; rewrite ?g_tower ?embed_f.
Qed. Qed.
Definition unfold (X : T) : F T := compl (unfold_chain X). Definition unfold (X : T) : oFunctor_apply F T := compl (unfold_chain X).
Instance unfold_ne : Proper (dist n ==> dist n) unfold. Global Instance unfold_ne : NonExpansive unfold.
Proof. Proof.
intros n X Y HXY. by rewrite /unfold (conv_compl n (unfold_chain X)) intros n X Y HXY. by rewrite /unfold (conv_compl n (unfold_chain X))
(conv_compl n (unfold_chain Y)) /= (HXY (S n)). (conv_compl n (unfold_chain Y)) /= (HXY (S n)).
Qed. Qed.
Program Definition fold (X : F T) : T := Program Definition fold (X : oFunctor_apply F T) : T :=
{| tower_car n := g n (map (embed' n,project n) X) |}. {| tower_car n := g n (map (embed' n,project n) X) |}.
Next Obligation. Next Obligation.
intros X k. apply (_ : Proper (() ==> ()) (g k)). intros X k. apply (_ : Proper (() ==> ()) (g k)).
rewrite g_S -cFunctor_compose. rewrite g_S -oFunctor_map_compose.
apply (contractive_proper map); split=> Y; [apply embed_f|apply g_tower]. apply (contractive_proper map); split=> Y; [apply embed_f|apply g_tower].
Qed. Qed.
Instance fold_ne : Proper (dist n ==> dist n) fold. Global Instance fold_ne : NonExpansive fold.
Proof. by intros n X Y HXY k; rewrite /fold /= HXY. Qed. Proof. by intros n X Y HXY k; rewrite /fold /= HXY. Qed.
Theorem result : solution F. Theorem result : solution F.
Proof. Proof using Type*.
apply (Solution F T (CofeMor unfold) (CofeMor fold)). refine (Solution F T _ _ (OfeIso (OfeMor fold) (OfeMor unfold) _ _)).
- move=> X /=. rewrite equiv_dist=> n k; rewrite /unfold /fold /=. - move=> X /=. rewrite equiv_dist=> n k; rewrite /unfold /fold /=.
rewrite -g_tower -(gg_tower _ n); apply (_ : Proper (_ ==> _) (g _)). rewrite -g_tower -(gg_tower _ n); apply (_ : Proper (_ ==> _) (g _)).
trans (map (ff n, gg n) (X (S (n + k)))). trans (map (ff n, gg n) (X (S (n + k)))).
{ rewrite /unfold (conv_compl n (unfold_chain X)). { rewrite /unfold (conv_compl n (unfold_chain X)).
rewrite -(chain_cauchy (unfold_chain X) n (S (n + k))) /=; last lia. rewrite -(chain_cauchy (unfold_chain X) n (S (n + k))) /=; last lia.
rewrite -(dist_le _ _ _ _ (f_tower (n + k) _)); last lia. rewrite -(dist_le _ _ _ _ (f_tower (n + k) _)); last lia.
rewrite f_S -!cFunctor_compose; apply (contractive_ne map); split=> Y. rewrite f_S -!oFunctor_map_compose; apply (contractive_ne map); split=> Y.
+ rewrite /embed' /= /embed_coerce. + rewrite /embed' /= /embed_coerce.
destruct (le_lt_dec _ _); simpl; [exfalso; lia|]. destruct (le_lt_dec _ _); simpl; [exfalso; lia|].
by rewrite (ff_ff _ (eq_refl (S n + (0 + k)))) /= gf. by rewrite (ff_ff _ (eq_refl (S n + (0 + k)))) /= gf.
...@@ -214,14 +238,14 @@ Proof. ...@@ -214,14 +238,14 @@ Proof.
assert ( i k (x : A (S i + k)) (H : S i + k = i + S k), assert ( i k (x : A (S i + k)) (H : S i + k = i + S k),
map (ff i, gg i) x gg i (coerce H x)) as map_ff_gg. map (ff i, gg i) x gg i (coerce H x)) as map_ff_gg.
{ intros i; induction i as [|i IH]; intros k' x H; simpl. { intros i; induction i as [|i IH]; intros k' x H; simpl.
{ by rewrite coerce_id cFunctor_id. } { by rewrite coerce_id oFunctor_map_id. }
rewrite cFunctor_compose g_coerce; apply IH. } rewrite oFunctor_map_compose g_coerce; apply IH. }
assert (H: S n + k = n + S k) by lia. assert (H: S n + k = n + S k) by lia.
rewrite (map_ff_gg _ _ _ H). rewrite (map_ff_gg _ _ _ H).
apply (_ : Proper (_ ==> _) (gg _)); by destruct H. apply (_ : Proper (_ ==> _) (gg _)); by destruct H.
- intros X; rewrite equiv_dist=> n /=. - intros X; rewrite equiv_dist=> n /=.
rewrite /unfold /= (conv_compl' n (unfold_chain (fold X))) /=. rewrite /unfold /= (conv_compl_S n (unfold_chain (fold X))) /=.
rewrite g_S -!cFunctor_compose -{2}[X]cFunctor_id. rewrite g_S -!oFunctor_map_compose -{2}[X]oFunctor_map_id.
apply (contractive_ne map); split => Y /=. apply (contractive_ne map); split => Y /=.
+ rewrite f_tower. apply dist_S. by rewrite embed_tower. + rewrite f_tower. apply dist_S. by rewrite embed_tower.
+ etrans; [apply embed_ne, equiv_dist, g_tower|apply embed_tower]. + etrans; [apply embed_ne, equiv_dist, g_tower|apply embed_tower].
......
From iris.algebra Require Export cmra. From iris.algebra Require Export cmra.
From iris.base_logic Require Import base_logic. From iris.algebra Require Import updates local_updates.
From iris.algebra Require Import local_updates. From iris.prelude Require Import options.
Local Arguments pcore _ _ !_ /. Local Arguments pcore _ _ !_ /.
Local Arguments cmra_pcore _ !_ /. Local Arguments cmra_pcore _ !_ /.
Local Arguments validN _ _ _ !_ /. Local Arguments validN _ _ _ _ !_ /.
Local Arguments valid _ _ !_ /. Local Arguments valid _ _ !_ /.
Local Arguments cmra_validN _ _ !_ /. Local Arguments cmra_validN _ _ _ !_ /.
Local Arguments cmra_valid _ !_ /. Local Arguments cmra_valid _ !_ /.
Inductive csum (A B : Type) := Inductive csum (A B : Type) :=
| Cinl : A csum A B | Cinl : A csum A B
| Cinr : B csum A B | Cinr : B csum A B
| CsumBot : csum A B. | CsumInvalid : csum A B.
Arguments Cinl {_ _} _. Global Arguments Cinl {_ _} _.
Arguments Cinr {_ _} _. Global Arguments Cinr {_ _} _.
Arguments CsumBot {_ _}. Global Arguments CsumInvalid {_ _}.
Instance maybe_Cinl {A B} : Maybe (@Cinl A B) := λ x, Global Instance: Params (@Cinl) 2 := {}.
Global Instance: Params (@Cinr) 2 := {}.
Global Instance: Params (@CsumInvalid) 2 := {}.
Global Instance maybe_Cinl {A B} : Maybe (@Cinl A B) := λ x,
match x with Cinl a => Some a | _ => None end. match x with Cinl a => Some a | _ => None end.
Instance maybe_Cinr {A B} : Maybe (@Cinr A B) := λ x, Global Instance maybe_Cinr {A B} : Maybe (@Cinr A B) := λ x,
match x with Cinr b => Some b | _ => None end. match x with Cinr b => Some b | _ => None end.
Section cofe. Section ofe.
Context {A B : cofeT}. Context {SI : sidx} {A B : ofe}.
Implicit Types a : A. Implicit Types a : A.
Implicit Types b : B. Implicit Types b : B.
(* Cofe *) (* Cofe *)
Inductive csum_equiv : Equiv (csum A B) := Inductive csum_equiv : Equiv (csum A B) :=
| Cinl_equiv a a' : a a' Cinl a Cinl a' | Cinl_equiv a a' : a a' Cinl a Cinl a'
| Cinlr_equiv b b' : b b' Cinr b Cinr b' | Cinr_equiv b b' : b b' Cinr b Cinr b'
| CsumBot_equiv : CsumBot CsumBot. | CsumInvalid_equiv : CsumInvalid CsumInvalid.
Existing Instance csum_equiv. Local Existing Instance csum_equiv.
Inductive csum_dist : Dist (csum A B) := Inductive csum_dist : Dist (csum A B) :=
| Cinl_dist n a a' : a {n} a' Cinl a {n} Cinl a' | Cinl_dist n a a' : a {n} a' Cinl a {n} Cinl a'
| Cinlr_dist n b b' : b {n} b' Cinr b {n} Cinr b' | Cinr_dist n b b' : b {n} b' Cinr b {n} Cinr b'
| CsumBot_dist n : CsumBot {n} CsumBot. | CsumInvalid_dist n : CsumInvalid {n} CsumInvalid.
Existing Instance csum_dist. Local Existing Instance csum_dist.
Global Instance Cinl_ne n : Proper (dist n ==> dist n) (@Cinl A B). Global Instance Cinl_ne : NonExpansive (@Cinl A B).
Proof. by constructor. Qed. Proof. by constructor. Qed.
Global Instance Cinl_proper : Proper (() ==> ()) (@Cinl A B). Global Instance Cinl_proper : Proper (() ==> ()) (@Cinl A B).
Proof. by constructor. Qed. Proof. by constructor. Qed.
...@@ -46,7 +51,7 @@ Global Instance Cinl_inj : Inj (≡) (≡) (@Cinl A B). ...@@ -46,7 +51,7 @@ Global Instance Cinl_inj : Inj (≡) (≡) (@Cinl A B).
Proof. by inversion_clear 1. Qed. Proof. by inversion_clear 1. Qed.
Global Instance Cinl_inj_dist n : Inj (dist n) (dist n) (@Cinl A B). Global Instance Cinl_inj_dist n : Inj (dist n) (dist n) (@Cinl A B).
Proof. by inversion_clear 1. Qed. Proof. by inversion_clear 1. Qed.
Global Instance Cinr_ne n : Proper (dist n ==> dist n) (@Cinr A B). Global Instance Cinr_ne : NonExpansive (@Cinr A B).
Proof. by constructor. Qed. Proof. by constructor. Qed.
Global Instance Cinr_proper : Proper (() ==> ()) (@Cinr A B). Global Instance Cinr_proper : Proper (() ==> ()) (@Cinr A B).
Proof. by constructor. Qed. Proof. by constructor. Qed.
...@@ -55,49 +60,92 @@ Proof. by inversion_clear 1. Qed. ...@@ -55,49 +60,92 @@ Proof. by inversion_clear 1. Qed.
Global Instance Cinr_inj_dist n : Inj (dist n) (dist n) (@Cinr A B). Global Instance Cinr_inj_dist n : Inj (dist n) (dist n) (@Cinr A B).
Proof. by inversion_clear 1. Qed. Proof. by inversion_clear 1. Qed.
Program Definition csum_chain_l (c : chain (csum A B)) (a : A) : chain A := Definition csum_ofe_mixin : OfeMixin (csum A B).
{| chain_car n := match c n return _ with Cinl a' => a' | _ => a end |}.
Next Obligation. intros c a n i ?; simpl. by destruct (chain_cauchy c n i). Qed.
Program Definition csum_chain_r (c : chain (csum A B)) (b : B) : chain B :=
{| chain_car n := match c n return _ with Cinr b' => b' | _ => b end |}.
Next Obligation. intros c b n i ?; simpl. by destruct (chain_cauchy c n i). Qed.
Instance csum_compl : Compl (csum A B) := λ c,
match c 0 with
| Cinl a => Cinl (compl (csum_chain_l c a))
| Cinr b => Cinr (compl (csum_chain_r c b))
| CsumBot => CsumBot
end.
Definition csum_cofe_mixin : CofeMixin (csum A B).
Proof. Proof.
split. split.
- intros mx my; split. - intros mx my; split.
+ by destruct 1; constructor; try apply equiv_dist. + by destruct 1; constructor; try apply equiv_dist.
+ intros Hxy; feed inversion (Hxy 0); subst; constructor; try done; + intros Hxy; oinversion (Hxy 0); subst; constructor; try done;
apply equiv_dist=> n; by feed inversion (Hxy n). apply equiv_dist=> n; by oinversion (Hxy n).
- intros n; split. - intros n; split.
+ by intros [|a|]; constructor. + by intros [|a|]; constructor.
+ by destruct 1; constructor. + by destruct 1; constructor.
+ destruct 1; inversion_clear 1; constructor; etrans; eauto. + destruct 1; inversion_clear 1; constructor; etrans; eauto.
- by inversion_clear 1; constructor; apply dist_S. - inversion_clear 1; intros Hlt; constructor; eauto using dist_le.
- intros n c; rewrite /compl /csum_compl. Qed.
feed inversion (chain_cauchy c 0 n); first auto with lia; constructor. Canonical Structure csumO : ofe := Ofe (csum A B) csum_ofe_mixin.
+ rewrite (conv_compl n (csum_chain_l c a')) /=. destruct (c n); naive_solver.
+ rewrite (conv_compl n (csum_chain_r c b')) /=. destruct (c n); naive_solver. Program Definition csum_chain_l (c : chain csumO) (a : A) : chain A :=
{| chain_car n := match c n return _ with Cinl a' => a' | _ => a end |}.
Next Obligation. intros c a n i ?; simpl. by destruct (chain_cauchy c n i). Qed.
Program Definition csum_chain_r (c : chain csumO) (b : B) : chain B :=
{| chain_car n := match c n return _ with Cinr b' => b' | _ => b end |}.
Next Obligation. intros c b n i ?; simpl. by destruct (chain_cauchy c n i). Qed.
Definition csum_compl `{!Cofe A, !Cofe B} : Compl csumO := λ c,
match c 0 with
| Cinl a => Cinl (compl (csum_chain_l c a))
| Cinr b => Cinr (compl (csum_chain_r c b))
| CsumInvalid => CsumInvalid
end.
Program Definition csum_bchain_l {α} (c : bchain csumO α) (a : A) : bchain A α :=
{| bchain_car β := match c β return _ with Cinl a' => a' | _ => a end |}.
Next Obligation.
intros α c a β γ Hle . simpl.
by destruct (bchain_cauchy _ c β γ Hle ).
Qed.
Program Definition csum_bchain_r {α} (c : bchain csumO α) (b : B) : bchain B α :=
{| bchain_car β := match c β return _ with Cinr b' => b' | _ => b end |}.
Next Obligation.
intros α c b β γ Hle . simpl.
by destruct (bchain_cauchy _ c β γ Hle ).
Qed.
Definition csum_lbcompl `{!Cofe A, !Cofe B} : LBCompl csumO := λ n Hn c,
match c _ (SIdx.limit_lt_0 _ Hn) with
| Cinl a => Cinl (lbcompl Hn (csum_bchain_l c a))
| Cinr b => Cinr (lbcompl Hn (csum_bchain_r c b))
| CsumInvalid => CsumInvalid
end.
Global Program Instance csum_cofe `{!Cofe A, !Cofe B} : Cofe csumO :=
{| compl := csum_compl; lbcompl := csum_lbcompl |}.
Next Obligation.
intros ?? n c; rewrite /compl /csum_compl.
oinversion (chain_cauchy c 0 n); first apply SIdx.le_0_l; constructor.
- rewrite (conv_compl n (csum_chain_l c _)) /=. destruct (c n); naive_solver.
- rewrite (conv_compl n (csum_chain_r c _)) /=. destruct (c n); naive_solver.
Qed.
Next Obligation.
intros ?? n Hn c m Hm. rewrite /lbcompl /csum_lbcompl.
oinversion (bchain_cauchy _ c 0 m (SIdx.limit_lt_0 _ Hn) Hm);
[apply SIdx.le_0_l|..]; f_equiv.
- rewrite (conv_lbcompl Hn (csum_bchain_l c _) Hm) /=.
destruct (c m); naive_solver.
- rewrite (conv_lbcompl Hn (csum_bchain_r c _) Hm) /=.
destruct (c m); naive_solver.
Qed. Qed.
Canonical Structure csumC : cofeT := CofeT (csum A B) csum_cofe_mixin. Next Obligation.
Global Instance csum_discrete : Discrete A Discrete B Discrete csumC. intros ?? n Hn c1 c2 m Hc. rewrite /lbcompl /csum_lbcompl.
Proof. by inversion_clear 3; constructor; apply (timeless _). Qed. destruct (Hc 0 (SIdx.limit_lt_0 _ Hn)); f_equiv.
- apply lbcompl_ne=> p Hp /=. by destruct (Hc p Hp).
- apply lbcompl_ne=> p Hp /=. by destruct (Hc p Hp).
Qed.
Global Instance csum_ofe_discrete :
OfeDiscrete A OfeDiscrete B OfeDiscrete csumO.
Proof. by inversion_clear 3; constructor; apply (discrete_0 _). Qed.
Global Instance csum_leibniz : Global Instance csum_leibniz :
LeibnizEquiv A LeibnizEquiv B LeibnizEquiv (csumC A B). LeibnizEquiv A LeibnizEquiv B LeibnizEquiv csumO.
Proof. by destruct 3; f_equal; apply leibniz_equiv. Qed. Proof. by destruct 3; f_equal; apply leibniz_equiv. Qed.
Global Instance Cinl_timeless a : Timeless a Timeless (Cinl a). Global Instance Cinl_discrete a : Discrete a Discrete (Cinl a).
Proof. by inversion_clear 2; constructor; apply (timeless _). Qed. Proof. by inversion_clear 2; constructor; apply (discrete_0 _). Qed.
Global Instance Cinr_timeless b : Timeless b Timeless (Cinr b). Global Instance Cinr_discrete b : Discrete b Discrete (Cinr b).
Proof. by inversion_clear 2; constructor; apply (timeless _). Qed. Proof. by inversion_clear 2; constructor; apply (discrete_0 _). Qed.
End cofe.
End ofe.
Arguments csumC : clear implicits. Global Arguments csumO {_} _ _.
(* Functor on COFEs *) (* Functor on COFEs *)
Definition csum_map {A A' B B'} (fA : A A') (fB : B B') Definition csum_map {A A' B B'} (fA : A A') (fB : B B')
...@@ -105,9 +153,9 @@ Definition csum_map {A A' B B'} (fA : A → A') (fB : B → B') ...@@ -105,9 +153,9 @@ Definition csum_map {A A' B B'} (fA : A → A') (fB : B → B')
match x with match x with
| Cinl a => Cinl (fA a) | Cinl a => Cinl (fA a)
| Cinr b => Cinr (fB b) | Cinr b => Cinr (fB b)
| CsumBot => CsumBot | CsumInvalid => CsumInvalid
end. end.
Instance: Params (@csum_map) 4. Global Instance: Params (@csum_map) 4 := {}.
Lemma csum_map_id {A B} (x : csum A B) : csum_map id id x = x. Lemma csum_map_id {A B} (x : csum A B) : csum_map id id x = x.
Proof. by destruct x. Qed. Proof. by destruct x. Qed.
...@@ -115,74 +163,98 @@ Lemma csum_map_compose {A A' A'' B B' B''} (f : A → A') (f' : A' → A'') ...@@ -115,74 +163,98 @@ Lemma csum_map_compose {A A' A'' B B' B''} (f : A → A') (f' : A' → A'')
(g : B B') (g' : B' B'') (x : csum A B) : (g : B B') (g' : B' B'') (x : csum A B) :
csum_map (f' f) (g' g) x = csum_map f' g' (csum_map f g x). csum_map (f' f) (g' g) x = csum_map f' g' (csum_map f g x).
Proof. by destruct x. Qed. Proof. by destruct x. Qed.
Lemma csum_map_ext {A A' B B' : cofeT} (f f' : A A') (g g' : B B') x : Lemma csum_map_ext {SI : sidx} {A A' B B' : ofe} (f f' : A A') (g g' : B B') x :
( x, f x f' x) ( x, g x g' x) csum_map f g x csum_map f' g' x. ( x, f x f' x) ( x, g x g' x) csum_map f g x csum_map f' g' x.
Proof. by destruct x; constructor. Qed. Proof. by destruct x; constructor. Qed.
Instance csum_map_cmra_ne {A A' B B' : cofeT} n : Global Instance csum_map_cmra_ne {SI : sidx} {A A' B B' : ofe} n :
Proper ((dist n ==> dist n) ==> (dist n ==> dist n) ==> dist n ==> dist n) Proper ((dist n ==> dist n) ==> (dist n ==> dist n) ==> dist n ==> dist n)
(@csum_map A A' B B'). (@csum_map A A' B B').
Proof. intros f f' Hf g g' Hg []; destruct 1; constructor; by apply Hf || apply Hg. Qed. Proof. intros f f' Hf g g' Hg []; destruct 1; constructor; by apply Hf || apply Hg. Qed.
Definition csumC_map {A A' B B'} (f : A -n> A') (g : B -n> B') : Definition csumO_map {SI : sidx} {A A' B B'} (f : A -n> A') (g : B -n> B') :
csumC A B -n> csumC A' B' := csumO A B -n> csumO A' B' :=
CofeMor (csum_map f g). OfeMor (csum_map f g).
Instance csumC_map_ne A A' B B' n : Global Instance csumO_map_ne {SI : sidx} A A' B B' :
Proper (dist n ==> dist n ==> dist n) (@csumC_map A A' B B'). NonExpansive2 (@csumO_map SI A A' B B').
Proof. by intros f f' Hf g g' Hg []; constructor. Qed. Proof. by intros n f f' Hf g g' Hg []; constructor. Qed.
Section cmra. Section cmra.
Context {A B : cmraT}. Context {SI : sidx} {A B : cmra}.
Implicit Types a : A. Implicit Types a : A.
Implicit Types b : B. Implicit Types b : B.
(* CMRA *) (* CMRA *)
Instance csum_valid : Valid (csum A B) := λ x, Local Instance csum_valid_instance : Valid (csum A B) := λ x,
match x with match x with
| Cinl a => a | Cinl a => a
| Cinr b => b | Cinr b => b
| CsumBot => False | CsumInvalid => False
end. end.
Instance csum_validN : ValidN (csum A B) := λ n x, Local Instance csum_validN_instance : ValidN (csum A B) := λ n x,
match x with match x with
| Cinl a => {n} a | Cinl a => {n} a
| Cinr b => {n} b | Cinr b => {n} b
| CsumBot => False | CsumInvalid => False
end. end.
Instance csum_pcore : PCore (csum A B) := λ x, Local Instance csum_pcore_instance : PCore (csum A B) := λ x,
match x with match x with
| Cinl a => Cinl <$> pcore a | Cinl a => Cinl <$> pcore a
| Cinr b => Cinr <$> pcore b | Cinr b => Cinr <$> pcore b
| CsumBot => Some CsumBot | CsumInvalid => Some CsumInvalid
end. end.
Instance csum_op : Op (csum A B) := λ x y, Local Instance csum_op_instance : Op (csum A B) := λ x y,
match x, y with match x, y with
| Cinl a, Cinl a' => Cinl (a a') | Cinl a, Cinl a' => Cinl (a a')
| Cinr b, Cinr b' => Cinr (b b') | Cinr b, Cinr b' => Cinr (b b')
| _, _ => CsumBot | _, _ => CsumInvalid
end. end.
Lemma Cinl_op a a' : Cinl a Cinl a' = Cinl (a a'). Lemma Cinl_op a a' : Cinl (a a') = Cinl a Cinl a'.
Proof. done. Qed. Proof. done. Qed.
Lemma Cinr_op b b' : Cinr b Cinr b' = Cinr (b b'). Lemma Cinr_op b b' : Cinr (b b') = Cinr b Cinr b'.
Proof. done. Qed.
Lemma Cinl_valid a : (Cinl a) a.
Proof. done. Qed.
Lemma Cinr_valid b : (Cinr b) b.
Proof. done. Qed. Proof. done. Qed.
Lemma csum_included x y : Lemma csum_included x y :
x y y = CsumBot ( a a', x = Cinl a y = Cinl a' a a') x y y = CsumInvalid ( a a', x = Cinl a y = Cinl a' a a')
( b b', x = Cinr b y = Cinr b' b b'). ( b b', x = Cinr b y = Cinr b' b b').
Proof. Proof.
split. split.
- intros [z Hy]; destruct x as [a|b|], z as [a'|b'|]; inversion_clear Hy; auto. - unfold included. intros [[a'|b'|] Hy]; destruct x as [a|b|];
+ right; left; eexists _, _; split_and!; eauto. eexists; eauto. inversion_clear Hy; eauto 10.
+ right; right; eexists _, _; split_and!; eauto. eexists; eauto. - intros [->|[(a&a'&->&->&c&?)|(b&b'&->&->&c&?)]].
+ destruct x; exists CsumInvalid; constructor.
+ exists (Cinl c); by constructor.
+ exists (Cinr c); by constructor.
Qed.
Lemma Cinl_included a a' : Cinl a Cinl a' a a'.
Proof. rewrite csum_included. naive_solver. Qed.
Lemma Cinr_included b b' : Cinr b Cinr b' b b'.
Proof. rewrite csum_included. naive_solver. Qed.
Lemma CsumInvalid_included x : x CsumInvalid.
Proof. exists CsumInvalid. by destruct x. Qed.
(** We register a hint for [CsumInvalid_included] below. *)
Lemma csum_includedN n x y :
x {n} y y = CsumInvalid ( a a', x = Cinl a y = Cinl a' a {n} a')
( b b', x = Cinr b y = Cinr b' b {n} b').
Proof.
split.
- unfold includedN. intros [[a'|b'|] Hy]; destruct x as [a|b|];
inversion_clear Hy; eauto 10.
- intros [->|[(a&a'&->&->&c&?)|(b&b'&->&->&c&?)]]. - intros [->|[(a&a'&->&->&c&?)|(b&b'&->&->&c&?)]].
+ destruct x; exists CsumBot; constructor. + destruct x; exists CsumInvalid; constructor.
+ exists (Cinl c); by constructor. + exists (Cinl c); by constructor.
+ exists (Cinr c); by constructor. + exists (Cinr c); by constructor.
Qed. Qed.
Lemma csum_cmra_mixin : CMRAMixin (csum A B). Lemma csum_cmra_mixin : CmraMixin (csum A B).
Proof. Proof.
split. split.
- intros n []; destruct 1; constructor; by cofe_subst. - intros [] n; destruct 1; constructor; by ofe_subst.
- intros ???? [n a a' Ha|n b b' Hb|n] [=]; subst; eauto. - intros ???? [n a a' Ha|n b b' Hb|n] [=]; subst; eauto.
+ destruct (pcore a) as [ca|] eqn:?; simplify_option_eq. + destruct (pcore a) as [ca|] eqn:?; simplify_option_eq.
destruct (cmra_pcore_ne n a a' ca) as (ca'&->&?); auto. destruct (cmra_pcore_ne n a a' ca) as (ca'&->&?); auto.
...@@ -190,9 +262,9 @@ Proof. ...@@ -190,9 +262,9 @@ Proof.
+ destruct (pcore b) as [cb|] eqn:?; simplify_option_eq. + destruct (pcore b) as [cb|] eqn:?; simplify_option_eq.
destruct (cmra_pcore_ne n b b' cb) as (cb'&->&?); auto. destruct (cmra_pcore_ne n b b' cb) as (cb'&->&?); auto.
exists (Cinr cb'); by repeat constructor. exists (Cinr cb'); by repeat constructor.
- intros ? [a|b|] [a'|b'|] H; inversion_clear H; cofe_subst; done. - intros ? [a|b|] [a'|b'|] H; inversion_clear H; ofe_subst; done.
- intros [a|b|]; rewrite /= ?cmra_valid_validN; naive_solver eauto using O. - pose 0. intros [a|b|]; rewrite /= ?cmra_valid_validN; naive_solver.
- intros n [a|b|]; simpl; auto using cmra_validN_S. - intros n m [a|b|]; simpl; eauto using cmra_validN_le.
- intros [a1|b1|] [a2|b2|] [a3|b3|]; constructor; by rewrite ?assoc. - intros [a1|b1|] [a2|b2|] [a3|b3|]; constructor; by rewrite ?assoc.
- intros [a1|b1|] [a2|b2|]; constructor; by rewrite 1?comm. - intros [a1|b1|] [a2|b2|]; constructor; by rewrite 1?comm.
- intros [a|b|] ? [=]; subst; auto. - intros [a|b|] ? [=]; subst; auto.
...@@ -202,11 +274,11 @@ Proof. ...@@ -202,11 +274,11 @@ Proof.
constructor; eauto using cmra_pcore_l. constructor; eauto using cmra_pcore_l.
- intros [a|b|] ? [=]; subst; auto. - intros [a|b|] ? [=]; subst; auto.
+ destruct (pcore a) as [ca|] eqn:?; simplify_option_eq. + destruct (pcore a) as [ca|] eqn:?; simplify_option_eq.
feed inversion (cmra_pcore_idemp a ca); repeat constructor; auto. oinversion (cmra_pcore_idemp a ca); repeat constructor; auto.
+ destruct (pcore b) as [cb|] eqn:?; simplify_option_eq. + destruct (pcore b) as [cb|] eqn:?; simplify_option_eq.
feed inversion (cmra_pcore_idemp b cb); repeat constructor; auto. oinversion (cmra_pcore_idemp b cb); repeat constructor; auto.
- intros x y ? [->|[(a&a'&->&->&?)|(b&b'&->&->&?)]]%csum_included [=]. - intros x y ? [->|[(a&a'&->&->&?)|(b&b'&->&->&?)]]%csum_included [=].
+ exists CsumBot. rewrite csum_included; eauto. + exists CsumInvalid. rewrite csum_included; eauto.
+ destruct (pcore a) as [ca|] eqn:?; simplify_option_eq. + destruct (pcore a) as [ca|] eqn:?; simplify_option_eq.
destruct (cmra_pcore_mono a a' ca) as (ca'&->&?); auto. destruct (cmra_pcore_mono a a' ca) as (ca'&->&?); auto.
exists (Cinl ca'). rewrite csum_included; eauto 10. exists (Cinl ca'). rewrite csum_included; eauto 10.
...@@ -215,58 +287,70 @@ Proof. ...@@ -215,58 +287,70 @@ Proof.
exists (Cinr cb'). rewrite csum_included; eauto 10. exists (Cinr cb'). rewrite csum_included; eauto 10.
- intros n [a1|b1|] [a2|b2|]; simpl; eauto using cmra_validN_op_l; done. - intros n [a1|b1|] [a2|b2|]; simpl; eauto using cmra_validN_op_l; done.
- intros n [a|b|] y1 y2 Hx Hx'. - intros n [a|b|] y1 y2 Hx Hx'.
+ destruct y1 as [a1|b1|], y2 as [a2|b2|]; inversion_clear Hx'. + destruct y1 as [a1|b1|], y2 as [a2|b2|]; try by exfalso; inversion Hx'.
destruct (cmra_extend n a a1 a2) as (z1&z2&?&?&?); auto. destruct (cmra_extend n a a1 a2) as (z1&z2&?&?&?); [done|apply (inj Cinl), Hx'|].
exists (Cinl z1), (Cinl z2). by repeat constructor. exists (Cinl z1), (Cinl z2). by repeat constructor.
+ destruct y1 as [a1|b1|], y2 as [a2|b2|]; inversion_clear Hx'. + destruct y1 as [a1|b1|], y2 as [a2|b2|]; try by exfalso; inversion Hx'.
destruct (cmra_extend n b b1 b2) as (z1&z2&?&?&?); auto. destruct (cmra_extend n b b1 b2) as (z1&z2&?&?&?); [done|apply (inj Cinr), Hx'|].
exists (Cinr z1), (Cinr z2). by repeat constructor. exists (Cinr z1), (Cinr z2). by repeat constructor.
+ by exists CsumBot, CsumBot; destruct y1, y2; inversion_clear Hx'. + by exists CsumInvalid, CsumInvalid; destruct y1, y2; inversion_clear Hx'.
Qed. Qed.
Canonical Structure csumR := Canonical Structure csumR := Cmra (csum A B) csum_cmra_mixin.
CMRAT (csum A B) csum_cofe_mixin csum_cmra_mixin.
Global Instance csum_cmra_discrete : Global Instance csum_cmra_discrete :
CMRADiscrete A CMRADiscrete B CMRADiscrete csumR. CmraDiscrete A CmraDiscrete B CmraDiscrete csumR.
Proof. Proof.
split; first apply _. split; first apply _.
by move=>[a|b|] HH /=; try apply cmra_discrete_valid. by move=>[a|b|] HH /=; try apply cmra_discrete_valid.
Qed. Qed.
Global Instance Cinl_persistent a : Persistent a Persistent (Cinl a). Global Instance Cinl_core_id a : CoreId a CoreId (Cinl a).
Proof. rewrite /Persistent /=. inversion_clear 1; by repeat constructor. Qed. Proof. rewrite /CoreId /=. inversion_clear 1; by repeat constructor. Qed.
Global Instance Cinr_persistent b : Persistent b Persistent (Cinr b). Global Instance Cinr_core_id b : CoreId b CoreId (Cinr b).
Proof. rewrite /Persistent /=. inversion_clear 1; by repeat constructor. Qed. Proof. rewrite /CoreId /=. inversion_clear 1; by repeat constructor. Qed.
Global Instance Cinl_exclusive a : Exclusive a Exclusive (Cinl a). Global Instance Cinl_exclusive a : Exclusive a Exclusive (Cinl a).
Proof. by move=> H[]? =>[/H||]. Qed. Proof. by move=> H[]? =>[/H||]. Qed.
Global Instance Cinr_exclusive b : Exclusive b Exclusive (Cinr b). Global Instance Cinr_exclusive b : Exclusive b Exclusive (Cinr b).
Proof. by move=> H[]? =>[|/H|]. Qed. Proof. by move=> H[]? =>[|/H|]. Qed.
Global Instance Cinl_cmra_homomorphism : CMRAHomomorphism Cinl. Global Instance Cinl_cancelable a : Cancelable a Cancelable (Cinl a).
Proof. split. apply _. done. Qed.
Global Instance Cinr_cmra_homomorphism : CMRAHomomorphism Cinr.
Proof. split. apply _. done. Qed.
(** Internalized properties *)
Lemma csum_equivI {M} (x y : csum A B) :
x y ⊣⊢ (match x, y with
| Cinl a, Cinl a' => a a'
| Cinr b, Cinr b' => b b'
| CsumBot, CsumBot => True
| _, _ => False
end : uPred M).
Proof. Proof.
uPred.unseal; do 2 split; first by destruct 1. move=> ?? [y|y|] [z|z|] ? EQ //; inversion_clear EQ.
by destruct x, y; try destruct 1; try constructor. constructor. by eapply (cancelableN a).
Qed.
Global Instance Cinr_cancelable b : Cancelable b Cancelable (Cinr b).
Proof.
move=> ?? [y|y|] [z|z|] ? EQ //; inversion_clear EQ.
constructor. by eapply (cancelableN b).
Qed.
Global Instance Cinl_id_free a : IdFree a IdFree (Cinl a).
Proof. intros ? [] ? EQ; inversion_clear EQ. by eapply id_free0_r. Qed.
Global Instance Cinr_id_free b : IdFree b IdFree (Cinr b).
Proof. intros ? [] ? EQ; inversion_clear EQ. by eapply id_free0_r. Qed.
(** Interaction with [option] *)
Lemma Some_csum_includedN x y n :
Some x {n} Some y
y = CsumInvalid
( a a', x = Cinl a y = Cinl a' Some a {n} Some a')
( b b', x = Cinr b y = Cinr b' Some b {n} Some b').
Proof.
repeat setoid_rewrite Some_includedN. rewrite csum_includedN. split.
- intros [Hxy|?]; [inversion Hxy|]; naive_solver.
- naive_solver by f_equiv.
Qed.
Lemma Some_csum_included x y :
Some x Some y
y = CsumInvalid
( a a', x = Cinl a y = Cinl a' Some a Some a')
( b b', x = Cinr b y = Cinr b' Some b Some b').
Proof.
repeat setoid_rewrite Some_included. rewrite csum_included. split.
- intros [Hxy|?]; [inversion Hxy|]; naive_solver.
- naive_solver by f_equiv.
Qed. Qed.
Lemma csum_validI {M} (x : csum A B) :
x ⊣⊢ (match x with
| Cinl a => a
| Cinr b => b
| CsumBot => False
end : uPred M).
Proof. uPred.unseal. by destruct x. Qed.
(** Updates *) (** Updates *)
Lemma csum_update_l (a1 a2 : A) : a1 ~~> a2 Cinl a1 ~~> Cinl a2. Lemma csum_update_l (a1 a2 : A) : a1 ~~> a2 Cinl a1 ~~> Cinl a2.
...@@ -308,7 +392,7 @@ Proof. ...@@ -308,7 +392,7 @@ Proof.
intros Hup n mf ? Ha1; simpl in *. intros Hup n mf ? Ha1; simpl in *.
destruct (Hup n (mf ≫= maybe Cinl)); auto. destruct (Hup n (mf ≫= maybe Cinl)); auto.
{ by destruct mf as [[]|]; inversion_clear Ha1. } { by destruct mf as [[]|]; inversion_clear Ha1. }
split. done. by destruct mf as [[]|]; inversion_clear Ha1; constructor. split; first done. by destruct mf as [[]|]; inversion_clear Ha1; constructor.
Qed. Qed.
Lemma csum_local_update_r (b1 b2 b1' b2' : B) : Lemma csum_local_update_r (b1 b2 b1' b2' : B) :
(b1,b2) ~l~> (b1',b2') (Cinr b1,Cinr b2) ~l~> (Cinr b1',Cinr b2'). (b1,b2) ~l~> (b1',b2') (Cinr b1,Cinr b2) ~l~> (Cinr b1',Cinr b2').
...@@ -316,42 +400,51 @@ Proof. ...@@ -316,42 +400,51 @@ Proof.
intros Hup n mf ? Ha1; simpl in *. intros Hup n mf ? Ha1; simpl in *.
destruct (Hup n (mf ≫= maybe Cinr)); auto. destruct (Hup n (mf ≫= maybe Cinr)); auto.
{ by destruct mf as [[]|]; inversion_clear Ha1. } { by destruct mf as [[]|]; inversion_clear Ha1. }
split. done. by destruct mf as [[]|]; inversion_clear Ha1; constructor. split; first done. by destruct mf as [[]|]; inversion_clear Ha1; constructor.
Qed. Qed.
End cmra. End cmra.
Arguments csumR : clear implicits. (* We use a [Hint Extern] with [apply:], instead of [Hint Immediate], to invoke
the new unification algorithm. The old unification algorithm sometimes gets
confused by going from [ucmra]'s to [cmra]'s and back. *)
Global Hint Extern 0 (_ CsumInvalid) => apply: CsumInvalid_included : core.
Global Arguments csumR {_} _ _.
(* Functor *) (* Functor *)
Instance csum_map_cmra_monotone {A A' B B' : cmraT} (f : A A') (g : B B') : Global Instance csum_map_cmra_morphism {SI : sidx} {A A' B B' : cmra}
CMRAMonotone f CMRAMonotone g CMRAMonotone (csum_map f g). (f : A A') (g : B B') :
CmraMorphism f CmraMorphism g CmraMorphism (csum_map f g).
Proof. Proof.
split; try apply _. split; try apply _.
- intros n [a|b|]; simpl; auto using cmra_monotone_validN. - intros n [a|b|]; simpl; auto using cmra_morphism_validN.
- intros x y; rewrite !csum_included. - move=> [a|b|]=>//=; rewrite -cmra_morphism_pcore; by destruct pcore.
intros [->|[(a&a'&->&->&?)|(b&b'&->&->&?)]]; simpl; - intros [xa|ya|] [xb|yb|]=>//=; by rewrite cmra_morphism_op.
eauto 10 using cmra_monotone.
Qed. Qed.
Program Definition csumRF (Fa Fb : rFunctor) : rFunctor := {| Program Definition csumRF {SI : sidx} (Fa Fb : rFunctor) : rFunctor := {|
rFunctor_car A B := csumR (rFunctor_car Fa A B) (rFunctor_car Fb A B); rFunctor_car A _ B _ := csumR (rFunctor_car Fa A B) (rFunctor_car Fb A B);
rFunctor_map A1 A2 B1 B2 fg := csumC_map (rFunctor_map Fa fg) (rFunctor_map Fb fg) rFunctor_map A1 _ A2 _ B1 _ B2 _ fg :=
csumO_map (rFunctor_map Fa fg) (rFunctor_map Fb fg)
|}. |}.
Next Obligation. Next Obligation.
by intros Fa Fb A1 A2 B1 B2 n f g Hfg; apply csumC_map_ne; try apply rFunctor_ne. intros ? Fa Fb A1 ? A2 ? B1 ? B2 ? n f g Hfg;
by apply csumO_map_ne; try apply rFunctor_map_ne.
Qed. Qed.
Next Obligation. Next Obligation.
intros Fa Fb A B x. rewrite /= -{2}(csum_map_id x). intros ? Fa Fb A ? B ? x. rewrite /= -{2}(csum_map_id x).
apply csum_map_ext=>y; apply rFunctor_id. apply csum_map_ext=>y; apply rFunctor_map_id.
Qed. Qed.
Next Obligation. Next Obligation.
intros Fa Fb A1 A2 A3 B1 B2 B3 f g f' g' x. rewrite /= -csum_map_compose. intros ? Fa Fb A1 ? A2 ? A3 ? B1 ? B2 ? B3 ? f g f' g' x.
apply csum_map_ext=>y; apply rFunctor_compose. rewrite /= -csum_map_compose.
apply csum_map_ext=>y; apply rFunctor_map_compose.
Qed. Qed.
Instance csumRF_contractive Fa Fb : Global Instance csumRF_contractive {SI : sidx} Fa Fb :
rFunctorContractive Fa rFunctorContractive Fb rFunctorContractive Fa rFunctorContractive Fb
rFunctorContractive (csumRF Fa Fb). rFunctorContractive (csumRF Fa Fb).
Proof. Proof.
by intros ?? A1 A2 B1 B2 n f g Hfg; apply csumC_map_ne; try apply rFunctor_contractive. intros ??? A1 ? A2 ? B1 ? B2 ? n f g Hfg.
by apply csumO_map_ne; try apply rFunctor_map_contractive.
Qed. Qed.
(** Camera of discardable fractions.
This is a generalisation of the fractional camera where elements can
represent both ownership of a fraction (as in the fractional camera) and the
knowledge that a fraction has been discarded.
Ownership of a fraction is denoted [DfracOwn q] and behaves identically to
[q] of the fractional camera.
Knowledge that a fraction has been discarded is denoted [DfracDiscarded].
This elements is its own core, making ownership persistent.
One can make a frame preserving update from _owning_ a fraction to _knowing_
that the fraction has been discarded.
Crucially, ownership over 1 is an exclusive element just as it is in the
fractional camera. Hence owning 1 implies that no fraction has been
discarded. Conversely, knowing that a fraction has been discarded implies
that no one can own 1. And, since discarding is an irreversible operation,
it also implies that no one can own 1 in the future *)
From stdpp Require Import countable.
From iris.algebra Require Export cmra.
From iris.algebra Require Import proofmode_classes updates frac.
From iris.prelude Require Import options.
(** An element of dfrac denotes ownership of a fraction, knowledge that a
fraction has been discarded, or both. Note that [DfracBoth] can be written
as [DfracOwn q ⋅ DfracDiscarded]. This should be used instead
of [DfracBoth] which is for internal use only. *)
Inductive dfrac :=
| DfracOwn : Qp dfrac
| DfracDiscarded : dfrac
| DfracBoth : Qp dfrac.
(* This notation is intended to be used as a component in other notations that
include discardable fractions. The notation provides shorthands for the
constructors and the commonly used full fraction. For an example
demonstrating how this can be used see the notation in [gen_heap.v]. *)
Declare Custom Entry dfrac.
Notation "{ dq }" := (dq) (in custom dfrac at level 1, dq constr).
Notation "□" := DfracDiscarded (in custom dfrac).
Notation "{# q }" := (DfracOwn q) (in custom dfrac at level 1, q constr).
Notation "" := (DfracOwn 1) (in custom dfrac).
Section dfrac.
Context {SI : sidx}.
Canonical Structure dfracO := leibnizO dfrac.
Implicit Types p q : Qp.
Implicit Types dp dq : dfrac.
Global Instance dfrac_inhabited : Inhabited dfrac := populate DfracDiscarded.
Global Instance dfrac_eq_dec : EqDecision dfrac.
Proof. solve_decision. Defined.
Global Instance dfrac_countable : Countable dfrac.
Proof.
set (enc dq := match dq with
| DfracOwn q => inl q
| DfracDiscarded => inr (inl ())
| DfracBoth q => inr (inr q)
end).
set (dec y := Some match y with
| inl q => DfracOwn q
| inr (inl ()) => DfracDiscarded
| inr (inr q) => DfracBoth q
end).
refine (inj_countable enc dec _). by intros [].
Qed.
Global Instance DfracOwn_inj : Inj (=) (=) DfracOwn.
Proof. by injection 1. Qed.
Global Instance DfracBoth_inj : Inj (=) (=) DfracBoth.
Proof. by injection 1. Qed.
(** An element is valid as long as the sum of its content is less than one. *)
Local Instance dfrac_valid_instance : Valid dfrac := λ dq,
match dq with
| DfracOwn q => q 1
| DfracDiscarded => True
| DfracBoth q => q < 1
end%Qp.
(** As in the fractional camera the core is undefined for elements denoting
ownership of a fraction. For elements denoting the knowledge that a fraction has
been discarded the core is the identity function. *)
Local Instance dfrac_pcore_instance : PCore dfrac := λ dq,
match dq with
| DfracOwn q => None
| DfracDiscarded => Some DfracDiscarded
| DfracBoth q => Some DfracDiscarded
end.
(** When elements are combined, ownership is added together and knowledge of
discarded fractions is preserved. *)
Local Instance dfrac_op_instance : Op dfrac := λ dq dp,
match dq, dp with
| DfracOwn q, DfracOwn q' => DfracOwn (q + q')
| DfracOwn q, DfracDiscarded => DfracBoth q
| DfracOwn q, DfracBoth q' => DfracBoth (q + q')
| DfracDiscarded, DfracOwn q' => DfracBoth q'
| DfracDiscarded, DfracDiscarded => DfracDiscarded
| DfracDiscarded, DfracBoth q' => DfracBoth q'
| DfracBoth q, DfracOwn q' => DfracBoth (q + q')
| DfracBoth q, DfracDiscarded => DfracBoth q
| DfracBoth q, DfracBoth q' => DfracBoth (q + q')
end.
Lemma dfrac_valid dq :
dq match dq with
| DfracOwn q => q 1
| DfracDiscarded => True
| DfracBoth q => q < 1
end%Qp.
Proof. done. Qed.
Lemma dfrac_op_own q p : DfracOwn p DfracOwn q = DfracOwn (p + q).
Proof. done. Qed.
Lemma dfrac_op_discarded :
DfracDiscarded DfracDiscarded = DfracDiscarded.
Proof. done. Qed.
Lemma dfrac_own_included q p : DfracOwn q DfracOwn p (q < p)%Qp.
Proof.
rewrite Qp.lt_sum. split.
- rewrite /included /op /dfrac_op_instance. intros [[o| |?] [= ->]]. by exists o.
- intros [o ->]. exists (DfracOwn o). by rewrite dfrac_op_own.
Qed.
(* [dfrac] does not have a unit so reflexivity is not for granted! *)
Lemma dfrac_discarded_included :
DfracDiscarded DfracDiscarded.
Proof. exists DfracDiscarded. done. Qed.
Definition dfrac_ra_mixin : RAMixin dfrac.
Proof.
split; try apply _.
- intros [?| |?] ? dq <-; intros [= <-]; eexists _; done.
- intros [?| |?] [?| |?] [?| |?];
rewrite /op /dfrac_op_instance 1?assoc_L 1?assoc_L; done.
- intros [?| |?] [?| |?];
rewrite /op /dfrac_op_instance 1?(comm_L Qp.add); done.
- intros [?| |?] dq; rewrite /pcore /dfrac_pcore_instance; intros [= <-];
rewrite /op /dfrac_op_instance; done.
- intros [?| |?] ? [= <-]; done.
- intros [?| |?] [?| |?] ? [[?| |?] [=]] [= <-]; eexists _; split; try done;
apply dfrac_discarded_included.
- intros [q| |q] [q'| |q']; rewrite /op /dfrac_op_instance /valid /dfrac_valid_instance //.
+ intros. trans (q + q')%Qp; [|done]. apply Qp.le_add_l.
+ apply Qp.lt_le_incl.
+ intros. trans (q + q')%Qp; [|by apply Qp.lt_le_incl]. apply Qp.le_add_l.
+ intros. trans (q + q')%Qp; [|done]. apply Qp.lt_add_l.
+ intros. trans (q + q')%Qp; [|done]. apply Qp.lt_add_l.
Qed.
Canonical Structure dfracR := discreteR dfrac dfrac_ra_mixin.
Global Instance dfrac_cmra_discrete : CmraDiscrete dfracR.
Proof. apply discrete_cmra_discrete. Qed.
Global Instance dfrac_full_exclusive : Exclusive (DfracOwn 1).
Proof.
intros [q| |q];
rewrite /op /cmra_op -cmra_discrete_valid_iff /valid /cmra_valid //=.
- apply Qp.not_add_le_l.
- move=> /Qp.lt_le_incl. apply Qp.not_add_le_l.
Qed.
Global Instance dfrac_cancelable q : Cancelable (DfracOwn q).
Proof.
apply: discrete_cancelable.
intros [q1| |q1][q2| |q2] _ [=]; simplify_eq/=; try done.
- by destruct (Qp.add_id_free q q2).
- by destruct (Qp.add_id_free q q1).
Qed.
Global Instance dfrac_own_id_free q : IdFree (DfracOwn q).
Proof. intros [q'| |q'] _ [=]. by apply (Qp.add_id_free q q'). Qed.
Global Instance dfrac_discarded_core_id : CoreId DfracDiscarded.
Proof. by constructor. Qed.
Lemma dfrac_valid_own p : DfracOwn p (p 1)%Qp.
Proof. done. Qed.
Lemma dfrac_valid_own_1 : DfracOwn 1.
Proof. done. Qed.
Lemma dfrac_valid_own_r dq q : (dq DfracOwn q) (q < 1)%Qp.
Proof.
destruct dq as [q'| |q']; [|done|].
- apply Qp.lt_le_trans, Qp.lt_add_r.
- intro Hlt. etrans; last apply Hlt. apply Qp.lt_add_r.
Qed.
Lemma dfrac_valid_own_l dq q : (DfracOwn q dq) (q < 1)%Qp.
Proof using SI. rewrite comm. apply dfrac_valid_own_r. Qed.
Lemma dfrac_valid_discarded : DfracDiscarded.
Proof. done. Qed.
Lemma dfrac_valid_own_discarded q :
(DfracOwn q DfracDiscarded) (q < 1)%Qp.
Proof. done. Qed.
Global Instance dfrac_is_op q q1 q2 :
IsOp q q1 q2
IsOp' (DfracOwn q) (DfracOwn q1) (DfracOwn q2).
Proof. rewrite /IsOp' /IsOp dfrac_op_own=>-> //. Qed.
(** Discarding a fraction is a frame preserving update. *)
Lemma dfrac_discard_update dq : dq ~~> DfracDiscarded.
Proof.
intros n [[q'| |q']|]; rewrite -!cmra_discrete_valid_iff //=.
- apply dfrac_valid_own_r.
- apply cmra_valid_op_r.
Qed.
Lemma dfrac_undiscard_update : DfracDiscarded ~~>: λ k, q, k = DfracOwn q.
Proof.
apply cmra_discrete_updateP. intros [[q'| |q']|].
- intros [qd Hqd]%Qp.lt_sum. exists (DfracOwn (qd/2)%Qp).
split; first by eexists. apply dfrac_valid_own.
rewrite Hqd Qp.add_comm -Qp.add_le_mono_l.
by apply Qp.div_le.
- intros _. exists (DfracOwn (1/2)); split; first by eexists.
by apply dfrac_valid_own_discarded.
- intros [qd Hqd]%Qp.lt_sum. exists (DfracOwn (qd/2)%Qp).
split; first by eexists. rewrite /= /op /valid /cmra_op /cmra_valid /=.
rewrite Hqd Qp.add_comm -Qp.add_lt_mono_l.
by apply Qp.div_lt.
- intros _. exists (DfracOwn 1); split; first by eexists.
by apply dfrac_valid_own.
Qed.
End dfrac.
From iris.algebra Require Export gmap coPset local_updates.
From iris.algebra Require Import reservation_map updates proofmode_classes.
From iris.prelude Require Import options.
(** The camera [dyn_reservation_map A] over a camera [A] extends [gmap positive
A] with a notion of "reservation tokens" for a (potentially infinite) set [E :
coPset] which represent the right to allocate a map entry at any position [k ∈
E]. Unlike [reservation_map], [dyn_reservation_map] supports dynamically
allocating these tokens, including infinite sets of them. This is useful when
syncing the keys of this map with another API that dynamically allocates names:
we can first reserve a fresh infinite set [E] of tokens here, then allocate a
new name *in [E]* with the other API (assuming it offers the usual "allocate a
fresh name in an infinite set" API), and then use our tokens to allocate the
same name here. In effect, we have performed synchronized allocation of the
same name across two maps, without the other API having to have dedicated
support for this.
The key connectives are [dyn_reservation_map_data k a] (the "points-to"
assertion of this map), which associates data [a : A] with a key [k : positive],
and [dyn_reservation_map_token E] (the reservation token), which says
that no data has been associated with the indices in the mask [E]. The important
properties of this camera are:
- The lemma [dyn_reservation_map_reserve] provides a frame-preserving update to
obtain ownership of [dyn_reservation_map_token E] for some fresh infinite [E].
- The lemma [dyn_reservation_map_alloc] provides a frame preserving update to
associate data to a key: [dyn_reservation_map_token E ~~> dyn_reservation_map_data k a]
provided [k ∈ E] and [✓ a].
In the future, it could be interesting to generalize this map to arbitrary key
types instead of hard-coding [positive]. *)
Record dyn_reservation_map (A : Type) := DynReservationMap {
dyn_reservation_map_data_proj : gmap positive A;
dyn_reservation_map_token_proj : coPset_disj
}.
Add Printing Constructor dyn_reservation_map.
Global Arguments DynReservationMap {_} _ _.
Global Arguments dyn_reservation_map_data_proj {_} _.
Global Arguments dyn_reservation_map_token_proj {_} _.
Global Instance: Params (@DynReservationMap) 1 := {}.
Global Instance: Params (@dyn_reservation_map_data_proj) 1 := {}.
Global Instance: Params (@dyn_reservation_map_token_proj) 1 := {}.
Definition dyn_reservation_map_data {SI : sidx} {A : cmra}
(k : positive) (a : A) : dyn_reservation_map A :=
DynReservationMap {[ k := a ]} ε.
Definition dyn_reservation_map_token {SI : sidx} {A : cmra}
(E : coPset) : dyn_reservation_map A :=
DynReservationMap (CoPset E).
Global Instance: Params (@dyn_reservation_map_data) 3 := {}.
(** We consruct the OFE and CMRA structure via an isomorphism with
[reservation_map]. *)
Section ofe.
Context {SI : sidx} {A : ofe}.
Implicit Types x y : dyn_reservation_map A.
Local Definition to_reservation_map x : reservation_map A :=
ReservationMap (dyn_reservation_map_data_proj x) (dyn_reservation_map_token_proj x).
Local Definition from_reservation_map (x : reservation_map A) : dyn_reservation_map A :=
DynReservationMap (reservation_map_data_proj x) (reservation_map_token_proj x).
Local Instance dyn_reservation_map_equiv : Equiv (dyn_reservation_map A) := λ x y,
dyn_reservation_map_data_proj x dyn_reservation_map_data_proj y
dyn_reservation_map_token_proj x = dyn_reservation_map_token_proj y.
Local Instance dyn_reservation_map_dist : Dist (dyn_reservation_map A) := λ n x y,
dyn_reservation_map_data_proj x {n} dyn_reservation_map_data_proj y
dyn_reservation_map_token_proj x = dyn_reservation_map_token_proj y.
Global Instance DynReservationMap_ne :
NonExpansive2 (@DynReservationMap A).
Proof. by split. Qed.
Global Instance DynReservationMap_proper :
Proper (() ==> (=) ==> ()) (@DynReservationMap A).
Proof. by split. Qed.
Global Instance dyn_reservation_map_data_proj_ne :
NonExpansive (@dyn_reservation_map_data_proj A).
Proof. by destruct 1. Qed.
Global Instance dyn_reservation_map_data_proj_proper :
Proper (() ==> ()) (@dyn_reservation_map_data_proj A).
Proof. by destruct 1. Qed.
Definition dyn_reservation_map_ofe_mixin : OfeMixin (dyn_reservation_map A).
Proof.
by apply (iso_ofe_mixin to_reservation_map).
Qed.
Canonical Structure dyn_reservation_mapO :=
Ofe (dyn_reservation_map A) dyn_reservation_map_ofe_mixin.
Global Instance DynReservationMap_discrete a b :
Discrete a Discrete b Discrete (DynReservationMap a b).
Proof. intros ?? [??] [??]; split; unfold_leibniz; by eapply discrete_0. Qed.
Global Instance dyn_reservation_map_ofe_discrete :
OfeDiscrete A OfeDiscrete dyn_reservation_mapO.
Proof. intros ? [??]; apply _. Qed.
End ofe.
Global Arguments dyn_reservation_mapO {_} _.
Section cmra.
Context {SI : sidx} {A : cmra}.
Implicit Types a b : A.
Implicit Types x y : dyn_reservation_map A.
Implicit Types k : positive.
Global Instance dyn_reservation_map_data_ne i :
NonExpansive (@dyn_reservation_map_data SI A i).
Proof. intros ? ???. rewrite /dyn_reservation_map_data. solve_proper. Qed.
Global Instance dyn_reservation_map_data_proper N :
Proper (() ==> ()) (@dyn_reservation_map_data SI A N).
Proof. solve_proper. Qed.
Global Instance dyn_reservation_map_data_discrete N a :
Discrete a Discrete (dyn_reservation_map_data N a).
Proof. intros. apply DynReservationMap_discrete; apply _. Qed.
Global Instance dyn_reservation_map_token_discrete E :
Discrete (@dyn_reservation_map_token SI A E).
Proof. intros. apply DynReservationMap_discrete; apply _. Qed.
Local Instance dyn_reservation_map_valid_instance : Valid (dyn_reservation_map A) := λ x,
match dyn_reservation_map_token_proj x with
| CoPset E =>
(dyn_reservation_map_data_proj x) set_infinite ( E)
(* dom (dyn_reservation_map_data_proj x) ⊥ E *)
( i, dyn_reservation_map_data_proj x !! i = None i E)
| CoPsetInvalid => False
end.
Global Arguments dyn_reservation_map_valid_instance !_ /.
Local Instance dyn_reservation_map_validN_instance : ValidN (dyn_reservation_map A) := λ n x,
match dyn_reservation_map_token_proj x with
| CoPset E =>
{n} (dyn_reservation_map_data_proj x) set_infinite ( E)
(* dom (dyn_reservation_map_data_proj x) ⊥ E *)
( i, dyn_reservation_map_data_proj x !! i = None i E)
| CoPsetInvalid => False
end.
Global Arguments dyn_reservation_map_validN_instance !_ /.
Local Instance dyn_reservation_map_pcore_instance : PCore (dyn_reservation_map A) := λ x,
Some (DynReservationMap (core (dyn_reservation_map_data_proj x)) ε).
Local Instance dyn_reservation_map_op_instance : Op (dyn_reservation_map A) := λ x y,
DynReservationMap (dyn_reservation_map_data_proj x dyn_reservation_map_data_proj y)
(dyn_reservation_map_token_proj x dyn_reservation_map_token_proj y).
Definition dyn_reservation_map_valid_eq :
valid = λ x, match dyn_reservation_map_token_proj x with
| CoPset E =>
(dyn_reservation_map_data_proj x) set_infinite ( E)
(* dom (dyn_reservation_map_data_proj x) ⊥ E *)
i, dyn_reservation_map_data_proj x !! i = None i E
| CoPsetInvalid => False
end := eq_refl _.
Definition dyn_reservation_map_validN_eq :
validN = λ n x, match dyn_reservation_map_token_proj x with
| CoPset E =>
{n} (dyn_reservation_map_data_proj x) set_infinite ( E)
(* dom (dyn_reservation_map_data_proj x) ⊥ E *)
i, dyn_reservation_map_data_proj x !! i = None i E
| CoPsetInvalid => False
end := eq_refl _.
Lemma dyn_reservation_map_included x y :
x y
dyn_reservation_map_data_proj x dyn_reservation_map_data_proj y
dyn_reservation_map_token_proj x dyn_reservation_map_token_proj y.
Proof.
split; [intros [[z1 z2] Hz]; split; [exists z1|exists z2]; apply Hz|].
intros [[z1 Hz1] [z2 Hz2]]; exists (DynReservationMap z1 z2); split; auto.
Qed.
Lemma dyn_reservation_map_data_proj_validN n x : {n} x {n} dyn_reservation_map_data_proj x.
Proof. by destruct x as [? [?|]]=> // -[??]. Qed.
Lemma dyn_reservation_map_token_proj_validN n x : {n} x {n} dyn_reservation_map_token_proj x.
Proof. by destruct x as [? [?|]]=> // -[??]. Qed.
Lemma dyn_reservation_map_cmra_mixin : CmraMixin (dyn_reservation_map A).
Proof.
apply (iso_cmra_mixin_restrict_validity from_reservation_map to_reservation_map); try done.
- intros n [m [E|]];
rewrite dyn_reservation_map_validN_eq reservation_map_validN_eq /=;
naive_solver.
- intros n [m1 [E1|]] [m2 [E2|]] [Hm ?]=> // -[?[??]]; split; simplify_eq/=.
+ by rewrite -Hm.
+ split; first done. intros i. by rewrite -(dist_None n) -Hm dist_None.
- pose 0.
intros [m [E|]]; rewrite dyn_reservation_map_valid_eq
dyn_reservation_map_validN_eq /= ?cmra_valid_validN; naive_solver.
- intros n m [r [E|]]; rewrite dyn_reservation_map_validN_eq /=;
naive_solver eauto using cmra_validN_le.
- intros n [m1 [E1|]] [m2 [E2|]]=> //=; rewrite dyn_reservation_map_validN_eq /=.
rewrite {1}/op /cmra_op /=. case_decide; last done.
intros [Hm [Hinf Hdisj]]; split; first by eauto using cmra_validN_op_l.
split.
+ rewrite ->difference_union_distr_r_L in Hinf.
eapply set_infinite_subseteq, Hinf. set_solver.
+ intros i. move: (Hdisj i). rewrite lookup_op.
case: (m1 !! i); case: (m2 !! i); set_solver.
Qed.
Canonical Structure dyn_reservation_mapR :=
Cmra (dyn_reservation_map A) dyn_reservation_map_cmra_mixin.
Global Instance dyn_reservation_map_cmra_discrete :
CmraDiscrete A CmraDiscrete dyn_reservation_mapR.
Proof.
split; first apply _.
intros [m [E|]]; rewrite dyn_reservation_map_validN_eq dyn_reservation_map_valid_eq //=.
by intros [?%cmra_discrete_valid ?].
Qed.
Local Instance dyn_reservation_map_empty_instance : Unit (dyn_reservation_map A) :=
DynReservationMap ε ε.
Lemma dyn_reservation_map_ucmra_mixin : UcmraMixin (dyn_reservation_map A).
Proof.
split; simpl.
- rewrite dyn_reservation_map_valid_eq /=. split; [apply ucmra_unit_valid|]. split.
+ rewrite difference_empty_L. apply top_infinite.
+ set_solver.
- split; simpl; [by rewrite left_id|by rewrite left_id_L].
- do 2 constructor; [apply (core_id_core _)|done].
Qed.
Canonical Structure dyn_reservation_mapUR :=
Ucmra (dyn_reservation_map A) dyn_reservation_map_ucmra_mixin.
Global Instance dyn_reservation_map_data_core_id k a :
CoreId a CoreId (dyn_reservation_map_data k a).
Proof. do 2 constructor; simpl; auto. apply core_id_core, _. Qed.
Lemma dyn_reservation_map_data_valid k a :
(dyn_reservation_map_data k a) a.
Proof.
rewrite dyn_reservation_map_valid_eq /= singleton_valid.
split; first naive_solver. intros Ha.
split; first done. split; last set_solver.
rewrite difference_empty_L. apply top_infinite.
Qed.
Lemma dyn_reservation_map_token_valid E :
(dyn_reservation_map_token E) set_infinite ( E).
Proof.
rewrite dyn_reservation_map_valid_eq /=. split; first naive_solver.
intros Hinf. do 2 (split; first done). by left.
Qed.
Lemma dyn_reservation_map_data_op k a b :
dyn_reservation_map_data k (a b) = dyn_reservation_map_data k a dyn_reservation_map_data k b.
Proof.
by rewrite {2}/op /dyn_reservation_map_op_instance
/dyn_reservation_map_data /= singleton_op left_id_L.
Qed.
Lemma dyn_reservation_map_data_mono k a b :
a b dyn_reservation_map_data k a dyn_reservation_map_data k b.
Proof. intros [c ->]. by rewrite dyn_reservation_map_data_op. Qed.
Global Instance dyn_reservation_map_data_is_op k a b1 b2 :
IsOp a b1 b2
IsOp' (dyn_reservation_map_data k a) (dyn_reservation_map_data k b1) (dyn_reservation_map_data k b2).
Proof. rewrite /IsOp' /IsOp=> ->. by rewrite dyn_reservation_map_data_op. Qed.
Lemma dyn_reservation_map_token_union E1 E2 :
E1 ## E2
dyn_reservation_map_token (E1 E2) = dyn_reservation_map_token E1 dyn_reservation_map_token E2.
Proof.
intros. by rewrite /op /dyn_reservation_map_op_instance
/dyn_reservation_map_token /= coPset_disj_union // left_id_L.
Qed.
Lemma dyn_reservation_map_token_difference E1 E2 :
E1 E2
dyn_reservation_map_token E2 = dyn_reservation_map_token E1 dyn_reservation_map_token (E2 E1).
Proof.
intros. rewrite -dyn_reservation_map_token_union; last set_solver.
by rewrite -union_difference_L.
Qed.
Lemma dyn_reservation_map_token_valid_op E1 E2 :
(dyn_reservation_map_token E1 dyn_reservation_map_token E2)
E1 ## E2 set_infinite ( (E1 E2)).
Proof.
split.
- rewrite dyn_reservation_map_valid_eq /= {1}/op /cmra_op /=. case_decide; last done.
naive_solver.
- intros [Hdisj Hinf]. rewrite -dyn_reservation_map_token_union //.
apply dyn_reservation_map_token_valid. done.
Qed.
Lemma dyn_reservation_map_reserve (Q : dyn_reservation_map A Prop) :
( E, set_infinite E Q (dyn_reservation_map_token E))
ε ~~>: Q.
Proof.
intros HQ. apply cmra_total_updateP=> n [mf [Ef|]];
rewrite left_id {1}dyn_reservation_map_validN_eq /=; last done.
intros [Hmap [Hinf Hdisj]].
(* Pick a fresh set disjoint from the existing tokens [Ef] and map [mf],
such that both that set [E1] and the remainder [E2] are infinite. *)
edestruct (coPset_split_infinite ( (Ef (gset_to_coPset $ dom mf)))) as
(E1 & E2 & HEunion & HEdisj & HE1inf & HE2inf).
{ rewrite -difference_difference_l_L.
by apply difference_infinite, gset_to_coPset_finite. }
exists (dyn_reservation_map_token E1).
split; first by apply HQ. clear HQ.
rewrite dyn_reservation_map_validN_eq /=.
rewrite coPset_disj_union; last set_solver.
split; first by rewrite left_id_L. split.
- eapply set_infinite_subseteq, HE2inf. set_solver.
- intros i. rewrite left_id_L. destruct (Hdisj i) as [?|Hi]; first by left.
destruct (mf !! i) as [p|] eqn:Hp; last by left.
apply elem_of_dom_2, elem_of_gset_to_coPset in Hp. right. set_solver.
Qed.
Lemma dyn_reservation_map_reserve' :
ε ~~>: λ x, E, set_infinite E x = dyn_reservation_map_token E.
Proof. eauto using dyn_reservation_map_reserve. Qed.
Lemma dyn_reservation_map_alloc E k a :
k E a dyn_reservation_map_token E ~~> dyn_reservation_map_data k a.
Proof.
intros ??. apply cmra_total_update=> n [mf [Ef|]] //.
rewrite dyn_reservation_map_validN_eq /= {1}/op {1}/cmra_op /=.
case_decide; last done.
rewrite !left_id_L. intros [Hmf [Hinf Hdisj]]; split; last split.
- destruct (Hdisj k) as [Hmfi|]; last set_solver.
intros j. rewrite lookup_op.
destruct (decide (k = j)) as [<-|].
+ rewrite Hmfi lookup_singleton right_id_L. by apply cmra_valid_validN.
+ by rewrite lookup_singleton_ne // left_id_L.
- eapply set_infinite_subseteq, Hinf. set_solver.
- intros j. destruct (decide (k = j)); first set_solver.
rewrite lookup_op lookup_singleton_ne //.
destruct (Hdisj j) as [Hmfi|?]; last set_solver. rewrite Hmfi; auto.
Qed.
Lemma dyn_reservation_map_updateP P (Q : dyn_reservation_map A Prop) k a :
a ~~>: P
( a', P a' Q (dyn_reservation_map_data k a'))
dyn_reservation_map_data k a ~~>: Q.
Proof.
intros Hup HP. apply cmra_total_updateP=> n [mf [Ef|]] //.
rewrite dyn_reservation_map_validN_eq /= left_id_L. intros [Hmf [Hinf Hdisj]].
destruct (Hup n (mf !! k)) as (a'&?&?).
{ move: (Hmf (k)).
by rewrite lookup_op lookup_singleton Some_op_opM. }
exists (dyn_reservation_map_data k a'); split; first by eauto.
rewrite /= left_id_L. split; last split.
- intros j. destruct (decide (k = j)) as [<-|].
+ by rewrite lookup_op lookup_singleton Some_op_opM.
+ rewrite lookup_op lookup_singleton_ne // left_id_L.
move: (Hmf j). rewrite lookup_op. eauto using cmra_validN_op_r.
- done.
- intros j. move: (Hdisj j).
rewrite !lookup_op !op_None !lookup_singleton_None. naive_solver.
Qed.
Lemma dyn_reservation_map_update k a b :
a ~~> b
dyn_reservation_map_data k a ~~> dyn_reservation_map_data k b.
Proof.
rewrite !cmra_update_updateP. eauto using dyn_reservation_map_updateP with subst.
Qed.
End cmra.
Global Arguments dyn_reservation_mapR {_} _.
Global Arguments dyn_reservation_mapUR {_} _.
From iris.algebra Require Export cmra. From iris.algebra Require Export cmra.
From iris.base_logic Require Import base_logic. From iris.prelude Require Import options.
Local Arguments validN _ _ _ !_ /.
Local Arguments validN _ _ _ _ !_ /.
Local Arguments valid _ _ !_ /. Local Arguments valid _ _ !_ /.
Inductive excl (A : Type) := Inductive excl (A : Type) :=
| Excl : A excl A | Excl : A excl A
| ExclBot : excl A. | ExclInvalid : excl A.
Arguments Excl {_} _. Global Arguments Excl {_} _.
Arguments ExclBot {_}. Global Arguments ExclInvalid {_}.
Global Instance: Params (@Excl) 1 := {}.
Global Instance: Params (@ExclInvalid) 1 := {}.
Notation excl' A := (option (excl A)). Notation excl' A := (option (excl A)).
Notation Excl' x := (Some (Excl x)). Notation Excl' x := (Some (Excl x)).
Notation ExclBot' := (Some ExclBot). Notation ExclInvalid' := (Some ExclInvalid).
Instance maybe_Excl {A} : Maybe (@Excl A) := λ x, Global Instance maybe_Excl {A} : Maybe (@Excl A) := λ x,
match x with Excl a => Some a | _ => None end. match x with Excl a => Some a | _ => None end.
Section excl. Section excl.
Context {A : cofeT}. Context {SI : sidx} {A : ofe}.
Implicit Types a b : A. Implicit Types a b : A.
Implicit Types x y : excl A. Implicit Types x y : excl A.
(* Cofe *) (* Cofe *)
Inductive excl_equiv : Equiv (excl A) := Inductive excl_equiv : Equiv (excl A) :=
| Excl_equiv a b : a b Excl a Excl b | Excl_equiv a b : a b Excl a Excl b
| ExclBot_equiv : ExclBot ExclBot. | ExclInvalid_equiv : ExclInvalid ExclInvalid.
Existing Instance excl_equiv. Local Existing Instance excl_equiv.
Inductive excl_dist : Dist (excl A) := Inductive excl_dist : Dist (excl A) :=
| Excl_dist a b n : a {n} b Excl a {n} Excl b | Excl_dist a b n : a {n} b Excl a {n} Excl b
| ExclBot_dist n : ExclBot {n} ExclBot. | ExclInvalid_dist n : ExclInvalid {n} ExclInvalid.
Existing Instance excl_dist. Local Existing Instance excl_dist.
Global Instance Excl_ne n : Proper (dist n ==> dist n) (@Excl A). Global Instance Excl_ne : NonExpansive (@Excl A).
Proof. by constructor. Qed. Proof. by constructor. Qed.
Global Instance Excl_proper : Proper (() ==> ()) (@Excl A). Global Instance Excl_proper : Proper (() ==> ()) (@Excl A).
Proof. by constructor. Qed. Proof. by constructor. Qed.
...@@ -40,76 +44,68 @@ Proof. by inversion_clear 1. Qed. ...@@ -40,76 +44,68 @@ Proof. by inversion_clear 1. Qed.
Global Instance Excl_dist_inj n : Inj (dist n) (dist n) (@Excl A). Global Instance Excl_dist_inj n : Inj (dist n) (dist n) (@Excl A).
Proof. by inversion_clear 1. Qed. Proof. by inversion_clear 1. Qed.
Program Definition excl_chain (c : chain (excl A)) (a : A) : chain A := Definition excl_ofe_mixin : OfeMixin (excl A).
{| chain_car n := match c n return _ with Excl y => y | _ => a end |}.
Next Obligation. intros c a n i ?; simpl. by destruct (chain_cauchy c n i). Qed.
Instance excl_compl : Compl (excl A) := λ c,
match c 0 with Excl a => Excl (compl (excl_chain c a)) | x => x end.
Definition excl_cofe_mixin : CofeMixin (excl A).
Proof. Proof.
split. apply (iso_ofe_mixin (maybe Excl)).
- intros x y; split; [by destruct 1; constructor; apply equiv_dist|]. - by intros [a|] [b|]; split; inversion_clear 1; constructor.
intros Hxy; feed inversion (Hxy 1); subst; constructor; apply equiv_dist. - by intros n [a|] [b|]; split; inversion_clear 1; constructor.
by intros n; feed inversion (Hxy n). Qed.
- intros n; split. Canonical Structure exclO : ofe := Ofe (excl A) excl_ofe_mixin.
+ by intros []; constructor.
+ by destruct 1; constructor. Global Instance excl_cofe `{!Cofe A} : Cofe exclO.
+ destruct 1; inversion_clear 1; constructor; etrans; eauto. Proof.
- by inversion_clear 1; constructor; apply dist_S. apply (iso_cofe (from_option Excl ExclInvalid) (maybe Excl)).
- intros n c; rewrite /compl /excl_compl. - by intros n [a|] [b|]; split; inversion_clear 1; constructor.
feed inversion (chain_cauchy c 0 n); first auto with lia; constructor. - by intros []; constructor.
rewrite (conv_compl n (excl_chain c _)) /=. destruct (c n); naive_solver.
Qed. Qed.
Canonical Structure exclC : cofeT := CofeT (excl A) excl_cofe_mixin.
Global Instance excl_discrete : Discrete A Discrete exclC. Global Instance excl_ofe_discrete : OfeDiscrete A OfeDiscrete exclO.
Proof. by inversion_clear 2; constructor; apply (timeless _). Qed. Proof. by inversion_clear 2; constructor; apply (discrete_0 _). Qed.
Global Instance excl_leibniz : LeibnizEquiv A LeibnizEquiv (excl A). Global Instance excl_leibniz : LeibnizEquiv A LeibnizEquiv (excl A).
Proof. by destruct 2; f_equal; apply leibniz_equiv. Qed. Proof. by destruct 2; f_equal; apply leibniz_equiv. Qed.
Global Instance Excl_timeless a : Timeless a Timeless (Excl a). Global Instance Excl_discrete a : Discrete a Discrete (Excl a).
Proof. by inversion_clear 2; constructor; apply (timeless _). Qed. Proof. by inversion_clear 2; constructor; apply (discrete_0 _). Qed.
Global Instance ExclBot_timeless : Timeless (@ExclBot A). Global Instance ExclInvalid_discrete : Discrete (@ExclInvalid A).
Proof. by inversion_clear 1; constructor. Qed. Proof. by inversion_clear 1; constructor. Qed.
(* CMRA *) (* CMRA *)
Instance excl_valid : Valid (excl A) := λ x, Local Instance excl_valid_instance : Valid (excl A) := λ x,
match x with Excl _ => True | ExclBot => False end. match x with Excl _ => True | ExclInvalid => False end.
Instance excl_validN : ValidN (excl A) := λ n x, Local Instance excl_validN_instance : ValidN (excl A) := λ n x,
match x with Excl _ => True | ExclBot => False end. match x with Excl _ => True | ExclInvalid => False end.
Instance excl_pcore : PCore (excl A) := λ _, None. Local Instance excl_pcore_instance : PCore (excl A) := λ _, None.
Instance excl_op : Op (excl A) := λ x y, ExclBot. Local Instance excl_op_instance : Op (excl A) := λ x y, ExclInvalid.
Lemma excl_cmra_mixin : CMRAMixin (excl A). Lemma excl_cmra_mixin : CmraMixin (excl A).
Proof. Proof.
split; try discriminate. split; try discriminate.
- by intros n []; destruct 1; constructor. - by intros [] n; destruct 1; constructor.
- by destruct 1; intros ?. - by destruct 1; intros ?.
- intros x; split. done. by move=> /(_ 0). - intros x; split; [done|]. by move=> /(_ 0).
- intros n [?|]; simpl; auto with lia. - intros n m [?|]; simpl; auto.
- by intros [?|] [?|] [?|]; constructor. - by intros [?|] [?|] [?|]; constructor.
- by intros [?|] [?|]; constructor. - by intros [?|] [?|]; constructor.
- by intros n [?|] [?|]. - by intros n [?|] [?|].
- intros n x [?|] [?|] ?; inversion_clear 1; eauto. - intros n x [?|] [?|] ? Hx; eexists _, _; inversion_clear Hx; eauto.
Qed.
Canonical Structure exclR := Cmra (excl A) excl_cmra_mixin.
Global Instance excl_cmra_discrete : OfeDiscrete A CmraDiscrete exclR.
Proof. split; first apply _. by intros []. Qed.
Lemma excl_included x y : x y y = ExclInvalid.
Proof.
split.
- destruct x, y; intros [[] Hxy]; by inv Hxy.
- intros ->. by exists ExclInvalid.
Qed. Qed.
Canonical Structure exclR := Lemma excl_includedN n x y : x {n} y y = ExclInvalid.
CMRAT (excl A) excl_cofe_mixin excl_cmra_mixin.
Global Instance excl_cmra_discrete : Discrete A CMRADiscrete exclR.
Proof. split. apply _. by intros []. Qed.
(** Internalized properties *)
Lemma excl_equivI {M} (x y : excl A) :
x y ⊣⊢ (match x, y with
| Excl a, Excl b => a b
| ExclBot, ExclBot => True
| _, _ => False
end : uPred M).
Proof. Proof.
uPred.unseal. do 2 split. by destruct 1. by destruct x, y; try constructor. split.
- destruct x, y; intros [[] Hxy]; by inv Hxy.
- intros ->. by exists ExclInvalid.
Qed. Qed.
Lemma excl_validI {M} (x : excl A) :
x ⊣⊢ (if x is ExclBot then False else True : uPred M).
Proof. uPred.unseal. by destruct x. Qed.
(** Exclusive *) (** Exclusive *)
Global Instance excl_exclusive x : Exclusive x. Global Instance excl_exclusive x : Exclusive x.
...@@ -121,60 +117,69 @@ Proof. by destruct mx. Qed. ...@@ -121,60 +117,69 @@ Proof. by destruct mx. Qed.
Lemma excl_validN_inv_r n mx a : {n} (mx Excl' a) mx = None. Lemma excl_validN_inv_r n mx a : {n} (mx Excl' a) mx = None.
Proof. by destruct mx. Qed. Proof. by destruct mx. Qed.
Lemma Excl_includedN n a b : Excl' a {n} Excl' b a {n} b. Lemma Excl_includedN n a b : Excl' a {n} Excl' b a {n} b.
Proof. by intros [[c|] Hb%(inj Some)]; inversion_clear Hb. Qed. Proof.
Lemma Excl_included a b : Excl' a Excl' b a b. split; [|by intros ->]. by intros [[c|] Hb%(inj Some)]; inversion_clear Hb.
Proof. by intros [[c|] Hb%(inj Some)]; inversion_clear Hb. Qed. Qed.
Lemma Excl_included a b : Excl' a Excl' b a b.
Proof.
split; [|by intros ->]. by intros [[c|] Hb%(inj Some)]; inversion_clear Hb.
Qed.
Lemma ExclInvalid_included ea : ea ExclInvalid.
Proof. by exists ExclInvalid. Qed.
End excl. End excl.
Arguments exclC : clear implicits. (* We use a [Hint Extern] with [apply:], instead of [Hint Immediate], to invoke
Arguments exclR : clear implicits. the new unification algorithm. The old unification algorithm sometimes gets
confused by going from [ucmra]'s to [cmra]'s and back. *)
Global Hint Extern 0 (_ ExclInvalid) => apply: ExclInvalid_included : core.
Global Arguments exclO {_} _.
Global Arguments exclR {_} _.
(* Functor *) (* Functor *)
Definition excl_map {A B} (f : A B) (x : excl A) : excl B := Definition excl_map {A B} (f : A B) (x : excl A) : excl B :=
match x with Excl a => Excl (f a) | ExclBot => ExclBot end. match x with Excl a => Excl (f a) | ExclInvalid => ExclInvalid end.
Lemma excl_map_id {A} (x : excl A) : excl_map id x = x. Lemma excl_map_id {A} (x : excl A) : excl_map id x = x.
Proof. by destruct x. Qed. Proof. by destruct x. Qed.
Lemma excl_map_compose {A B C} (f : A B) (g : B C) (x : excl A) : Lemma excl_map_compose {A B C} (f : A B) (g : B C) (x : excl A) :
excl_map (g f) x = excl_map g (excl_map f x). excl_map (g f) x = excl_map g (excl_map f x).
Proof. by destruct x. Qed. Proof. by destruct x. Qed.
Lemma excl_map_ext {A B : cofeT} (f g : A B) x : Lemma excl_map_ext {SI : sidx} {A B : ofe} (f g : A B) x :
( x, f x g x) excl_map f x excl_map g x. ( x, f x g x) excl_map f x excl_map g x.
Proof. by destruct x; constructor. Qed. Proof. by destruct x; constructor. Qed.
Instance excl_map_ne {A B : cofeT} n : Global Instance excl_map_ne {SI : sidx} {A B : ofe} n :
Proper ((dist n ==> dist n) ==> dist n ==> dist n) (@excl_map A B). Proper ((dist n ==> dist n) ==> dist n ==> dist n) (@excl_map A B).
Proof. by intros f f' Hf; destruct 1; constructor; apply Hf. Qed. Proof. by intros f f' Hf; destruct 1; constructor; apply Hf. Qed.
Instance excl_map_cmra_monotone {A B : cofeT} (f : A B) : Global Instance excl_map_cmra_morphism {SI : sidx} {A B : ofe} (f : A B) :
( n, Proper (dist n ==> dist n) f) CMRAMonotone (excl_map f). NonExpansive f CmraMorphism (excl_map f).
Proof. Proof. split; try done; try apply _. by intros n [a|]. Qed.
split; try apply _. Definition exclO_map {SI : sidx} {A B} (f : A -n> B) : exclO A -n> exclO B :=
- by intros n [a|]. OfeMor (excl_map f).
- intros x y [z Hy]; exists (excl_map f z); apply equiv_dist=> n. Global Instance exclO_map_ne {SI : sidx} A B : NonExpansive (@exclO_map SI A B).
move: Hy=> /equiv_dist /(_ n) ->; by destruct x, z. Proof. by intros n f f' Hf []; constructor; apply Hf. Qed.
Qed.
Definition exclC_map {A B} (f : A -n> B) : exclC A -n> exclC B := Program Definition exclRF {SI : sidx} (F : oFunctor) : rFunctor := {|
CofeMor (excl_map f). rFunctor_car A _ B _ := (exclR (oFunctor_car F A B));
Instance exclC_map_ne A B n : Proper (dist n ==> dist n) (@exclC_map A B). rFunctor_map A1 _ A2 _ B1 _ B2 _ fg := exclO_map (oFunctor_map F fg)
Proof. by intros f f' Hf []; constructor; apply Hf. Qed.
Program Definition exclRF (F : cFunctor) : rFunctor := {|
rFunctor_car A B := (exclR (cFunctor_car F A B));
rFunctor_map A1 A2 B1 B2 fg := exclC_map (cFunctor_map F fg)
|}. |}.
Next Obligation. Next Obligation.
intros F A1 A2 B1 B2 n x1 x2 ??. by apply exclC_map_ne, cFunctor_ne. intros ? F A1 ? A2 ? B1 ? B2 ? n x1 x2 ??.
by apply exclO_map_ne, oFunctor_map_ne.
Qed. Qed.
Next Obligation. Next Obligation.
intros F A B x; simpl. rewrite -{2}(excl_map_id x). intros ? F A ? B ? x; simpl. rewrite -{2}(excl_map_id x).
apply excl_map_ext=>y. by rewrite cFunctor_id. apply excl_map_ext=>y. by rewrite oFunctor_map_id.
Qed. Qed.
Next Obligation. Next Obligation.
intros F A1 A2 A3 B1 B2 B3 f g f' g' x; simpl. rewrite -excl_map_compose. intros ? F A1 ? A2 ? A3 ? B1 ? B2 ? B3 ? f g f' g' x; simpl.
apply excl_map_ext=>y; apply cFunctor_compose. rewrite -excl_map_compose.
apply excl_map_ext=>y; apply oFunctor_map_compose.
Qed. Qed.
Instance exclRF_contractive F : Global Instance exclRF_contractive {SI : sidx} F :
cFunctorContractive F rFunctorContractive (exclRF F). oFunctorContractive F rFunctorContractive (exclRF F).
Proof. Proof.
intros A1 A2 B1 B2 n x1 x2 ??. by apply exclC_map_ne, cFunctor_contractive. intros A1 ? A2 ? B1 ? B2 ? n x1 x2 ??.
by apply exclO_map_ne, oFunctor_map_contractive.
Qed. Qed.
(** This file provides a version of the fractional camera whose elements are
in the internal (0,1] of the rational numbers.
Notice that this camera could in principle be obtained by restricting the
validity of the unbounded fractional camera [ufrac]. *)
From iris.algebra Require Export cmra.
From iris.algebra Require Import proofmode_classes.
From iris.prelude Require Import options.
(** Since the standard (0,1] fractional camera is used more often, we define
[frac] through a [Notation] instead of a [Definition]. That way, Coq infers the
[frac] camera by default when using the [Qp] type. *)
Notation frac := Qp (only parsing).
Section frac.
Context {SI : sidx}.
Canonical Structure fracO := leibnizO frac.
Local Instance frac_valid_instance : Valid frac := λ x, (x 1)%Qp.
Local Instance frac_pcore_instance : PCore frac := λ _, None.
Local Instance frac_op_instance : Op frac := λ x y, (x + y)%Qp.
Lemma frac_valid p : p (p 1)%Qp.
Proof. done. Qed.
Lemma frac_valid_1 : 1%Qp.
Proof. done. Qed.
Lemma frac_op p q : p q = (p + q)%Qp.
Proof. done. Qed.
Lemma frac_included p q : p q (p < q)%Qp.
Proof. by rewrite Qp.lt_sum. Qed.
Corollary frac_included_weak p q : p q (p q)%Qp.
Proof. rewrite frac_included. apply Qp.lt_le_incl. Qed.
Definition frac_ra_mixin : RAMixin frac.
Proof.
split; try apply _; try done.
intros p q. rewrite !frac_valid frac_op=> ?.
trans (p + q)%Qp; last done. apply Qp.le_add_l.
Qed.
Canonical Structure fracR := discreteR frac frac_ra_mixin.
Global Instance frac_cmra_discrete : CmraDiscrete fracR.
Proof. apply discrete_cmra_discrete. Qed.
Global Instance frac_full_exclusive : Exclusive 1%Qp.
Proof. intros p. apply Qp.not_add_le_l. Qed.
Global Instance frac_cancelable (q : frac) : Cancelable q.
Proof. intros n p1 p2 _. apply (inj (Qp.add q)). Qed.
Global Instance frac_id_free (q : frac) : IdFree q.
Proof. intros p _. apply Qp.add_id_free. Qed.
(* This one has a higher precendence than [is_op_op] so we get a [+] instead
of an [⋅]. *)
Global Instance frac_is_op q1 q2 : IsOp (q1 + q2)%Qp q1 q2 | 10.
Proof. done. Qed.
Global Instance is_op_frac q : IsOp' q (q/2)%Qp (q/2)%Qp.
Proof. by rewrite /IsOp' /IsOp frac_op Qp.div_2. Qed.
End frac.
From stdpp Require Import finite.
From iris.algebra Require Export cmra.
From iris.algebra Require Import updates.
From iris.prelude Require Import options.
Definition discrete_fun_insert {SI : sidx} `{!EqDecision A} {B : A ofe}
(x : A) (y : B x) (f : discrete_fun B) : discrete_fun B := λ x',
match decide (x = x') with left H => eq_rect _ B y _ H | right _ => f x' end.
Global Instance: Params (@discrete_fun_insert) 6 := {}.
Definition discrete_fun_singleton {SI : sidx} `{!EqDecision A} {B : A ucmra}
(x : A) (y : B x) : discrete_fun B := discrete_fun_insert x y ε.
Global Instance: Params (@discrete_fun_singleton) 6 := {}.
Section ofe.
Context {SI : sidx} `{!EqDecision A} {B : A ofe}.
Implicit Types x : A.
Implicit Types f g : discrete_fun B.
Global Instance discrete_funO_ofe_discrete :
( i, OfeDiscrete (B i)) OfeDiscrete (discrete_funO B).
Proof. intros HB f f' Heq i. apply HB, Heq. Qed.
(** Properties of discrete_fun_insert. *)
Global Instance discrete_fun_insert_ne x :
NonExpansive2 (discrete_fun_insert (B:=B) x).
Proof.
intros n y1 y2 ? f1 f2 ? x'; rewrite /discrete_fun_insert.
by destruct (decide _) as [[]|].
Qed.
Global Instance discrete_fun_insert_proper x :
Proper (() ==> () ==> ()) (discrete_fun_insert (B:=B) x) := ne_proper_2 _.
Lemma discrete_fun_lookup_insert f x y : (discrete_fun_insert x y f) x = y.
Proof.
rewrite /discrete_fun_insert; destruct (decide _) as [Hx|]; last done.
by rewrite (proof_irrel Hx eq_refl).
Qed.
Lemma discrete_fun_lookup_insert_ne f x x' y :
x x' (discrete_fun_insert x y f) x' = f x'.
Proof. by rewrite /discrete_fun_insert; destruct (decide _). Qed.
Global Instance discrete_fun_insert_discrete f x y :
Discrete f Discrete y Discrete (discrete_fun_insert x y f).
Proof.
intros ?? g Heq x'; destruct (decide (x = x')) as [->|].
- rewrite discrete_fun_lookup_insert.
apply: discrete. by rewrite -(Heq x') discrete_fun_lookup_insert.
- rewrite discrete_fun_lookup_insert_ne //.
apply: discrete. by rewrite -(Heq x') discrete_fun_lookup_insert_ne.
Qed.
End ofe.
Section cmra.
Context {SI : sidx} `{!EqDecision A} {B : A ucmra}.
Implicit Types x : A.
Implicit Types f g : discrete_fun B.
Global Instance discrete_funR_cmra_discrete:
( i, CmraDiscrete (B i)) CmraDiscrete (discrete_funR B).
Proof. intros HB. split; [apply _|]. intros x Hv i. apply HB, Hv. Qed.
Global Instance discrete_fun_singleton_ne x :
NonExpansive (discrete_fun_singleton x : B x _).
Proof.
intros n y1 y2 ?; apply discrete_fun_insert_ne; [done|].
by apply equiv_dist.
Qed.
Global Instance discrete_fun_singleton_proper x :
Proper (() ==> ()) (discrete_fun_singleton x) := ne_proper _.
Lemma discrete_fun_lookup_singleton x (y : B x) : (discrete_fun_singleton x y) x = y.
Proof. by rewrite /discrete_fun_singleton discrete_fun_lookup_insert. Qed.
Lemma discrete_fun_lookup_singleton_ne x x' (y : B x) :
x x' (discrete_fun_singleton x y) x' = ε.
Proof. intros; by rewrite /discrete_fun_singleton discrete_fun_lookup_insert_ne. Qed.
Global Instance discrete_fun_singleton_discrete x (y : B x) :
( i, Discrete (ε : B i)) Discrete y Discrete (discrete_fun_singleton x y).
Proof. apply _. Qed.
Lemma discrete_fun_singleton_validN n x (y : B x) : {n} discrete_fun_singleton x y {n} y.
Proof.
split; [by move=>/(_ x); rewrite discrete_fun_lookup_singleton|].
move=>Hx x'; destruct (decide (x = x')) as [->|];
rewrite ?discrete_fun_lookup_singleton ?discrete_fun_lookup_singleton_ne //.
by apply ucmra_unit_validN.
Qed.
Lemma discrete_fun_singleton_valid x (y : B x) : discrete_fun_singleton x y y.
Proof.
by rewrite !cmra_valid_validN; setoid_rewrite discrete_fun_singleton_validN.
Qed.
Lemma discrete_fun_singleton_unit x : discrete_fun_singleton x (ε : B x) ε.
Proof.
move=>x'; destruct (decide (x = x')) as [->|];
by rewrite ?discrete_fun_lookup_singleton ?discrete_fun_lookup_singleton_ne.
Qed.
Lemma discrete_fun_singleton_core x (y : B x) :
core (discrete_fun_singleton x y) discrete_fun_singleton x (core y).
Proof.
move=>x'; destruct (decide (x = x')) as [->|];
by rewrite discrete_fun_lookup_core ?discrete_fun_lookup_singleton
?discrete_fun_lookup_singleton_ne // (core_id_core _).
Qed.
Global Instance discrete_fun_singleton_core_id x (y : B x) :
CoreId y CoreId (discrete_fun_singleton x y).
Proof. by rewrite !core_id_total discrete_fun_singleton_core=> ->. Qed.
Lemma discrete_fun_singleton_op (x : A) (y1 y2 : B x) :
discrete_fun_singleton x y1 discrete_fun_singleton x y2 discrete_fun_singleton x (y1 y2).
Proof.
intros x'; destruct (decide (x' = x)) as [->|].
- by rewrite discrete_fun_lookup_op !discrete_fun_lookup_singleton.
- by rewrite discrete_fun_lookup_op !discrete_fun_lookup_singleton_ne // left_id.
Qed.
Lemma discrete_fun_insert_updateP x (P : B x Prop) (Q : discrete_fun B Prop) g y1 :
y1 ~~>: P ( y2, P y2 Q (discrete_fun_insert x y2 g))
discrete_fun_insert x y1 g ~~>: Q.
Proof.
intros Hy1 HP; apply cmra_total_updateP.
intros n gf Hg. destruct (Hy1 n (Some (gf x))) as (y2&?&?).
{ move: (Hg x). by rewrite discrete_fun_lookup_op discrete_fun_lookup_insert. }
exists (discrete_fun_insert x y2 g); split; [auto|].
intros x'; destruct (decide (x' = x)) as [->|];
rewrite discrete_fun_lookup_op ?discrete_fun_lookup_insert //; [].
move: (Hg x'). by rewrite discrete_fun_lookup_op !discrete_fun_lookup_insert_ne.
Qed.
Lemma discrete_fun_insert_updateP' x (P : B x Prop) g y1 :
y1 ~~>: P
discrete_fun_insert x y1 g ~~>: λ g', y2, g' = discrete_fun_insert x y2 g P y2.
Proof. eauto using discrete_fun_insert_updateP. Qed.
Lemma discrete_fun_insert_update g x y1 y2 :
y1 ~~> y2 discrete_fun_insert x y1 g ~~> discrete_fun_insert x y2 g.
Proof.
rewrite !cmra_update_updateP; eauto using discrete_fun_insert_updateP with subst.
Qed.
Lemma discrete_fun_singleton_updateP x (P : B x Prop) (Q : discrete_fun B Prop) y1 :
y1 ~~>: P ( y2, P y2 Q (discrete_fun_singleton x y2))
discrete_fun_singleton x y1 ~~>: Q.
Proof. rewrite /discrete_fun_singleton; eauto using discrete_fun_insert_updateP. Qed.
Lemma discrete_fun_singleton_updateP' x (P : B x Prop) y1 :
y1 ~~>: P
discrete_fun_singleton x y1 ~~>: λ g, y2, g = discrete_fun_singleton x y2 P y2.
Proof. eauto using discrete_fun_singleton_updateP. Qed.
Lemma discrete_fun_singleton_update x (y1 y2 : B x) :
y1 ~~> y2 discrete_fun_singleton x y1 ~~> discrete_fun_singleton x y2.
Proof. eauto using discrete_fun_insert_update. Qed.
Lemma discrete_fun_singleton_updateP_empty x (P : B x Prop) (Q : discrete_fun B Prop) :
ε ~~>: P ( y2, P y2 Q (discrete_fun_singleton x y2)) ε ~~>: Q.
Proof.
intros Hx HQ; apply cmra_total_updateP.
intros n gf Hg. destruct (Hx n (Some (gf x))) as (y2&?&?); first apply Hg.
exists (discrete_fun_singleton x y2); split; [by apply HQ|].
intros x'; destruct (decide (x' = x)) as [->|].
- by rewrite discrete_fun_lookup_op discrete_fun_lookup_singleton.
- rewrite discrete_fun_lookup_op discrete_fun_lookup_singleton_ne //; by apply Hg.
Qed.
Lemma discrete_fun_singleton_updateP_empty' x (P : B x Prop) :
ε ~~>: P ε ~~>: λ g, y2, g = discrete_fun_singleton x y2 P y2.
Proof. eauto using discrete_fun_singleton_updateP_empty. Qed.
Lemma discrete_fun_singleton_update_empty x (y : B x) :
ε ~~> y ε ~~> discrete_fun_singleton x y.
Proof.
rewrite !cmra_update_updateP;
eauto using discrete_fun_singleton_updateP_empty with subst.
Qed.
Lemma discrete_fun_updateP `{!Finite A} f P Q :
( a, f a ~~>: P a) ( f', ( a, P a (f' a)) Q f') f ~~>: Q.
Proof.
repeat setoid_rewrite cmra_total_updateP. intros Hf HP n h Hh.
destruct (finite_choice
(λ a y, P a y {n} (y (h a)))) as [f' Hf']; first naive_solver.
naive_solver.
Qed.
Lemma discrete_fun_updateP' `{!Finite A} f P :
( a, f a ~~>: P a) f ~~>: λ f', a, P a (f' a).
Proof. eauto using discrete_fun_updateP. Qed.
Lemma discrete_fun_update f g :
( a, f a ~~> g a) f ~~> g.
Proof.
repeat setoid_rewrite cmra_total_update.
intros Hfg n h Hh a. apply (Hfg a), Hh.
Qed.
End cmra.
From stdpp Require Export list gmap.
From iris.algebra Require Export list cmra.
From iris.algebra Require Import gset.
From iris.algebra Require Import updates local_updates proofmode_classes big_op.
From iris.prelude Require Import options.
Section ofe.
Context {SI : sidx} `{Countable K} {A : ofe}.
Implicit Types m : gmap K A.
Implicit Types i : K.
Local Instance gmap_dist : Dist (gmap K A) := λ n m1 m2,
i, m1 !! i {n} m2 !! i.
Definition gmap_ofe_mixin : OfeMixin (gmap K A).
Proof.
split.
- intros m1 m2; split.
+ by intros Hm n k; apply equiv_dist.
+ intros Hm k; apply equiv_dist; intros n; apply Hm.
- intros n; split.
+ by intros m k.
+ by intros m1 m2 ? k.
+ by intros m1 m2 m3 ?? k; trans (m2 !! k).
- intros n n' m1 m2 ? k ?. by eapply dist_le.
Qed.
Canonical Structure gmapO : ofe := Ofe (gmap K A) gmap_ofe_mixin.
Program Definition gmap_chain (c : chain gmapO) (k : K) : chain (optionO A) :=
{| chain_car n := c n !! k |}.
Next Obligation. intros c k n p Hp. by apply c. Qed.
Program Definition gmap_bchain {n} (c : bchain gmapO n)
(k : K) : bchain (optionO A) n :=
{| bchain_car p Hp := c p Hp !! k |}.
Next Obligation. intros n c k p γ Hpγ Hp ; by apply c. Qed.
Definition gmap_compl `{!Cofe A} : Compl gmapO := λ c,
map_imap (λ i _, compl (gmap_chain c i)) (c 0).
Definition gmap_lbcompl `{!Cofe A} : LBCompl gmapO := λ n Hn c,
map_imap (λ i _, lbcompl Hn (gmap_bchain c i)) (c _ (SIdx.limit_lt_0 _ Hn)).
Global Program Instance gmap_cofe `{!Cofe A} : Cofe gmapO :=
{| compl := gmap_compl; lbcompl := gmap_lbcompl |}.
Next Obligation.
intros ? n c k. rewrite /gmap_compl map_lookup_imap.
oinversion (λ H, chain_cauchy c 0 n H k); simplify_option_eq;
[apply SIdx.le_0_l| |done].
rewrite conv_compl /=. by apply reflexive_eq.
Qed.
Next Obligation.
intros ? n Hn c p Hp k. rewrite /lbcompl /gmap_lbcompl.
rewrite map_lookup_imap.
oinversion (bchain_cauchy _ c _ p (SIdx.limit_lt_0 _ Hn) Hp _ k);
simplify_option_eq; [apply SIdx.le_0_l| |done].
rewrite (conv_lbcompl _ _ Hp) /=. by apply reflexive_eq.
Qed.
Next Obligation.
intros ? n Hn c1 c2 p Hc k. rewrite /gmap_lbcompl !map_lookup_imap.
oinversion (Hc _ (SIdx.limit_lt_0 _ Hn) k); simpl; eauto.
apply lbcompl_ne=> ??. apply Hc.
Qed.
Global Instance gmap_ofe_discrete : OfeDiscrete A OfeDiscrete gmapO.
Proof. intros ? m m' ? i. by apply (discrete_0 _). Qed.
(* why doesn't this go automatic? *)
Global Instance gmapO_leibniz: LeibnizEquiv A LeibnizEquiv gmapO.
Proof. intros; change (LeibnizEquiv (gmap K A)); apply _. Qed.
Global Instance lookup_ne k : NonExpansive (lookup k : gmap K A option A).
Proof. by intros n m1 m2. Qed.
Global Instance lookup_total_ne `{!Inhabited A} k :
NonExpansive (lookup_total k : gmap K A A).
Proof. intros n m1 m2. rewrite !lookup_total_alt. by intros ->. Qed.
Global Instance partial_alter_ne n :
Proper ((dist n ==> dist n) ==> (=) ==> dist n ==> dist n)
(partial_alter (M:=gmap K A)).
Proof.
by intros f1 f2 Hf i ? <- m1 m2 Hm j; destruct (decide (i = j)) as [->|];
rewrite ?lookup_partial_alter ?lookup_partial_alter_ne //;
try apply Hf; apply lookup_ne.
Qed.
Global Instance insert_ne i : NonExpansive2 (insert (M:=gmap K A) i).
Proof. intros n x y ? m m' ? j; apply partial_alter_ne; by try constructor. Qed.
Global Instance singleton_ne i : NonExpansive (singletonM i : A gmap K A).
Proof. by intros ????; apply insert_ne. Qed.
Global Instance delete_ne i : NonExpansive (delete (M:=gmap K A) i).
Proof.
intros n m m' ? j; destruct (decide (i = j)); simplify_map_eq;
[by constructor|by apply lookup_ne].
Qed.
Global Instance alter_ne (f : A A) (k : K) n :
Proper (dist n ==> dist n) f Proper (dist n ==> dist n) (alter (M := gmap K A) f k).
Proof. intros ? m m' Hm k'. by apply partial_alter_ne; [solve_proper|..]. Qed.
Global Instance gmap_empty_discrete : Discrete ( : gmap K A).
Proof.
intros m Hm i; specialize (Hm i); rewrite lookup_empty in Hm |- *.
inversion_clear Hm; constructor.
Qed.
Global Instance gmap_lookup_discrete m i : Discrete m Discrete (m !! i).
Proof.
intros ? [x|] Hx; [|by symmetry; apply: discrete].
assert (m {0} <[i:=x]> m)
by (by symmetry in Hx; inversion Hx; ofe_subst; rewrite insert_id).
by rewrite (discrete_0 m (<[i:=x]>m)) // lookup_insert.
Qed.
Global Instance gmap_insert_discrete m i x :
Discrete x Discrete m Discrete (<[i:=x]>m).
Proof.
intros ?? m' Hm j; destruct (decide (i = j)); simplify_map_eq.
{ by apply: discrete; rewrite -Hm lookup_insert. }
by apply: discrete; rewrite -Hm lookup_insert_ne.
Qed.
Global Instance gmap_singleton_discrete i x :
Discrete x Discrete ({[ i := x ]} : gmap K A).
Proof. rewrite /singletonM /map_singleton. apply _. Qed.
Lemma insert_idN n m i x :
m !! i {n} Some x <[i:=x]>m {n} m.
Proof. intros (y'&?&->)%dist_Some_inv_r'. by rewrite insert_id. Qed.
Global Instance gmap_dom_ne n :
Proper (({n}@{gmap K A}) ==> (=)) dom.
Proof. intros m1 m2 Hm. apply set_eq=> k. by rewrite !elem_of_dom Hm. Qed.
End ofe.
Global Instance map_seq_ne {SI : sidx} {A : ofe} start :
NonExpansive (map_seq (M:=gmap nat A) start).
Proof.
intros n l1 l2 Hl. revert start.
induction Hl; intros; simpl; repeat (done || f_equiv).
Qed.
Global Arguments gmapO {_} _ {_ _} _.
Global Instance merge_ne {SI : sidx} `{Countable K} {A B C : ofe} n :
Proper ((dist (A:=option A) n ==> dist (A:=option B) n ==> dist (A:=option C) n) ==>
dist n ==> dist n ==> ({n}@{gmap K C})) merge.
Proof.
intros ?? Hf ?? Hm1 ?? Hm2 i. rewrite !lookup_merge.
destruct (Hm1 i), (Hm2 i); try apply Hf; by constructor.
Qed.
Global Instance union_with_ne {SI : sidx} `{Countable K} {A : ofe} n :
Proper ((dist n ==> dist n ==> dist n) ==>
dist n ==> dist n ==> dist n) (union_with (M:=gmap K A)).
Proof.
intros ?? Hf ?? Hm1 ?? Hm2 i; apply (merge_ne _ _); auto.
by do 2 destruct 1; first [apply Hf | constructor].
Qed.
Global Instance map_fmap_ne {SI : sidx} `{Countable K} {A B : ofe} (f : A B) n :
Proper (dist n ==> dist n) f
Proper (dist n ==> ({n}@{gmap K B})) (fmap f).
Proof. intros ? m m' ? k; rewrite !lookup_fmap. by repeat f_equiv. Qed.
Global Instance map_zip_with_ne {SI : sidx}
`{Countable K} {A B C : ofe} (f : A B C) n :
Proper (dist n ==> dist n ==> dist n) f
Proper (dist n ==> dist n ==> dist n) (map_zip_with (M:=gmap K) f).
Proof.
intros Hf m1 m1' Hm1 m2 m2' Hm2. apply merge_ne; try done.
destruct 1; destruct 1; repeat f_equiv; constructor || done.
Qed.
Global Instance gmap_union_ne {SI : sidx} `{Countable K} {A : ofe} :
NonExpansive2 (union (A:=gmap K A)).
Proof. intros n. apply union_with_ne. by constructor. Qed.
Global Instance gmap_disjoint_ne {SI : sidx} `{Countable K} {A : ofe} n :
Proper (dist n ==> dist n ==> iff) (map_disjoint (M:=gmap K) (A:=A)).
Proof.
intros m1 m1' Hm1 m2 m2' Hm2; split;
intros Hm i; specialize (Hm i); by destruct (Hm1 i), (Hm2 i).
Qed.
Lemma gmap_union_dist_eq {SI : sidx} `{Countable K} {A : ofe}
(m m1 m2 : gmap K A) n :
m {n} m1 m2 m1' m2', m = m1' m2' m1' {n} m1 m2' {n} m2.
Proof.
split; last first.
{ by intros (m1'&m2'&->&<-&<-). }
intros Hm.
exists (filter (λ '(l,_), is_Some (m1 !! l)) m),
(m2 m1 filter (λ '(l,_), is_Some (m2 !! l)) m).
split_and!; [apply map_eq|..]; intros k; move: (Hm k);
rewrite ?lookup_union ?lookup_intersection !map_lookup_filter;
case _ : (m !! k)=> [x|] /=; case _ : (m1 !! k)=> [x1|] /=;
case _ : (m2 !! k)=> [x2|] /=; by inversion 1.
Qed.
Lemma big_opM_ne_2 {SI : sidx} {M: ofe} {o: M M M}
`{!Monoid o} `{Countable K} {A : ofe} (f g : K A M) m1 m2 n :
m1 {n} m2
( k y1 y2,
m1 !! k = Some y1 m2 !! k = Some y2 y1 {n} y2 f k y1 {n} g k y2)
([^o map] k y m1, f k y) {n} ([^o map] k y m2, g k y).
Proof.
intros Hl Hf. apply big_opM_gen_proper_2; try (apply _ || done).
{ by intros ?? ->. }
{ apply monoid_ne. }
intros k. assert (m1 !! k {n} m2 !! k) as Hlk by (by f_equiv).
destruct (m1 !! k) eqn:?, (m2 !! k) eqn:?; inversion Hlk; naive_solver.
Qed.
(* CMRA *)
Section cmra.
Context {SI : sidx} `{Countable K} {A : cmra}.
Implicit Types m : gmap K A.
Local Instance gmap_unit_instance : Unit (gmap K A) := ( : gmap K A).
Local Instance gmap_op_instance : Op (gmap K A) := merge op.
Local Instance gmap_pcore_instance : PCore (gmap K A) := λ m, Some (omap pcore m).
Local Instance gmap_valid_instance : Valid (gmap K A) := λ m, i, (m !! i).
Local Instance gmap_validN_instance : ValidN (gmap K A) := λ n m, i, {n} (m !! i).
Lemma gmap_op m1 m2 : m1 m2 = merge op m1 m2.
Proof. done. Qed.
Lemma lookup_op m1 m2 i : (m1 m2) !! i = m1 !! i m2 !! i.
Proof. rewrite lookup_merge. by destruct (m1 !! i), (m2 !! i). Qed.
Lemma lookup_core m i : core m !! i = core (m !! i).
Proof. by apply lookup_omap. Qed.
Lemma lookup_includedN n (m1 m2 : gmap K A) : m1 {n} m2 i, m1 !! i {n} m2 !! i.
Proof.
split; [by intros [m Hm] i; exists (m !! i); rewrite -lookup_op Hm|].
revert m2. induction m1 as [|i x m Hi IH] using map_ind=> m2 Hm.
{ exists m2. by rewrite left_id. }
destruct (IH (delete i m2)) as [m2' Hm2'].
{ intros j. move: (Hm j); destruct (decide (i = j)) as [->|].
- intros _. rewrite Hi. apply: ucmra_unit_leastN.
- rewrite lookup_insert_ne // lookup_delete_ne //. }
destruct (Hm i) as [my Hi']; simplify_map_eq.
exists (partial_alter (λ _, my) i m2')=>j; destruct (decide (i = j)) as [->|].
- by rewrite Hi' lookup_op lookup_insert lookup_partial_alter.
- move: (Hm2' j). by rewrite !lookup_op lookup_delete_ne //
lookup_insert_ne // lookup_partial_alter_ne.
Qed.
(* [m1 ≼ m2] is not equivalent to [∀ n, m1 ≼{n} m2],
so there is no good way to reuse the above proof. *)
Lemma lookup_included (m1 m2 : gmap K A) : m1 m2 i, m1 !! i m2 !! i.
Proof.
split; [by intros [m Hm] i; exists (m !! i); rewrite -lookup_op Hm|].
revert m2. induction m1 as [|i x m Hi IH] using map_ind=> m2 Hm.
{ exists m2. by rewrite left_id. }
destruct (IH (delete i m2)) as [m2' Hm2'].
{ intros j. move: (Hm j); destruct (decide (i = j)) as [->|].
- intros _. rewrite Hi. apply: ucmra_unit_least.
- rewrite lookup_insert_ne // lookup_delete_ne //. }
destruct (Hm i) as [my Hi']; simplify_map_eq.
exists (partial_alter (λ _, my) i m2')=>j; destruct (decide (i = j)) as [->|].
- by rewrite Hi' lookup_op lookup_insert lookup_partial_alter.
- move: (Hm2' j). by rewrite !lookup_op lookup_delete_ne //
lookup_insert_ne // lookup_partial_alter_ne.
Qed.
Lemma gmap_cmra_mixin : CmraMixin (gmap K A).
Proof.
apply cmra_total_mixin.
- eauto.
- intros n m1 m2 m3 Hm i; by rewrite !lookup_op (Hm i).
- intros n m1 m2 Hm i; by rewrite !lookup_core (Hm i).
- intros n m1 m2 Hm ? i; by rewrite -(Hm i).
- intros m; split.
+ by intros ? n i; apply cmra_valid_validN.
+ intros Hm i; apply cmra_valid_validN=> n; apply Hm.
- intros n m Hm i ? ?; eauto using cmra_validN_le.
- by intros m1 m2 m3 i; rewrite !lookup_op assoc.
- by intros m1 m2 i; rewrite !lookup_op comm.
- intros m i. by rewrite lookup_op lookup_core cmra_core_l.
- intros m i. by rewrite !lookup_core cmra_core_idemp.
- intros m1 m2; rewrite !lookup_included=> Hm i.
rewrite !lookup_core. by apply cmra_core_mono.
- intros n m1 m2 Hm i; apply cmra_validN_op_l with (m2 !! i).
by rewrite -lookup_op.
- intros n m y1 y2 Hm Heq.
refine ((λ FUN, _) (λ i, cmra_extend n (m !! i) (y1 !! i) (y2 !! i) (Hm i) _));
last by rewrite -lookup_op.
exists (map_imap (λ i _, projT1 (FUN i)) y1).
exists (map_imap (λ i _, proj1_sig (projT2 (FUN i))) y2).
split; [|split]=>i; rewrite ?lookup_op !map_lookup_imap;
destruct (FUN i) as (z1i&z2i&Hmi&Hz1i&Hz2i)=>/=.
+ destruct (y1 !! i), (y2 !! i); inversion Hz1i; inversion Hz2i; subst=>//.
+ revert Hz1i. case: (y1!!i)=>[?|] //.
+ revert Hz2i. case: (y2!!i)=>[?|] //.
Qed.
Canonical Structure gmapR := Cmra (gmap K A) gmap_cmra_mixin.
Global Instance gmap_cmra_discrete : CmraDiscrete A CmraDiscrete gmapR.
Proof. split; [apply _|]. intros m ? i. by apply: cmra_discrete_valid. Qed.
Lemma gmap_ucmra_mixin : UcmraMixin (gmap K A).
Proof.
split.
- by intros i; rewrite lookup_empty.
- by intros m i; rewrite /= lookup_op lookup_empty (left_id_L None _).
- constructor=> i. by rewrite lookup_omap lookup_empty.
Qed.
Canonical Structure gmapUR := Ucmra (gmap K A) gmap_ucmra_mixin.
Global Instance gmap_op_empty_l_L : LeftId (=@{gmap K A}) op.
Proof. apply _. Qed.
Global Instance gmap_op_empty_r : RightId (=@{gmap K A}) op.
Proof. apply _. Qed.
End cmra.
Global Arguments gmapR {_} _ {_ _} _.
Global Arguments gmapUR {_} _ {_ _} _.
Section properties.
Context {SI : sidx} `{Countable K} {A : cmra}.
Implicit Types m : gmap K A.
Implicit Types i : K.
Implicit Types x y : A.
Global Instance lookup_op_homomorphism i :
MonoidHomomorphism op op () (lookup i : gmap K A option A).
Proof.
split; [split|]; try apply _.
- intros m1 m2; by rewrite lookup_op.
- done.
Qed.
Lemma lookup_opM m1 mm2 i : (m1 ? mm2) !! i = m1 !! i (mm2 ≫= (.!! i)).
Proof. destruct mm2; by rewrite /= ?lookup_op ?right_id_L. Qed.
Lemma lookup_validN_Some n m i x : {n} m m !! i {n} Some x {n} x.
Proof. by move=> /(_ i) Hm Hi; move:Hm; rewrite Hi. Qed.
Lemma lookup_valid_Some m i x : m m !! i Some x x.
Proof. move=> Hm Hi. move:(Hm i). by rewrite Hi. Qed.
Lemma insert_validN n m i x : {n} x {n} m {n} <[i:=x]>m.
Proof. by intros ?? j; destruct (decide (i = j)); simplify_map_eq. Qed.
Lemma insert_valid m i x : x m <[i:=x]>m.
Proof. by intros ?? j; destruct (decide (i = j)); simplify_map_eq. Qed.
Lemma singleton_validN n i x : {n} ({[ i := x ]} : gmap K A) {n} x.
Proof.
split.
- move=>/(_ i); by simplify_map_eq.
- intros. apply insert_validN; first done. apply: ucmra_unit_validN.
Qed.
Lemma singleton_valid i x : ({[ i := x ]} : gmap K A) x.
Proof. rewrite !cmra_valid_validN. by setoid_rewrite singleton_validN. Qed.
Lemma delete_validN n m i : {n} m {n} (delete i m).
Proof. intros Hm j; destruct (decide (i = j)); by simplify_map_eq. Qed.
Lemma delete_valid m i : m (delete i m).
Proof. intros Hm j; destruct (decide (i = j)); by simplify_map_eq. Qed.
Lemma insert_singleton_op m i x : m !! i = None <[i:=x]> m = {[ i := x ]} m.
Proof.
intros Hi; apply map_eq=> j; destruct (decide (i = j)) as [->|].
- by rewrite lookup_op lookup_insert lookup_singleton Hi right_id_L.
- by rewrite lookup_op lookup_insert_ne // lookup_singleton_ne // left_id_L.
Qed.
Lemma singleton_core (i : K) (x : A) cx :
pcore x = Some cx core {[ i := x ]} =@{gmap K A} {[ i := cx ]}.
Proof. apply omap_singleton_Some. Qed.
Lemma singleton_core' (i : K) (x : A) cx :
pcore x Some cx core {[ i := x ]} ≡@{gmap K A} {[ i := cx ]}.
Proof.
intros (cx'&?&<-)%Some_equiv_eq. by rewrite (singleton_core _ _ cx').
Qed.
Lemma singleton_core_total `{!CmraTotal A} (i : K) (x : A) :
core {[ i := x ]} =@{gmap K A} {[ i := core x ]}.
Proof. apply singleton_core. rewrite cmra_pcore_core //. Qed.
Lemma singleton_op (i : K) (x y : A) :
{[ i := x ]} {[ i := y ]} =@{gmap K A} {[ i := x y ]}.
Proof. by apply (merge_singleton _ _ _ x y). Qed.
Global Instance singleton_is_op i a a1 a2 :
IsOp a a1 a2 IsOp' ({[ i := a ]} : gmap K A) {[ i := a1 ]} {[ i := a2 ]}.
Proof. rewrite /IsOp' /IsOp=> ->. by rewrite -singleton_op. Qed.
Lemma gmap_core_id m : ( i x, m !! i = Some x CoreId x) CoreId m.
Proof.
intros Hcore; apply core_id_total=> i.
rewrite lookup_core. destruct (m !! i) as [x|] eqn:Hix; rewrite Hix; [|done].
by eapply Hcore.
Qed.
Global Instance gmap_core_id' m : ( x : A, CoreId x) CoreId m.
Proof. auto using gmap_core_id. Qed.
Global Instance gmap_singleton_core_id i (x : A) :
CoreId x CoreId {[ i := x ]}.
Proof. intros. by apply core_id_total, singleton_core'. Qed.
Lemma singleton_includedN_l n m i x :
{[ i := x ]} {n} m y, m !! i {n} Some y Some x {n} Some y.
Proof.
split.
- move=> [m' /(_ i)]; rewrite lookup_op lookup_singleton=> Hi.
exists (x ? m' !! i). rewrite -Some_op_opM.
split; first done. apply cmra_includedN_l.
- intros (y&Hi&[mz Hy]). exists (partial_alter (λ _, mz) i m).
intros j; destruct (decide (i = j)) as [->|].
+ by rewrite lookup_op lookup_singleton lookup_partial_alter Hi.
+ by rewrite lookup_op lookup_singleton_ne// lookup_partial_alter_ne// left_id.
Qed.
(* We do not have [x ≼ y ↔ ∀ n, x ≼{n} y], so we cannot use the previous lemma *)
Lemma singleton_included_l m i x :
{[ i := x ]} m y, m !! i Some y Some x Some y.
Proof.
split.
- move=> [m' /(_ i)]; rewrite lookup_op lookup_singleton.
exists (x ? m' !! i). by rewrite -Some_op_opM.
- intros (y&Hi&[mz Hy]). exists (partial_alter (λ _, mz) i m).
intros j; destruct (decide (i = j)) as [->|].
+ by rewrite lookup_op lookup_singleton lookup_partial_alter Hi.
+ by rewrite lookup_op lookup_singleton_ne// lookup_partial_alter_ne// left_id.
Qed.
Lemma singleton_included_exclusive_l m i x :
Exclusive x m
{[ i := x ]} m m !! i Some x.
Proof.
intros ? Hm. rewrite singleton_included_l. split; last by eauto.
intros (y&?&->%(Some_included_exclusive _)); eauto using lookup_valid_Some.
Qed.
Lemma singleton_included i x y :
{[ i := x ]} ({[ i := y ]} : gmap K A) Some x Some y.
Proof.
rewrite singleton_included_l. split.
- intros (y'&Hi&?). rewrite lookup_insert in Hi. by rewrite Hi.
- intros ?. exists y. by rewrite lookup_insert.
Qed.
Lemma singleton_included_total `{!CmraTotal A} i x y :
{[ i := x ]} ({[ i := y ]} : gmap K A) x y.
Proof. rewrite singleton_included Some_included_total. done. Qed.
Lemma singleton_included_mono i x y :
x y {[ i := x ]} ({[ i := y ]} : gmap K A).
Proof. intros Hincl. apply singleton_included, Some_included_mono. done. Qed.
Global Instance singleton_cancelable i x :
Cancelable (Some x) Cancelable {[ i := x ]}.
Proof.
intros ? n m1 m2 Hv EQ j. move: (Hv j) (EQ j). rewrite !lookup_op.
destruct (decide (i = j)) as [->|].
- rewrite lookup_singleton. by apply cancelableN.
- by rewrite lookup_singleton_ne // !(left_id None _).
Qed.
Global Instance gmap_cancelable (m : gmap K A) :
( x : A, IdFree x) ( x : A, Cancelable x) Cancelable m.
Proof.
intros ?? n m1 m2 ?? i. apply (cancelableN (m !! i)); by rewrite -!lookup_op.
Qed.
Lemma insert_op m1 m2 i x y :
<[i:=x y]>(m1 m2) = <[i:=x]>m1 <[i:=y]>m2.
Proof. by rewrite (insert_merge () m1 m2 i (x y) x y). Qed.
Lemma insert_updateP (P : A Prop) (Q : gmap K A Prop) m i x :
x ~~>: P
( y, P y Q (<[i:=y]>m))
<[i:=x]>m ~~>: Q.
Proof.
intros Hx%option_updateP' HP; apply cmra_total_updateP=> n mf Hm.
destruct (Hx n (Some (mf !! i))) as ([y|]&?&?); try done.
{ by generalize (Hm i); rewrite lookup_op; simplify_map_eq. }
exists (<[i:=y]> m); split; first by auto.
intros j; move: (Hm j)=>{Hm}; rewrite !lookup_op=>Hm.
destruct (decide (i = j)); simplify_map_eq/=; auto.
Qed.
Lemma insert_updateP' (P : A Prop) m i x :
x ~~>: P <[i:=x]>m ~~>: λ m', y, m' = <[i:=y]>m P y.
Proof. eauto using insert_updateP. Qed.
Lemma insert_update m i x y : x ~~> y <[i:=x]>m ~~> <[i:=y]>m.
Proof. rewrite !cmra_update_updateP; eauto using insert_updateP with subst. Qed.
Lemma singleton_updateP (P : A Prop) (Q : gmap K A Prop) i x :
x ~~>: P ( y, P y Q {[ i := y ]}) {[ i := x ]} ~~>: Q.
Proof. apply insert_updateP. Qed.
Lemma singleton_updateP' (P : A Prop) i x :
x ~~>: P {[ i := x ]} ~~>: λ m, y, m = {[ i := y ]} P y.
Proof. apply insert_updateP'. Qed.
Lemma singleton_update i (x y : A) : x ~~> y {[ i := x ]} ~~> {[ i := y ]}.
Proof. apply insert_update. Qed.
Lemma delete_update m i : m ~~> delete i m.
Proof.
apply cmra_total_update=> n mf Hm j; destruct (decide (i = j)); subst.
- move: (Hm j). rewrite !lookup_op lookup_delete left_id.
apply cmra_validN_op_r.
- move: (Hm j). by rewrite !lookup_op lookup_delete_ne.
Qed.
Lemma gmap_op_union m1 m2 : m1 ## m2 m1 m2 = m1 m2.
Proof.
intros Hm. apply map_eq=> k. specialize (Hm k).
rewrite lookup_op lookup_union. by destruct (m1 !! k), (m2 !! k).
Qed.
Lemma gmap_op_valid0_disjoint m1 m2 :
{0} (m1 m2) ( k x, m1 !! k = Some x Exclusive x) m1 ## m2.
Proof.
unfold Exclusive. intros Hvalid Hexcl k.
specialize (Hvalid k). rewrite lookup_op in Hvalid. specialize (Hexcl k).
destruct (m1 !! k), (m2 !! k); [|done..].
rewrite -Some_op Some_validN in Hvalid. naive_solver.
Qed.
Lemma gmap_op_valid_disjoint m1 m2 :
(m1 m2) ( k x, m1 !! k = Some x Exclusive x) m1 ## m2.
Proof. move=> /cmra_valid_validN /(_ 0). apply gmap_op_valid0_disjoint. Qed.
Lemma dom_op m1 m2 : dom (m1 m2) = dom m1 dom m2.
Proof.
apply set_eq=> i; rewrite elem_of_union !elem_of_dom.
unfold is_Some; setoid_rewrite lookup_op.
destruct (m1 !! i), (m2 !! i); naive_solver.
Qed.
Lemma dom_included m1 m2 : m1 m2 dom m1 dom m2.
Proof.
rewrite lookup_included=>? i; rewrite !elem_of_dom. by apply is_Some_included.
Qed.
Section freshness.
Local Set Default Proof Using "Type*".
Context `{!Infinite K}.
Lemma alloc_updateP_strong_dep (Q : gmap K A Prop) (I : K Prop) m (f : K A) :
pred_infinite I
( i, m !! i = None I i (f i))
( i, m !! i = None I i Q (<[i:=f i]>m)) m ~~>: Q.
Proof.
move=> /(pred_infinite_set I (C:=gset K)) HP ? HQ.
apply cmra_total_updateP. intros n mf Hm.
destruct (HP (dom (m mf))) as [i [Hi1 Hi2]].
assert (m !! i = None).
{ eapply not_elem_of_dom. revert Hi2.
rewrite dom_op not_elem_of_union. naive_solver. }
exists (<[i:=f i]>m); split.
- by apply HQ.
- rewrite insert_singleton_op //.
rewrite -assoc -insert_singleton_op; last by eapply not_elem_of_dom.
apply insert_validN; [apply cmra_valid_validN|]; auto.
Qed.
Lemma alloc_updateP_strong (Q : gmap K A Prop) (I : K Prop) m x :
pred_infinite I
x ( i, m !! i = None I i Q (<[i:=x]>m)) m ~~>: Q.
Proof.
move=> HP ? HQ. eapply (alloc_updateP_strong_dep _ _ _ (λ _, x)); eauto.
Qed.
Lemma alloc_updateP (Q : gmap K A Prop) m x :
x ( i, m !! i = None Q (<[i:=x]>m)) m ~~>: Q.
Proof.
move=>??.
eapply (alloc_updateP_strong _ (λ _, True));
eauto using pred_infinite_True.
Qed.
Lemma alloc_updateP_cofinite (Q : gmap K A Prop) (J : gset K) m x :
x ( i, m !! i = None i J Q (<[i:=x]>m)) m ~~>: Q.
Proof.
eapply alloc_updateP_strong.
apply (pred_infinite_set (C:=gset K)).
intros E. exists (fresh (J E)).
apply not_elem_of_union, is_fresh.
Qed.
(* Variants without the universally quantified Q, for use in case that is an evar. *)
Lemma alloc_updateP_strong_dep' m (f : K A) (I : K Prop) :
pred_infinite I
( i, m !! i = None I i (f i))
m ~~>: λ m', i, I i m' = <[i:=f i]>m m !! i = None.
Proof. eauto using alloc_updateP_strong_dep. Qed.
Lemma alloc_updateP_strong' m x (I : K Prop) :
pred_infinite I
x m ~~>: λ m', i, I i m' = <[i:=x]>m m !! i = None.
Proof. eauto using alloc_updateP_strong. Qed.
Lemma alloc_updateP' m x :
x m ~~>: λ m', i, m' = <[i:=x]>m m !! i = None.
Proof. eauto using alloc_updateP. Qed.
Lemma alloc_updateP_cofinite' m x (J : gset K) :
x m ~~>: λ m', i, i J m' = <[i:=x]>m m !! i = None.
Proof. eauto using alloc_updateP_cofinite. Qed.
End freshness.
Lemma alloc_unit_singleton_updateP (P : A Prop) (Q : gmap K A Prop) u i :
u LeftId () u ()
u ~~>: P ( y, P y Q {[ i := y ]}) ~~>: Q.
Proof.
intros ?? Hx HQ. apply cmra_total_updateP=> n gf Hg.
destruct (Hx n (gf !! i)) as (y&?&Hy).
{ move:(Hg i). rewrite !left_id.
case: (gf !! i)=>[x|]; rewrite /= ?left_id //.
intros; by apply cmra_valid_validN. }
exists {[ i := y ]}; split; first by auto.
intros i'; destruct (decide (i' = i)) as [->|].
- rewrite lookup_op lookup_singleton.
move:Hy; case: (gf !! i)=>[x|]; rewrite /= ?right_id //.
- move:(Hg i'). by rewrite !lookup_op lookup_singleton_ne // !left_id.
Qed.
Lemma alloc_unit_singleton_updateP' (P: A Prop) u i :
u LeftId () u ()
u ~~>: P ~~>: λ m, y, m = {[ i := y ]} P y.
Proof. eauto using alloc_unit_singleton_updateP. Qed.
Lemma alloc_unit_singleton_update (u : A) i (y : A) :
u LeftId () u () u ~~> y (∅:gmap K A) ~~> {[ i := y ]}.
Proof.
rewrite !cmra_update_updateP;
eauto using alloc_unit_singleton_updateP with subst.
Qed.
Lemma gmap_local_update m1 m2 m1' m2' :
( i, (m1 !! i, m2 !! i) ~l~> (m1' !! i, m2' !! i))
(m1, m2) ~l~> (m1', m2').
Proof.
intros Hupd. apply local_update_unital=> n mf Hmv Hm.
apply forall_and_distr=> i. rewrite lookup_op -cmra_opM_fmap_Some.
apply Hupd; simpl; first done. by rewrite Hm lookup_op cmra_opM_fmap_Some.
Qed.
Lemma alloc_local_update m1 m2 i x :
m1 !! i = None x (m1,m2) ~l~> (<[i:=x]>m1, <[i:=x]>m2).
Proof.
intros Hi ?. apply gmap_local_update=> j.
destruct (decide (i = j)) as [->|]; last by rewrite !lookup_insert_ne.
rewrite !lookup_insert Hi. by apply alloc_option_local_update.
Qed.
Lemma alloc_singleton_local_update m i x :
m !! i = None x (m,) ~l~> (<[i:=x]>m, {[ i:=x ]}).
Proof. apply alloc_local_update. Qed.
Lemma insert_local_update m1 m2 i x y x' y' :
m1 !! i = Some x m2 !! i = Some y
(x, y) ~l~> (x', y')
(m1, m2) ~l~> (<[i:=x']>m1, <[i:=y']>m2).
Proof.
intros Hi1 Hi2 Hup. apply gmap_local_update=> j.
destruct (decide (i = j)) as [->|]; last by rewrite !lookup_insert_ne.
rewrite !lookup_insert Hi1 Hi2. by apply option_local_update.
Qed.
Lemma singleton_local_update_any m i y x' y' :
( x, m !! i = Some x (x, y) ~l~> (x', y'))
(m, {[ i := y ]}) ~l~> (<[i:=x']>m, {[ i := y' ]}).
Proof.
intros. apply gmap_local_update=> j.
destruct (decide (i = j)) as [->|]; last by rewrite !lookup_insert_ne.
rewrite !lookup_singleton lookup_insert.
destruct (m !! j); first by eauto using option_local_update.
apply local_update_total_valid0=> _ _ /option_includedN; naive_solver.
Qed.
Lemma singleton_local_update m i x y x' y' :
m !! i = Some x
(x, y) ~l~> (x', y')
(m, {[ i := y ]}) ~l~> (<[i:=x']>m, {[ i := y' ]}).
Proof.
intros Hmi ?. apply singleton_local_update_any.
intros x2. rewrite Hmi=>[=<-]. done.
Qed.
Lemma delete_local_update m1 m2 i x `{!Exclusive x} :
m2 !! i = Some x (m1, m2) ~l~> (delete i m1, delete i m2).
Proof.
intros Hi. apply gmap_local_update=> j.
destruct (decide (i = j)) as [->|]; last by rewrite !lookup_delete_ne.
rewrite !lookup_delete Hi. by apply delete_option_local_update.
Qed.
Lemma delete_singleton_local_update m i x `{!Exclusive x} :
(m, {[ i := x ]}) ~l~> (delete i m, ).
Proof.
rewrite -(delete_singleton i x).
by eapply delete_local_update, lookup_singleton.
Qed.
Lemma delete_local_update_cancelable m1 m2 i mx `{!Cancelable mx} :
m1 !! i mx m2 !! i mx
(m1, m2) ~l~> (delete i m1, delete i m2).
Proof.
intros Hi1 Hi2. apply gmap_local_update=> j.
destruct (decide (i = j)) as [->|]; last by rewrite !lookup_delete_ne.
rewrite !lookup_delete Hi1 Hi2. by apply delete_option_local_update_cancelable.
Qed.
Lemma delete_singleton_local_update_cancelable m i x `{!Cancelable (Some x)} :
m !! i Some x (m, {[ i := x ]}) ~l~> (delete i m, ).
Proof.
intros. rewrite -(delete_singleton i x).
apply (delete_local_update_cancelable m _ i (Some x));
[done|by rewrite lookup_singleton].
Qed.
Lemma gmap_fmap_mono {B : cmra} (f : A B) m1 m2 :
Proper (() ==> ()) f
( x y, x y f x f y) m1 m2 fmap f m1 fmap f m2.
Proof.
intros ??. rewrite !lookup_included=> Hm i.
rewrite !lookup_fmap. by apply option_fmap_mono.
Qed.
Lemma big_opM_singletons m :
([^op map] k x m, {[ k := x ]}) = m.
Proof.
(* We are breaking the big_opM abstraction here. The reason is that [map_ind]
is too weak: we need an induction principle that visits all the keys in the
right order, namely the order in which they appear in map_to_list. Here,
we achieve this by unfolding [big_opM] and doing induction over that list
instead. *)
rewrite big_op.big_opM_unseal /big_op.big_opM_def -{2}(list_to_map_to_list m).
assert (NoDup (map_to_list m).*1) as Hnodup by apply NoDup_fst_map_to_list.
revert Hnodup. induction (map_to_list m) as [|[k x] l IH]; csimpl; first done.
intros [??]%NoDup_cons. rewrite IH //.
rewrite insert_singleton_op ?not_elem_of_list_to_map_1 //.
Qed.
Lemma big_opS_gset_to_gmap (X : gset K) (a : A) :
([^op set] x X, {[ x := a ]}) gset_to_gmap a X.
Proof.
induction X as [|x X ? IH] using set_ind_L.
{ rewrite big_opS_empty gset_to_gmap_empty //. }
rewrite big_opS_insert //.
rewrite gset_to_gmap_union_singleton.
rewrite insert_singleton_op; [|by rewrite lookup_gset_to_gmap_None].
by rewrite IH.
Qed.
Lemma big_opS_gset_to_gmap_L `{!LeibnizEquiv A} (X : gset K) (a : A) :
([^op set] x X, {[ x := a ]}) = gset_to_gmap a X.
Proof. apply leibniz_equiv, big_opS_gset_to_gmap. Qed.
End properties.
Section unital_properties.
Context {SI : sidx} `{Countable K} {A : ucmra}.
Implicit Types m : gmap K A.
Implicit Types i : K.
Implicit Types x y : A.
Lemma insert_alloc_local_update m1 m2 i x x' y' :
m1 !! i = Some x m2 !! i = None
(x, ε) ~l~> (x', y')
(m1, m2) ~l~> (<[i:=x']>m1, <[i:=y']>m2).
Proof.
intros Hi1 Hi2 Hup. apply local_update_unital=> n mf Hm1v Hm.
assert (mf !! i {n} Some x) as Hif.
{ move: (Hm i). by rewrite lookup_op Hi1 Hi2 left_id. }
destruct (Hup n (mf !! i)) as [Hx'v Hx'eq].
{ move: (Hm1v i). by rewrite Hi1. }
{ by rewrite Hif -(inj_iff Some) -Some_op_opM -Some_op left_id. }
split.
- by apply insert_validN.
- simpl in Hx'eq. by rewrite -(insert_idN n mf i x) // -insert_op -Hm Hx'eq Hif.
Qed.
End unital_properties.
(** Functor *)
Global Instance gmap_fmap_ne {SI : sidx} `{Countable K} {A B : ofe} (f : A B) n :
Proper (dist n ==> dist n) f Proper (dist n ==> dist (A:=gmap K _) n) (fmap f).
Proof. by intros ? m m' Hm k; rewrite !lookup_fmap; apply option_fmap_ne. Qed.
Lemma gmap_fmap_ne_ext {SI : sidx} `{Countable K}
{A : Type} {B : ofe} (f1 f2 : A B) (m : gmap K A) n :
( i x, m !! i = Some x f1 x {n} f2 x)
f1 <$> m {n} f2 <$> m.
Proof.
move => Hf i. rewrite !lookup_fmap.
destruct (m !! i) eqn:?; constructor; by eauto.
Qed.
Global Instance gmap_fmap_cmra_morphism {SI : sidx} `{Countable K} {A B : cmra}
(f : A B) `{!CmraMorphism f} : CmraMorphism (fmap f : gmap K A gmap K B).
Proof.
split; try apply _.
- by intros n m ? i; rewrite lookup_fmap; apply (cmra_morphism_validN _).
- intros m. apply Some_proper=>i. rewrite lookup_fmap !lookup_omap lookup_fmap.
case: (m!!i)=> //= ?. apply cmra_morphism_pcore, _.
- intros m1 m2 i. by rewrite lookup_op !lookup_fmap lookup_op cmra_morphism_op.
Qed.
Definition gmapO_map {SI : sidx} `{Countable K} {A B: ofe} (f: A -n> B) :
gmapO K A -n> gmapO K B := OfeMor (fmap f : gmapO K A gmapO K B).
Global Instance gmapO_map_ne {SI : sidx} `{Countable K} {A B: ofe} :
NonExpansive (@gmapO_map _ K _ _ A B).
Proof.
intros n f g Hf m k; rewrite /= !lookup_fmap.
destruct (_ !! k) eqn:?; simpl; constructor; apply Hf.
Qed.
Program Definition gmapOF {SI : sidx} K `{Countable K}
(F : oFunctor) : oFunctor := {|
oFunctor_car A _ B _ := gmapO K (oFunctor_car F A B);
oFunctor_map A1 _ A2 _ B1 _ B2 _ fg := gmapO_map (oFunctor_map F fg)
|}.
Next Obligation.
intros ? K ?? F A1 ? A2 ? B1 ? B2 ? n f g Hfg.
by apply gmapO_map_ne, oFunctor_map_ne.
Qed.
Next Obligation.
intros ? K ?? F A ? B ? m; simpl in *. rewrite /= -{2}(map_fmap_id m).
apply: map_fmap_equiv_ext=>y ??; apply oFunctor_map_id.
Qed.
Next Obligation.
intros ? K ?? F A1 ? A2 ? A3 ? B1 ? B2 ? B3 ? f g f' g' m.
rewrite /= -map_fmap_compose.
apply: map_fmap_equiv_ext=>y ??; apply oFunctor_map_compose.
Qed.
Global Instance gmapOF_contractive {SI : sidx} `{Countable K} F :
oFunctorContractive F oFunctorContractive (gmapOF K F).
Proof.
by intros ? ? A1 ? A2 ? B1 ? B2 ? n f g Hfg;
apply gmapO_map_ne, oFunctor_map_contractive.
Qed.
Program Definition gmapURF {SI : sidx} K `{Countable K}
(F : rFunctor) : urFunctor := {|
urFunctor_car A _ B _ := gmapUR K (rFunctor_car F A B);
urFunctor_map A1 _ A2 _ B1 _ B2 _ fg := gmapO_map (rFunctor_map F fg)
|}.
Next Obligation.
by intros ? K ?? F A1 ? A2 ? B1 ? B2 ? n f g Hfg;
apply gmapO_map_ne, rFunctor_map_ne.
Qed.
Next Obligation.
intros ? K ?? F A ? B ? m. rewrite /= -{2}(map_fmap_id m).
apply: map_fmap_equiv_ext=> y ??. apply rFunctor_map_id.
Qed.
Next Obligation.
intros ? K ?? F A1 ? A2 ? A3 ? B1 ? B2 ? B3 ? f g f' g' m.
rewrite /= -map_fmap_compose.
apply: map_fmap_equiv_ext=> y ??. apply rFunctor_map_compose.
Qed.
Global Instance gmapURF_contractive {SI : sidx} `{Countable K} F :
rFunctorContractive F urFunctorContractive (gmapURF K F).
Proof.
intros ? A1 ? A2 ? B1 ? B2 ? n f g Hfg;
by apply gmapO_map_ne, rFunctor_map_contractive.
Qed.
Program Definition gmapRF {SI : sidx} K `{Countable K}
(F : rFunctor) : rFunctor := {|
rFunctor_car A _ B _ := gmapR K (rFunctor_car F A B);
rFunctor_map A1 _ A2 _ B1 _ B2 _ fg := gmapO_map (rFunctor_map F fg)
|}.
Solve Obligations with apply @gmapURF.
Global Instance gmapRF_contractive {SI : sidx} `{Countable K} F :
rFunctorContractive F rFunctorContractive (gmapRF K F).
Proof. apply gmapURF_contractive. Qed.
From stdpp Require Export sets gmultiset countable.
From iris.algebra Require Export cmra.
From iris.algebra Require Import updates local_updates big_op.
From iris.prelude Require Import options.
(* The multiset union CMRA *)
Section gmultiset.
Context `{Countable K} {SI : sidx}.
Implicit Types X Y : gmultiset K.
Canonical Structure gmultisetO := discreteO (gmultiset K).
Local Instance gmultiset_valid_instance : Valid (gmultiset K) := λ _, True.
Local Instance gmultiset_validN_instance : ValidN (gmultiset K) := λ _ _, True.
Local Instance gmultiset_unit_instance : Unit (gmultiset K) := ( : gmultiset K).
Local Instance gmultiset_op_instance : Op (gmultiset K) := disj_union.
Local Instance gmultiset_pcore_instance : PCore (gmultiset K) := λ X, Some ∅.
Lemma gmultiset_op X Y : X Y = X Y.
Proof. done. Qed.
Lemma gmultiset_core X : core X = ∅.
Proof. done. Qed.
Lemma gmultiset_included X Y : X Y X Y.
Proof.
split.
- intros [Z ->%leibniz_equiv].
rewrite gmultiset_op. apply gmultiset_disj_union_subseteq_l.
- intros ->%gmultiset_disj_union_difference. by exists (Y X).
Qed.
Lemma gmultiset_ra_mixin : RAMixin (gmultiset K).
Proof.
apply ra_total_mixin; eauto.
- by intros X Y Z ->%leibniz_equiv.
- by intros X Y ->%leibniz_equiv.
- solve_proper.
- intros X1 X2 X3. by rewrite !gmultiset_op assoc_L.
- intros X1 X2. by rewrite !gmultiset_op comm_L.
- intros X. by rewrite gmultiset_core left_id.
- intros X1 X2 HX. rewrite !gmultiset_core. exists ∅.
by rewrite left_id.
Qed.
Canonical Structure gmultisetR := discreteR (gmultiset K) gmultiset_ra_mixin.
Global Instance gmultiset_cmra_discrete : CmraDiscrete gmultisetR.
Proof. apply discrete_cmra_discrete. Qed.
Lemma gmultiset_ucmra_mixin : UcmraMixin (gmultiset K).
Proof.
split; [done | | done]. intros X.
by rewrite gmultiset_op left_id_L.
Qed.
Canonical Structure gmultisetUR := Ucmra (gmultiset K) gmultiset_ucmra_mixin.
Global Instance gmultiset_cancelable X : Cancelable X.
Proof.
apply: discrete_cancelable=> Y Z _ ?. fold_leibniz. by apply (inj (X ⊎.)).
Qed.
Lemma gmultiset_opM X mY : X ? mY = X default mY.
Proof. destruct mY; by rewrite /= ?right_id_L. Qed.
Lemma gmultiset_update X Y : X ~~> Y.
Proof. done. Qed.
Lemma gmultiset_local_update X Y X' Y' : X Y' = X' Y (X,Y) ~l~> (X', Y').
Proof.
intros HXY. rewrite local_update_unital_discrete=> Z' _. intros ->%leibniz_equiv.
split; first done. apply leibniz_equiv_iff, (inj (. Y)).
rewrite -HXY !gmultiset_op.
by rewrite -(comm_L _ Y) (comm_L _ Y') assoc_L.
Qed.
Lemma gmultiset_local_update_alloc X Y X' : (X,Y) ~l~> (X X', Y X').
Proof. apply gmultiset_local_update. by rewrite (comm_L _ Y) assoc_L. Qed.
Lemma gmultiset_local_update_dealloc X Y X' :
X' Y (X,Y) ~l~> (X X', Y X').
Proof.
intros ->%gmultiset_disj_union_difference. apply local_update_total_valid.
intros _ _ ->%gmultiset_included%gmultiset_disj_union_difference.
apply gmultiset_local_update. apply gmultiset_eq=> x.
repeat (rewrite multiplicity_difference || rewrite multiplicity_disj_union).
lia.
Qed.
Lemma big_opMS_singletons X :
([^op mset] x X, {[+ x +]}) = X.
Proof.
induction X as [|x X IH] using gmultiset_ind.
- rewrite big_opMS_empty. done.
- unfold_leibniz. rewrite big_opMS_disj_union // big_opMS_singleton IH //.
Qed.
End gmultiset.
Global Arguments gmultisetO _ {_ _ _}.
Global Arguments gmultisetR _ {_ _ _}.
Global Arguments gmultisetUR _ {_ _ _}.
From stdpp Require Export sets gmap mapset.
From iris.algebra Require Export cmra. From iris.algebra Require Export cmra.
From iris.algebra Require Import updates local_updates. From iris.algebra Require Import updates local_updates big_op.
From iris.prelude Require Export collections gmap mapset. From iris.prelude Require Import options.
(* The union CMRA *) (* The union CMRA *)
Section gset. Section gset.
Context `{Countable K}. Context {SI : sidx} `{Countable K}.
Implicit Types X Y : gset K. Implicit Types X Y : gset K.
Canonical Structure gsetC := leibnizC (gset K). Canonical Structure gsetO := discreteO (gset K).
Instance gset_valid : Valid (gset K) := λ _, True. Local Instance gset_valid_instance : Valid (gset K) := λ _, True.
Instance gset_op : Op (gset K) := union. Local Instance gset_unit_instance : Unit (gset K) := ( : gset K).
Instance gset_pcore : PCore (gset K) := λ X, Some X. Local Instance gset_op_instance : Op (gset K) := union.
Local Instance gset_pcore_instance : PCore (gset K) := λ X, Some X.
Lemma gset_op_union X Y : X Y = X Y. Lemma gset_op X Y : X Y = X Y.
Proof. done. Qed. Proof. done. Qed.
Lemma gset_core_self X : core X = X. Lemma gset_core X : core X = X.
Proof. done. Qed. Proof. done. Qed.
Lemma gset_included X Y : X Y X Y. Lemma gset_included X Y : X Y X Y.
Proof. Proof.
split. split.
- intros [Z ->]. rewrite gset_op_union. set_solver. - intros [Z ->]. rewrite gset_op. set_solver.
- intros (Z&->&?)%subseteq_disjoint_union_L. by exists Z. - intros (Z&->&?)%subseteq_disjoint_union_L. by exists Z.
Qed. Qed.
Lemma gset_ra_mixin : RAMixin (gset K). Lemma gset_ra_mixin : RAMixin (gset K).
Proof. Proof.
apply ra_total_mixin; eauto. apply ra_total_mixin; apply _ || eauto; [].
- solve_proper. intros X. by rewrite gset_core idemp_L.
- solve_proper.
- solve_proper.
- intros X1 X2 X3. by rewrite !gset_op_union assoc_L.
- intros X1 X2. by rewrite !gset_op_union comm_L.
- intros X. by rewrite gset_core_self idemp_L.
Qed. Qed.
Canonical Structure gsetR := discreteR (gset K) gset_ra_mixin. Canonical Structure gsetR := discreteR (gset K) gset_ra_mixin.
Lemma gset_ucmra_mixin : UCMRAMixin (gset K). Global Instance gset_cmra_discrete : CmraDiscrete gsetR.
Proof. split. done. intros X. by rewrite gset_op_union left_id_L. done. Qed. Proof. apply discrete_cmra_discrete. Qed.
Canonical Structure gsetUR :=
discreteUR (gset K) gset_ra_mixin gset_ucmra_mixin. Lemma gset_ucmra_mixin : UcmraMixin (gset K).
Proof. split; [ done | | done ]. intros X. by rewrite gset_op left_id_L. Qed.
Canonical Structure gsetUR := Ucmra (gset K) gset_ucmra_mixin.
Lemma gset_opM X mY : X ? mY = X from_option id mY. Lemma gset_opM X mY : X ? mY = X default mY.
Proof. destruct mY; by rewrite /= ?right_id_L. Qed. Proof using SI. destruct mY; by rewrite /= ?right_id_L. Qed.
Lemma gset_update X Y : X ~~> Y. Lemma gset_update X Y : X ~~> Y.
Proof. done. Qed. Proof. done. Qed.
...@@ -51,40 +50,58 @@ Section gset. ...@@ -51,40 +50,58 @@ Section gset.
Proof. Proof.
intros (Z&->&?)%subseteq_disjoint_union_L. intros (Z&->&?)%subseteq_disjoint_union_L.
rewrite local_update_unital_discrete=> Z' _ /leibniz_equiv_iff->. rewrite local_update_unital_discrete=> Z' _ /leibniz_equiv_iff->.
split. done. rewrite gset_op_union. set_solver. split; [done|]. rewrite gset_op. set_solver.
Qed. Qed.
Global Instance gset_persistent X : Persistent X. Global Instance gset_core_id X : CoreId X.
Proof. by apply persistent_total; rewrite gset_core_self. Qed. Proof. by apply core_id_total; rewrite gset_core. Qed.
Lemma big_opS_singletons X :
([^op set] x X, {[ x ]}) = X.
Proof.
induction X as [|x X Hx IH] using set_ind_L.
- rewrite big_opS_empty. done.
- unfold_leibniz. rewrite big_opS_insert // IH //.
Qed.
(** Add support [X ≼ Y] to [set_solver]. (We get support for [⋅] for free
because it is definitionally equal to [∪]). *)
Global Instance set_unfold_gset_included X Y Q :
SetUnfold (X Y) Q SetUnfold (X Y) Q.
Proof. intros [?]; constructor. by rewrite gset_included. Qed.
End gset. End gset.
Arguments gsetR _ {_ _}. Global Arguments gsetO {_} _ {_ _}.
Arguments gsetUR _ {_ _}. Global Arguments gsetR {_} _ {_ _}.
Global Arguments gsetUR {_} _ {_ _}.
(* The disjoint union CMRA *) (* The disjoint union CMRA *)
Inductive gset_disj K `{Countable K} := Inductive gset_disj K `{Countable K} :=
| GSet : gset K gset_disj K | GSet : gset K gset_disj K
| GSetBot : gset_disj K. | GSetInvalid : gset_disj K.
Arguments GSet {_ _ _} _. Global Arguments GSet {_ _ _} _.
Arguments GSetBot {_ _ _}. Global Arguments GSetInvalid {_ _ _}.
Section gset_disj. Section gset_disj.
Context `{Countable K}. Context {SI : sidx} `{Countable K}.
Arguments op _ _ !_ !_ /. Local Arguments op _ _ !_ !_ /.
Arguments cmra_op _ !_ !_ /. Local Arguments cmra_op _ !_ !_ /.
Arguments ucmra_op _ !_ !_ /. Local Arguments ucmra_op _ !_ !_ /.
Global Instance GSet_inj : Inj (=@{gset K}) (=) GSet.
Proof. intros ???. naive_solver. Qed.
Canonical Structure gset_disjC := leibnizC (gset_disj K). Canonical Structure gset_disjO := leibnizO (gset_disj K).
Instance gset_disj_valid : Valid (gset_disj K) := λ X, Local Instance gset_disj_valid_instance : Valid (gset_disj K) := λ X,
match X with GSet _ => True | GSetBot => False end. match X with GSet _ => True | GSetInvalid => False end.
Instance gset_disj_empty : Empty (gset_disj K) := GSet ∅. Local Instance gset_disj_unit_instance : Unit (gset_disj K) := GSet ∅.
Instance gset_disj_op : Op (gset_disj K) := λ X Y, Local Instance gset_disj_op_instance : Op (gset_disj K) := λ X Y,
match X, Y with match X, Y with
| GSet X, GSet Y => if decide (X Y) then GSet (X Y) else GSetBot | GSet X, GSet Y => if decide (X ## Y) then GSet (X Y) else GSetInvalid
| _, _ => GSetBot | _, _ => GSetInvalid
end. end.
Instance gset_disj_pcore : PCore (gset_disj K) := λ _, Some . Local Instance gset_disj_pcore_instance : PCore (gset_disj K) := λ _, Some ε.
Ltac gset_disj_solve := Ltac gset_disj_solve :=
repeat (simpl || case_decide); repeat (simpl || case_decide);
...@@ -97,11 +114,11 @@ Section gset_disj. ...@@ -97,11 +114,11 @@ Section gset_disj.
- intros (Z&->&?)%subseteq_disjoint_union_L. - intros (Z&->&?)%subseteq_disjoint_union_L.
exists (GSet Z). gset_disj_solve. exists (GSet Z). gset_disj_solve.
Qed. Qed.
Lemma gset_disj_valid_inv_l X Y : (GSet X Y) Y', Y = GSet Y' X Y'. Lemma gset_disj_valid_inv_l X Y : (GSet X Y) Y', Y = GSet Y' X ## Y'.
Proof. destruct Y; repeat (simpl || case_decide); by eauto. Qed. Proof. destruct Y; repeat (simpl || case_decide); by eauto. Qed.
Lemma gset_disj_union X Y : X Y GSet X GSet Y = GSet (X Y). Lemma gset_disj_union X Y : X ## Y GSet X GSet Y = GSet (X Y).
Proof. intros. by rewrite /= decide_True. Qed. Proof. intros. by rewrite /= decide_True. Qed.
Lemma gset_disj_valid_op X Y : (GSet X GSet Y) X Y. Lemma gset_disj_valid_op X Y : (GSet X GSet Y) X ## Y.
Proof. simpl. case_decide; by split. Qed. Proof. simpl. case_decide; by split. Qed.
Lemma gset_disj_ra_mixin : RAMixin (gset_disj K). Lemma gset_disj_ra_mixin : RAMixin (gset_disj K).
...@@ -118,19 +135,22 @@ Section gset_disj. ...@@ -118,19 +135,22 @@ Section gset_disj.
Qed. Qed.
Canonical Structure gset_disjR := discreteR (gset_disj K) gset_disj_ra_mixin. Canonical Structure gset_disjR := discreteR (gset_disj K) gset_disj_ra_mixin.
Lemma gset_disj_ucmra_mixin : UCMRAMixin (gset_disj K). Global Instance gset_disj_cmra_discrete : CmraDiscrete gset_disjR.
Proof. apply discrete_cmra_discrete. Qed.
Lemma gset_disj_ucmra_mixin : UcmraMixin (gset_disj K).
Proof. split; try apply _ || done. intros [X|]; gset_disj_solve. Qed. Proof. split; try apply _ || done. intros [X|]; gset_disj_solve. Qed.
Canonical Structure gset_disjUR := Canonical Structure gset_disjUR := Ucmra (gset_disj K) gset_disj_ucmra_mixin.
discreteUR (gset_disj K) gset_disj_ra_mixin gset_disj_ucmra_mixin.
Arguments op _ _ _ _ : simpl never. Local Arguments op _ _ _ _ : simpl never.
Lemma gset_disj_alloc_updateP_strong P (Q : gset_disj K Prop) X : Lemma gset_disj_alloc_updateP_strong P (Q : gset_disj K Prop) X :
( Y, X Y j, j Y P j) ( Y, X Y j, j Y P j)
( i, i X P i Q (GSet ({[i]} X))) GSet X ~~>: Q. ( i, i X P i Q (GSet ({[i]} X)))
GSet X ~~>: Q.
Proof. Proof.
intros Hfresh HQ. intros Hfresh HQ.
apply cmra_discrete_updateP=> ? /gset_disj_valid_inv_l [Y [->?]]. apply cmra_discrete_total_updateP=> ? /gset_disj_valid_inv_l [Y [->?]].
destruct (Hfresh (X Y)) as (i&?&?); first set_solver. destruct (Hfresh (X Y)) as (i&?&?); first set_solver.
exists (GSet ({[ i ]} X)); split. exists (GSet ({[ i ]} X)); split.
- apply HQ; set_solver by eauto. - apply HQ; set_solver by eauto.
...@@ -154,13 +174,14 @@ Section gset_disj. ...@@ -154,13 +174,14 @@ Section gset_disj.
Proof. eauto using gset_disj_alloc_empty_updateP_strong. Qed. Proof. eauto using gset_disj_alloc_empty_updateP_strong. Qed.
Section fresh_updates. Section fresh_updates.
Context `{Fresh K (gset K), !FreshSpec K (gset K)}. Local Set Default Proof Using "Type*".
Context `{!Infinite K}.
Lemma gset_disj_alloc_updateP (Q : gset_disj K Prop) X : Lemma gset_disj_alloc_updateP (Q : gset_disj K Prop) X :
( i, i X Q (GSet ({[i]} X))) GSet X ~~>: Q. ( i, i X Q (GSet ({[i]} X))) GSet X ~~>: Q.
Proof. Proof.
intro; eapply gset_disj_alloc_updateP_strong with (λ _, True); eauto. intro; eapply gset_disj_alloc_updateP_strong with (λ _, True); eauto.
intros Y ?; exists (fresh Y); eauto using is_fresh. intros Y ?; exists (fresh Y). split; [|done]. apply is_fresh.
Qed. Qed.
Lemma gset_disj_alloc_updateP' X : Lemma gset_disj_alloc_updateP' X :
GSet X ~~>: λ Y, i, Y = GSet ({[ i ]} X) i X. GSet X ~~>: λ Y, i, Y = GSet ({[ i ]} X) i X.
...@@ -176,7 +197,7 @@ Section gset_disj. ...@@ -176,7 +197,7 @@ Section gset_disj.
End fresh_updates. End fresh_updates.
Lemma gset_disj_dealloc_local_update X Y : Lemma gset_disj_dealloc_local_update X Y :
(GSet X, GSet Y) ~l~> (GSet (X Y), ). (GSet X, GSet Y) ~l~> (GSet (X Y), GSet ).
Proof. Proof.
apply local_update_total_valid=> _ _ /gset_disj_included HYX. apply local_update_total_valid=> _ _ /gset_disj_included HYX.
rewrite local_update_unital_discrete=> -[Xf|] _ /leibniz_equiv_iff //=. rewrite local_update_unital_discrete=> -[Xf|] _ /leibniz_equiv_iff //=.
...@@ -185,7 +206,7 @@ Section gset_disj. ...@@ -185,7 +206,7 @@ Section gset_disj.
difference_diag_L !left_id_L difference_disjoint_L. difference_diag_L !left_id_L difference_disjoint_L.
Qed. Qed.
Lemma gset_disj_dealloc_empty_local_update X Z : Lemma gset_disj_dealloc_empty_local_update X Z :
(GSet Z GSet X, GSet Z) ~l~> (GSet X,). (GSet Z GSet X, GSet Z) ~l~> (GSet X, GSet ).
Proof. Proof.
apply local_update_total_valid=> /gset_disj_valid_op HZX _ _. apply local_update_total_valid=> /gset_disj_valid_op HZX _ _.
assert (X = (Z X) Z) as HX by set_solver. assert (X = (Z X) Z) as HX by set_solver.
...@@ -194,29 +215,44 @@ Section gset_disj. ...@@ -194,29 +215,44 @@ Section gset_disj.
Lemma gset_disj_dealloc_op_local_update X Y Z : Lemma gset_disj_dealloc_op_local_update X Y Z :
(GSet Z GSet X, GSet Z GSet Y) ~l~> (GSet X,GSet Y). (GSet Z GSet X, GSet Z GSet Y) ~l~> (GSet X,GSet Y).
Proof. Proof.
rewrite -{2}(left_id _ (GSet Y)). rewrite -{2}(left_id ε _ (GSet Y)).
apply op_local_update_frame, gset_disj_dealloc_empty_local_update. apply op_local_update_frame, gset_disj_dealloc_empty_local_update.
Qed. Qed.
Lemma gset_disj_alloc_op_local_update X Y Z : Lemma gset_disj_alloc_op_local_update X Y Z :
Z X (GSet X,GSet Y) ~l~> (GSet Z GSet X, GSet Z GSet Y). Z ## X (GSet X,GSet Y) ~l~> (GSet Z GSet X, GSet Z GSet Y).
Proof. Proof.
intros. apply op_local_update_discrete. by rewrite gset_disj_valid_op. intros. apply op_local_update_discrete. by rewrite gset_disj_valid_op.
Qed. Qed.
Lemma gset_disj_alloc_local_update X Y Z : Lemma gset_disj_alloc_local_update X Y Z :
Z X (GSet X,GSet Y) ~l~> (GSet (Z X), GSet (Z Y)). Z ## X (GSet X,GSet Y) ~l~> (GSet (Z X), GSet (Z Y)).
Proof. Proof.
intros. apply local_update_total_valid=> _ _ /gset_disj_included ?. intros. apply local_update_total_valid=> _ _ /gset_disj_included ?.
rewrite -!gset_disj_union //; last set_solver. rewrite -!gset_disj_union //; last set_solver.
auto using gset_disj_alloc_op_local_update. auto using gset_disj_alloc_op_local_update.
Qed. Qed.
Lemma gset_disj_alloc_empty_local_update X Z : Lemma gset_disj_alloc_empty_local_update X Z :
Z X (GSet X, GSet ) ~l~> (GSet (Z X), GSet Z). Z ## X (GSet X, GSet ) ~l~> (GSet (Z X), GSet Z).
Proof. Proof.
intros. rewrite -{2}(right_id_L _ union Z). intros. rewrite -{2}(right_id_L _ union Z).
apply gset_disj_alloc_local_update; set_solver. apply gset_disj_alloc_local_update; set_solver.
Qed. Qed.
(** Add some basic support for [GSet X = GSet Y], [GSet X ≼ GSet Y], and
[✓ (GSet X ⋅ GSet Y)] to [set_solver]. There are probably more cases we could
cover (e.g., involving [GSetInvalid], or nesting of [⋅]), but it is not clear
these are useful in practice, nor how to handle them effectively. *)
Global Instance set_unfold_gset_eq (X Y : gset K) Q :
SetUnfold (X = Y) Q SetUnfold (GSet X = GSet Y) Q.
Proof. intros [?]; constructor. by rewrite (inj_iff _). Qed.
Global Instance set_unfold_gset_disj_included (X Y : gset K) Q :
SetUnfold (X Y) Q SetUnfold (GSet X GSet Y) Q.
Proof. intros [?]; constructor. by rewrite gset_disj_included. Qed.
Global Instance set_unfold_gset_disj_valid_op (X Y : gset K) Q :
SetUnfold (X ## Y) Q SetUnfold ( (GSet X GSet Y)) Q.
Proof. intros [?]; constructor. by rewrite gset_disj_valid_op. Qed.
End gset_disj. End gset_disj.
Arguments gset_disjR _ {_ _}. Global Arguments gset_disjO {_} _ {_ _}.
Arguments gset_disjUR _ {_ _}. Global Arguments gset_disjR {_} _ {_ _}.
Global Arguments gset_disjUR {_} _ {_ _}.
From iris.algebra Require Export dfrac agree updates local_updates.
From iris.algebra Require Import proofmode_classes.
From iris.prelude Require Import options.
Definition dfrac_agreeR {SI : sidx} (A : ofe) : cmra := prodR dfracR (agreeR A).
Definition to_dfrac_agree {SI : sidx} {A : ofe} (d : dfrac) (a : A) : dfrac_agreeR A :=
(d, to_agree a).
Global Instance: Params (@to_dfrac_agree) 3 := {}.
(** To make it easier to work with the [Qp] version of this, we also provide
[to_frac_agree] and appropriate lemmas. *)
Definition to_frac_agree {SI : sidx} {A : ofe} (q : Qp) (a : A) : dfrac_agreeR A :=
to_dfrac_agree (DfracOwn q) a.
Global Instance: Params (@to_frac_agree) 3 := {}.
Section lemmas.
Context {SI : sidx} {A : ofe}.
Implicit Types (q : Qp) (d : dfrac) (a : A).
Global Instance to_dfrac_agree_ne d : NonExpansive (@to_dfrac_agree SI A d).
Proof. solve_proper. Qed.
Global Instance to_dfrac_agree_proper d :
Proper (() ==> ()) (@to_dfrac_agree SI A d).
Proof. solve_proper. Qed.
Global Instance to_dfrac_agree_exclusive a :
Exclusive (to_dfrac_agree (DfracOwn 1) a).
Proof. apply _. Qed.
Global Instance to_dfrac_agree_discrete d a : Discrete a Discrete (to_dfrac_agree d a).
Proof. apply _. Qed.
Global Instance to_dfrac_agree_injN n :
Inj2 (dist n) (dist n) (dist n) (@to_dfrac_agree SI A).
Proof. by intros d1 a1 d2 a2 [??%(inj to_agree)]. Qed.
Global Instance to_dfrac_agree_inj : Inj2 () () () (@to_dfrac_agree SI A).
Proof. by intros d1 a1 d2 a2 [??%(inj to_agree)]. Qed.
Lemma dfrac_agree_op d1 d2 a :
to_dfrac_agree (d1 d2) a to_dfrac_agree d1 a to_dfrac_agree d2 a.
Proof. rewrite /to_dfrac_agree -pair_op agree_idemp //. Qed.
Lemma frac_agree_op q1 q2 a :
to_frac_agree (q1 + q2) a to_frac_agree q1 a to_frac_agree q2 a.
Proof. rewrite -dfrac_agree_op. done. Qed.
Lemma dfrac_agree_op_valid d1 a1 d2 a2 :
(to_dfrac_agree d1 a1 to_dfrac_agree d2 a2)
(d1 d2) a1 a2.
Proof.
rewrite /to_dfrac_agree -pair_op pair_valid to_agree_op_valid. done.
Qed.
Lemma dfrac_agree_op_valid_L `{!LeibnizEquiv A} d1 a1 d2 a2 :
(to_dfrac_agree d1 a1 to_dfrac_agree d2 a2)
(d1 d2) a1 = a2.
Proof. unfold_leibniz. apply dfrac_agree_op_valid. Qed.
Lemma dfrac_agree_op_validN d1 a1 d2 a2 n :
{n} (to_dfrac_agree d1 a1 to_dfrac_agree d2 a2)
(d1 d2) a1 {n} a2.
Proof.
rewrite /to_dfrac_agree -pair_op pair_validN to_agree_op_validN. done.
Qed.
Lemma frac_agree_op_valid q1 a1 q2 a2 :
(to_frac_agree q1 a1 to_frac_agree q2 a2)
(q1 + q2 1)%Qp a1 a2.
Proof. apply dfrac_agree_op_valid. Qed.
Lemma frac_agree_op_valid_L `{!LeibnizEquiv A} q1 a1 q2 a2 :
(to_frac_agree q1 a1 to_frac_agree q2 a2)
(q1 + q2 1)%Qp a1 = a2.
Proof. apply dfrac_agree_op_valid_L. Qed.
Lemma frac_agree_op_validN q1 a1 q2 a2 n :
{n} (to_frac_agree q1 a1 to_frac_agree q2 a2)
(q1 + q2 1)%Qp a1 {n} a2.
Proof. apply dfrac_agree_op_validN. Qed.
Lemma dfrac_agree_included d1 a1 d2 a2 :
to_dfrac_agree d1 a1 to_dfrac_agree d2 a2
(d1 d2) a1 a2.
Proof. by rewrite pair_included to_agree_included. Qed.
Lemma dfrac_agree_included_L `{!LeibnizEquiv A} d1 a1 d2 a2 :
to_dfrac_agree d1 a1 to_dfrac_agree d2 a2
(d1 d2) a1 = a2.
Proof. unfold_leibniz. apply dfrac_agree_included. Qed.
Lemma dfrac_agree_includedN d1 a1 d2 a2 n :
to_dfrac_agree d1 a1 {n} to_dfrac_agree d2 a2
(d1 d2) a1 {n} a2.
Proof.
by rewrite pair_includedN -cmra_discrete_included_iff
to_agree_includedN.
Qed.
Lemma frac_agree_included q1 a1 q2 a2 :
to_frac_agree q1 a1 to_frac_agree q2 a2
(q1 < q2)%Qp a1 a2.
Proof. by rewrite dfrac_agree_included dfrac_own_included. Qed.
Lemma frac_agree_included_L `{!LeibnizEquiv A} q1 a1 q2 a2 :
to_frac_agree q1 a1 to_frac_agree q2 a2
(q1 < q2)%Qp a1 = a2.
Proof. by rewrite dfrac_agree_included_L dfrac_own_included. Qed.
Lemma frac_agree_includedN q1 a1 q2 a2 n :
to_frac_agree q1 a1 {n} to_frac_agree q2 a2
(q1 < q2)%Qp a1 {n} a2.
Proof. by rewrite dfrac_agree_includedN dfrac_own_included. Qed.
(** While [cmra_update_exclusive] takes care of most updates, it is not sufficient
for this one since there is no abstraction-preserving way to rewrite
[to_dfrac_agree d1 v1 ⋅ to_dfrac_agree d2 v2] into something simpler. *)
Lemma dfrac_agree_update_2 d1 d2 a1 a2 a' :
d1 d2 = DfracOwn 1
to_dfrac_agree d1 a1 to_dfrac_agree d2 a2 ~~>
to_dfrac_agree d1 a' to_dfrac_agree d2 a'.
Proof.
intros Hq. rewrite -pair_op Hq.
apply cmra_update_exclusive.
rewrite dfrac_agree_op_valid Hq //.
Qed.
Lemma frac_agree_update_2 q1 q2 a1 a2 a' :
(q1 + q2 = 1)%Qp
to_frac_agree q1 a1 to_frac_agree q2 a2 ~~>
to_frac_agree q1 a' to_frac_agree q2 a'.
Proof. intros Hq. apply dfrac_agree_update_2. rewrite dfrac_op_own Hq //. Qed.
Lemma dfrac_agree_persist d a :
to_dfrac_agree d a ~~> to_dfrac_agree DfracDiscarded a.
Proof.
rewrite /to_dfrac_agree. apply prod_update; last done.
simpl. apply dfrac_discard_update.
Qed.
Lemma dfrac_agree_unpersist a :
to_dfrac_agree DfracDiscarded a ~~>: λ k, q, k = to_dfrac_agree (DfracOwn q) a.
Proof.
rewrite /to_dfrac_agree. eapply prod_updateP; first apply dfrac_undiscard_update.
{ by eapply cmra_update_updateP. }
naive_solver.
Qed.
End lemmas.
Definition dfrac_agreeRF {SI : sidx} (F : oFunctor) : rFunctor :=
prodRF (constRF dfracR) (agreeRF F).
Global Instance dfrac_agreeRF_contractive {SI : sidx} F :
oFunctorContractive F rFunctorContractive (dfrac_agreeRF F).
Proof. apply _. Qed.
Global Typeclasses Opaque to_dfrac_agree.
(* [to_frac_agree] is deliberately transparent to reuse the [to_dfrac_agree] instances *)
From iris.algebra Require Export auth excl updates.
From iris.algebra Require Import local_updates.
From iris.prelude Require Import options.
(** Authoritative CMRA where the fragment is exclusively owned.
This is effectively a single "ghost variable" with two views, the frament [◯E a]
and the authority [●E a]. *)
Definition excl_authR {SI : sidx} (A : ofe) : cmra :=
authR (optionUR (exclR A)).
Definition excl_authUR {SI : sidx} (A : ofe) : ucmra :=
authUR (optionUR (exclR A)).
Definition excl_auth_auth {SI : sidx} {A : ofe} (a : A) : excl_authR A :=
(Some (Excl a)).
Definition excl_auth_frag {SI : sidx} {A : ofe} (a : A) : excl_authR A :=
(Some (Excl a)).
Global Typeclasses Opaque excl_auth_auth excl_auth_frag.
Global Instance: Params (@excl_auth_auth) 2 := {}.
Global Instance: Params (@excl_auth_frag) 3 := {}.
Notation "●E a" := (excl_auth_auth a) (at level 10).
Notation "◯E a" := (excl_auth_frag a) (at level 10).
Section excl_auth.
Context {SI : sidx} {A : ofe}.
Implicit Types a b : A.
Global Instance excl_auth_auth_ne : NonExpansive (@excl_auth_auth SI A).
Proof. solve_proper. Qed.
Global Instance excl_auth_auth_proper :
Proper (() ==> ()) (@excl_auth_auth SI A).
Proof. solve_proper. Qed.
Global Instance excl_auth_frag_ne : NonExpansive (@excl_auth_frag SI A).
Proof. solve_proper. Qed.
Global Instance excl_auth_frag_proper :
Proper (() ==> ()) (@excl_auth_frag SI A).
Proof. solve_proper. Qed.
Global Instance excl_auth_auth_discrete a : Discrete a Discrete (E a).
Proof. intros; apply auth_auth_discrete; [apply Some_discrete|]; apply _. Qed.
Global Instance excl_auth_frag_discrete a : Discrete a Discrete (E a).
Proof. intros; apply auth_frag_discrete, Some_discrete; apply _. Qed.
Lemma excl_auth_validN n a : {n} (E a E a).
Proof. by rewrite auth_both_validN. Qed.
Lemma excl_auth_valid a : (E a E a).
Proof. intros. by apply auth_both_valid_2. Qed.
Lemma excl_auth_agreeN n a b : {n} (E a E b) a {n} b.
Proof.
rewrite auth_both_validN /= => -[Hincl Hvalid].
move: Hincl=> /Some_includedN_exclusive /(_ I) ?. by apply (inj Excl).
Qed.
Lemma excl_auth_agree a b : (E a E b) a b.
Proof.
intros. apply equiv_dist=> n. by apply excl_auth_agreeN, cmra_valid_validN.
Qed.
Lemma excl_auth_agree_L `{!LeibnizEquiv A} a b : (E a E b) a = b.
Proof. intros. by apply leibniz_equiv, excl_auth_agree. Qed.
Lemma excl_auth_auth_op_validN n a b : {n} (E a E b) False.
Proof. apply auth_auth_op_validN. Qed.
Lemma excl_auth_auth_op_valid a b : (E a E b) False.
Proof. apply auth_auth_op_valid. Qed.
Lemma excl_auth_frag_op_validN n a b : {n} (E a E b) False.
Proof. by rewrite -auth_frag_op auth_frag_validN. Qed.
Lemma excl_auth_frag_op_valid a b : (E a E b) False.
Proof. by rewrite -auth_frag_op auth_frag_valid. Qed.
Lemma excl_auth_update a b a' : E a E b ~~> E a' E a'.
Proof.
intros. by apply auth_update, option_local_update, exclusive_local_update.
Qed.
End excl_auth.
Definition excl_authURF {SI : sidx} (F : oFunctor) : urFunctor :=
authURF (optionURF (exclRF F)).
Global Instance excl_authURF_contractive {SI : sidx} F :
oFunctorContractive F urFunctorContractive (excl_authURF F).
Proof. apply _. Qed.
Definition excl_authRF {SI : sidx} (F : oFunctor) : rFunctor :=
authRF (optionURF (exclRF F)).
Global Instance excl_authRF_contractive {SI : sidx} F :
oFunctorContractive F rFunctorContractive (excl_authRF F).
Proof. apply _. Qed.
From iris.algebra Require Export frac auth updates local_updates.
From iris.algebra Require Import proofmode_classes.
From iris.prelude Require Import options.
(** Authoritative CMRA where the NON-authoritative parts can be fractional.
This CMRA allows the original non-authoritative element [◯ a] to be split into
fractional parts [◯F{q} a]. Using [◯F a ≡ ◯F{1} a] is in effect the same as
using the original [◯ a]. Currently, however, this CMRA hides the ability to
split the authoritative part into fractions.
*)
Definition frac_authR {SI : sidx} (A : cmra) : cmra :=
authR (optionUR (prodR fracR A)).
Definition frac_authUR {SI : sidx} (A : cmra) : ucmra :=
authUR (optionUR (prodR fracR A)).
Definition frac_auth_auth {SI : sidx} {A : cmra} (x : A) : frac_authR A :=
(Some (1%Qp,x)).
Definition frac_auth_frag {SI : sidx}
{A : cmra} (q : frac) (x : A) : frac_authR A :=
(Some (q,x)).
Global Typeclasses Opaque frac_auth_auth frac_auth_frag.
Global Instance: Params (@frac_auth_auth) 2 := {}.
Global Instance: Params (@frac_auth_frag) 3 := {}.
Notation "●F a" := (frac_auth_auth a) (at level 10).
Notation "◯F{ q } a" := (frac_auth_frag q a) (at level 10, format "◯F{ q } a").
Notation "◯F a" := (frac_auth_frag 1 a) (at level 10).
Section frac_auth.
Context {SI : sidx} {A : cmra}.
Implicit Types a b : A.
Global Instance frac_auth_auth_ne : NonExpansive (@frac_auth_auth SI A).
Proof. solve_proper. Qed.
Global Instance frac_auth_auth_proper :
Proper (() ==> ()) (@frac_auth_auth SI A).
Proof. solve_proper. Qed.
Global Instance frac_auth_frag_ne q : NonExpansive (@frac_auth_frag SI A q).
Proof. solve_proper. Qed.
Global Instance frac_auth_frag_proper q :
Proper (() ==> ()) (@frac_auth_frag SI A q).
Proof. solve_proper. Qed.
Global Instance frac_auth_auth_discrete a : Discrete a Discrete (F a).
Proof. intros; apply auth_auth_discrete; [apply Some_discrete|]; apply _. Qed.
Global Instance frac_auth_frag_discrete q a : Discrete a Discrete (F{q} a).
Proof. intros; apply auth_frag_discrete, Some_discrete; apply _. Qed.
Lemma frac_auth_validN n a : {n} a {n} (F a F a).
Proof. by rewrite auth_both_validN. Qed.
Lemma frac_auth_valid a : a (F a F a).
Proof. intros. by apply auth_both_valid_2. Qed.
Lemma frac_auth_agreeN n a b : {n} (F a F b) a {n} b.
Proof.
rewrite auth_both_validN /= => -[Hincl Hvalid].
by move: Hincl=> /Some_includedN_exclusive /(_ Hvalid ) [??].
Qed.
Lemma frac_auth_agree a b : (F a F b) a b.
Proof.
intros. apply equiv_dist=> n. by apply frac_auth_agreeN, cmra_valid_validN.
Qed.
Lemma frac_auth_agree_L `{!LeibnizEquiv A} a b : (F a F b) a = b.
Proof. intros. by apply leibniz_equiv, frac_auth_agree. Qed.
Lemma frac_auth_includedN n q a b : {n} (F a F{q} b) Some b {n} Some a.
Proof. by rewrite auth_both_validN /= => -[/Some_pair_includedN [_ ?] _]. Qed.
Lemma frac_auth_included `{!CmraDiscrete A} q a b :
(F a F{q} b) Some b Some a.
Proof. by rewrite auth_both_valid_discrete /= => -[/Some_pair_included [_ ?] _]. Qed.
Lemma frac_auth_includedN_total `{!CmraTotal A} n q a b :
{n} (F a F{q} b) b {n} a.
Proof. intros. by eapply Some_includedN_total, frac_auth_includedN. Qed.
Lemma frac_auth_included_total `{!CmraDiscrete A, !CmraTotal A} q a b :
(F a F{q} b) b a.
Proof. intros. by eapply Some_included_total, frac_auth_included. Qed.
Lemma frac_auth_auth_validN n a : {n} (F a) {n} a.
Proof.
rewrite auth_auth_dfrac_validN Some_validN. split.
- by intros [?[]].
- intros. by split.
Qed.
Lemma frac_auth_auth_valid a : (F a) a.
Proof. rewrite !cmra_valid_validN. by setoid_rewrite frac_auth_auth_validN. Qed.
Lemma frac_auth_frag_validN n q a : {n} (F{q} a) (q 1)%Qp {n} a.
Proof. by rewrite /frac_auth_frag auth_frag_validN. Qed.
Lemma frac_auth_frag_valid q a : (F{q} a) (q 1)%Qp a.
Proof. by rewrite /frac_auth_frag auth_frag_valid. Qed.
Lemma frac_auth_frag_op q1 q2 a1 a2 : F{q1+q2} (a1 a2) F{q1} a1 F{q2} a2.
Proof. done. Qed.
Lemma frac_auth_frag_op_validN n q1 q2 a b :
{n} (F{q1} a F{q2} b) (q1 + q2 1)%Qp {n} (a b).
Proof. by rewrite -frac_auth_frag_op frac_auth_frag_validN. Qed.
Lemma frac_auth_frag_op_valid q1 q2 a b :
(F{q1} a F{q2} b) (q1 + q2 1)%Qp (a b).
Proof. by rewrite -frac_auth_frag_op frac_auth_frag_valid. Qed.
Global Instance frac_auth_is_op (q q1 q2 : frac) (a a1 a2 : A) :
IsOp q q1 q2 IsOp a a1 a2 IsOp' (F{q} a) (F{q1} a1) (F{q2} a2).
Proof. by rewrite /IsOp' /IsOp=> /leibniz_equiv_iff -> ->. Qed.
Global Instance frac_auth_is_op_core_id (q q1 q2 : frac) (a : A) :
CoreId a IsOp q q1 q2 IsOp' (F{q} a) (F{q1} a) (F{q2} a).
Proof.
rewrite /IsOp' /IsOp=> ? /leibniz_equiv_iff ->.
by rewrite -frac_auth_frag_op -core_id_dup.
Qed.
Lemma frac_auth_update q a b a' b' :
(a,b) ~l~> (a',b') F a F{q} b ~~> F a' F{q} b'.
Proof.
intros. by apply auth_update, option_local_update, prod_local_update_2.
Qed.
Lemma frac_auth_update_1 a b a' : a' F a F b ~~> F a' F a'.
Proof.
intros. by apply auth_update, option_local_update, exclusive_local_update.
Qed.
End frac_auth.
Definition frac_authURF {SI : sidx} (F : rFunctor) : urFunctor :=
authURF (optionURF (prodRF (constRF fracR) F)).
Global Instance frac_authURF_contractive {SI : sidx} F :
rFunctorContractive F urFunctorContractive (frac_authURF F).
Proof. apply _. Qed.
Definition frac_authRF {SI : sidx} (F : rFunctor) : rFunctor :=
authRF (optionURF (prodRF (constRF fracR) F)).
Global Instance frac_authRF_contractive {SI : sidx} F :
rFunctorContractive F rFunctorContractive (frac_authRF F).
Proof. apply _. Qed.
From Coq.QArith Require Import Qcanon.
From iris.algebra Require Export view gmap frac dfrac.
From iris.algebra Require Import local_updates proofmode_classes big_op.
From iris.prelude Require Import options.
(** * CMRA for a "view of a gmap".
The authoritative element [gmap_view_auth] is any [gmap K V]. The fragments
[gmap_view_frag] represent ownership of a single key in that map. Ownership is
governed by a discardable fraction, which provides the possibiltiy to obtain
persistent read-only ownership of a key.
The key frame-preserving updates are [gmap_view_alloc] to allocate a new key,
[gmap_view_update] to update a key given full ownership of the corresponding
fragment, and [gmap_view_persist] to make a key read-only by discarding any
fraction of the corresponding fragment. Crucially, the latter does not require
owning the authoritative element.
NOTE: The API surface for [gmap_view] is experimental and subject to change. We
plan to add notations for authoritative elements and fragments, and hope to
support arbitrary maps as fragments. *)
Local Definition gmap_view_fragUR {SI : sidx}
(K : Type) `{Countable K} (V : cmra) : ucmra :=
gmapUR K (prodR dfracR V).
(** View relation. *)
Section rel.
Context {SI : sidx} (K : Type) `{Countable K} (V : cmra).
Implicit Types (m : gmap K V) (k : K) (v : V) (n : SI).
Implicit Types (f : gmap K (dfrac * V)).
(* If we exactly followed [auth], we'd write something like [f ≼{n} m ∧ ✓{n} m],
which is equivalent to:
[map_Forall (λ k fv, ∃ v, m !! k = Some v ∧ Some fv ≼{n} Some v ∧ ✓{n} v) f].
(Note the use of [Some] in the inclusion; the elementwise RA might not have a
unit and we want a reflexive relation!) However, [f] and [m] do not have the
same type, so this definition does not type-check: the fractions have been
erased from the authoritative [m]. So we additionally quantify over the erased
fraction [dq] and [(dq, v)] becomes the authoritative value.
An alternative definition one might consider is to replace the erased fraction
by a hard-coded [DfracOwn 1], the biggest possible fraction. That would not
work: we would end up with [Some dv ≼{n} Some (DfracOwn 1, v)] but that cannot
be satisfied if [dv.1 = DfracDiscarded], a case that we definitely want to
allow!
It is possible that [∀ k, ∃ dq, let auth := (pair dq) <$> m !! k in ✓{n} auth
∧ f !! k ≼{n} auth] would also work, but now the proofs are all done already. ;)
The two are probably equivalent, with a proof similar to [lookup_includedN]? *)
Local Definition gmap_view_rel_raw n m f :=
map_Forall (λ k fv,
v dq, m !! k = Some v {n} (dq, v) (Some fv {n} Some (dq, v))) f.
Local Lemma gmap_view_rel_raw_mono n1 n2 m1 m2 f1 f2 :
gmap_view_rel_raw n1 m1 f1
m1 {n2} m2
f2 {n2} f1
(n2 n1)%sidx
gmap_view_rel_raw n2 m2 f2.
Proof.
intros Hrel Hm Hf Hn k [dqa va] Hk.
(* For some reason applying the lemma in [Hf] does not work... *)
destruct (lookup_includedN n2 f2 f1) as [Hf' _].
specialize (Hf' Hf k). clear Hf. rewrite Hk in Hf'.
destruct (Some_includedN_is_Some _ _ _ Hf') as [[q' va'] Heq]. rewrite Heq in Hf'.
specialize (Hrel _ _ Heq) as (v & dq & Hm1 & [Hvval Hdqval] & Hvincl). simpl in *.
specialize (Hm k).
edestruct (dist_Some_inv_l _ _ _ _ Hm Hm1) as (v' & Hm2 & Hv).
eexists. exists dq. split; first done. split.
{ split; first done. simpl. rewrite -Hv. eapply cmra_validN_le; done. }
rewrite -Hv. etrans; first exact Hf'.
apply: cmra_includedN_le; done.
Qed.
Local Lemma gmap_view_rel_raw_valid n m f :
gmap_view_rel_raw n m f {n} f.
Proof.
intros Hrel k. destruct (f !! k) as [[dqa va]|] eqn:Hf; rewrite Hf; last done.
specialize (Hrel _ _ Hf) as (v & dq & Hmval & Hvval & Hvincl). simpl in *.
eapply cmra_validN_includedN. 2:done. done.
Qed.
Local Lemma gmap_view_rel_raw_unit n :
m, gmap_view_rel_raw n m ε.
Proof. exists ∅. apply: map_Forall_empty. Qed.
Local Canonical Structure gmap_view_rel :
view_rel (gmapO K V) (gmap_view_fragUR K V) :=
ViewRel gmap_view_rel_raw gmap_view_rel_raw_mono
gmap_view_rel_raw_valid gmap_view_rel_raw_unit.
Local Lemma gmap_view_rel_exists n f :
( m, gmap_view_rel n m f) {n} f.
Proof.
split.
{ intros [m Hrel]. eapply gmap_view_rel_raw_valid, Hrel. }
intros Hf.
cut ( m, gmap_view_rel n m f k, f !! k = None m !! k = None).
{ naive_solver. }
induction f as [|k [dq v] f Hk' IH] using map_ind.
{ exists ∅. split; [|done]. apply: map_Forall_empty. }
move: (Hf k). rewrite lookup_insert=> -[/= ??].
destruct IH as (m & Hm & Hdom).
{ intros k'. destruct (decide (k = k')) as [->|?]; [by rewrite Hk'|].
move: (Hf k'). by rewrite lookup_insert_ne. }
exists (<[k:=v]> m).
rewrite /gmap_view_rel /= /gmap_view_rel_raw map_Forall_insert //=. split_and!.
- exists v, dq. split; first by rewrite lookup_insert.
split; first by split. done.
- eapply map_Forall_impl; [apply Hm|]; simpl.
intros k' [dq' ag'] (v'&?&?&?). exists v'.
rewrite lookup_insert_ne; naive_solver.
- intros k'. rewrite !lookup_insert_None. naive_solver.
Qed.
Local Lemma gmap_view_rel_unit n m : gmap_view_rel n m ε.
Proof. apply: map_Forall_empty. Qed.
Local Lemma gmap_view_rel_discrete :
CmraDiscrete V ViewRelDiscrete gmap_view_rel.
Proof.
intros ? n m f Hrel k [df va] Hk.
destruct (Hrel _ _ Hk) as (v & dq & Hm & Hvval & Hvincl).
exists v, dq. split; first done.
split; first by apply cmra_discrete_valid_iff_0.
rewrite -cmra_discrete_included_iff_0. done.
Qed.
End rel.
Local Existing Instance gmap_view_rel_discrete.
(** [gmap_view] is a notation to give canonical structure search the chance
to infer the right instances (see [auth]). *)
Notation gmap_view K V := (view (gmap_view_rel_raw K V)).
Definition gmap_viewO {SI : sidx} (K : Type) `{Countable K} (V : cmra) : ofe :=
viewO (gmap_view_rel K V).
Definition gmap_viewR {SI : sidx} (K : Type) `{Countable K} (V : cmra) : cmra :=
viewR (gmap_view_rel K V).
Definition gmap_viewUR {SI : sidx} (K : Type) `{Countable K} (V : cmra) : ucmra :=
viewUR (gmap_view_rel K V).
Section definitions.
Context {SI : sidx} `{Countable K} {V : cmra}.
Definition gmap_view_auth (dq : dfrac) (m : gmap K V) : gmap_viewR K V :=
V{dq} m.
Definition gmap_view_frag (k : K) (dq : dfrac) (v : V) : gmap_viewR K V :=
V {[k := (dq, v)]}.
End definitions.
Section lemmas.
Context {SI : sidx} `{Countable K} {V : cmra}.
Implicit Types (m : gmap K V) (k : K) (q : Qp) (dq : dfrac) (v : V).
Global Instance : Params (@gmap_view_auth) 6 := {}.
Global Instance gmap_view_auth_ne dq :
NonExpansive (gmap_view_auth (K:=K) (V:=V) dq).
Proof. solve_proper. Qed.
Global Instance gmap_view_auth_proper dq :
Proper (() ==> ()) (gmap_view_auth (K:=K) (V:=V) dq).
Proof. apply ne_proper, _. Qed.
Global Instance : Params (@gmap_view_frag) 7 := {}.
Global Instance gmap_view_frag_ne k oq : NonExpansive (gmap_view_frag (V:=V) k oq).
Proof. solve_proper. Qed.
Global Instance gmap_view_frag_proper k oq :
Proper (() ==> ()) (gmap_view_frag (V:=V) k oq).
Proof. apply ne_proper, _. Qed.
(* Helper lemmas *)
Local Lemma gmap_view_rel_lookup n m k dq v :
gmap_view_rel K V n m {[k := (dq, v)]}
v' dq', m !! k = Some v' {n} (dq', v') Some (dq, v) {n} Some (dq', v').
Proof.
split.
- intros Hrel.
edestruct (Hrel k) as (v' & dq' & Hlookup & Hval & Hinc).
{ rewrite lookup_singleton. done. }
simpl in *. eexists _, _. split_and!; done.
- intros (v' & dq' & Hlookup & Hval & ?) j [df va].
destruct (decide (k = j)) as [<-|Hne]; last by rewrite lookup_singleton_ne.
rewrite lookup_singleton. intros [= <- <-]. simpl.
exists v', dq'. split_and!; by rewrite ?Hv'.
Qed.
(** Composition and validity *)
Lemma gmap_view_auth_dfrac_op dp dq m :
gmap_view_auth (dp dq) m
gmap_view_auth dp m gmap_view_auth dq m.
Proof. by rewrite /gmap_view_auth view_auth_dfrac_op. Qed.
Global Instance gmap_view_auth_dfrac_is_op dq dq1 dq2 m :
IsOp dq dq1 dq2
IsOp' (gmap_view_auth dq m) (gmap_view_auth dq1 m) (gmap_view_auth dq2 m).
Proof. rewrite /gmap_view_auth. apply _. Qed.
Lemma gmap_view_auth_dfrac_op_invN n dp m1 dq m2 :
{n} (gmap_view_auth dp m1 gmap_view_auth dq m2) m1 {n} m2.
Proof. apply view_auth_dfrac_op_invN. Qed.
Lemma gmap_view_auth_dfrac_op_inv dp m1 dq m2 :
(gmap_view_auth dp m1 gmap_view_auth dq m2) m1 m2.
Proof. apply view_auth_dfrac_op_inv. Qed.
Lemma gmap_view_auth_dfrac_validN m n dq : {n} gmap_view_auth dq m dq.
Proof.
rewrite view_auth_dfrac_validN. intuition eauto using gmap_view_rel_unit.
Qed.
Lemma gmap_view_auth_dfrac_valid m dq : gmap_view_auth dq m dq.
Proof.
rewrite view_auth_dfrac_valid. intuition eauto using gmap_view_rel_unit.
Qed.
Lemma gmap_view_auth_valid m : gmap_view_auth (DfracOwn 1) m.
Proof. rewrite gmap_view_auth_dfrac_valid. done. Qed.
Lemma gmap_view_auth_dfrac_op_validN n dq1 dq2 m1 m2 :
{n} (gmap_view_auth dq1 m1 gmap_view_auth dq2 m2) (dq1 dq2) m1 {n} m2.
Proof.
rewrite view_auth_dfrac_op_validN. intuition eauto using gmap_view_rel_unit.
Qed.
Lemma gmap_view_auth_dfrac_op_valid dq1 dq2 m1 m2 :
(gmap_view_auth dq1 m1 gmap_view_auth dq2 m2) (dq1 dq2) m1 m2.
Proof.
rewrite view_auth_dfrac_op_valid. intuition eauto using gmap_view_rel_unit.
Qed.
Lemma gmap_view_auth_op_validN n m1 m2 :
{n} (gmap_view_auth (DfracOwn 1) m1 gmap_view_auth (DfracOwn 1) m2) False.
Proof. apply view_auth_op_validN. Qed.
Lemma gmap_view_auth_op_valid m1 m2 :
(gmap_view_auth (DfracOwn 1) m1 gmap_view_auth (DfracOwn 1) m2) False.
Proof. apply view_auth_op_valid. Qed.
Lemma gmap_view_frag_validN n k dq v : {n} gmap_view_frag k dq v dq {n} v.
Proof.
rewrite view_frag_validN gmap_view_rel_exists singleton_validN pair_validN.
naive_solver.
Qed.
Lemma gmap_view_frag_valid k dq v : gmap_view_frag k dq v dq v.
Proof.
rewrite cmra_valid_validN. setoid_rewrite gmap_view_frag_validN.
rewrite cmra_valid_validN. pose 0. naive_solver.
Qed.
Lemma gmap_view_frag_op k dq1 dq2 v1 v2 :
gmap_view_frag k (dq1 dq2) (v1 v2)
gmap_view_frag k dq1 v1 gmap_view_frag k dq2 v2.
Proof. rewrite -view_frag_op singleton_op -pair_op //. Qed.
Lemma gmap_view_frag_add k q1 q2 v1 v2 :
gmap_view_frag k (DfracOwn (q1 + q2)) (v1 v2)
gmap_view_frag k (DfracOwn q1) v1 gmap_view_frag k (DfracOwn q2) v2.
Proof. rewrite -gmap_view_frag_op. done. Qed.
Lemma gmap_view_frag_op_validN n k dq1 dq2 v1 v2 :
{n} (gmap_view_frag k dq1 v1 gmap_view_frag k dq2 v2)
(dq1 dq2) {n} (v1 v2).
Proof.
rewrite view_frag_validN gmap_view_rel_exists singleton_op singleton_validN.
by rewrite -pair_op pair_validN.
Qed.
Lemma gmap_view_frag_op_valid k dq1 dq2 v1 v2 :
(gmap_view_frag k dq1 v1 gmap_view_frag k dq2 v2)
(dq1 dq2) (v1 v2).
Proof.
rewrite view_frag_valid. setoid_rewrite gmap_view_rel_exists.
rewrite -cmra_valid_validN singleton_op singleton_valid.
by rewrite -pair_op pair_valid.
Qed.
Lemma gmap_view_both_dfrac_validN n dp m k dq v :
{n} (gmap_view_auth dp m gmap_view_frag k dq v)
v' dq', dp m !! k = Some v' {n} (dq', v')
Some (dq, v) {n} Some (dq', v').
Proof.
rewrite /gmap_view_auth /gmap_view_frag.
rewrite view_both_dfrac_validN gmap_view_rel_lookup. naive_solver.
Qed.
Lemma gmap_view_both_validN n dp m k v :
{n} (gmap_view_auth dp m gmap_view_frag k (DfracOwn 1) v)
dp {n} v m !! k {n} Some v.
Proof.
rewrite gmap_view_both_dfrac_validN. split.
- intros [Hdp (v' & dq' & Hlookup & Hvalid & Hincl)].
split; first done. rewrite Hlookup.
destruct (Some_includedN_exclusive _ _ _ Hincl Hvalid) as [_ Heq].
simpl in Heq. split.
+ rewrite pair_validN in Hvalid. rewrite Heq. naive_solver.
+ by rewrite Heq.
- intros (Hdp & Hval & Hlookup).
apply dist_Some_inv_r' in Hlookup as [v' [Hlookup Heq]].
exists v', (DfracOwn 1). do 2 (split; [done|]). split.
+ rewrite pair_validN. by rewrite -Heq.
+ by apply: Some_includedN_refl.
Qed.
(** The backwards direction here does not hold: if [dq = DfracOwn 1] but
[v ≠ v'], we have to find a suitable erased fraction [dq'] to satisfy the view
relation, but there is no way to satisfy [Some (DfracOwn 1, v) ≼{n} Some (dq', v')]
for any [dq']. The "if and only if" version of this lemma would have to
involve some extra condition like [dq = DfracOwn 1 → v = v'], or phrased
more like the view relation itself: [∃ dq', ✓ dq' ∧ Some (v, dq) ≼{n} Some (v', dq')]. *)
Lemma gmap_view_both_dfrac_validN_total `{!CmraTotal V} n dp m k dq v :
{n} (gmap_view_auth dp m gmap_view_frag k dq v)
v', dp dq m !! k = Some v' {n} v' v {n} v'.
Proof.
rewrite gmap_view_both_dfrac_validN.
intros (v' & dq' & Hdp & Hlookup & Hvalid & Hincl).
exists v'. split; first done. split.
- eapply (cmra_valid_Some_included dq'); first by apply Hvalid.
eapply cmra_discrete_included_iff.
eapply Some_pair_includedN_l. done.
- split; first done. split; first apply Hvalid.
move:Hincl=> /Some_pair_includedN_r /Some_includedN_total. done.
Qed.
(** Without [CmraDiscrete], we cannot do much better than [∀ n, <same as above>].
This is because both the [dq'] and the witness for the [≼{n}] can be different for
each step-index. It is totally possible that at low step-indices, [v] has a frame
(and [dq' > dq]) while at higher step-indices, [v] has no frame (and [dq' = dq]). *)
Lemma gmap_view_both_dfrac_valid_discrete `{!CmraDiscrete V} dp m k dq v :
(gmap_view_auth dp m gmap_view_frag k dq v)
v' dq', dp m !! k = Some v'
(dq', v')
Some (dq, v) Some (dq', v').
Proof.
rewrite cmra_valid_validN. setoid_rewrite gmap_view_both_dfrac_validN. split.
- intros Hvalid. specialize (Hvalid 0).
destruct Hvalid as (v' & dq' & Hdp & Hlookup & Hvalid & Hincl).
exists v', dq'. do 2 (split; first done).
split; first by apply cmra_discrete_valid.
by apply: cmra_discrete_included_r.
- intros (v' & dq' & Hdp & Hlookup & Hvalid & Hincl) n.
exists v', dq'. do 2 (split; first done).
split; first by apply cmra_valid_validN.
by apply: cmra_included_includedN.
Qed.
(** The backwards direction here does not hold: if [dq = DfracOwn 1] but
[v ≠ v'], we have to find a suitable erased fraction [dq'] to satisfy the view
relation, but there is no way to satisfy [Some (DfracOwn 1, v) ≼ Some (dq', v')]
for any [dq']. The "if and only if" version of this lemma would have to
involve some extra condition like [dq = DfracOwn 1 → v = v'], or phrased
more like the view relation itself: [∃ dq', ✓ dq' ∧ Some (v, dq) ≼ Some (v', dq')]. *)
Lemma gmap_view_both_dfrac_valid_discrete_total
`{!CmraDiscrete V, !CmraTotal V} dp m k dq v :
(gmap_view_auth dp m gmap_view_frag k dq v)
v', dp dq m !! k = Some v' v' v v'.
Proof.
rewrite gmap_view_both_dfrac_valid_discrete.
intros (v' & dq' & Hdp & Hlookup & Hvalid & Hincl).
exists v'. split; first done. split.
- eapply (cmra_valid_Some_included dq'); first by apply Hvalid.
eapply Some_pair_included_l. done.
- split; first done. split; first apply Hvalid.
move:Hincl=> /Some_pair_included_r /Some_included_total. done.
Qed.
(** On the other hand, this one holds for all CMRAs, not just discrete ones. *)
Lemma gmap_view_both_valid m dp k v :
(gmap_view_auth dp m gmap_view_frag k (DfracOwn 1) v)
dp v m !! k Some v.
Proof.
rewrite cmra_valid_validN. setoid_rewrite gmap_view_both_validN. split.
- intros Hvalid. split; last split.
+ eapply (Hvalid 0).
+ apply cmra_valid_validN. intros n. eapply Hvalid.
+ apply equiv_dist. intros n. eapply Hvalid.
- intros (Hdp & Hval & Hlookup). intros n.
split; first done. split.
+ apply cmra_valid_validN. done.
+ rewrite Hlookup. done.
Qed.
(** Frame-preserving updates *)
Lemma gmap_view_alloc m k dq v :
m !! k = None
dq
v
gmap_view_auth (DfracOwn 1) m ~~>
gmap_view_auth (DfracOwn 1) (<[k := v]> m) gmap_view_frag k dq v.
Proof.
intros Hfresh Hdq Hval. apply view_update_alloc=>n bf Hrel j [df va] /=.
rewrite lookup_op. destruct (decide (j = k)) as [->|Hne].
- assert (bf !! k = None) as Hbf.
{ destruct (bf !! k) as [[df' va']|] eqn:Hbf; last done.
specialize (Hrel _ _ Hbf). destruct Hrel as (v' & dq' & Hm & _).
exfalso. rewrite Hm in Hfresh. done. }
rewrite lookup_singleton Hbf right_id.
intros [= <- <-]. eexists _, _.
rewrite lookup_insert. split; first done.
split; last by apply: Some_includedN_refl.
split; first done. by eapply cmra_valid_validN.
- rewrite lookup_singleton_ne; last done.
rewrite left_id=>Hbf.
specialize (Hrel _ _ Hbf). destruct Hrel as (v' & ? & Hm & ?).
eexists _, _. split; last done.
rewrite lookup_insert_ne //.
Qed.
Lemma gmap_view_alloc_big m m' dq :
m' ## m
dq
map_Forall (λ k v, v) m'
gmap_view_auth (DfracOwn 1) m ~~>
gmap_view_auth (DfracOwn 1) (m' m)
([^op map] kv m', gmap_view_frag k dq v).
Proof.
intros ?? Hm'.
induction m' as [|k v m' ? IH] using map_ind; decompose_map_disjoint.
{ rewrite big_opM_empty left_id_L right_id. done. }
apply map_Forall_insert in Hm' as [??]; last done.
rewrite IH //. rewrite big_opM_insert // assoc.
apply cmra_update_op; last done.
rewrite -insert_union_l. apply (gmap_view_alloc _ k dq); [|done..].
by apply lookup_union_None.
Qed.
Lemma gmap_view_delete m k v :
gmap_view_auth (DfracOwn 1) m gmap_view_frag k (DfracOwn 1) v ~~>
gmap_view_auth (DfracOwn 1) (delete k m).
Proof.
apply view_update_dealloc=>n bf Hrel j [df va] Hbf /=.
destruct (decide (j = k)) as [->|Hne].
- edestruct (Hrel k) as (v' & dq' & ? & Hval & Hincl).
{ rewrite lookup_op Hbf lookup_singleton -Some_op. done. }
eapply (cmra_validN_Some_includedN _ _ _ Hval) in Hincl as Hval'.
exfalso. clear Hval Hincl.
rewrite pair_validN /= in Hval'.
apply: dfrac_full_exclusive. apply Hval'.
- edestruct (Hrel j) as (v' & ? & ? & ?).
{ rewrite lookup_op lookup_singleton_ne // Hbf. done. }
eexists v', _. split; last done.
rewrite lookup_delete_ne //.
Qed.
Lemma gmap_view_delete_big m m' :
gmap_view_auth (DfracOwn 1) m
([^op map] kv m', gmap_view_frag k (DfracOwn 1) v) ~~>
gmap_view_auth (DfracOwn 1) (m m').
Proof.
induction m' as [|k v m' ? IH] using map_ind.
{ rewrite right_id_L big_opM_empty right_id //. }
rewrite big_opM_insert //.
rewrite [gmap_view_frag _ _ _ _]comm assoc IH gmap_view_delete.
rewrite -delete_difference. done.
Qed.
(** We do not use [local_update] ([~l~>]) in the premise because we also want
to expose the role of the fractions. *)
Lemma gmap_view_update m k dq v mv' v' dq' :
( n mv f,
m !! k = Some mv
{n} ((dq, v) ? f)
mv {n} v ? (snd <$> f)
{n} ((dq', v') ? f) mv' {n} v' ? (snd <$> f))
gmap_view_auth (DfracOwn 1) m gmap_view_frag k dq v ~~>
gmap_view_auth (DfracOwn 1) (<[k := mv']> m) gmap_view_frag k dq' v'.
Proof.
intros Hup. apply view_update=> n bf Hrel j [df va].
rewrite lookup_op.
destruct (decide (j = k)) as [->|Hne]; last first.
{ (* prove that other keys are unaffected *)
simplify_map_eq. rewrite lookup_singleton_ne //.
(* FIXME simplify_map_eq should have done this *)
rewrite left_id. intros Hbf.
edestruct (Hrel j) as (mva & mdf & Hlookup & Hval & Hincl).
{ rewrite lookup_op lookup_singleton_ne // left_id //. }
naive_solver. }
simplify_map_eq. rewrite lookup_singleton.
(* FIXME simplify_map_eq should have done this *)
intros Hbf.
edestruct (Hrel k) as (mv & mdf & Hlookup & Hval & Hincl).
{ rewrite lookup_op lookup_singleton // Some_op_opM //. }
rewrite Some_includedN_opM in Hincl.
destruct Hincl as [f' Hincl]. rewrite cmra_opM_opM_assoc in Hincl.
set f := bf !! k f'. (* the complete frame *)
change (bf !! k f') with f in Hincl.
specialize (Hup n mv f). destruct Hup as (Hval' & Hincl').
{ done. }
{ rewrite -Hincl. done. }
{ destruct Hincl as [_ Hincl]. simpl in Hincl. rewrite Hincl.
by destruct f. }
eexists mv', (dq' ? (fst <$> f)). split; first done.
rewrite -Hbf. clear Hbf. split.
- rewrite Hincl'. destruct Hval'. by destruct f.
- rewrite Some_op_opM. rewrite Some_includedN_opM.
exists f'. rewrite Hincl'.
rewrite cmra_opM_opM_assoc. change (bf !! k f') with f.
by destruct f.
Qed.
(** This derived version cannot exploit [dq = DfracOwn 1]. *)
Lemma gmap_view_update_local m k dq mv v mv' v' :
m !! k = Some mv
(mv, v) ~l~> (mv', v')
gmap_view_auth (DfracOwn 1) m gmap_view_frag k dq v ~~>
gmap_view_auth (DfracOwn 1) (<[k := mv']> m) gmap_view_frag k dq v'.
Proof.
intros Hlookup Hup. apply gmap_view_update.
intros n mv0 f Hmv0 Hval Hincl.
rewrite Hlookup in Hmv0. injection Hmv0 as [= <-].
specialize (Hup n (snd <$> f)). destruct Hup as (Hval' & Hincl').
{ rewrite Hincl. destruct Hval. by destruct f. }
{ simpl. done. }
split; last done. split.
- destruct Hval. by destruct f.
- simpl in *. replace (((dq, v') ? f).2) with (v' ? (snd <$> f)).
2:{ by destruct f. }
rewrite -Hincl'. done.
Qed.
Lemma gmap_view_replace m k v v' :
v'
gmap_view_auth (DfracOwn 1) m gmap_view_frag k (DfracOwn 1) v ~~>
gmap_view_auth (DfracOwn 1) (<[k := v']> m) gmap_view_frag k (DfracOwn 1) v'.
Proof.
(* There would be a simple proof via delete-then-insert... but we use this as a
sanity check to make sure the update lemma is strong enough. *)
intros Hval'. apply gmap_view_update.
intros n mv f Hlookup Hval Hincl.
destruct f; simpl.
{ apply exclusiveN_l in Hval; first done. apply _. }
split; last done.
split; first done. simpl. apply cmra_valid_validN. done.
Qed.
Lemma gmap_view_replace_big m m0 m1 :
dom m0 = dom m1
map_Forall (λ k v, v) m1
gmap_view_auth (DfracOwn 1) m
([^op map] kv m0, gmap_view_frag k (DfracOwn 1) v) ~~>
gmap_view_auth (DfracOwn 1) (m1 m)
([^op map] kv m1, gmap_view_frag k (DfracOwn 1) v).
Proof.
intros Hdom%eq_sym. revert m1 Hdom.
induction m0 as [|k v m0 Hnotdom IH] using map_ind; intros m1 Hdom Hval.
{ rewrite dom_empty_L in Hdom.
apply dom_empty_iff_L in Hdom as ->.
rewrite left_id_L big_opM_empty. done. }
rewrite dom_insert_L in Hdom.
assert (k dom m1) as Hindom by set_solver.
apply elem_of_dom in Hindom as [v' Hlookup].
rewrite big_opM_insert //.
rewrite [gmap_view_frag _ _ _ _]comm assoc.
rewrite (IH (delete k m1)); last first.
{ by apply map_Forall_delete. }
{ rewrite dom_delete_L Hdom.
apply not_elem_of_dom in Hnotdom. set_solver -Hdom. }
rewrite -assoc [_ gmap_view_frag _ _ _]comm assoc.
rewrite (gmap_view_replace _ _ _ v').
2:{ eapply Hval. done. }
rewrite (big_opM_delete _ m1 k v') // -assoc.
rewrite insert_union_r; last by rewrite lookup_delete.
rewrite union_delete_insert //.
Qed.
Lemma gmap_view_auth_persist dq m :
gmap_view_auth dq m ~~> gmap_view_auth DfracDiscarded m.
Proof. apply view_update_auth_persist. Qed.
Lemma gmap_view_auth_unpersist m :
gmap_view_auth DfracDiscarded m ~~>: λ a, q, a = gmap_view_auth (DfracOwn q) m.
Proof. apply view_updateP_auth_unpersist. Qed.
Local Lemma gmap_view_frag_dfrac k dq P v :
dq ~~>: P
gmap_view_frag k dq v ~~>: λ a, dq', a = gmap_view_frag k dq' v P dq'.
Proof.
intros Hdq.
eapply cmra_updateP_weaken;
[apply view_updateP_frag
with (P := λ b', dq', V b' = gmap_view_frag k dq' v P dq')
|naive_solver].
intros m n bf Hrel.
destruct (Hrel k ((dq, v) ? bf !! k)) as (v' & dq' & Hlookup & Hval & Hincl).
{ by rewrite lookup_op lookup_singleton Some_op_opM. }
rewrite Some_includedN_opM in Hincl.
destruct Hincl as [f' Hincl]. rewrite cmra_opM_opM_assoc in Hincl.
set f := bf !! k f'. (* the complete frame *)
change (bf !! k f') with f in Hincl.
destruct (Hdq n (option_map fst f)) as (dq'' & HPdq'' & Hvdq'').
{ destruct Hincl as [Heq _]. simpl in Heq. rewrite Heq in Hval.
destruct Hval as [Hval _]. by destruct f. }
eexists. split; first by exists dq''.
intros j [df va] Heq.
destruct (decide (k = j)) as [->|Hne].
- rewrite lookup_op lookup_singleton in Heq.
eexists v', (dq'' ? (fst <$> f)).
split; first done. split.
+ split; last by apply Hval. simpl. done.
+ rewrite -Heq. exists f'.
rewrite -assoc. change (bf !! j f') with f.
destruct Hincl as [_ Hincl]. simpl in Hincl. rewrite Hincl.
by destruct f.
- rewrite lookup_op lookup_singleton_ne // left_id in Heq.
eapply Hrel. rewrite lookup_op lookup_singleton_ne // left_id Heq //.
Qed.
Lemma gmap_view_frag_persist k dq v :
gmap_view_frag k dq v ~~> gmap_view_frag k DfracDiscarded v.
Proof.
eapply (cmra_update_lift_updateP (λ dq, gmap_view_frag k dq v)).
- intros. by apply gmap_view_frag_dfrac.
- apply dfrac_discard_update.
Qed.
Lemma gmap_view_frag_unpersist k v :
gmap_view_frag k DfracDiscarded v ~~>:
λ a, q, a = gmap_view_frag k (DfracOwn q) v.
Proof.
eapply cmra_updateP_weaken.
{ apply gmap_view_frag_dfrac, dfrac_undiscard_update. }
naive_solver.
Qed.
(** Typeclass instances *)
Global Instance gmap_view_frag_core_id k dq v :
CoreId dq CoreId v CoreId (gmap_view_frag k dq v).
Proof. apply _. Qed.
Global Instance gmap_view_cmra_discrete :
CmraDiscrete V CmraDiscrete (gmap_viewR K V).
Proof. apply _. Qed.
Global Instance gmap_view_frag_mut_is_op dq dq1 dq2 k v v1 v2 :
IsOp dq dq1 dq2
IsOp v v1 v2
IsOp' (gmap_view_frag k dq v) (gmap_view_frag k dq1 v1) (gmap_view_frag k dq2 v2).
Proof. rewrite /IsOp' /IsOp => -> ->. apply gmap_view_frag_op. Qed.
End lemmas.
(** Functor *)
Program Definition gmap_viewURF {SI : sidx} (K : Type) `{Countable K}
(F : rFunctor) : urFunctor := {|
urFunctor_car A _ B _ := gmap_viewUR K (rFunctor_car F A B);
urFunctor_map A1 _ A2 _ B1 _ B2 _ fg :=
viewO_map (rel:=gmap_view_rel K (rFunctor_car F A1 B1))
(rel':=gmap_view_rel K (rFunctor_car F A2 B2))
(gmapO_map (K:=K) (rFunctor_map F fg))
(gmapO_map (K:=K) (prodO_map cid (rFunctor_map F fg)))
|}.
Next Obligation.
intros ? K ?? F A1 ? A2 ? B1 ? B2 ? n f g Hfg.
apply viewO_map_ne.
- apply gmapO_map_ne, rFunctor_map_ne. done.
- apply gmapO_map_ne. apply prodO_map_ne; first done.
apply rFunctor_map_ne. done.
Qed.
Next Obligation.
intros ? K ? ? F A ? B ? m; simpl in *. rewrite -{2}(view_map_id m).
apply (view_map_ext _ _ _ _)=> y.
- rewrite /= -{2}(map_fmap_id y).
apply: map_fmap_equiv_ext=>k ??.
apply rFunctor_map_id.
- rewrite /= -{2}(map_fmap_id y).
apply: map_fmap_equiv_ext=>k [df va] ?.
split; first done. simpl.
apply rFunctor_map_id.
Qed.
Next Obligation.
intros ? K ?? F A1 ? A2 ? A3 ? B1 ? B2 ? B3 ? f g f' g' m; simpl in *.
rewrite -view_map_compose.
apply (view_map_ext _ _ _ _)=> y.
- rewrite /= -map_fmap_compose.
apply: map_fmap_equiv_ext=>k ??.
apply rFunctor_map_compose.
- rewrite /= -map_fmap_compose.
apply: map_fmap_equiv_ext=>k [df va] ?.
split; first done. simpl.
apply rFunctor_map_compose.
Qed.
Next Obligation.
intros ? K ?? F A1 ? A2 ? B1 ? B2 ? fg; simpl.
(* [apply] does not work, probably the usual unification probem (Coq #6294) *)
eapply @view_map_cmra_morphism; [apply _..|]=> n m f.
intros Hrel k [df va] Hf. move: Hf.
rewrite !lookup_fmap.
destruct (f !! k) as [[df' va']|] eqn:Hfk; rewrite Hfk; last done.
simpl=>[= <- <-].
specialize (Hrel _ _ Hfk). simpl in Hrel.
destruct Hrel as (v & dq & Hlookup & Hval & Hincl).
eexists (rFunctor_map F fg v), dq.
rewrite Hlookup. split; first done. split.
- split; first by apply Hval. simpl. apply: cmra_morphism_validN. apply Hval.
- destruct Hincl as [[[fdq fv]|] Hincl].
+ apply: Some_includedN_mono. rewrite -Some_op in Hincl.
apply (inj _) in Hincl. rewrite -pair_op in Hincl.
exists (fdq, rFunctor_map F fg fv). rewrite -pair_op.
split; first apply Hincl. rewrite -cmra_morphism_op.
simpl. f_equiv. apply Hincl.
+ exists None. rewrite right_id in Hincl. apply (inj _) in Hincl.
rewrite right_id. f_equiv. split; first apply Hincl.
simpl. f_equiv. apply Hincl.
Qed.
Global Instance gmap_viewURF_contractive {SI : sidx} `{Countable K} F :
rFunctorContractive F urFunctorContractive (gmap_viewURF K F).
Proof.
intros ? A1 ? A2 ? B1 ? B2 ? n f g Hfg.
apply viewO_map_ne.
- apply gmapO_map_ne. apply rFunctor_map_contractive. done.
- apply gmapO_map_ne. apply prodO_map_ne; first done.
apply rFunctor_map_contractive. done.
Qed.
Program Definition gmap_viewRF {SI : sidx} (K : Type) `{Countable K}
(F : rFunctor) : rFunctor := {|
rFunctor_car A _ B _ := gmap_viewR K (rFunctor_car F A B);
rFunctor_map A1 _ A2 _ B1 _ B2 _ fg :=
viewO_map (rel:=gmap_view_rel K (rFunctor_car F A1 B1))
(rel':=gmap_view_rel K (rFunctor_car F A2 B2))
(gmapO_map (K:=K) (rFunctor_map F fg))
(gmapO_map (K:=K) (prodO_map cid (rFunctor_map F fg)))
|}.
Solve Obligations with apply @gmap_viewURF.
Global Instance gmap_viewRF_contractive {SI : sidx} `{Countable K} F :
rFunctorContractive F rFunctorContractive (gmap_viewRF K F).
Proof. apply gmap_viewURF_contractive. Qed.
Global Typeclasses Opaque gmap_view_auth gmap_view_frag.