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 6580 additions and 118 deletions
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.
(** RA for monotone partial bijections.
This RA is a view where the authoritative element is a partial bijection between
types [A] and [B] and the fragments are subrels of the bijection. The data for
the bijection is represented as a set of pairs [A * B], and the view relation
enforces when an authoritative element is valid it is a bijection (that is, it
is deterministic as a function from [A → option B] and [B → option A]).
The fragments compose by set union, which means that fragments are their own
core, ownership of a fragment is persistent, and the authoritative element can
only grow (in that it can only map more pairs [(a,b)]). *)
(* [algebra.view] needs to be exported for the canonical instances *)
From iris.algebra Require Export view gset.
From iris.algebra Require Import updates.
From iris.prelude Require Import options.
Section gset_bijective.
Context {SI : sidx} `{Countable A, Countable B}.
Implicit Types (a : A) (b : B).
(** [gset_bijective] states that for a graph [L] of [(a, b)] pairs, [L] maps
from [A] to [B] and back deterministically. The key property characterizing
[gset_bijective] is [gset_bijective_eq_iff]. *)
Definition gset_bijective (L : gset (A * B)) :=
a b, (a, b) L
( b', (a, b') L b' = b) ( a', (a', b) L a' = a).
(* Properties of [gset_bijective]. *)
Lemma gset_bijective_empty : gset_bijective ( : gset (A * B)).
Proof. by intros ?? []%not_elem_of_empty. Qed.
(* a bijective graph [L] can be extended with a new mapping [(a,b)] as long as
neither [a] nor [b] is currently mapped to anything. *)
Lemma gset_bijective_extend L a b :
gset_bijective L
( b', (a, b') L)
( a', (a', b) L)
gset_bijective ({[(a, b)]} L).
Proof. rewrite /gset_bijective. set_solver. Qed.
Lemma gset_bijective_eq_iff L (a1 a2 : A) (b1 b2 : B) :
gset_bijective L
(a1, b1) L
(a2, b2) L
a1 = a2 b1 = b2.
Proof. rewrite /gset_bijective. set_solver. Qed.
Lemma gset_bijective_pair a1 b1 a2 b2 :
gset_bijective {[(a1, b1); (a2, b2)]}
(a1 = a2 b1 = b2).
Proof. rewrite /gset_bijective. set_solver. Qed.
Lemma subseteq_gset_bijective L L' :
gset_bijective L
L' L
gset_bijective L'.
Proof. rewrite /gset_bijective. set_solver. Qed.
End gset_bijective.
Section gset_bij_view_rel.
Context {SI : sidx} `{Countable A, Countable B}.
Implicit Types (bijL : gset (A * B)) (L : gsetUR (A * B)).
Local Definition gset_bij_view_rel_raw (n : SI) bijL L: Prop :=
L bijL gset_bijective bijL.
Local Lemma gset_bij_view_rel_raw_mono n1 n2 bijL1 bijL2 L1 L2 :
gset_bij_view_rel_raw n1 bijL1 L1
bijL1 {n2} bijL2
L2 {n2} L1
(n2 n1)%sidx
gset_bij_view_rel_raw n2 bijL2 L2.
Proof.
intros [??] <-%(discrete_iff _ _)%leibniz_equiv ?%gset_included _.
split; [|done]. by trans L1.
Qed.
Local Lemma gset_bij_view_rel_raw_valid n bijL L :
gset_bij_view_rel_raw n bijL L {n}L.
Proof. by intros _. Qed.
Local Lemma gset_bij_view_rel_raw_unit n :
bijL, gset_bij_view_rel_raw n bijL ε.
Proof. exists ∅. split; eauto using gset_bijective_empty. Qed.
Canonical Structure gset_bij_view_rel : view_rel (gsetO (A * B)) (gsetUR (A * B)) :=
ViewRel gset_bij_view_rel_raw gset_bij_view_rel_raw_mono
gset_bij_view_rel_raw_valid gset_bij_view_rel_raw_unit.
Global Instance gset_bij_view_rel_discrete : ViewRelDiscrete gset_bij_view_rel.
Proof. intros n bijL L [??]. split; auto. Qed.
Local Lemma gset_bij_view_rel_iff n bijL L :
gset_bij_view_rel n bijL L L bijL gset_bijective bijL.
Proof. done. Qed.
End gset_bij_view_rel.
Definition gset_bij {SI : sidx} A B `{Countable A, Countable B} :=
view (gset_bij_view_rel_raw (A:=A) (B:=B)).
Definition gset_bijO {SI : sidx} A B `{Countable A, Countable B} : ofe :=
viewO (gset_bij_view_rel (A:=A) (B:=B)).
Definition gset_bijR {SI : sidx} A B `{Countable A, Countable B} : cmra :=
viewR (gset_bij_view_rel (A:=A) (B:=B)).
Definition gset_bijUR {SI : sidx} A B `{Countable A, Countable B} : ucmra :=
viewUR (gset_bij_view_rel (A:=A) (B:=B)).
Definition gset_bij_auth {SI : sidx} `{Countable A, Countable B}
(dq : dfrac) (L : gset (A * B)) : gset_bij A B := V{dq} L V L.
Definition gset_bij_elem {SI : sidx} `{Countable A, Countable B}
(a : A) (b : B) : gset_bij A B := V {[ (a, b) ]}.
Section gset_bij.
Context {SI : sidx} `{Countable A, Countable B}.
Implicit Types (a : A) (b : B).
Implicit Types (L : gset (A*B)).
Implicit Types dq : dfrac.
Global Instance gset_bij_elem_core_id a b : CoreId (gset_bij_elem a b).
Proof. apply _. Qed.
Lemma gset_bij_auth_dfrac_op dq1 dq2 L :
gset_bij_auth dq1 L gset_bij_auth dq2 L gset_bij_auth (dq1 dq2) L.
Proof.
rewrite /gset_bij_auth view_auth_dfrac_op.
rewrite (comm _ (V{dq2} _)) -!assoc (assoc _ (V _)).
by rewrite -core_id_dup (comm _ (V _)).
Qed.
Lemma gset_bij_auth_dfrac_valid dq L : gset_bij_auth dq L dq gset_bijective L.
Proof.
rewrite /gset_bij_auth view_both_dfrac_valid.
setoid_rewrite gset_bij_view_rel_iff. pose 0. naive_solver.
Qed.
Lemma gset_bij_auth_valid L : gset_bij_auth (DfracOwn 1) L gset_bijective L.
Proof. rewrite gset_bij_auth_dfrac_valid. naive_solver by done. Qed.
Lemma gset_bij_auth_empty_dfrac_valid dq : gset_bij_auth (A:=A) (B:=B) dq dq.
Proof.
rewrite gset_bij_auth_dfrac_valid. naive_solver eauto using gset_bijective_empty.
Qed.
Lemma gset_bij_auth_empty_valid : gset_bij_auth (A:=A) (B:=B) (DfracOwn 1) ∅.
Proof. by apply gset_bij_auth_empty_dfrac_valid. Qed.
Lemma gset_bij_auth_dfrac_op_valid dq1 dq2 L1 L2 :
(gset_bij_auth dq1 L1 gset_bij_auth dq2 L2)
(dq1 dq2) L1 = L2 gset_bijective L1.
Proof.
rewrite /gset_bij_auth (comm _ (V{dq2} _)) -!assoc (assoc _ (V _)).
rewrite -view_frag_op (comm _ (V _)) assoc. split.
- move=> /cmra_valid_op_l /view_auth_dfrac_op_valid.
setoid_rewrite gset_bij_view_rel_iff. pose 0. naive_solver.
- intros (?&->&?). rewrite -core_id_dup -view_auth_dfrac_op.
apply view_both_dfrac_valid. setoid_rewrite gset_bij_view_rel_iff. naive_solver.
Qed.
Lemma gset_bij_auth_op_valid L1 L2 :
(gset_bij_auth (DfracOwn 1) L1 gset_bij_auth (DfracOwn 1) L2) False.
Proof. rewrite gset_bij_auth_dfrac_op_valid. naive_solver. Qed.
Lemma bij_both_dfrac_valid dq L a b :
(gset_bij_auth dq L gset_bij_elem a b) dq gset_bijective L (a, b) L.
Proof.
rewrite /gset_bij_auth /gset_bij_elem -assoc -view_frag_op view_both_dfrac_valid.
setoid_rewrite gset_bij_view_rel_iff. pose 0. set_solver.
Qed.
Lemma bij_both_valid L a b :
(gset_bij_auth (DfracOwn 1) L gset_bij_elem a b) gset_bijective L (a, b) L.
Proof. rewrite bij_both_dfrac_valid. naive_solver by done. Qed.
Lemma gset_bij_elem_agree a1 b1 a2 b2 :
(gset_bij_elem a1 b1 gset_bij_elem a2 b2) (a1 = a2 b1 = b2).
Proof.
rewrite /gset_bij_elem -view_frag_op gset_op view_frag_valid.
setoid_rewrite gset_bij_view_rel_iff. intros. apply gset_bijective_pair.
pose 0. naive_solver eauto using subseteq_gset_bijective.
Qed.
Lemma bij_view_included dq L a b :
(a,b) L gset_bij_elem a b gset_bij_auth dq L.
Proof.
intros. etrans; [|apply cmra_included_r].
apply view_frag_mono, gset_included. set_solver.
Qed.
Lemma gset_bij_auth_extend {L} a b :
( b', (a, b') L) ( a', (a', b) L)
gset_bij_auth (DfracOwn 1) L ~~> gset_bij_auth (DfracOwn 1) ({[(a, b)]} L).
Proof.
intros. apply view_update=> n bijL.
rewrite !gset_bij_view_rel_iff gset_op.
set_solver by eauto using gset_bijective_extend.
Qed.
End gset_bij.
From iris.algebra Require Export auth.
From iris.algebra Require Import numbers updates.
From iris.prelude Require Import options.
(** Authoritative CMRA over [max_Z]. The authoritative element is a
monotonically increasing [Z], while a fragment is a lower bound. *)
Definition mono_Z {SI: sidx} := auth (option max_ZR).
Definition mono_ZR {SI: sidx} := authR (optionUR max_ZR).
Definition mono_ZUR {SI: sidx} := authUR (optionUR max_ZR).
(** [mono_Z_auth] is the authoritative element. The definition includes the
fragment at the same value so that lemma [mono_Z_included], which states that
[mono_Z_lb n ≼ mono_Z_auth dq n], holds. Without this trick, a frame-preserving
update lemma would be required instead. *)
Definition mono_Z_auth {SI : sidx} (dq : dfrac) (n : Z) : mono_Z :=
{dq} (Some (MaxZ n)) (Some (MaxZ n)).
Definition mono_Z_lb {SI : sidx} (n : Z) : mono_Z := (Some (MaxZ n)).
Notation "●MZ dq a" := (mono_Z_auth dq a)
(at level 20, dq custom dfrac at level 1, format "●MZ dq a").
Notation "◯MZ a" := (mono_Z_lb a) (at level 20).
Section mono_Z.
Context {SI : sidx}.
Implicit Types (n : Z).
Local Open Scope Z_scope.
Global Instance mono_Z_lb_core_id n : CoreId (MZ n).
Proof. apply _. Qed.
Global Instance mono_Z_auth_core_id l : CoreId (MZ l).
Proof. apply _. Qed.
Lemma mono_Z_auth_dfrac_op dq1 dq2 n :
MZ{dq1 dq2} n MZ{dq1} n MZ{dq2} n.
Proof.
rewrite /mono_Z_auth auth_auth_dfrac_op.
rewrite (comm _ ({dq2} _)) -!assoc (assoc _ ( _)).
by rewrite -core_id_dup (comm _ ( _)).
Qed.
Lemma mono_Z_lb_op n1 n2 :
MZ (n1 `max` n2) = MZ n1 MZ n2.
Proof. rewrite -auth_frag_op -Some_op max_Z_op //. Qed.
Lemma mono_Z_auth_lb_op dq n :
MZ{dq} n MZ{dq} n MZ n.
Proof.
rewrite /mono_Z_auth /mono_Z_lb.
rewrite -!assoc -auth_frag_op -Some_op max_Z_op.
rewrite Z.max_id //.
Qed.
Global Instance mono_Z_auth_dfrac_is_op dq dq1 dq2 n :
IsOp dq dq1 dq2 IsOp' (MZ{dq} n) (MZ{dq1} n) (MZ{dq2} n).
Proof. rewrite /IsOp' /IsOp=> ->. rewrite mono_Z_auth_dfrac_op //. Qed.
Global Instance mono_Z_lb_max_is_op n n1 n2 :
IsOp (MaxZ n) (MaxZ n1) (MaxZ n2) IsOp' (MZ n) (MZ n1) (MZ n2).
Proof. rewrite /IsOp' /IsOp /mono_Z_lb=> ->. done. Qed.
(** rephrasing of [mono_Z_lb_op] useful for weakening a fragment to a
smaller lower-bound *)
Lemma mono_Z_lb_op_le_l n n' :
n' n
MZ n = MZ n' MZ n.
Proof. intros. rewrite -mono_Z_lb_op Z.max_r //. Qed.
Lemma mono_Z_auth_dfrac_valid dq n : ( MZ{dq} n) dq.
Proof.
rewrite /mono_Z_auth auth_both_dfrac_valid_discrete /=. naive_solver.
Qed.
Lemma mono_Z_auth_valid n : MZ n.
Proof. by apply auth_both_valid. Qed.
Lemma mono_Z_auth_dfrac_op_valid dq1 dq2 n1 n2 :
(MZ{dq1} n1 MZ{dq2} n2) (dq1 dq2) n1 = n2.
Proof.
rewrite /mono_Z_auth (comm _ ({dq2} _)) -!assoc (assoc _ ( _)).
rewrite -auth_frag_op (comm _ ( _)) assoc. split.
- move=> /cmra_valid_op_l /auth_auth_dfrac_op_valid. naive_solver.
- intros [? ->]. rewrite -core_id_dup -auth_auth_dfrac_op.
by apply auth_both_dfrac_valid_discrete.
Qed.
Lemma mono_Z_auth_op_valid n1 n2 :
(MZ n1 MZ n2) False.
Proof. rewrite mono_Z_auth_dfrac_op_valid. naive_solver. Qed.
Lemma mono_Z_both_dfrac_valid dq n m :
(MZ{dq} n MZ m) dq m n.
Proof.
rewrite /mono_Z_auth /mono_Z_lb -assoc -auth_frag_op.
rewrite auth_both_dfrac_valid_discrete Some_included_total max_Z_included /=.
naive_solver lia.
Qed.
Lemma mono_Z_both_valid n m :
(MZ n MZ m) m n.
Proof. rewrite mono_Z_both_dfrac_valid dfrac_valid_own. naive_solver. Qed.
Lemma mono_Z_lb_mono n1 n2 : n1 n2 MZ n1 MZ n2.
Proof.
intros. by apply auth_frag_mono, Some_included_total, max_Z_included.
Qed.
Lemma mono_Z_included dq n : MZ n MZ{dq} n.
Proof. apply: cmra_included_r. Qed.
Lemma mono_Z_update {n} n' :
n n' MZ n ~~> MZ n'.
Proof.
intros. rewrite /mono_Z_auth /mono_Z_lb.
by apply auth_update, option_local_update, max_Z_local_update.
Qed.
Lemma mono_Z_auth_persist n dq :
MZ{dq} n ~~> MZ n.
Proof.
intros. rewrite /mono_Z_auth /mono_Z_lb.
eapply cmra_update_op_proper; last done.
eapply auth_update_auth_persist.
Qed.
Lemma mono_Z_auth_unpersist n :
MZ n ~~>: λ k, q, k = MZ{# q} n.
Proof. eapply auth_updateP_both_unpersist. Qed.
End mono_Z.
Global Typeclasses Opaque mono_Z_auth mono_Z_lb.
(** Authoritative CMRA of append-only lists, where the fragment represents a
snap-shot of the list, and the authoritative element can only grow by
appending. *)
From iris.algebra Require Export auth dfrac max_prefix_list.
From iris.algebra Require Import updates local_updates proofmode_classes.
From iris.prelude Require Import options.
Definition mono_listR {SI : sidx} (A : ofe) : cmra := authR (max_prefix_listUR A).
Definition mono_listUR {SI : sidx} (A : ofe) : ucmra := authUR (max_prefix_listUR A).
Definition mono_list_auth {SI : sidx}
{A : ofe} (q : dfrac) (l : list A) : mono_listR A :=
{q} (to_max_prefix_list l) (to_max_prefix_list l).
Definition mono_list_lb {SI : sidx} {A : ofe} (l : list A) : mono_listR A :=
(to_max_prefix_list l).
Global Instance: Params (@mono_list_auth) 3 := {}.
Global Instance: Params (@mono_list_lb) 2 := {}.
Global Typeclasses Opaque mono_list_auth mono_list_lb.
Notation "●ML dq l" := (mono_list_auth dq l)
(at level 20, dq custom dfrac at level 1, format "●ML dq l").
Notation "◯ML l" := (mono_list_lb l) (at level 20).
Section mono_list_props.
Context {SI : sidx} {A : ofe}.
Implicit Types l : list A.
Implicit Types q : frac.
Implicit Types dq : dfrac.
(** Setoid properties *)
Global Instance mono_list_auth_ne dq : NonExpansive (@mono_list_auth SI A dq).
Proof. solve_proper. Qed.
Global Instance mono_list_auth_proper dq :
Proper (() ==> ()) (@mono_list_auth SI A dq).
Proof. solve_proper. Qed.
Global Instance mono_list_lb_ne : NonExpansive (@mono_list_lb SI A).
Proof. solve_proper. Qed.
Global Instance mono_list_lb_proper : Proper (() ==> ()) (@mono_list_lb SI A).
Proof. solve_proper. Qed.
Global Instance mono_list_lb_dist_inj n :
Inj (dist n) (dist n) (@mono_list_lb SI A).
Proof. rewrite /mono_list_lb. by intros ?? ?%(inj _)%(inj _). Qed.
Global Instance mono_list_lb_inj : Inj () () (@mono_list_lb SI A).
Proof. rewrite /mono_list_lb. by intros ?? ?%(inj _)%(inj _). Qed.
(** * Operation *)
Global Instance mono_list_lb_core_id l : CoreId (ML l).
Proof. rewrite /mono_list_lb. apply _. Qed.
Global Instance mono_list_auth_core_id l : CoreId (ML l).
Proof. rewrite /mono_list_auth. apply _. Qed.
Lemma mono_list_auth_dfrac_op dq1 dq2 l :
ML{dq1 dq2} l ML{dq1} l ML{dq2} l.
Proof.
rewrite /mono_list_auth auth_auth_dfrac_op.
rewrite (comm _ ({dq2} _)) -!assoc (assoc _ ( _)).
by rewrite -core_id_dup (comm _ ( _)).
Qed.
Lemma mono_list_lb_op_l l1 l2 : l1 `prefix_of` l2 ML l1 ML l2 ML l2.
Proof. intros ?. by rewrite /mono_list_lb -auth_frag_op to_max_prefix_list_op_l. Qed.
Lemma mono_list_lb_op_r l1 l2 : l1 `prefix_of` l2 ML l2 ML l1 ML l2.
Proof. intros ?. by rewrite /mono_list_lb -auth_frag_op to_max_prefix_list_op_r. Qed.
Lemma mono_list_auth_lb_op dq l : ML{dq} l ML{dq} l ML l.
Proof.
by rewrite /mono_list_auth /mono_list_lb -!assoc -auth_frag_op -core_id_dup.
Qed.
Global Instance mono_list_auth_dfrac_is_op dq dq1 dq2 l :
IsOp dq dq1 dq2 IsOp' (ML{dq} l) (ML{dq1} l) (ML{dq2} l).
Proof. rewrite /IsOp' /IsOp=> ->. rewrite mono_list_auth_dfrac_op //. Qed.
(** * Validity *)
Lemma mono_list_auth_dfrac_validN n dq l : {n} (ML{dq} l) dq.
Proof.
rewrite /mono_list_auth auth_both_dfrac_validN.
naive_solver apply to_max_prefix_list_validN.
Qed.
Lemma mono_list_auth_validN n l : {n} (ML l).
Proof. by apply mono_list_auth_dfrac_validN. Qed.
Lemma mono_list_auth_dfrac_valid dq l : (ML{dq} l) dq.
Proof.
rewrite /mono_list_auth auth_both_dfrac_valid.
naive_solver apply to_max_prefix_list_valid.
Qed.
Lemma mono_list_auth_valid l : (ML l).
Proof. by apply mono_list_auth_dfrac_valid. Qed.
Lemma mono_list_auth_dfrac_op_validN n dq1 dq2 l1 l2 :
{n} (ML{dq1} l1 ML{dq2} l2) (dq1 dq2) l1 {n} l2.
Proof.
rewrite /mono_list_auth (comm _ ({dq2} _)) -!assoc (assoc _ ( _)).
rewrite -auth_frag_op (comm _ ( _)) assoc. split.
- move=> /cmra_validN_op_l /auth_auth_dfrac_op_validN.
rewrite (inj_iff to_max_prefix_list). naive_solver.
- intros [? ->]. rewrite -core_id_dup -auth_auth_dfrac_op auth_both_dfrac_validN.
naive_solver apply to_max_prefix_list_validN.
Qed.
Lemma mono_list_auth_op_validN n l1 l2 : {n} (ML l1 ML l2) False.
Proof. rewrite mono_list_auth_dfrac_op_validN. naive_solver. Qed.
Lemma mono_list_auth_dfrac_op_valid dq1 dq2 l1 l2 :
(ML{dq1} l1 ML{dq2} l2) (dq1 dq2) l1 l2.
Proof.
rewrite cmra_valid_validN equiv_dist.
setoid_rewrite mono_list_auth_dfrac_op_validN. pose 0. naive_solver.
Qed.
Lemma mono_list_auth_op_valid l1 l2 : (ML l1 ML l2) False.
Proof. rewrite mono_list_auth_dfrac_op_valid. naive_solver. Qed.
Lemma mono_list_auth_dfrac_op_valid_L `{!LeibnizEquiv A} dq1 dq2 l1 l2 :
(ML{dq1} l1 ML{dq2} l2) (dq1 dq2) l1 = l2.
Proof. unfold_leibniz. apply mono_list_auth_dfrac_op_valid. Qed.
Lemma mono_list_both_dfrac_validN n dq l1 l2 :
{n} (ML{dq} l1 ML l2) dq l, l1 {n} l2 ++ l.
Proof.
rewrite /mono_list_auth /mono_list_lb -assoc
-auth_frag_op auth_both_dfrac_validN -to_max_prefix_list_includedN.
f_equiv; split.
- intros [Hincl _]. etrans; [apply: cmra_includedN_r|done].
- intros. split; [|by apply to_max_prefix_list_validN].
rewrite {2}(core_id_dup (to_max_prefix_list l1)). by f_equiv.
Qed.
Lemma mono_list_both_validN n l1 l2 :
{n} (ML l1 ML l2) l, l1 {n} l2 ++ l.
Proof. rewrite mono_list_both_dfrac_validN. split; [naive_solver|done]. Qed.
Lemma mono_list_both_dfrac_valid dq l1 l2 :
(ML{dq} l1 ML l2) dq l, l1 l2 ++ l.
Proof.
rewrite /mono_list_auth /mono_list_lb -assoc -auth_frag_op
auth_both_dfrac_valid -max_prefix_list_included_includedN
-to_max_prefix_list_included.
f_equiv; split.
- intros [Hincl _]. etrans; [apply: cmra_included_r|done].
- intros. split; [|by apply to_max_prefix_list_valid].
rewrite {2}(core_id_dup (to_max_prefix_list l1)). by f_equiv.
Qed.
Lemma mono_list_both_valid l1 l2 :
(ML l1 ML l2) l, l1 l2 ++ l.
Proof. rewrite mono_list_both_dfrac_valid. split; [naive_solver|done]. Qed.
Lemma mono_list_both_dfrac_valid_L `{!LeibnizEquiv A} dq l1 l2 :
(ML{dq} l1 ML l2) dq l2 `prefix_of` l1.
Proof. rewrite /prefix. rewrite mono_list_both_dfrac_valid. naive_solver. Qed.
Lemma mono_list_both_valid_L `{!LeibnizEquiv A} l1 l2 :
(ML l1 ML l2) l2 `prefix_of` l1.
Proof. rewrite /prefix. rewrite mono_list_both_valid. naive_solver. Qed.
Lemma mono_list_lb_op_validN n l1 l2 :
{n} (ML l1 ML l2) ( l, l2 {n} l1 ++ l) ( l, l1 {n} l2 ++ l).
Proof. by rewrite auth_frag_op_validN to_max_prefix_list_op_validN. Qed.
Lemma mono_list_lb_op_valid l1 l2 :
(ML l1 ML l2) ( l, l2 l1 ++ l) ( l, l1 l2 ++ l).
Proof. by rewrite auth_frag_op_valid to_max_prefix_list_op_valid. Qed.
Lemma mono_list_lb_op_valid_L `{!LeibnizEquiv A} l1 l2 :
(ML l1 ML l2) l1 `prefix_of` l2 l2 `prefix_of` l1.
Proof. rewrite mono_list_lb_op_valid / prefix. naive_solver. Qed.
Lemma mono_list_lb_op_valid_1_L `{!LeibnizEquiv A} l1 l2 :
(ML l1 ML l2) l1 `prefix_of` l2 l2 `prefix_of` l1.
Proof. by apply mono_list_lb_op_valid_L. Qed.
Lemma mono_list_lb_op_valid_2_L `{!LeibnizEquiv A} l1 l2 :
l1 `prefix_of` l2 l2 `prefix_of` l1 (ML l1 ML l2).
Proof. by apply mono_list_lb_op_valid_L. Qed.
Lemma mono_list_lb_mono l1 l2 : l1 `prefix_of` l2 ML l1 ML l2.
Proof. intros. exists (ML l2). by rewrite mono_list_lb_op_l. Qed.
Lemma mono_list_included dq l : ML l ML{dq} l.
Proof. apply cmra_included_r. Qed.
(** * Update *)
Lemma mono_list_update {l1} l2 : l1 `prefix_of` l2 ML l1 ~~> ML l2.
Proof. intros ?. by apply auth_update, max_prefix_list_local_update. Qed.
Lemma mono_list_auth_persist dq l : ML{dq} l ~~> ML l.
Proof.
rewrite /mono_list_auth. apply cmra_update_op; [|done].
by apply auth_update_auth_persist.
Qed.
Lemma mono_list_auth_unpersist l :
ML l ~~>: λ k, q, k = ML{#q} l.
Proof. eapply auth_updateP_both_unpersist. Qed.
End mono_list_props.
Definition mono_listURF {SI : sidx} (F : oFunctor) : urFunctor :=
authURF (max_prefix_listURF F).
Global Instance mono_listURF_contractive {SI : sidx} F :
oFunctorContractive F urFunctorContractive (mono_listURF F).
Proof. apply _. Qed.
Definition mono_listRF {SI : sidx} (F : oFunctor) : rFunctor :=
authRF (max_prefix_listURF F).
Global Instance mono_listRF_contractive {SI : sidx} F :
oFunctorContractive F rFunctorContractive (mono_listRF F).
Proof. apply _. Qed.
From iris.algebra Require Export auth.
From iris.algebra Require Import numbers updates.
From iris.prelude Require Import options.
(** Authoritative CMRA over [max_nat]. The authoritative element is a
monotonically increasing [nat], while a fragment is a lower bound. *)
Definition mono_nat {SI : sidx} := auth max_natUR.
Definition mono_natR {SI : sidx} := authR max_natUR.
Definition mono_natUR {SI : sidx} := authUR max_natUR.
(** [mono_nat_auth] is the authoritative element. The definition includes the
fragment at the same value so that lemma [mono_nat_included], which states that
[mono_nat_lb n ≼ mono_nat_auth dq n], holds. Without this trick, a
frame-preserving update lemma would be required instead. *)
Definition mono_nat_auth {SI : sidx} (dq : dfrac) (n : nat) : mono_nat :=
{dq} MaxNat n MaxNat n.
Definition mono_nat_lb {SI : sidx} (n : nat) : mono_nat := MaxNat n.
Notation "●MN dq a" := (mono_nat_auth dq a)
(at level 20, dq custom dfrac at level 1, format "●MN dq a").
Notation "◯MN a" := (mono_nat_lb a) (at level 20).
Section mono_nat.
Context {SI : sidx}.
Implicit Types (n : nat).
Global Instance mono_nat_lb_core_id n : CoreId (MN n).
Proof. apply _. Qed.
Global Instance mono_nat_auth_core_id l : CoreId (MN l).
Proof. apply _. Qed.
Lemma mono_nat_auth_dfrac_op dq1 dq2 n :
MN{dq1 dq2} n MN{dq1} n MN{dq2} n.
Proof.
rewrite /mono_nat_auth auth_auth_dfrac_op.
rewrite (comm _ ({dq2} _)) -!assoc (assoc _ ( _)).
by rewrite -core_id_dup (comm _ ( _)).
Qed.
Lemma mono_nat_lb_op n1 n2 :
MN (n1 `max` n2) = MN n1 MN n2.
Proof. rewrite -auth_frag_op max_nat_op //. Qed.
Lemma mono_nat_auth_lb_op dq n :
MN{dq} n MN{dq} n MN n.
Proof.
rewrite /mono_nat_auth /mono_nat_lb.
rewrite -!assoc -auth_frag_op max_nat_op.
rewrite Nat.max_id //.
Qed.
Global Instance mono_nat_auth_dfrac_is_op dq dq1 dq2 n :
IsOp dq dq1 dq2 IsOp' (MN{dq} n) (MN{dq1} n) (MN{dq2} n).
Proof. rewrite /IsOp' /IsOp=> ->. rewrite mono_nat_auth_dfrac_op //. Qed.
Global Instance mono_nat_lb_max_is_op n n1 n2 :
IsOp (MaxNat n) (MaxNat n1) (MaxNat n2) IsOp' (MN n) (MN n1) (MN n2).
Proof. rewrite /IsOp' /IsOp /mono_nat_lb=> ->. done. Qed.
(** rephrasing of [mono_nat_lb_op] useful for weakening a fragment to a
smaller lower-bound *)
Lemma mono_nat_lb_op_le_l n n' :
n' n
MN n = MN n' MN n.
Proof. intros. rewrite -mono_nat_lb_op Nat.max_r //. Qed.
Lemma mono_nat_auth_dfrac_valid dq n : ( MN{dq} n) dq.
Proof.
rewrite /mono_nat_auth auth_both_dfrac_valid_discrete /=. naive_solver.
Qed.
Lemma mono_nat_auth_valid n : MN n.
Proof. by apply auth_both_valid. Qed.
Lemma mono_nat_auth_dfrac_op_valid dq1 dq2 n1 n2 :
(MN{dq1} n1 MN{dq2} n2) (dq1 dq2) n1 = n2.
Proof.
rewrite /mono_nat_auth (comm _ ({dq2} _)) -!assoc (assoc _ ( _)).
rewrite -auth_frag_op (comm _ ( _)) assoc. split.
- move=> /cmra_valid_op_l /auth_auth_dfrac_op_valid. naive_solver.
- intros [? ->]. rewrite -core_id_dup -auth_auth_dfrac_op.
by apply auth_both_dfrac_valid_discrete.
Qed.
Lemma mono_nat_auth_op_valid n1 n2 :
(MN n1 MN n2) False.
Proof. rewrite mono_nat_auth_dfrac_op_valid. naive_solver. Qed.
Lemma mono_nat_both_dfrac_valid dq n m :
(MN{dq} n MN m) dq m n.
Proof.
rewrite /mono_nat_auth /mono_nat_lb -assoc -auth_frag_op.
rewrite auth_both_dfrac_valid_discrete max_nat_included /=.
naive_solver lia.
Qed.
Lemma mono_nat_both_valid n m :
(MN n MN m) m n.
Proof. rewrite mono_nat_both_dfrac_valid dfrac_valid_own. naive_solver. Qed.
Lemma mono_nat_lb_mono n1 n2 : n1 n2 MN n1 MN n2.
Proof. intros. by apply auth_frag_mono, max_nat_included. Qed.
Lemma mono_nat_included dq n : MN n MN{dq} n.
Proof. apply cmra_included_r. Qed.
Lemma mono_nat_update {n} n' :
n n' MN n ~~> MN n'.
Proof.
intros. rewrite /mono_nat_auth /mono_nat_lb.
by apply auth_update, max_nat_local_update.
Qed.
Lemma mono_nat_auth_persist n dq :
MN{dq} n ~~> MN n.
Proof.
intros. rewrite /mono_nat_auth /mono_nat_lb.
eapply cmra_update_op_proper; last done.
eapply auth_update_auth_persist.
Qed.
Lemma mono_nat_auth_unpersist n :
MN n ~~>: λ k, q, k = MN{# q} n.
Proof. eapply auth_updateP_both_unpersist. Qed.
End mono_nat.
Global Typeclasses Opaque mono_nat_auth mono_nat_lb.
......@@ -9,133 +9,186 @@ difference:
if we have the authoritative element we can always increase its fraction
and allocate a new fragment.
<<<
✓ (a ⋅ b) → ●U{p} a ~~> ●U{p + q} (a ⋅ b) ⋅ ◯U{q} b
>>>
<<
✓ (a ⋅ b) → ●U_p a ~~> ●U_(p + q) (a ⋅ b) ⋅ ◯U_q b
>>
- We no longer have the [◯U{1} a] is the exclusive fragmental element (cf.
[frac_auth_frag_validN_op_1_l]).
- We no longer have the [◯U_1 a] is an exclusive fragmental element. That is,
while [◯F{1} a ⋅ ◯F{q} b] is vacuously false, [◯U_1 a ⋅ ◯U_q2 b] is not.
*)
From Coq Require Import QArith Qcanon.
From iris.algebra Require Export auth frac updates local_updates.
From iris.algebra Require Import ufrac proofmode_classes.
From iris.prelude Require Import options.
Definition ufrac_authR (A : cmraT) : cmraT :=
Definition ufrac_authR {SI : sidx} (A : cmra) : cmra :=
authR (optionUR (prodR ufracR A)).
Definition ufrac_authUR (A : cmraT) : ucmraT :=
Definition ufrac_authUR {SI : sidx} (A : cmra) : ucmra :=
authUR (optionUR (prodR ufracR A)).
(** Note in the signature of [ufrac_auth_auth] and [ufrac_auth_frag] we use
[q : Qp] instead of [q : ufrac]. This way, the API does not expose that [ufrac]
is used internally. This is quite important, as there are two canonical camera
instances with carrier [Qp], namely [fracR] and [ufracR]. When writing things
like [ufrac_auth_auth q a ∧ ✓{q}] we want Coq to infer the type of [q] as [Qp]
like [ufrac_auth_auth q a ∧ ✓ q] we want Coq to infer the type of [q] as [Qp]
such that the [✓] of the default [fracR] camera is used, and not the [✓] of
the [ufracR] camera. *)
Definition ufrac_auth_auth {A : cmraT} (q : Qp) (x : A) : ufrac_authR A :=
Definition ufrac_auth_auth {SI : sidx} {A : cmra}
(q : Qp) (x : A) : ufrac_authR A :=
(Some (q : ufracR,x)).
Definition ufrac_auth_frag {A : cmraT} (q : Qp) (x : A) : ufrac_authR A :=
Definition ufrac_auth_frag {SI : sidx} {A : cmra}
(q : Qp) (x : A) : ufrac_authR A :=
(Some (q : ufracR,x)).
Typeclasses Opaque ufrac_auth_auth ufrac_auth_frag.
Global Typeclasses Opaque ufrac_auth_auth ufrac_auth_frag.
Instance: Params (@ufrac_auth_auth) 2 := {}.
Instance: Params (@ufrac_auth_frag) 2 := {}.
Global Instance: Params (@ufrac_auth_auth) 3 := {}.
Global Instance: Params (@ufrac_auth_frag) 3 := {}.
Notation "●U{ q } a" := (ufrac_auth_auth q a) (at level 10, format "●U{ q } a").
Notation "◯U{ q } a" := (ufrac_auth_frag q a) (at level 10, format "◯U{ q } a").
Notation "●U_ q a" := (ufrac_auth_auth q a) (at level 10, q at level 9, format "●U_ q a").
Notation "◯U_ q a" := (ufrac_auth_frag q a) (at level 10, q at level 9, format "◯U_ q a").
Section ufrac_auth.
Context {A : cmraT}.
Context {SI : sidx} {A : cmra}.
Implicit Types a b : A.
Global Instance ufrac_auth_auth_ne q : NonExpansive (@ufrac_auth_auth A q).
Global Instance ufrac_auth_auth_ne q : NonExpansive (@ufrac_auth_auth SI A q).
Proof. solve_proper. Qed.
Global Instance ufrac_auth_auth_proper q : Proper (() ==> ()) (@ufrac_auth_auth A q).
Global Instance ufrac_auth_auth_proper q :
Proper (() ==> ()) (@ufrac_auth_auth SI A q).
Proof. solve_proper. Qed.
Global Instance ufrac_auth_frag_ne q : NonExpansive (@ufrac_auth_frag A q).
Global Instance ufrac_auth_frag_ne q : NonExpansive (@ufrac_auth_frag SI A q).
Proof. solve_proper. Qed.
Global Instance ufrac_auth_frag_proper q : Proper (() ==> ()) (@ufrac_auth_frag A q).
Global Instance ufrac_auth_frag_proper q :
Proper (() ==> ()) (@ufrac_auth_frag SI A q).
Proof. solve_proper. Qed.
Global Instance ufrac_auth_auth_discrete q a : Discrete a Discrete (U{q} a).
Global Instance ufrac_auth_auth_discrete q a : Discrete a Discrete (U_q a).
Proof. intros. apply auth_auth_discrete; [apply Some_discrete|]; apply _. Qed.
Global Instance ufrac_auth_frag_discrete q a : Discrete a Discrete (U{q} a).
Global Instance ufrac_auth_frag_discrete q a : Discrete a Discrete (U_q a).
Proof. intros. apply auth_frag_discrete; apply Some_discrete; apply _. Qed.
Lemma ufrac_auth_validN n a p : {n} a {n} (U{p} a U{p} a).
Lemma ufrac_auth_validN n a p : {n} a {n} (U_p a U_p a).
Proof. by rewrite auth_both_validN. Qed.
Lemma ufrac_auth_valid p a : a (U{p} a U{p} a).
Lemma ufrac_auth_valid p a : a (U_p a U_p a).
Proof. intros. by apply auth_both_valid_2. Qed.
Lemma ufrac_auth_agreeN n p a b : {n} (U{p} a U{p} b) a {n} b.
Lemma ufrac_auth_agreeN n p a b : {n} (U_p a U_p b) a {n} b.
Proof.
rewrite auth_both_validN=> -[Hincl Hvalid].
move: Hincl=> /Some_includedN=> -[[_ ? //]|[[[p' ?] ?] [/=]]].
rewrite -discrete_iff leibniz_equiv_iff. rewrite ufrac_op'=> [/Qp_eq/=].
rewrite -{1}(Qcplus_0_r p)=> /(inj (Qcplus p))=> ?; by subst.
rewrite auth_both_validN=> -[/Some_includedN [[_ ? //]|Hincl] _].
move: Hincl=> /pair_includedN=> -[/ufrac_included Hincl _].
by destruct (irreflexivity (<)%Qp p).
Qed.
Lemma ufrac_auth_agree p a b : (U{p} a U{p} b) a b.
Lemma ufrac_auth_agree p a b : (U_p a U_p b) a b.
Proof.
intros. apply equiv_dist=> n. by eapply ufrac_auth_agreeN, cmra_valid_validN.
Qed.
Lemma ufrac_auth_agreeL `{!LeibnizEquiv A} p a b : (U{p} a U{p} b) a = b.
Lemma ufrac_auth_agree_L `{!LeibnizEquiv A} p a b : (U_p a U_p b) a = b.
Proof. intros. by eapply leibniz_equiv, ufrac_auth_agree. Qed.
Lemma ufrac_auth_includedN n p q a b : {n} (U{p} a U{q} b) Some b {n} Some a.
Lemma ufrac_auth_includedN n p q a b : {n} (U_p a U_q b) Some b {n} Some a.
Proof. by rewrite auth_both_validN=> -[/Some_pair_includedN [_ ?] _]. Qed.
Lemma ufrac_auth_included `{CmraDiscrete A} q p a b :
(U{p} a U{q} b) Some b Some a.
Proof. rewrite auth_both_valid=> -[/Some_pair_included [_ ?] _] //. Qed.
Lemma ufrac_auth_includedN_total `{CmraTotal A} n q p a b :
{n} (U{p} a U{q} b) b {n} a.
Lemma ufrac_auth_included `{!CmraDiscrete A} q p a b :
(U_p a U_q b) Some b Some a.
Proof. rewrite auth_both_valid_discrete=> -[/Some_pair_included [_ ?] _] //. Qed.
Lemma ufrac_auth_includedN_total `{!CmraTotal A} n q p a b :
{n} (U_p a U_q b) b {n} a.
Proof. intros. by eapply Some_includedN_total, ufrac_auth_includedN. Qed.
Lemma ufrac_auth_included_total `{CmraDiscrete A, CmraTotal A} q p a b :
(U{p} a U{q} b) b a.
Lemma ufrac_auth_included_total `{!CmraDiscrete A, !CmraTotal A} q p a b :
(U_p a U_q b) b a.
Proof. intros. by eapply Some_included_total, ufrac_auth_included. Qed.
Lemma ufrac_auth_auth_validN n q a : {n} (U{q} a) {n} a.
Lemma ufrac_auth_auth_validN n q a : {n} (U_q a) {n} a.
Proof.
rewrite auth_auth_frac_validN Some_validN. split.
by intros [?[]]. intros. by split.
rewrite auth_auth_dfrac_validN Some_validN. split.
- by intros [?[]].
- intros. by split.
Qed.
Lemma ufrac_auth_auth_valid q a : (U{q} a) a.
Lemma ufrac_auth_auth_valid q a : (U_q a) a.
Proof. rewrite !cmra_valid_validN. by setoid_rewrite ufrac_auth_auth_validN. Qed.
Lemma ufrac_auth_frag_validN n q a : {n} (U{q} a) {n} a.
Proof. rewrite auth_frag_validN. split. by intros [??]. by split. Qed.
Lemma ufrac_auth_frag_valid q a : (U{q} a) a.
Proof. rewrite auth_frag_valid. split. by intros [??]. by split. Qed.
Lemma ufrac_auth_frag_validN n q a : {n} (U_q a) {n} a.
Proof.
rewrite auth_frag_validN. split.
- by intros [??].
- by split.
Qed.
Lemma ufrac_auth_frag_valid q a : (U_q a) a.
Proof.
rewrite auth_frag_valid. split.
- by intros [??].
- by split.
Qed.
Lemma ufrac_auth_frag_op q1 q2 a1 a2 : U{q1+q2} (a1 a2) U{q1} a1 U{q2} a2.
Lemma ufrac_auth_frag_op q1 q2 a1 a2 : U_(q1+q2) (a1 a2) U_q1 a1 U_q2 a2.
Proof. done. Qed.
Global Instance is_op_ufrac_auth q q1 q2 a a1 a2 :
IsOp q q1 q2 IsOp a a1 a2 IsOp' (U{q} a) (U{q1} a1) (U{q2} a2).
Lemma ufrac_auth_frag_op_validN n q1 q2 a b :
{n} (U_q1 a U_q2 b) {n} (a b).
Proof. by rewrite -ufrac_auth_frag_op ufrac_auth_frag_validN. Qed.
Lemma ufrac_auth_frag_op_valid q1 q2 a b :
(U_q1 a U_q2 b) (a b).
Proof. by rewrite -ufrac_auth_frag_op ufrac_auth_frag_valid. Qed.
Global Instance ufrac_auth_is_op q q1 q2 a a1 a2 :
IsOp q q1 q2 IsOp a a1 a2 IsOp' (U_q a) (U_q1 a1) (U_q2 a2).
Proof. by rewrite /IsOp' /IsOp=> /leibniz_equiv_iff -> ->. Qed.
Global Instance is_op_ufrac_auth_core_id q q1 q2 a :
CoreId a IsOp q q1 q2 IsOp' (U{q} a) (U{q1} a) (U{q2} a).
Global Instance ufrac_auth_is_op_core_id q q1 q2 a :
CoreId a IsOp q q1 q2 IsOp' (U_q a) (U_q1 a) (U_q2 a).
Proof.
rewrite /IsOp' /IsOp=> ? /leibniz_equiv_iff ->.
by rewrite -ufrac_auth_frag_op -core_id_dup.
Qed.
Lemma ufrac_auth_update p q a b a' b' :
(a,b) ~l~> (a',b') U{p} a U{q} b ~~> U{p} a' U{q} b'.
(a,b) ~l~> (a',b') U_p a U_q b ~~> U_p a' U_q b'.
Proof.
intros. apply: auth_update.
apply: option_local_update. by apply: prod_local_update_2.
Qed.
Lemma ufrac_auth_update_surplus p q a b :
(a b) U{p} a ~~> U{p + q} (a b) U{q} b.
(a b) U_p a ~~> U_(p+q) (a b) U_q b.
Proof.
intros Hconsistent. apply: auth_update_alloc.
intros n m; simpl; intros [Hvalid1 Hvalid2] Heq.
intros Hconsistent. apply: auth_update_alloc. apply local_update_unital.
intros n mpa; simpl; intros [Hvalid1 Hvalid2] Heq; simplify_eq/=.
split.
- split; by apply cmra_valid_validN.
- rewrite pair_op Some_op Heq comm.
destruct m; simpl; [rewrite left_id | rewrite right_id]; done.
destruct mpa; simpl; [rewrite left_id | rewrite right_id]; done.
Qed.
Lemma ufrac_auth_update_surplus_cancel p q a b :
Cancelable b
U_(p+q) (a b) U_q b ~~> U_p a.
Proof.
intros. apply: auth_update_dealloc. apply local_update_unital.
intros n mpa; simpl; intros [Hvalid1 Hvalid2] Heq; simplify_eq/=.
split.
{ split; by eapply cmra_validN_op_l. }
rewrite (comm _ p) (comm _ a) in Heq.
destruct mpa as [[p' a']|]; simplify_eq/=.
- rewrite -Some_op in Heq.
apply (inj _) in Heq as [Hp%leibniz_equiv Ha]; simplify_eq/=.
apply (inj _) in Hp as ->.
apply (cancelableN _) in Ha; last by rewrite comm.
by repeat constructor.
- rewrite right_id in Heq.
apply (inj _) in Heq as [Hp%leibniz_equiv _]; simplify_eq/=.
by apply Qp.add_id_free in Hp.
Qed.
End ufrac_auth.
Definition ufrac_authURF {SI : sidx} (F : rFunctor) : urFunctor :=
authURF (optionURF (prodRF (constRF ufracR) F)).
Global Instance ufrac_authURF_contractive {SI : sidx} F :
rFunctorContractive F urFunctorContractive (ufrac_authURF F).
Proof. apply _. Qed.
Definition ufrac_authRF {SI : sidx} (F : rFunctor) : rFunctor :=
authRF (optionURF (prodRF (constRF ufracR) F)).
Global Instance ufrac_authRF_contractive {SI : sidx} F :
rFunctorContractive F rFunctorContractive (ufrac_authRF F).
Proof. apply _. Qed.
From stdpp Require Export list.
From iris.algebra Require Export ofe.
From iris.algebra Require Import big_op.
From iris.prelude Require Import options.
Section ofe.
Context {SI : sidx} {A : ofe}.
Implicit Types l : list A.
Local Instance list_dist : Dist (list A) := λ n, Forall2 (dist n).
Lemma list_dist_Forall2 n l k : l {n} k Forall2 (dist n) l k.
Proof. done. Qed.
Lemma list_dist_lookup n l1 l2 : l1 {n} l2 i, l1 !! i {n} l2 !! i.
Proof. setoid_rewrite option_dist_Forall2. apply Forall2_lookup. Qed.
Global Instance cons_ne : NonExpansive2 (@cons A) := _.
Global Instance app_ne : NonExpansive2 (@app A) := _.
Global Instance length_ne n : Proper (dist n ==> (=)) (@length A) := _.
Global Instance tail_ne : NonExpansive (@tail A) := _.
Global Instance take_ne n : NonExpansive (@take A n) := _.
Global Instance drop_ne n : NonExpansive (@drop A n) := _.
Global Instance head_ne : NonExpansive (head (A:=A)).
Proof. destruct 1; by constructor. Qed.
Global Instance list_lookup_ne i : NonExpansive (lookup (M:=list A) i).
Proof. intros ????. by apply option_dist_Forall2, Forall2_lookup. Qed.
Global Instance list_lookup_total_ne `{!Inhabited A} i :
NonExpansive (lookup_total (M:=list A) i).
Proof. intros ???. rewrite !list_lookup_total_alt. by intros ->. Qed.
Global Instance list_alter_ne n :
Proper ((dist n ==> dist n) ==> (=) ==> dist n ==> dist n) (alter (M:=list A)) := _.
Global Instance list_insert_ne i : NonExpansive2 (insert (M:=list A) i) := _.
Global Instance list_inserts_ne i : NonExpansive2 (@list_inserts A i) := _.
Global Instance list_delete_ne i : NonExpansive (delete (M:=list A) i) := _.
Global Instance option_list_ne : NonExpansive (@option_list A).
Proof. intros ????; by apply Forall2_option_list, option_dist_Forall2. Qed.
Global Instance list_filter_ne n P `{ x, Decision (P x)} :
Proper (dist n ==> iff) P
Proper (dist n ==> dist n) (filter (B:=list A) P) := _.
Global Instance replicate_ne n : NonExpansive (@replicate A n) := _.
Global Instance reverse_ne : NonExpansive (@reverse A) := _.
Global Instance last_ne : NonExpansive (@last A).
Proof. intros ????; by apply option_dist_Forall2, Forall2_last. Qed.
Global Instance resize_ne n : NonExpansive2 (@resize A n) := _.
Global Instance cons_dist_inj n :
Inj2 (dist n) (dist n) (dist n) (@cons A).
Proof. by inversion_clear 1. Qed.
Lemma nil_dist_eq n l : l {n} [] l = [].
Proof. split; by inversion 1. Qed.
Lemma cons_dist_eq n l k y :
l {n} y :: k x l', x {n} y l' {n} k l = x :: l'.
Proof. apply Forall2_cons_inv_r. Qed.
Lemma app_dist_eq n l k1 k2 :
l {n} k1 ++ k2 k1' k2', l = k1' ++ k2' k1' {n} k1 k2' {n} k2.
Proof. rewrite list_dist_Forall2 Forall2_app_inv_r. naive_solver. Qed.
Lemma list_singleton_dist_eq n l x :
l {n} [x] x', l = [x'] x' {n} x.
Proof.
split; [|by intros (?&->&->)].
intros (?&?&?&->%Forall2_nil_inv_r&->)%list_dist_Forall2%Forall2_cons_inv_r.
eauto.
Qed.
Definition list_ofe_mixin : OfeMixin (list A).
Proof.
split.
- intros l k. rewrite list_equiv_Forall2 -Forall2_forall.
split; induction 1; constructor; intros; try apply equiv_dist; auto.
- apply _.
- rewrite /dist /list_dist. eauto using Forall2_impl, dist_le.
Qed.
Canonical Structure listO := Ofe (list A) list_ofe_mixin.
(** To define [compl : chain (list A) → list A] we make use of the fact that
given a given chain [c0, c1, c2, ...] of lists, the list [c0] completely
determines the shape (i.e., the length) of all lists in the chain. So, the
[compl] operation is defined by structural recursion on [c0], and takes the
completion of the elements of all lists in the chain point-wise. We use [head]
and [tail] as the inverse of [cons]. *)
Fixpoint list_compl_go `{!Cofe A} (c0 : list A) (c : chain listO) : listO :=
match c0 with
| [] => []
| x :: c0 => compl (chain_map (default x head) c) ::
list_compl_go c0 (chain_map tail c)
end.
Fixpoint list_lbcompl_go `{!Cofe A} (c0 : list A) {n} Hn (c : bchain listO n) : listO :=
match c0 with
| [] => []
| x :: c0 => lbcompl Hn (bchain_map (default x head) c) ::
list_lbcompl_go c0 Hn (bchain_map tail c)
end.
Global Program Instance list_cofe `{!Cofe A} : Cofe listO := {
compl c := list_compl_go (c 0) c;
lbcompl n Hn c := list_lbcompl_go (c _ (SIdx.limit_lt_0 _ Hn)) Hn c
}.
Next Obligation.
intros ? n c; simpl.
assert (c 0 {0} c n) as Hc0 by (symmetry; apply chain_cauchy, SIdx.le_0_l).
revert Hc0. generalize (c 0)=> c0. revert c.
induction c0 as [|x c0 IH]=> c Hc0 /=.
{ by inversion Hc0. }
apply symmetry, cons_dist_eq in Hc0 as (x' & xs' & Hx & Hc0 & Hcn).
rewrite Hcn. f_equiv.
- rewrite conv_compl.
by rewrite /chain_map //= Hcn.
- rewrite IH /chain_map /= ?Hcn //.
Qed.
Next Obligation.
intros ? n Hn c m Hm; simpl.
assert (c _ (SIdx.limit_lt_0 _ Hn) {0} c m Hm) as Hc0
by (symmetry; apply bchain_cauchy, SIdx.le_0_l).
revert Hc0. generalize (c _ (SIdx.limit_lt_0 _ Hn))=> c0. revert c.
induction c0 as [|x c0 IH]=> c Hc0 /=.
{ by inversion Hc0. }
apply symmetry in Hc0.
apply cons_dist_eq in Hc0 as (x' & xs' & Hx & Hc0 & Hcn).
rewrite Hcn. f_equiv.
- rewrite (conv_lbcompl _ _ Hm). by rewrite /bchain_map //= Hcn.
- rewrite IH /bchain_map /= ?Hcn //.
Qed.
Next Obligation.
intros ? n Hn c1 c2 m Hc; simpl.
specialize (Hc _ (SIdx.limit_lt_0 _ Hn)) as Heq.
move: Heq. generalize (c1 _ (SIdx.limit_lt_0 _ Hn)) as c0 => c0.
generalize (c2 _ (SIdx.limit_lt_0 _ Hn)) as d0 => d0 Heq.
induction Heq as [ | ???? Heq1 Heq2 IH ] in c1, c2, Hc |-*; simpl; f_equiv.
- apply lbcompl_ne=> γ . by rewrite /bchain_map //= Hc Heq1.
- apply IH=> γ . by rewrite /bchain_map /= Hc.
Qed.
Global Instance list_ofe_discrete : OfeDiscrete A OfeDiscrete listO.
Proof. induction 2; constructor; try apply (discrete_0 _); auto. Qed.
Global Instance nil_discrete : Discrete (@nil A).
Proof. inversion_clear 1; constructor. Qed.
Global Instance cons_discrete x l : Discrete x Discrete l Discrete (x :: l).
Proof. intros ??; inversion_clear 1; constructor; by apply discrete_0. Qed.
Lemma dist_Permutation n l1 l2 l3 :
l1 {n} l2 l2 l3 l2', l1 l2' l2' {n} l3.
Proof.
intros Hequiv Hperm. revert l1 Hequiv.
induction Hperm as [|x l2 l3 _ IH|x y l2|l2 l3 l4 _ IH1 _ IH2]; intros l1.
- intros ?. by exists l1.
- intros (x'&l2'&?&(l2''&?&?)%IH&->)%cons_dist_eq.
exists (x' :: l2''). by repeat constructor.
- intros (y'&?&?&(x'&l2'&?&?&->)%cons_dist_eq&->)%cons_dist_eq.
exists (x' :: y' :: l2'). by repeat constructor.
- intros (l2'&?&(l3'&?&?)%IH2)%IH1. exists l3'. split; [by etrans|done].
Qed.
End ofe.
Global Arguments listO {SI} _.
(** Non-expansiveness of higher-order list functions and big-ops *)
Global Instance list_fmap_ne {SI : sidx} {A B : ofe} n :
Proper ((dist n ==> dist n) ==> ({n}@{list A}) ==> ({n}@{list B})) fmap.
Proof. intros f1 f2 Hf l1 l2 Hl; by eapply Forall2_fmap, Forall2_impl; eauto. Qed.
Global Instance list_omap_ne {SI : sidx} {A B : ofe} n :
Proper ((dist n ==> dist n) ==> ({n}@{list A}) ==> ({n}@{list B})) omap.
Proof.
intros f1 f2 Hf. induction 1 as [|x1 x2 l1 l2 Hx Hl]; csimpl; [constructor|].
destruct (Hf _ _ Hx); [f_equiv|]; auto.
Qed.
Global Instance imap_ne {SI : sidx} {A B : ofe} n :
Proper (pointwise_relation _ ((dist n ==> dist n)) ==> dist n ==> dist n)
(imap (A:=A) (B:=B)).
Proof.
intros f1 f2 Hf l1 l2 Hl. revert f1 f2 Hf.
induction Hl as [|x1 x2 l1 l2 ?? IH]; intros f1 f2 Hf; simpl; [constructor|].
f_equiv; [by apply Hf|]. apply IH. intros i y1 y2 Hy. by apply Hf.
Qed.
Global Instance list_bind_ne {SI : sidx} {A B : ofe} n :
Proper ((dist n ==> dist n) ==> ({n}@{list B}) ==> ({n}@{list A})) mbind.
Proof. intros f1 f2 Hf. induction 1; csimpl; [constructor|f_equiv; auto]. Qed.
Global Instance list_join_ne {SI : sidx} {A : ofe} n :
Proper (dist n ==> ({n}@{list A})) mjoin.
Proof. induction 1; simpl; [constructor|solve_proper]. Qed.
Global Instance zip_with_ne {SI : sidx} {A B C : ofe} n :
Proper ((dist n ==> dist n ==> dist n) ==> dist n ==> dist n ==> dist n)
(zip_with (A:=A) (B:=B) (C:=C)).
Proof.
intros f1 f2 Hf.
induction 1; destruct 1; simpl; [constructor..|f_equiv; try apply Hf; auto].
Qed.
Global Instance list_fmap_dist_inj {SI : sidx} {A B : ofe} (f : A B) n :
Inj ({n}) ({n}) f Inj ({n}@{list A}) ({n}@{list B}) (fmap f).
Proof. apply list_fmap_inj. Qed.
Lemma big_opL_ne_2 {SI : sidx} {M : ofe}
{o : M M M} `{!Monoid o} {A : ofe} (f g : nat A M) l1 l2 n :
l1 {n} l2
( k y1 y2,
l1 !! k = Some y1 l2 !! k = Some y2 y1 {n} y2 f k y1 {n} g k y2)
([^o list] k y l1, f k y) {n} ([^o list] k y l2, g k y).
Proof.
intros Hl Hf. apply big_opL_gen_proper_2; try (apply _ || done).
{ apply monoid_ne. }
intros k. assert (l1 !! k {n} l2 !! k) as Hlk by (by f_equiv).
destruct (l1 !! k) eqn:?, (l2 !! k) eqn:?; inversion Hlk; naive_solver.
Qed.
(** Functor *)
Lemma list_fmap_ext_ne {SI : sidx} {A} {B : ofe} (f g : A B) (l : list A) n :
( x, f x {n} g x) f <$> l {n} g <$> l.
Proof. intros Hf. by apply Forall2_fmap, Forall_Forall2_diag, Forall_true. Qed.
Definition listO_map {SI : sidx} {A B} (f : A -n> B) : listO A -n> listO B :=
OfeMor (fmap f : listO A listO B).
Global Instance listO_map_ne {SI : sidx} A B : NonExpansive (@listO_map SI A B).
Proof. intros n f g ? l. by apply list_fmap_ext_ne. Qed.
Program Definition listOF {SI : sidx} (F : oFunctor) : oFunctor := {|
oFunctor_car A _ B _ := listO (oFunctor_car F A B);
oFunctor_map A1 _ A2 _ B1 _ B2 _ fg := listO_map (oFunctor_map F fg)
|}.
Next Obligation.
by intros ? F A1 ? A2 ? B1 ? B2 ? n f g Hfg; apply listO_map_ne, oFunctor_map_ne.
Qed.
Next Obligation.
intros ? F A ? B ? x. rewrite /= -{2}(list_fmap_id x).
apply list_fmap_equiv_ext=>???. apply oFunctor_map_id.
Qed.
Next Obligation.
intros ? F A1 ? A2 ? A3 ? B1 ? B2 ? B3 ? f g f' g' x. rewrite /= -list_fmap_compose.
apply list_fmap_equiv_ext=>???; apply oFunctor_map_compose.
Qed.
Global Instance listOF_contractive {SI : sidx} F :
oFunctorContractive F oFunctorContractive (listOF F).
Proof.
intros ? A1 ? A2 ? B1 ? B2 ? n f g Hfg;
by apply listO_map_ne, oFunctor_map_contractive.
Qed.
From iris.algebra Require Export cmra.
Set Default Proof Using "Type".
From iris.prelude Require Import options.
(** * Local updates *)
Definition local_update {A : cmraT} (x y : A * A) := n mz,
Definition local_update {SI : sidx} {A : cmra} (x y : A * A) := n mz,
{n} x.1 x.1 {n} x.2 ? mz {n} y.1 y.1 {n} y.2 ? mz.
Instance: Params (@local_update) 1 := {}.
Global Instance: Params (@local_update) 2 := {}.
Infix "~l~>" := local_update (at level 70).
Section updates.
Context {A : cmraT}.
Context {SI : sidx} {A : cmra}.
Implicit Types x y : A.
Global Instance local_update_proper :
Proper (() ==> () ==> iff) (@local_update A).
Proper (() ==> () ==> iff) (@local_update SI A).
Proof. unfold local_update. by repeat intro; setoid_subst. Qed.
Global Instance local_update_preorder : PreOrder (@local_update A).
Global Instance local_update_preorder : PreOrder (@local_update SI A).
Proof. split; unfold local_update; red; naive_solver. Qed.
Lemma exclusive_local_update `{!Exclusive y} x x' : x' (x,y) ~l~> (x',x').
......@@ -72,45 +72,43 @@ Section updates.
Proof.
rewrite /local_update /=. setoid_rewrite <-cmra_discrete_valid_iff.
setoid_rewrite <-(λ n, discrete_iff n x).
setoid_rewrite <-(λ n, discrete_iff n x'). naive_solver eauto using 0.
setoid_rewrite <-(λ n, discrete_iff n x'). pose 0. naive_solver.
Qed.
Lemma local_update_valid0 x y x' y' :
({0} x {0} y x {0} y y {0} x (x,y) ~l~> (x',y'))
({0} x {0} y Some y {0} Some x (x,y) ~l~> (x',y'))
(x,y) ~l~> (x',y').
Proof.
intros Hup n mz Hmz Hz; simpl in *. apply Hup; auto.
- by apply (cmra_validN_le n); last lia.
- apply (cmra_validN_le n); last lia.
- by apply (cmra_validN_le n), SIdx.le_0_l.
- apply (cmra_validN_le n), SIdx.le_0_l.
move: Hmz; rewrite Hz. destruct mz; simpl; eauto using cmra_validN_op_l.
- destruct mz as [z|].
+ right. exists z. apply dist_le with n; auto with lia.
+ left. apply dist_le with n; auto with lia.
- eapply (cmra_includedN_le n), SIdx.le_0_l.
apply Some_includedN_opM. eauto.
Qed.
Lemma local_update_valid `{!CmraDiscrete A} x y x' y' :
( x y x y y x (x,y) ~l~> (x',y')) (x,y) ~l~> (x',y').
( x y Some y Some x (x,y) ~l~> (x',y')) (x,y) ~l~> (x',y').
Proof.
rewrite !(cmra_discrete_valid_iff 0)
(cmra_discrete_included_iff 0) (discrete_iff 0).
rewrite !(cmra_discrete_valid_iff 0) (cmra_discrete_included_iff 0).
apply local_update_valid0.
Qed.
Lemma local_update_total_valid0 `{!CmraTotal A} x y x' y' :
({0} x {0} y y {0} x (x,y) ~l~> (x',y')) (x,y) ~l~> (x',y').
({0} x {0} y y {0} x (x,y) ~l~> (x',y')) (x,y) ~l~> (x',y').
Proof.
intros Hup. apply local_update_valid0=> ?? [Hx|?]; apply Hup; auto.
by rewrite Hx.
intros Hup. apply local_update_valid0=> ?? Hincl. apply Hup; auto.
by apply Some_includedN_total.
Qed.
Lemma local_update_total_valid `{!CmraTotal A, !CmraDiscrete A} x y x' y' :
( x y y x (x,y) ~l~> (x',y')) (x,y) ~l~> (x',y').
Proof.
rewrite !(cmra_discrete_valid_iff 0) (cmra_discrete_included_iff 0).
rewrite !(cmra_discrete_valid_iff 0) (cmra_discrete_included_iff 0).
apply local_update_total_valid0.
Qed.
End updates.
Section updates_unital.
Context {A : ucmraT}.
Context {SI : sidx} {A : ucmra}.
Implicit Types x y : A.
Lemma local_update_unital x y x' y' :
......@@ -137,70 +135,102 @@ Section updates_unital.
Proof. rewrite -{2}(right_id ε op x). by apply cancel_local_update. Qed.
End updates_unital.
(** * Product *)
Lemma prod_local_update {A B : cmraT} (x y x' y' : A * B) :
(x.1,y.1) ~l~> (x'.1,y'.1) (x.2,y.2) ~l~> (x'.2,y'.2)
(x,y) ~l~> (x',y').
Proof.
intros Hup1 Hup2 n mz [??] [??]; simpl in *.
destruct (Hup1 n (fst <$> mz)); [done|by destruct mz|].
destruct (Hup2 n (snd <$> mz)); [done|by destruct mz|].
by destruct mz.
Qed.
Lemma prod_local_update' {A B : cmraT} (x1 y1 x1' y1' : A) (x2 y2 x2' y2' : B) :
(x1,y1) ~l~> (x1',y1') (x2,y2) ~l~> (x2',y2')
((x1,x2),(y1,y2)) ~l~> ((x1',x2'),(y1',y2')).
Proof. intros. by apply prod_local_update. Qed.
Lemma prod_local_update_1 {A B : cmraT} (x1 y1 x1' y1' : A) (x2 y2 : B) :
(x1,y1) ~l~> (x1',y1') ((x1,x2),(y1,y2)) ~l~> ((x1',x2),(y1',y2)).
Proof. intros. by apply prod_local_update. Qed.
Lemma prod_local_update_2 {A B : cmraT} (x1 y1 : A) (x2 y2 x2' y2' : B) :
(x2,y2) ~l~> (x2',y2') ((x1,x2),(y1,y2)) ~l~> ((x1,x2'),(y1,y2')).
Proof. intros. by apply prod_local_update. Qed.
(** * Option *)
(* TODO: Investigate whether we can use these in proving the very similar local
updates on finmaps. *)
Lemma option_local_update {A : cmraT} (x y x' y' : A) :
(x, y) ~l~> (x',y')
(Some x, Some y) ~l~> (Some x', Some y').
Proof.
intros Hup. apply local_update_unital=>n mz Hxv Hx; simpl in *.
destruct (Hup n mz); first done.
{ destruct mz as [?|]; inversion_clear Hx; auto. }
split; first done. destruct mz as [?|]; constructor; auto.
Qed.
Lemma alloc_option_local_update {A : cmraT} (x : A) y :
x
(None, y) ~l~> (Some x, Some x).
Proof.
move=>Hx. apply local_update_unital=> n z _ /= Heq. split.
{ rewrite Some_validN. apply cmra_valid_validN. done. }
destruct z as [z|]; last done. destruct y; inversion Heq.
Qed.
Lemma delete_option_local_update {A : cmraT} (x : option A) (y : A) :
Exclusive y (x, Some y) ~l~> (None, None).
Proof.
move=>Hex. apply local_update_unital=>n z /= Hy Heq. split; first done.
destruct z as [z|]; last done. exfalso.
move: Hy. rewrite Heq /= -Some_op=>Hy. eapply Hex.
eapply cmra_validN_le. eapply Hy. lia.
Qed.
(** * Natural numbers *)
Lemma nat_local_update (x y x' y' : nat) :
x + y' = x' + y (x,y) ~l~> (x',y').
Proof.
intros ??; apply local_update_unital_discrete=> z _.
compute -[minus plus]; lia.
Qed.
Lemma mnat_local_update (x y x' : mnatUR) :
x x' (x,y) ~l~> (x',x').
Proof.
intros ??; apply local_update_unital_discrete=> z _.
compute -[max]; lia.
Qed.
Section updates_unit.
Context {SI : sidx}.
(** * Unit *)
Lemma unit_local_update (x y x' y' : unit) : (x, y) ~l~> (x', y').
Proof. destruct x,y,x',y'; reflexivity. Qed.
End updates_unit.
Section updates_discrete_fun.
Context {SI : sidx}.
(** * Dependently-typed functions over a discrete domain *)
Lemma discrete_fun_local_update {A} {B : A ucmra} (f g f' g' : discrete_fun B) :
( x : A, (f x, g x) ~l~> (f' x, g' x))
(f, g) ~l~> (f', g').
Proof.
setoid_rewrite local_update_unital. intros Hupd n h Hf Hg.
split=> x; eapply Hupd, Hg; eauto.
Qed.
End updates_discrete_fun.
Section updates_product.
Context {SI : sidx}.
(** * Product *)
Lemma prod_local_update {A B : cmra} (x y x' y' : A * B) :
(x.1,y.1) ~l~> (x'.1,y'.1) (x.2,y.2) ~l~> (x'.2,y'.2)
(x,y) ~l~> (x',y').
Proof.
intros Hup1 Hup2 n mz [??] [??]; simpl in *.
destruct (Hup1 n (fst <$> mz)); [done|by destruct mz|].
destruct (Hup2 n (snd <$> mz)); [done|by destruct mz|].
by destruct mz.
Qed.
Lemma prod_local_update' {A B : cmra} (x1 y1 x1' y1' : A) (x2 y2 x2' y2' : B) :
(x1,y1) ~l~> (x1',y1') (x2,y2) ~l~> (x2',y2')
((x1,x2),(y1,y2)) ~l~> ((x1',x2'),(y1',y2')).
Proof. intros. by apply prod_local_update. Qed.
Lemma prod_local_update_1 {A B : cmra} (x1 y1 x1' y1' : A) (x2 y2 : B) :
(x1,y1) ~l~> (x1',y1') ((x1,x2),(y1,y2)) ~l~> ((x1',x2),(y1',y2)).
Proof. intros. by apply prod_local_update. Qed.
Lemma prod_local_update_2 {A B : cmra} (x1 y1 : A) (x2 y2 x2' y2' : B) :
(x2,y2) ~l~> (x2',y2') ((x1,x2),(y1,y2)) ~l~> ((x1,x2'),(y1,y2')).
Proof. intros. by apply prod_local_update. Qed.
End updates_product.
Section updates_option.
Context {SI : sidx}.
(** * Option *)
(* TODO: Investigate whether we can use these in proving the very similar local
updates on finmaps. *)
Lemma option_local_update {A : cmra} (x y x' y' : A) :
(x, y) ~l~> (x',y')
(Some x, Some y) ~l~> (Some x', Some y').
Proof.
intros Hup. apply local_update_unital=>n mz Hxv Hx; simpl in *.
destruct (Hup n mz); first done.
{ destruct mz as [?|]; inversion_clear Hx; auto. }
split; first done. destruct mz as [?|]; constructor; auto.
Qed.
Lemma option_local_update_None {A: ucmra} (x x' y': A):
(x, ε) ~l~> (x', y') ->
(Some x, None) ~l~> (Some x', Some y').
Proof.
intros Hup. apply local_update_unital=> n mz.
rewrite left_id=> ? <-.
destruct (Hup n (Some x)); simpl in *; first done.
- by rewrite left_id.
- split; first done. rewrite -Some_op. by constructor.
Qed.
Lemma alloc_option_local_update {A : cmra} (x : A) y :
x
(None, y) ~l~> (Some x, Some x).
Proof.
move=>Hx. apply local_update_unital=> n z _ /= Heq. split.
{ rewrite Some_validN. apply cmra_valid_validN. done. }
destruct z as [z|]; last done. destruct y; inversion Heq.
Qed.
Lemma delete_option_local_update {A : cmra} (x : option A) (y : A) :
Exclusive y (x, Some y) ~l~> (None, None).
Proof.
move=>Hex. apply local_update_unital=>n z /= Hy Heq. split; first done.
destruct z as [z|]; last done. exfalso.
move: Hy. rewrite Heq /= -Some_op=>Hy. eapply Hex.
eapply cmra_validN_le, SIdx.le_0_l. eapply Hy.
Qed.
Lemma delete_option_local_update_cancelable {A : cmra} (mx : option A) :
Cancelable mx (mx, mx) ~l~> (None, None).
Proof.
intros ?. apply local_update_unital=>n mf /= Hmx Heq. split; first done.
rewrite left_id. eapply (cancelableN mx); by rewrite right_id_L.
Qed.
End updates_option.
(** Defines an RA on lists whose composition is only defined when one operand is
a prefix of the other. The result is the longer list.
In particular, the core is the identity function for all elements. *)
From iris.algebra Require Export agree list gmap updates.
From iris.algebra Require Import local_updates proofmode_classes.
From iris.prelude Require Import options.
Definition max_prefix_list (A : Type) := gmap nat (agree A).
Definition max_prefix_listR {SI : sidx} (A : ofe) := gmapUR nat (agreeR A).
Definition max_prefix_listUR {SI : sidx} (A : ofe) := gmapUR nat (agreeR A).
Definition to_max_prefix_list {A} (l : list A) : gmap nat (agree A) :=
to_agree <$> map_seq 0 l.
Global Instance: Params (@to_max_prefix_list) 1 := {}.
Global Typeclasses Opaque to_max_prefix_list.
Section max_prefix_list.
Context {SI : sidx} {A : ofe}.
Implicit Types l : list A.
Global Instance to_max_prefix_list_ne : NonExpansive (@to_max_prefix_list A).
Proof. solve_proper. Qed.
Global Instance to_max_prefix_list_proper :
Proper (() ==> ()) (@to_max_prefix_list A).
Proof. solve_proper. Qed.
Global Instance to_max_prefix_list_dist_inj n :
Inj (dist n) (dist n) (@to_max_prefix_list A).
Proof.
rewrite /to_max_prefix_list. intros l1 l2 Hl. apply list_dist_lookup=> i.
move: (Hl i). rewrite !lookup_fmap !lookup_map_seq Nat.sub_0_r.
rewrite !option_guard_True; [|lia..].
destruct (l1 !! i), (l2 !! i); inversion_clear 1;
constructor; by apply (inj to_agree).
Qed.
Global Instance to_max_prefix_list_inj : Inj () () (@to_max_prefix_list A).
Proof.
intros l1 l2. rewrite !equiv_dist=> ? n. by apply (inj to_max_prefix_list).
Qed.
Global Instance mono_list_lb_core_id (m : max_prefix_list A) : CoreId m := _.
Lemma to_max_prefix_list_valid l : to_max_prefix_list l.
Proof.
intros i. rewrite /to_max_prefix_list lookup_fmap.
by destruct (map_seq 0 l !! i).
Qed.
Lemma to_max_prefix_list_validN n l : {n} to_max_prefix_list l.
Proof. apply cmra_valid_validN, to_max_prefix_list_valid. Qed.
Local Lemma to_max_prefix_list_app l1 l2 :
to_max_prefix_list (l1 ++ l2)
to_max_prefix_list l1 (to_agree <$> map_seq (length l1) l2).
Proof.
rewrite /to_max_prefix_list map_seq_app=> i /=. rewrite lookup_op !lookup_fmap.
destruct (map_seq 0 l1 !! i) as [x|] eqn:Hl1; simpl; last first.
{ by rewrite lookup_union_r // left_id. }
rewrite (lookup_union_Some_l _ _ _ x) //=.
assert (map_seq (M:=gmap nat A) (length l1) l2 !! i = None) as ->.
{ apply lookup_map_seq_None.
apply lookup_map_seq_Some in Hl1 as [_ ?%lookup_lt_Some]. lia. }
by rewrite /= right_id.
Qed.
Lemma to_max_prefix_list_op_l l1 l2 :
l1 `prefix_of` l2
to_max_prefix_list l1 to_max_prefix_list l2 to_max_prefix_list l2.
Proof. intros [l ->]. by rewrite to_max_prefix_list_app assoc -core_id_dup. Qed.
Lemma to_max_prefix_list_op_r l1 l2 :
l1 `prefix_of` l2
to_max_prefix_list l2 to_max_prefix_list l1 to_max_prefix_list l2.
Proof. intros. by rewrite comm to_max_prefix_list_op_l. Qed.
Lemma max_prefix_list_included_includedN (ml1 ml2 : max_prefix_list A) :
ml1 ml2 n, ml1 {n} ml2.
Proof.
split; [intros; by apply: cmra_included_includedN|].
intros Hincl. exists ml2. apply equiv_dist=> n. destruct (Hincl n) as [l ->].
by rewrite assoc -core_id_dup.
Qed.
Local Lemma to_max_prefix_list_includedN_aux n l1 l2 :
to_max_prefix_list l1 {n} to_max_prefix_list l2
l2 {n} l1 ++ drop (length l1) l2.
Proof.
rewrite lookup_includedN=> Hincl. apply list_dist_lookup=> i.
rewrite lookup_app. move: (Hincl i).
rewrite /to_max_prefix_list !lookup_fmap !lookup_map_seq Nat.sub_0_r.
rewrite !option_guard_True; [|lia..].
rewrite option_includedN_total fmap_None.
intros [Hi|(?&?&(a2&->&->)%fmap_Some&(a1&->&->)%fmap_Some&Ha)].
- rewrite lookup_drop Hi. apply lookup_ge_None in Hi. f_equiv; lia.
- f_equiv. symmetry. by apply to_agree_includedN.
Qed.
Lemma to_max_prefix_list_includedN n l1 l2 :
to_max_prefix_list l1 {n} to_max_prefix_list l2 l, l2 {n} l1 ++ l.
Proof.
split.
- intros. eexists. by apply to_max_prefix_list_includedN_aux.
- intros [l ->]. rewrite to_max_prefix_list_app. apply: cmra_includedN_l.
Qed.
Lemma to_max_prefix_list_included l1 l2 :
to_max_prefix_list l1 to_max_prefix_list l2 l, l2 l1 ++ l.
Proof.
split.
- intros. eexists. apply equiv_dist=> n.
apply to_max_prefix_list_includedN_aux. by apply: cmra_included_includedN.
- intros [l ->]. rewrite to_max_prefix_list_app. eauto.
Qed.
Lemma to_max_prefix_list_included_L `{!LeibnizEquiv A} l1 l2 :
to_max_prefix_list l1 to_max_prefix_list l2 l1 `prefix_of` l2.
Proof. rewrite to_max_prefix_list_included /prefix. naive_solver. Qed.
Local Lemma to_max_prefix_list_op_validN_aux n l1 l2 :
length l1 length l2
{n} (to_max_prefix_list l1 to_max_prefix_list l2)
l2 {n} l1 ++ drop (length l1) l2.
Proof.
intros Hlen Hvalid. apply list_dist_lookup=> i. move: (Hvalid i).
rewrite /to_max_prefix_list lookup_op !lookup_fmap !lookup_map_seq Nat.sub_0_r.
rewrite !option_guard_True; [|lia..].
intros ?. rewrite lookup_app.
destruct (l1 !! i) as [x1|] eqn:Hi1, (l2 !! i) as [x2|] eqn:Hi2; simpl in *.
- f_equiv. symmetry. by apply to_agree_op_validN.
- apply lookup_lt_Some in Hi1; apply lookup_ge_None in Hi2. lia.
- apply lookup_ge_None in Hi1. rewrite lookup_drop -Hi2. f_equiv; lia.
- apply lookup_ge_None in Hi1. rewrite lookup_drop -Hi2. f_equiv; lia.
Qed.
Lemma to_max_prefix_list_op_validN n l1 l2 :
{n} (to_max_prefix_list l1 to_max_prefix_list l2)
( l, l2 {n} l1 ++ l) ( l, l1 {n} l2 ++ l).
Proof.
split.
- destruct (decide (length l1 length l2)).
+ left. eexists. by eapply to_max_prefix_list_op_validN_aux.
+ right. eexists. eapply to_max_prefix_list_op_validN_aux; [lia|by rewrite comm].
- intros [[l ->]|[l ->]].
+ rewrite to_max_prefix_list_op_l; last by apply prefix_app_r.
apply to_max_prefix_list_validN.
+ rewrite to_max_prefix_list_op_r; last by apply prefix_app_r.
apply to_max_prefix_list_validN.
Qed.
Lemma to_max_prefix_list_op_valid l1 l2 :
(to_max_prefix_list l1 to_max_prefix_list l2)
( l, l2 l1 ++ l) ( l, l1 l2 ++ l).
Proof.
split.
- destruct (decide (length l1 length l2)).
+ left. eexists. apply equiv_dist=> n'.
by eapply to_max_prefix_list_op_validN_aux, cmra_valid_validN.
+ right. eexists. apply equiv_dist=> n'.
by eapply to_max_prefix_list_op_validN_aux,
cmra_valid_validN; [lia|by rewrite comm].
- intros [[l ->]|[l ->]].
+ rewrite to_max_prefix_list_op_l; last by apply prefix_app_r.
apply to_max_prefix_list_valid.
+ rewrite to_max_prefix_list_op_r; last by apply prefix_app_r.
apply to_max_prefix_list_valid.
Qed.
Lemma to_max_prefix_list_op_valid_L `{!LeibnizEquiv A} l1 l2 :
(to_max_prefix_list l1 to_max_prefix_list l2)
l1 `prefix_of` l2 l2 `prefix_of` l1.
Proof. rewrite to_max_prefix_list_op_valid /prefix. naive_solver. Qed.
Lemma max_prefix_list_local_update l1 l2 :
l1 `prefix_of` l2
(to_max_prefix_list l1, to_max_prefix_list l1) ~l~>
(to_max_prefix_list l2, to_max_prefix_list l2).
Proof.
intros [l ->]. rewrite to_max_prefix_list_app (comm _ (to_max_prefix_list l1)).
apply op_local_update=> n _. rewrite comm -to_max_prefix_list_app.
apply to_max_prefix_list_validN.
Qed.
End max_prefix_list.
Definition max_prefix_listURF {SI : sidx} (F : oFunctor) : urFunctor :=
gmapURF nat (agreeRF F).
Global Instance max_prefix_listURF_contractive {SI : sidx} F :
oFunctorContractive F urFunctorContractive (max_prefix_listURF F).
Proof. apply _. Qed.
Definition max_prefix_listRF {SI : sidx} (F : oFunctor) : rFunctor :=
gmapRF nat (agreeRF F).
Global Instance max_prefix_listRF_contractive {SI : sidx} F :
oFunctorContractive F rFunctorContractive (max_prefix_listRF F).
Proof. apply _. Qed.
From iris.algebra Require Export ofe.
Set Default Proof Using "Type".
From iris.prelude Require Import options.
(** The Monoid class that is used for generic big operators in the file
[algebra/big_op]. The operation is an argument because we want to have multiple
......@@ -17,24 +17,24 @@ we do not have a canonical structure for setoids, we do not go that way.
Note that we do not declare any of the projections as type class instances. That
is because we only need them in the [big_op] file, and nowhere else. Hence, we
declare these instances locally there to avoid them being used elsewhere. *)
Class Monoid {M : ofeT} (o : M M M) := {
Class Monoid {SI : sidx} {M : ofe} (o : M M M) := {
monoid_unit : M;
monoid_ne : NonExpansive2 o;
monoid_assoc : Assoc () o;
monoid_comm : Comm () o;
monoid_left_id : LeftId () monoid_unit o;
}.
Lemma monoid_proper `{Monoid M o} : Proper (() ==> () ==> ()) o.
Lemma monoid_proper {SI : sidx} `{Monoid M o} : Proper (() ==> () ==> ()) o.
Proof. apply ne_proper_2, monoid_ne. Qed.
Lemma monoid_right_id `{Monoid M o} : RightId () monoid_unit o.
Lemma monoid_right_id {SI : sidx} `{Monoid M o} : RightId () monoid_unit o.
Proof. intros x. etrans; [apply monoid_comm|apply monoid_left_id]. Qed.
(** The [Homomorphism] classes give rise to generic lemmas about big operators
commuting with each other. We also consider a [WeakMonoidHomomorphism] which
does not necesarrily commute with unit; an example is the [own] connective: we
does not necessarily commute with unit; an example is the [own] connective: we
only have `True ==∗ own γ ∅`, not `True ↔ own γ ∅`. *)
Class WeakMonoidHomomorphism {M1 M2 : ofeT}
(o1 : M1 M1 M1) (o2 : M2 M2 M2) `{Monoid M1 o1, Monoid M2 o2}
Class WeakMonoidHomomorphism {SI : sidx} {M1 M2 : ofe}
(o1 : M1 M1 M1) (o2 : M2 M2 M2) `{!Monoid o1, !Monoid o2}
(R : relation M2) (f : M1 M2) := {
monoid_homomorphism_rel_po : PreOrder R;
monoid_homomorphism_rel_proper : Proper (() ==> () ==> iff) R;
......@@ -43,13 +43,13 @@ Class WeakMonoidHomomorphism {M1 M2 : ofeT}
monoid_homomorphism x y : R (f (o1 x y)) (o2 (f x) (f y))
}.
Class MonoidHomomorphism {M1 M2 : ofeT}
(o1 : M1 M1 M1) (o2 : M2 M2 M2) `{Monoid M1 o1, Monoid M2 o2}
Class MonoidHomomorphism {SI : sidx} {M1 M2 : ofe}
(o1 : M1 M1 M1) (o2 : M2 M2 M2) `{!Monoid o1, !Monoid o2}
(R : relation M2) (f : M1 M2) := {
monoid_homomorphism_weak :> WeakMonoidHomomorphism o1 o2 R f;
#[global] monoid_homomorphism_weak :: WeakMonoidHomomorphism o1 o2 R f;
monoid_homomorphism_unit : R (f monoid_unit) monoid_unit
}.
Lemma weak_monoid_homomorphism_proper
Lemma weak_monoid_homomorphism_proper {SI : sidx}
`{WeakMonoidHomomorphism M1 M2 o1 o2 R f} : Proper (() ==> ()) f.
Proof. apply ne_proper, monoid_homomorphism_ne. Qed.
From iris.algebra Require Export cmra.
From iris.algebra Require Import updates local_updates.
From iris.prelude Require Import options.
(** Given a preorder [R] on a type [A] we construct the "monotone" resource
algebra [mra R] and an injection [to_mra : A → mra R] such that:
[R x y] iff [to_mra x ≼ to_mra y]
Here, [≼] is the extension order of the [mra R] resource algebra. This is
exactly what the lemma [to_mra_included] shows.
This resource algebra is useful for reasoning about monotonicity. See the
following paper for more details ([to_mra] is called "principal"):
Reasoning About Monotonicity in Separation Logic
Amin Timany and Lars Birkedal
in Certified Programs and Proofs (CPP) 2021
Note that unlike most Iris algebra constructions [mra A] works on [A : Type],
not on [A : ofe]. See the comment at [mraO] below for more information. If [A]
has an [Equiv A] (i.e., is a setoid), there are some results at the bottom of
this file. *)
Record mra {A} (R : relation A) := { mra_car : list A }.
Definition to_mra {A} {R : relation A} (a : A) : mra R :=
{| mra_car := [a] |}.
Global Arguments mra_car {_ _} _.
Section mra.
Context {SI : sidx} {A} {R : relation A}.
Implicit Types a b : A.
Implicit Types x y : mra R.
(** Pronounced [a] is below [x]. *)
Local Definition mra_below (a : A) (x : mra R) := b, b mra_car x R a b.
Local Lemma mra_below_to_mra a b : mra_below a (to_mra b) R a b.
Proof. set_solver. Qed.
(* OFE *)
Local Instance mra_equiv : Equiv (mra R) := λ x y,
a, mra_below a x mra_below a y.
Local Instance mra_equiv_equiv : Equivalence mra_equiv.
Proof. unfold mra_equiv; split; intros ?; naive_solver. Qed.
(** Generalizing [mra A] to [A : ofe] and [R : A -n> A -n> siProp] is not
obvious. It is not clear what axioms to impose on [R] for the "extension
axiom" to hold:
cmra_extend :
x ≡{n}≡ y1 ⋅ y2 →
∃ z1 z2, x ≡ z1 ⋅ z2 ∧ y1 ≡{n}≡ z1 ∧ y2 ≡{n}≡ z2
To prove this, assume ([⋅] is defined as [++], see [mra_op]):
x ≡{n}≡ y1 ++ y2
When defining [dist] as the step-indexed version of [mra_equiv], this means:
∀ n' a, n' ≤ n →
mra_below a x n' ↔ mra_below a y1 n' ∨ mra_below a y2 n'
From this assumption it is not clear how to obtain witnesses [z1] and [z2]. *)
Canonical Structure mraO := discreteO (mra R).
(* CMRA *)
Local Instance mra_valid : Valid (mra R) := λ x, True.
Local Instance mra_validN : ValidN (mra R) := λ n x, True.
Local Program Instance mra_op : Op (mra R) := λ x y,
{| mra_car := mra_car x ++ mra_car y |}.
Local Instance mra_pcore : PCore (mra R) := Some.
Lemma mra_cmra_mixin : CmraMixin (mra R).
Proof.
apply discrete_cmra_mixin; first apply _.
apply ra_total_mixin; try done.
- (* [Proper] of [op] *) intros x y z Hyz a. specialize (Hyz a). set_solver.
- (* [Proper] of [core] *) apply _.
- (* [Assoc] *) intros x y z a. set_solver.
- (* [Comm] *) intros x y a. set_solver.
- (* [core x ⋅ x ≡ x] *) intros x a. set_solver.
Qed.
Canonical Structure mraR : cmra := Cmra (mra R) mra_cmra_mixin.
Global Instance mra_cmra_total : CmraTotal mraR.
Proof. rewrite /CmraTotal; eauto. Qed.
Global Instance mra_core_id x : CoreId x.
Proof. by constructor. Qed.
Global Instance mra_cmra_discrete : CmraDiscrete mraR.
Proof. split; last done. intros ? ?; done. Qed.
Local Instance mra_unit : Unit (mra R) := {| mra_car := [] |}.
Lemma auth_ucmra_mixin : UcmraMixin (mra R).
Proof. split; done. Qed.
Canonical Structure mraUR := Ucmra (mra R) auth_ucmra_mixin.
(* Laws *)
Lemma mra_idemp x : x x x.
Proof. intros a. set_solver. Qed.
Lemma mra_included x y : x y y x y.
Proof using SI.
split; [|by intros ?; exists y].
intros [z ->]; rewrite assoc mra_idemp; done.
Qed.
Lemma to_mra_R_op `{!Transitive R} a b :
R a b
to_mra a to_mra b to_mra b.
Proof. intros Hab c. set_solver. Qed.
Lemma to_mra_included `{!PreOrder R} a b :
to_mra a to_mra b R a b.
Proof.
split.
- move=> [z Hz]. specialize (Hz a). set_solver.
- intros ?; exists (to_mra b). by rewrite to_mra_R_op.
Qed.
Lemma mra_local_update_grow `{!Transitive R} a x b:
R a b
(to_mra a, x) ~l~> (to_mra b, to_mra b).
Proof.
intros Hana. apply local_update_unital_discrete=> z _ Habz.
split; first done. intros c. specialize (Habz c). set_solver.
Qed.
Lemma mra_local_update_get_frag `{!PreOrder R} a b:
R b a
(to_mra a, ε) ~l~> (to_mra a, to_mra b).
Proof.
intros Hana. apply local_update_unital_discrete=> z _.
rewrite left_id. intros <-. split; first done.
apply mra_included; by apply to_mra_included.
Qed.
End mra.
Global Arguments mraO {_ _} _.
Global Arguments mraR {_ _} _.
Global Arguments mraUR {_ _} _.
(** If [R] is a partial order, relative to a reflexive relation [S] on the
carrier [A], then [to_mra] is proper and injective. The theory for
arbitrary relations [S] is overly general, so we do not declare the results
as instances. Below we provide instances for [S] being [=] and [≡]. *)
Section mra_over_rel.
Context {SI : sidx} {A} {R : relation A} (S : relation A).
Implicit Types a b : A.
Implicit Types x y : mra R.
Lemma to_mra_rel_proper :
Reflexive S
Proper (S ==> S ==> iff) R
Proper (S ==> (≡@{mra R})) (to_mra).
Proof. intros ? HR a1 a2 Ha b. rewrite !mra_below_to_mra. by apply HR. Qed.
Lemma to_mra_rel_inj :
Reflexive R
AntiSymm S R
Inj S (≡@{mra R}) (to_mra).
Proof.
intros ?? a b Hab. move: (Hab a) (Hab b). rewrite !mra_below_to_mra.
intros. apply (anti_symm R); naive_solver.
Qed.
End mra_over_rel.
Global Instance to_mra_inj {SI : sidx} {A} {R : relation A} :
Reflexive R
AntiSymm (=) R
Inj (=) (≡@{mra R}) (to_mra) | 0. (* Lower cost than [to_mra_equiv_inj] *)
Proof. intros. by apply (to_mra_rel_inj (=)). Qed.
Global Instance to_mra_proper {SI : sidx} `{Equiv A} {R : relation A} :
Reflexive (≡@{A})
Proper (() ==> () ==> iff) R
Proper (() ==> (≡@{mra R})) (to_mra).
Proof. intros. by apply (to_mra_rel_proper ()). Qed.
Global Instance to_mra_equiv_inj {SI : sidx} `{Equiv A} {R : relation A} :
Reflexive R
AntiSymm () R
Inj () (≡@{mra R}) (to_mra) | 1.
Proof. intros. by apply (to_mra_rel_inj ()). Qed.
From iris.algebra Require Export cmra local_updates proofmode_classes.
From iris.prelude Require Import options.
(** ** Natural numbers with [add] as the operation. *)
Section nat.
Context {SI : sidx}.
Local Instance nat_valid_instance : Valid nat := λ x, True.
Local Instance nat_validN_instance : ValidN nat := λ n x, True.
Local Instance nat_pcore_instance : PCore nat := λ x, Some 0.
Local Instance nat_op_instance : Op nat := plus.
Definition nat_op x y : x y = x + y := eq_refl.
Lemma nat_included (x y : nat) : x y x y.
Proof. by rewrite Nat.le_sum. Qed.
Lemma nat_ra_mixin : RAMixin nat.
Proof.
apply ra_total_mixin; try by eauto.
- solve_proper.
- intros x y z. apply Nat.add_assoc.
- intros x y. apply Nat.add_comm.
- by exists 0.
Qed.
Canonical Structure natR : cmra := discreteR nat nat_ra_mixin.
Global Instance nat_cmra_discrete : CmraDiscrete natR.
Proof. apply discrete_cmra_discrete. Qed.
Local Instance nat_unit_instance : Unit nat := 0.
Lemma nat_ucmra_mixin : UcmraMixin nat.
Proof. split; apply _ || done. Qed.
Canonical Structure natUR : ucmra := Ucmra nat nat_ucmra_mixin.
Global Instance nat_cancelable (x : nat) : Cancelable x.
Proof. by intros ???? ?%Nat.add_cancel_l. Qed.
Lemma nat_local_update (x y x' y' : nat) :
(** Morally [x - y = x' - y']: the difference between auth and frag must
stay the same with this update. Written using [+] due to underflow. *)
x + y' = x' + y (x,y) ~l~> (x',y').
Proof.
intros ??; apply local_update_unital_discrete=> z _.
compute -[minus plus]; lia.
Qed.
(* This one has a higher precendence than [is_op_op] so we get a [+] instead
of an [⋅]. *)
Global Instance nat_is_op (n1 n2 : nat) : IsOp (n1 + n2) n1 n2.
Proof. done. Qed.
End nat.
(** ** Natural numbers with [max] as the operation. *)
Record max_nat := MaxNat { max_nat_car : nat }.
Add Printing Constructor max_nat.
Canonical Structure max_natO {SI : sidx} := leibnizO max_nat.
Section max_nat.
Context {SI : sidx}.
Local Instance max_nat_unit_instance : Unit max_nat := MaxNat 0.
Local Instance max_nat_valid_instance : Valid max_nat := λ x, True.
Local Instance max_nat_validN_instance : ValidN max_nat := λ n x, True.
Local Instance max_nat_pcore_instance : PCore max_nat := Some.
Local Instance max_nat_op_instance : Op max_nat := λ n m, MaxNat (max_nat_car n `max` max_nat_car m).
Definition max_nat_op x y : MaxNat x MaxNat y = MaxNat (x `max` y) := eq_refl.
Lemma max_nat_included x y : x y max_nat_car x max_nat_car y.
Proof.
split.
- intros [z ->]. simpl. lia.
- exists y. rewrite /op /max_nat_op_instance. rewrite Nat.max_r; last lia. by destruct y.
Qed.
Lemma max_nat_ra_mixin : RAMixin max_nat.
Proof.
apply ra_total_mixin; apply _ || eauto.
- intros [x] [y] [z]. repeat rewrite max_nat_op. by rewrite Nat.max_assoc.
- intros [x] [y]. by rewrite max_nat_op Nat.max_comm.
- intros [x]. by rewrite max_nat_op Nat.max_id.
Qed.
Canonical Structure max_natR : cmra := discreteR max_nat max_nat_ra_mixin.
Global Instance max_nat_cmra_discrete : CmraDiscrete max_natR.
Proof. apply discrete_cmra_discrete. Qed.
Lemma max_nat_ucmra_mixin : UcmraMixin max_nat.
Proof. split; try apply _ || done. intros [x]. done. Qed.
Canonical Structure max_natUR : ucmra := Ucmra max_nat max_nat_ucmra_mixin.
Global Instance max_nat_core_id (x : max_nat) : CoreId x.
Proof. by constructor. Qed.
Lemma max_nat_local_update (x y x' : max_nat) :
max_nat_car x max_nat_car x' (x,y) ~l~> (x',x').
Proof.
move: x y x' => [x] [y] [y'] /= ?.
rewrite local_update_unital_discrete=> [[z]] _.
rewrite 2!max_nat_op. intros [= ?].
split; first done. apply f_equal. lia.
Qed.
Global Instance : IdemP (=@{max_nat}) ().
Proof. intros [x]. rewrite max_nat_op. apply f_equal. lia. Qed.
Global Instance max_nat_is_op (x y : nat) :
IsOp (MaxNat (x `max` y)) (MaxNat x) (MaxNat y).
Proof. done. Qed.
End max_nat.
(** ** Natural numbers with [min] as the operation. *)
Record min_nat := MinNat { min_nat_car : nat }.
Add Printing Constructor min_nat.
Canonical Structure min_natO {SI : sidx} := leibnizO min_nat.
Section min_nat.
Context {SI : sidx}.
Local Instance min_nat_valid_instance : Valid min_nat := λ x, True.
Local Instance min_nat_validN_instance : ValidN min_nat := λ n x, True.
Local Instance min_nat_pcore_instance : PCore min_nat := Some.
Local Instance min_nat_op_instance : Op min_nat := λ n m, MinNat (min_nat_car n `min` min_nat_car m).
Definition min_nat_op_min x y : MinNat x MinNat y = MinNat (x `min` y) := eq_refl.
Lemma min_nat_included (x y : min_nat) : x y min_nat_car y min_nat_car x.
Proof.
split.
- intros [z ->]. simpl. lia.
- exists y. rewrite /op /min_nat_op_instance. rewrite Nat.min_r; last lia. by destruct y.
Qed.
Lemma min_nat_ra_mixin : RAMixin min_nat.
Proof.
apply ra_total_mixin; apply _ || eauto.
- intros [x] [y] [z]. repeat rewrite min_nat_op_min. by rewrite Nat.min_assoc.
- intros [x] [y]. by rewrite min_nat_op_min Nat.min_comm.
- intros [x]. by rewrite min_nat_op_min Nat.min_id.
Qed.
Canonical Structure min_natR : cmra := discreteR min_nat min_nat_ra_mixin.
Global Instance min_nat_cmra_discrete : CmraDiscrete min_natR.
Proof. apply discrete_cmra_discrete. Qed.
Global Instance min_nat_core_id (x : min_nat) : CoreId x.
Proof. by constructor. Qed.
Lemma min_nat_local_update (x y x' : min_nat) :
min_nat_car x' min_nat_car x (x,y) ~l~> (x',x').
Proof.
move: x y x' => [x] [y] [x'] /= ?.
rewrite local_update_discrete. move=> [[z]|] _ /=; last done.
rewrite 2!min_nat_op_min. intros [= ?].
split; first done. apply f_equal. lia.
Qed.
Global Instance : LeftAbsorb (=) (MinNat 0) ().
Proof. done. Qed.
Global Instance : RightAbsorb (=) (MinNat 0) ().
Proof. intros [x]. by rewrite min_nat_op_min Nat.min_0_r. Qed.
Global Instance : IdemP (=@{min_nat}) ().
Proof. intros [x]. rewrite min_nat_op_min. apply f_equal. lia. Qed.
Global Instance min_nat_is_op (x y : nat) :
IsOp (MinNat (x `min` y)) (MinNat x) (MinNat y).
Proof. done. Qed.
End min_nat.
(** ** Positive integers with [Pos.add] as the operation. *)
Section positive.
Context {SI : sidx}.
Local Instance pos_valid_instance : Valid positive := λ x, True.
Local Instance pos_validN_instance : ValidN positive := λ n x, True.
Local Instance pos_pcore_instance : PCore positive := λ x, None.
Local Instance pos_op_instance : Op positive := Pos.add.
Definition pos_op_add x y : x y = (x + y)%positive := eq_refl.
Lemma pos_included (x y : positive) : x y (x < y)%positive.
Proof. by rewrite Pos.lt_sum. Qed.
Lemma pos_ra_mixin : RAMixin positive.
Proof.
split; try by eauto.
- by intros ??? ->.
- intros ???. apply Pos.add_assoc.
- intros ??. apply Pos.add_comm.
Qed.
Canonical Structure positiveR : cmra := discreteR positive pos_ra_mixin.
Global Instance pos_cmra_discrete : CmraDiscrete positiveR.
Proof. apply discrete_cmra_discrete. Qed.
Global Instance pos_cancelable (x : positive) : Cancelable x.
Proof. intros n y z ??. by eapply Pos.add_reg_l, leibniz_equiv. Qed.
Global Instance pos_id_free (x : positive) : IdFree x.
Proof.
intros y ??. apply (Pos.add_no_neutral x y). rewrite Pos.add_comm.
by apply leibniz_equiv.
Qed.
(* This one has a higher precendence than [is_op_op] so we get a [+] instead
of an [⋅]. *)
Global Instance pos_is_op (x y : positive) : IsOp (x + y)%positive x y.
Proof. done. Qed.
End positive.
(** ** Integers (positive and negative) with [Z.add] as the operation. *)
Section Z.
Context {SI : sidx}.
Local Open Scope Z_scope.
Local Instance Z_valid_instance : Valid Z := λ x, True.
Local Instance Z_validN_instance : ValidN Z := λ n x, True.
Local Instance Z_pcore_instance : PCore Z := λ x, Some 0.
Local Instance Z_op_instance : Op Z := Z.add.
Definition Z_op x y : x y = x + y := eq_refl.
Lemma Z_ra_mixin : RAMixin Z.
Proof.
apply ra_total_mixin; try by eauto.
- solve_proper.
- intros x y z. apply Z.add_assoc.
- intros x y. apply Z.add_comm.
- by exists 0.
Qed.
Canonical Structure ZR : cmra := discreteR Z Z_ra_mixin.
Global Instance Z_cmra_discrete : CmraDiscrete ZR.
Proof. apply discrete_cmra_discrete. Qed.
Local Instance Z_unit_instance : Unit Z := 0.
Lemma Z_ucmra_mixin : UcmraMixin Z.
Proof. split; apply _ || done. Qed.
Canonical Structure ZUR : ucmra := Ucmra Z Z_ucmra_mixin.
Global Instance Z_cancelable (x : Z) : Cancelable x.
Proof. by intros ???? ?%Z.add_cancel_l. Qed.
(** The difference between auth and frag must stay the same with this update. *)
Lemma Z_local_update (x y x' y' : Z) :
x - y = x' - y' (x,y) ~l~> (x',y').
Proof.
intros. rewrite local_update_unital_discrete=> z _.
compute -[Z.sub Z.add]; lia.
Qed.
(* This one has a higher precendence than [is_op_op] so we get a [+] instead
of an [⋅]. *)
Global Instance Z_is_op (n1 n2 : Z) : IsOp (n1 + n2) n1 n2.
Proof. done. Qed.
End Z.
(** ** Integers (positive and negative) with [Z.max] as the operation. *)
Record max_Z := MaxZ { max_Z_car : Z }.
Add Printing Constructor max_Z.
Canonical Structure max_ZO {SI : sidx} := leibnizO max_Z.
Section max_Z.
Context {SI : sidx}.
Local Open Scope Z_scope.
Local Instance max_Z_unit_instance : Unit max_Z := MaxZ 0.
Local Instance max_Z_valid_instance : Valid max_Z := λ x, True.
Local Instance max_Z_validN_instance : ValidN max_Z := λ n x, True.
Local Instance max_Z_pcore_instance : PCore max_Z := Some.
Local Instance max_Z_op_instance : Op max_Z := λ n m,
MaxZ (max_Z_car n `max` max_Z_car m).
Definition max_Z_op x y : MaxZ x MaxZ y = MaxZ (x `max` y) := eq_refl.
Lemma max_Z_included x y : x y max_Z_car x max_Z_car y.
Proof.
split.
- intros [z ->]. simpl. lia.
- exists y. rewrite /op /max_Z_op_instance. rewrite Z.max_r; last lia.
by destruct y.
Qed.
Lemma max_Z_ra_mixin : RAMixin max_Z.
Proof.
apply ra_total_mixin; apply _ || eauto.
- intros [x] [y] [z]. repeat rewrite max_Z_op. by rewrite Z.max_assoc.
- intros [x] [y]. by rewrite max_Z_op Z.max_comm.
- intros [x]. by rewrite max_Z_op Z.max_id.
Qed.
Canonical Structure max_ZR : cmra := discreteR max_Z max_Z_ra_mixin.
Global Instance max_Z_cmra_total : CmraTotal max_ZR.
Proof. intros x. eauto. Qed.
Global Instance max_Z_cmra_discrete : CmraDiscrete max_ZR.
Proof. apply discrete_cmra_discrete. Qed.
Global Instance max_Z_core_id (x : max_Z) : CoreId x.
Proof. by constructor. Qed.
Lemma max_Z_local_update (x y x' : max_Z) :
max_Z_car x max_Z_car x' (x,y) ~l~> (x',x').
Proof.
move: x y x' => [x] [y] [y'] /= ?.
rewrite local_update_discrete=> [[[z]|]] //= _ [?].
split; first done. rewrite max_Z_op. f_equal. lia.
Qed.
Global Instance : IdemP (=@{max_Z}) ().
Proof. intros [x]. rewrite max_Z_op. apply f_equal. lia. Qed.
Global Instance max_Z_is_op (x y : nat) :
IsOp (MaxZ (x `max` y)) (MaxZ x) (MaxZ y).
Proof. done. Qed.
End max_Z.
From iris.algebra Require Export base.
Set Default Proof Using "Type".
Set Primitive Projections.
(** This files defines (a shallow embedding of) the category of OFEs:
Complete ordered families of equivalences. This is a cartesian closed
category, and mathematically speaking, the entire development lives
in this category. However, we will generally prefer to work with raw
Coq functions plus some registered Proper instances for non-expansiveness.
This makes writing such functions much easier. It turns out that it many
cases, we do not even need non-expansiveness.
*)
From iris.prelude Require Export prelude.
From iris.prelude Require Import options.
From iris.algebra Require Export stepindex.
Local Set Primitive Projections.
Local Open Scope sidx_scope.
(** This files defines (a shallow embedding of) the category of OFEs: Complete
ordered families of equivalences. This is a cartesian closed category, and
mathematically speaking, the entire development lives in this category. However,
we will generally prefer to work with raw Coq functions plus some registered
[Proper] instances for non-expansiveness. This makes writing such functions much
easier. It turns out that in many cases, we do not even need non-expansiveness. *)
(** Unbundled version *)
Class Dist A := dist : nat relation A.
Instance: Params (@dist) 3 := {}.
Class Dist {SI : sidx} A := dist : SI relation A.
Global Hint Mode Dist - ! : typeclass_instances.
Global Instance: Params (@dist) 4 := {}.
Notation "x ≡{ n }≡ y" := (dist n x y)
(at level 70, n at next level, format "x ≡{ n }≡ y").
Notation "x ≡{ n }@{ A }≡ y" := (dist (A:=A) n x y)
(at level 70, n at next level, only parsing).
Notation "(≡{ n }≡)" := (dist n) (only parsing).
Notation "(≡{ n }@{ A }≡)" := (dist (A:=A) n) (only parsing).
Notation "( x ≡{ n }≡.)" := (dist n x) (only parsing).
Notation "(.≡{ n }≡ y )" := (λ x, x {n} y) (only parsing).
Hint Extern 0 (_ {_} _) => reflexivity : core.
Hint Extern 0 (_ {_} _) => symmetry; assumption : core.
Global Hint Extern 0 (_ {_} _) => reflexivity : core.
Global Hint Extern 0 (_ {_} _) => symmetry; assumption : core.
Notation NonExpansive f := ( n, Proper (dist n ==> dist n) f).
Notation NonExpansive2 f := ( n, Proper (dist n ==> dist n ==> dist n) f).
Notation NonExpansive3 f :=
( n, Proper (dist n ==> dist n ==> dist n ==> dist n) f).
Notation NonExpansive4 f :=
( n, Proper (dist n ==> dist n ==> dist n ==> dist n ==> dist n) f).
Tactic Notation "ofe_subst" ident(x) :=
repeat match goal with
| _ => progress simplify_eq/=
| H:@dist ?A ?d ?n x _ |- _ => setoid_subst_aux (@dist A d n) x
| H:@dist ?A ?d ?n _ x |- _ => symmetry in H;setoid_subst_aux (@dist A d n) x
| H : @dist ?SI ?A ?d ?n x _ |- _ => setoid_subst_aux (@dist SI A d n) x
| H : @dist ?SI ?A ?d ?n _ x |- _ =>
symmetry in H; setoid_subst_aux (@dist SI A d n) x
end.
Tactic Notation "ofe_subst" :=
repeat match goal with
| _ => progress simplify_eq/=
| H:@dist ?A ?d ?n ?x _ |- _ => setoid_subst_aux (@dist A d n) x
| H:@dist ?A ?d ?n _ ?x |- _ => symmetry in H;setoid_subst_aux (@dist A d n) x
| H : @dist ?SI ?A ?d ?n ?x _ |- _ => setoid_subst_aux (@dist SI A d n) x
| H : @dist ?SI ?A ?d ?n _ ?x |- _ =>
symmetry in H; setoid_subst_aux (@dist SI A d n) x
end.
Record OfeMixin A `{Equiv A, Dist A} := {
mixin_equiv_dist x y : x y n, x {n} y;
mixin_dist_equivalence n : Equivalence (dist n);
mixin_dist_S n x y : x {S n} y x {n} y
Record OfeMixin {SI : sidx} A `{Equiv A, !Dist A} := {
mixin_equiv_dist (x y : A) : x y n, x {n} y;
mixin_dist_equivalence n : Equivalence (@dist SI A _ n);
mixin_dist_le n m (x y : A) : x {n} y m n x {m} y
}.
(** Bundled version *)
Structure ofeT := OfeT {
Structure ofe {SI : sidx} := Ofe {
ofe_car :> Type;
ofe_equiv : Equiv ofe_car;
ofe_dist : Dist ofe_car;
ofe_mixin : OfeMixin ofe_car
}.
Arguments OfeT _ {_ _} _.
Add Printing Constructor ofeT.
Hint Extern 0 (Equiv _) => eapply (@ofe_equiv _) : typeclass_instances.
Hint Extern 0 (Dist _) => eapply (@ofe_dist _) : typeclass_instances.
Arguments ofe_car : simpl never.
Arguments ofe_equiv : simpl never.
Arguments ofe_dist : simpl never.
Arguments ofe_mixin : simpl never.
Global Arguments Ofe {_} _ {_ _} _.
Add Printing Constructor ofe.
(* FIXME(Coq #6294) : we need the new unification algorithm here. *)
Global Hint Extern 0 (Equiv _) => refine (ofe_equiv _); shelve : typeclass_instances.
Global Hint Extern 0 (Dist _) => refine (ofe_dist _); shelve : typeclass_instances.
Global Arguments ofe_car : simpl never.
Global Arguments ofe_equiv : simpl never.
Global Arguments ofe_dist : simpl never.
Global Arguments ofe_mixin : simpl never.
(** When declaring instances of subclasses of OFE (like CMRAs and unital CMRAs)
we need Coq to *infer* the canonical OFE instance of a given type and take the
mixin out of it. This makes sure we do not use two different OFE instances in
different places (see for example the constructors [CmraT] and [UcmraT] in the
different places (see for example the constructors [Cmra] and [Ucmra] in the
file [cmra.v].)
In order to infer the OFE instance, we use the definition [ofe_mixin_of'] which
......@@ -80,69 +93,163 @@ The notation [ofe_mixin_of A] that we define on top of [ofe_mixin_of' A id]
hides the [id] and normalizes the mixin to head normal form. The latter is to
ensure that we do not end up with redundant canonical projections to the mixin,
i.e. them all being of the shape [ofe_mixin_of' A id]. *)
Definition ofe_mixin_of' A {Ac : ofeT} (f : Ac A) : OfeMixin Ac := ofe_mixin Ac.
Definition ofe_mixin_of' {SI : sidx} A {Ac : ofe} (f : Ac A) : OfeMixin Ac := ofe_mixin Ac.
Notation ofe_mixin_of A :=
ltac:(let H := eval hnf in (ofe_mixin_of' A id) in exact H) (only parsing).
(** Lifting properties from the mixin *)
Section ofe_mixin.
Context {A : ofeT}.
Context {SI : sidx} {A : ofe}.
Implicit Types x y : A.
Lemma equiv_dist x y : x y n, x {n} y.
Proof. apply (mixin_equiv_dist _ (ofe_mixin A)). Qed.
Global Instance dist_equivalence n : Equivalence (@dist A _ n).
Global Instance dist_equivalence n : Equivalence (@dist SI A _ n).
Proof. apply (mixin_dist_equivalence _ (ofe_mixin A)). Qed.
Lemma dist_S n x y : x {S n} y x {n} y.
Proof. apply (mixin_dist_S _ (ofe_mixin A)). Qed.
Lemma dist_le n m x y : x {n} y m n x {m} y.
Proof. apply (mixin_dist_le _ (ofe_mixin A)). Qed.
End ofe_mixin.
Hint Extern 1 (_ {_} _) => apply equiv_dist; assumption : core.
Global Hint Extern 1 (_ {_} _) => apply equiv_dist; assumption : core.
(** Discrete OFEs and discrete OFE elements *)
Class Discrete {A : ofeT} (x : A) := discrete y : x {0} y x y.
Arguments discrete {_} _ {_} _ _.
Hint Mode Discrete + ! : typeclass_instances.
Instance: Params (@Discrete) 1 := {}.
Class Discrete {SI : sidx} {A : ofe} (x : A) :=
discrete_0 y : x {0} y x y.
Global Arguments discrete_0 {_ _} _ {_} _ _.
Global Hint Mode Discrete - + ! : typeclass_instances.
Global Instance: Params (@Discrete) 2 := {}.
Class OfeDiscrete (A : ofeT) := ofe_discrete_discrete (x : A) :> Discrete x.
Class OfeDiscrete {SI : sidx} (A : ofe) :=
#[global] ofe_discrete_discrete (x : A) :: Discrete x.
Global Hint Mode OfeDiscrete - ! : typeclass_instances.
(** OFEs with a completion *)
Record chain (A : ofeT) := {
chain_car :> nat A;
chain_cauchy n i : n i chain_car i {n} chain_car n
(** A (converging) "chain" is a sequence of type [A] such that from the n-th
position onwards, everything is n-equal. *)
Record chain {SI : sidx} (A : ofe) := {
chain_car :> SI A;
chain_cauchy n m: n m chain_car m {n} chain_car n
}.
Global Arguments chain_car {_ _} _ _.
Global Arguments chain_cauchy {_ _} _ _ _ _.
(** A "bounded chain" up to step index [n] is similar to a "chain", except that
the sequence has length [n], i.e., it is only defined for indices strictly less
than [n]. *)
Record bchain {SI : sidx} (A : ofe) (n : SI) := {
bchain_car :> m, m < n A;
bchain_cauchy m p Hm Hp : m p bchain_car p Hp {m} bchain_car m Hm
}.
Arguments chain_car {_} _ _.
Arguments chain_cauchy {_} _ _ _ _.
Global Arguments bchain_car {_ _} _ _ _.
Global Arguments bchain_cauchy {_ _} _ _ _ _ _.
Program Definition chain_map {A B : ofeT} (f : A B)
`{!NonExpansive f} (c : chain A) : chain B :=
Program Definition chain_map {SI : sidx} {A B : ofe} (f : A B)
`{NonExpansive f} (c : chain A) : chain B :=
{| chain_car n := f (c n) |}.
Next Obligation. by intros A B f Hf c n i ?; apply Hf, chain_cauchy. Qed.
Next Obligation. by intros SI A B f Hf c n i ?; apply Hf, chain_cauchy. Qed.
Program Definition bchain_map {SI : sidx} {A B : ofe} (f : A B)
`{NonExpansive f} {n} (c : bchain A n) : bchain B n :=
{| bchain_car m Hm := f (c m Hm) |}.
Next Obligation.
by intros SI A B f Hf n c m p ? Hm Hp; apply Hf, bchain_cauchy.
Qed.
(** We define a _complete_ OFE, COFE for short. Roughly speaking, the power of
a COFE (over an OFE) is to allow us to compute fixpoints. We want to compute two
different kinds of fixpoints:
1. Fixpoints inside of COFEs. For various kinds of recursive definitions inside
COFEs (e.g., the Iris weakest precondition or a logical relation with
recursive types), we want to compute the fixpoint of a function [f : A → A]
where [A] is a COFE. We can do so if [f] is contractive, using a variant of
Banach's fixpoint theorem. The construction of this fixpoint is given by
[fixpoint] below.
2. Fixpoints on COFEs. For step-indexed types in Iris (e.g., [iProp]), we have
to solve a recursive domain equation on COFEs. The construction of this
fixpoint for natural numbers as the step-index type is given in
[cofe_solver.v].
A COFE extends an OFE [A] with two additional operations:
1. [compl: chain A → A], which takes a chain [c] of elements from [A] and maps
them to a limit element [compl c],
2. [lbcompl: ∀ n, SIdx.limit n → bchain A n → A], which takes a
bounded chain [c] of elements from [A] and maps them to a limit element
[lbcompl c]. The chain is bounded in the sense that its domain ranges from
0 to (but not including) the limit index [n]. (Later we will define
[bcompl] which completes chains that are bounded by any index, not
necessarily a limit index.)
We will see the need for the two different limit operations of a COFE below
as part of the [fixpoint] construction. A more detailed explanation can be
found in the Iris Reference and the Transfinite Iris Documentation (see
https://iris-project.org/pdfs/2021-pldi-transfinite-iris-final-appendix.pdf). *)
(** These notations [Compl] and [LBCompl] are convenient to define instances
(e.g., [ofe_mor_compl] without having to repeat the type. The notation [BCompl]
will be used as the type for [bcompl], completing [bchain] with an arbitrary
bound [n]. *)
Notation Compl A := (chain A%type A).
Class Cofe (A : ofeT) := {
Notation BCompl A := ( n, bchain A%type n A).
Notation LBCompl A := ( n, SIdx.limit n bchain A%type n A).
Class Cofe {SI : sidx} (A : ofe) := {
compl : Compl A;
lbcompl : LBCompl A;
conv_compl n c : compl c {n} c n;
conv_lbcompl {n} Hn (c : bchain A n) {m} Hm : lbcompl n Hn c {m} c m Hm;
(** The bounded limit operation is non-expansive: for chains agreeing up to
the limit index, the bounded limits agree *)
lbcompl_ne n {Hn} (c1 c2 : bchain A n) m :
( p (Hp : p < n), c1 p Hp {m} c2 p Hp)
lbcompl n Hn c1 {m} lbcompl n Hn c2
}.
Arguments compl : simpl never.
Hint Mode Cofe ! : typeclass_instances.
Global Arguments compl : simpl never.
Global Arguments lbcompl {_ _ _ _} : simpl never.
Global Hint Mode Cofe - ! : typeclass_instances.
Lemma compl_chain_map `{Cofe A, Cofe B} (f : A B) c `(NonExpansive f) :
Lemma compl_chain_map {SI : sidx} `{!Cofe A, !Cofe B} (f : A B) c
`(!NonExpansive f) :
compl (chain_map f c) f (compl c).
Proof. apply equiv_dist=>n. by rewrite !conv_compl. Qed.
Program Definition chain_const {A : ofeT} (a : A) : chain A :=
Program Definition chain_const {SI : sidx} {A : ofe} (a : A) : chain A :=
{| chain_car n := a |}.
Next Obligation. by intros A a n i _. Qed.
Next Obligation. by intros SI A a n i _. Qed.
Lemma compl_chain_const {A : ofeT} `{!Cofe A} (a : A) :
Lemma compl_chain_const {SI : sidx} `{!Cofe A} (a : A) :
compl (chain_const a) a.
Proof. apply equiv_dist=>n. by rewrite conv_compl. Qed.
Proof. apply equiv_dist=> n. by rewrite conv_compl. Qed.
Lemma compl_bchain_map {SI : sidx} `{!Cofe A, !Cofe B} (f : A B)
n Hn (c : bchain A n) `(!NonExpansive f) m :
m < n lbcompl Hn (bchain_map f c) {m} f (lbcompl Hn c).
Proof. intros Hm. by rewrite !(conv_lbcompl _ _ Hm). Qed.
Program Definition bchain_const {SI : sidx} {A : ofe} (a : A) n : bchain A n :=
{| bchain_car m _ := a |}.
Next Obligation. by intros SI A a n m p Hm Hp Hle. Qed.
Program Definition bchain_le {SI : sidx} `{!Cofe A} {n}
(c : bchain A n) {m} (Hm : m n) : bchain A m :=
{| bchain_car m' Hm' := c m' (SIdx.lt_le_trans _ _ _ Hm' Hm) |}.
Next Obligation.
intros SI A ? n c m Hm p1 p2 Hp Hp1 Hp2; simpl. by apply (bchain_cauchy n c).
Qed.
Lemma lbcompl_bchain_le {SI : sidx} `{!Cofe A} n Hn m Hm (Hmn : m n)
(c : bchain A n) p :
p < m lbcompl Hm (bchain_le c Hmn) {p} lbcompl Hn c.
Proof.
intros Hp. rewrite (conv_lbcompl _ _ Hp).
rewrite conv_lbcompl. by rewrite /bchain_le /=.
Qed.
(** General properties *)
(** General properties of OFEs *)
Section ofe.
Context {A : ofeT}.
Context {SI : sidx} {A : ofe}.
Implicit Types x y : A.
Global Instance ofe_equivalence : Equivalence (() : relation A).
Proof.
split.
......@@ -150,187 +257,489 @@ Section ofe.
- by intros x y; rewrite !equiv_dist.
- by intros x y z; rewrite !equiv_dist; intros; trans y.
Qed.
Global Instance dist_ne n : Proper (dist n ==> dist n ==> iff) (@dist A _ n).
Global Instance dist_ne n : Proper (dist n ==> dist n ==> iff) (@dist SI A _ n).
Proof.
intros x1 x2 ? y1 y2 ?; split; intros.
- by trans x1; [|trans y1].
- by trans x2; [|trans y2].
Qed.
Global Instance dist_proper n : Proper (() ==> () ==> iff) (@dist A _ n).
Global Instance dist_proper n : Proper (() ==> () ==> iff) (@dist SI A _ n).
Proof.
by move => x1 x2 /equiv_dist Hx y1 y2 /equiv_dist Hy; rewrite (Hx n) (Hy n).
Qed.
Global Instance dist_proper_2 n x : Proper (() ==> iff) (dist n x).
Proof. by apply dist_proper. Qed.
Global Instance Discrete_proper : Proper (() ==> iff) (@Discrete A).
Global Instance Discrete_proper : Proper (() ==> iff) (@Discrete SI A).
Proof. intros x y Hxy. rewrite /Discrete. by setoid_rewrite Hxy. Qed.
Lemma dist_le n n' x y : x {n} y n' n x {n'} y.
Proof. induction 2; eauto using dist_S. Qed.
Lemma dist_lt n m x y : x {n} y m < n x {m} y.
Proof. eauto using dist_le, SIdx.lt_le_incl. Qed.
Lemma dist_le' n n' x y : n' n x {n} y x {n'} y.
Proof. intros; eauto using dist_le. Qed.
Instance ne_proper {B : ofeT} (f : A B) `{!NonExpansive f} :
Proper (() ==> ()) f | 100.
Proof. eauto using dist_le. Qed.
Lemma dist_S n x y : x {Sᵢ n} y x {n} y.
Proof. intros H. eapply dist_lt; eauto using SIdx.lt_succ_diag_r. Qed.
(** [ne_proper] and [ne_proper_2] are not instances to improve efficiency of
type class search during setoid rewriting.
Local Instances of [NonExpansive{,2}] are hence accompanied by instances of
[Proper] built using these lemmas. *)
Lemma ne_proper {B : ofe} (f : A B) `{!NonExpansive f} :
Proper (() ==> ()) f.
Proof. by intros x1 x2; rewrite !equiv_dist; intros Hx n; rewrite (Hx n). Qed.
Instance ne_proper_2 {B C : ofeT} (f : A B C) `{!NonExpansive2 f} :
Proper (() ==> () ==> ()) f | 100.
Lemma ne_proper_2 {B C : ofe} (f : A B C) `{!NonExpansive2 f} :
Proper (() ==> () ==> ()) f.
Proof.
unfold Proper, respectful; setoid_rewrite equiv_dist.
by intros x1 x2 Hx y1 y2 Hy n; rewrite (Hx n) (Hy n).
Qed.
Lemma conv_compl' `{Cofe A} n (c : chain A) : compl c {n} c (S n).
Lemma conv_compl_le `{!Cofe A} n m (c : chain A) : n m compl c {n} c m.
Proof.
transitivity (c n); first by apply conv_compl. symmetry.
apply chain_cauchy. lia.
transitivity (c n); first by rewrite conv_compl.
symmetry. by rewrite chain_cauchy.
Qed.
Lemma conv_compl_S `{!Cofe A} n (c : chain A) : compl c {n} c (Sᵢ n).
Proof. apply conv_compl_le, SIdx.le_succ_diag_r. Qed.
Lemma discrete_iff n (x : A) `{!Discrete x} y : x y x {n} y.
Proof.
split; intros; auto. apply (discrete _), dist_le with n; auto with lia.
split; intros; [by auto|].
apply (discrete_0 _), dist_le with n; eauto using SIdx.le_0_l.
Qed.
Lemma discrete_iff_0 n (x : A) `{!Discrete x} y : x {0} y x {n} y.
Lemma discrete_iff_0 n (x : A) `{!Discrete x} y : x {0} y x {n} y.
Proof. by rewrite -!discrete_iff. Qed.
Lemma discrete n (x : A) `{!Discrete x} y : x {n} y x y.
Proof. intros. eapply discrete_iff; done. Qed.
Global Instance ofe_discrete_subrelation `{!OfeDiscrete A} n :
@SolveProperSubrelation A (dist n) ().
Proof. intros ???. apply: discrete. done. Qed.
Global Instance ofe_leibniz_subrelation `{!OfeDiscrete A, !LeibnizEquiv A} n :
@SolveProperSubrelation A (dist n) (=).
Proof. intros ?? EQ. unfold_leibniz. apply (is_solve_proper_subrelation EQ). Qed.
End ofe.
(** Contractive functions *)
Definition dist_later `{Dist A} (n : nat) (x y : A) : Prop :=
match n with 0 => True | S n => x {n} y end.
Arguments dist_later _ _ !_ _ _ /.
(** Defined as a record to avoid eager unfolding. *)
Record dist_later {SI : sidx} `{!Dist A} (n : SI) (x y : A) : Prop :=
{ dist_later_lt : m, m < n x {m} y }.
Global Instance dist_later_equivalence (A : ofeT) n : Equivalence (@dist_later A _ n).
Proof. destruct n as [|n]. by split. apply dist_equivalence. Qed.
Section dist_later.
Context {SI : sidx} {A : ofe}.
Implicit Types x y : A.
Lemma dist_dist_later {A : ofeT} n (x y : A) : dist n x y dist_later n x y.
Proof. intros Heq. destruct n; first done. exact: dist_S. Qed.
Global Instance dist_later_equivalence n : Equivalence (@dist_later SI A _ n).
Proof.
split.
- intros ?; by split.
- intros ?? [H]; split; intros ??; by rewrite H.
- intros ??? [H1] [H2]; split; intros ??; by rewrite H1 ?H2.
Qed.
Lemma dist_later_dist {A : ofeT} n (x y : A) : dist_later (S n) x y dist n x y.
Proof. done. Qed.
Lemma dist_dist_later n x y : dist n x y dist_later n x y.
Proof. intros. split; eauto using dist_lt. Qed.
Lemma dist_later_dist_lt n m (x y : A) : m < n dist_later n x y dist m x y.
Proof. intros ? H; by apply H. Qed.
Lemma dist_later_0 x y : dist_later 0 x y.
Proof. split. intros ? []%SIdx.nlt_0_r. Qed.
Lemma dist_later_S n x y : x {n} y dist_later (Sᵢ n) x y.
Proof.
split.
- intros Hn; split; intros m Hm%SIdx.lt_succ_r. by eapply dist_le.
- intros Hdist. by apply Hdist, SIdx.lt_succ_r.
Qed.
End dist_later.
(* We don't actually need this lemma (as our tactics deal with this through
other means), but technically speaking, this is the reason why
pre-composing a non-expansive function to a contractive function
preserves contractivity. *)
Lemma ne_dist_later {A B : ofeT} (f : A B) :
Lemma ne_dist_later {SI : sidx} {A B : ofe} (f : A B) :
NonExpansive f n, Proper (dist_later n ==> dist_later n) f.
Proof. intros Hf [|n]; last exact: Hf. hnf. by intros. Qed.
Proof. intros Hf ??? [H]; split; intros ??; by eapply Hf, H. Qed.
Notation Contractive f := ( n, Proper (dist_later n ==> dist n) f).
Instance const_contractive {A B : ofeT} (x : A) : Contractive (@const A B x).
Global Instance const_contractive {SI : sidx} {A B : ofe} (x : A) :
Contractive (@const A B x).
Proof. by intros n y1 y2. Qed.
Section contractive.
Local Set Default Proof Using "Type*".
Context {A B : ofeT} (f : A B) `{!Contractive f}.
Context {SI : sidx} {A B : ofe} (f : A B) `{!Contractive f}.
Implicit Types x y : A.
Lemma contractive_0 x y : f x {0} f y.
Lemma contractive_0 x y : f x {0} f y.
Proof. by apply (_ : Contractive f), dist_later_0. Qed.
Lemma contractive_dist_later_dist n x y : dist_later n x y f x {n} f y.
Proof. by apply (_ : Contractive f). Qed.
Lemma contractive_S n x y : x {n} y f x {S n} f y.
Proof. intros. by apply (_ : Contractive f). Qed.
Lemma contractive_S n x y : x {n} y f x {S n} f y.
Proof. intros. by apply contractive_dist_later_dist, dist_later_S. Qed.
Global Instance contractive_ne : NonExpansive f | 100.
Proof. by intros n x y ?; apply dist_S, contractive_S. Qed.
Proof.
intros n x y ?; eapply (dist_lt (Sᵢ n)), SIdx.lt_succ_diag_r.
eapply contractive_dist_later_dist. split.
intros ??%SIdx.lt_succ_r. by eapply dist_le.
Qed.
Global Instance contractive_proper : Proper (() ==> ()) f | 100.
Proof. apply (ne_proper _). Qed.
End contractive.
Ltac f_contractive :=
Lemma dist_pointwise_lt {SI : sidx} {A} {B : ofe} n m (f g : A B):
m < n
pointwise_relation A (dist_later n) f g
pointwise_relation A (dist m) f g.
Proof. intros Hlt Hp a. by apply Hp. Qed.
(** The tactic [f_contractive] can be used to prove contractiveness or
non-expansiveness of a function [f]. Inside of the proof of
contractiveness/non-expansiveness, if the current goal is
[g x1 ... xn ≡{i}≡ g y1 ... yn]
for a contractive function [g] (that is used inside of the body of [f]),
then the tactic will try to find a suitable [Contractive] instance for [g]
and apply it. Currently, the tactic only supports one (i.e., [n = 1]) and
two (i.e., [n = 2]) arguments. As a result of applying the [Contractive]
instance for [g], one of the goals will be [dist_later i xi yi] and the tactic
will try to simplify or solve the goal. By simplify we mean that it will
turn hypotheses [dist_later] into [dist].
The tactic [f_contractive] is implemented using
1. [f_contractive_prepare] which looks up a [Contractive] looks at which
function is being applied on both sides of a [dist], looks up the
[Contractive] instance (or the equivalent for two arguments) and applies it.
2. [dist_later_intro] introduces the resulting goals with [dist_later n x y]. *)
Ltac f_contractive_prepare :=
match goal with
| |- ?f _ {_} ?f _ => simple apply (_ : Proper (dist_later _ ==> dist _) f)
| |- ?f _ _ {_} ?f _ _ => simple apply (_ : Proper (dist_later _ ==> _ ==> dist _) f)
| |- ?f _ _ {_} ?f _ _ => simple apply (_ : Proper (_ ==> dist_later _ ==> dist _) f)
end.
(** For the goal [dist_later n x y], the tactic [dist_later_intro as m Hm]
introduces a smaller step-index [Hm : m < n] and tries to lower assumptions in
the context to [m] where possible. The arguments [m] and [Hm] can be omitted,
in which case a fresh identifier is used. *)
Tactic Notation "dist_later_intro" "as" ident(idxName) ident(ltName) :=
match goal with
| |- ?f _ {_} ?f _ => simple apply (_ : Proper (dist_later _ ==> _) f)
| |- ?f _ _ {_} ?f _ _ => simple apply (_ : Proper (dist_later _ ==> _ ==> _) f)
| |- ?f _ _ {_} ?f _ _ => simple apply (_ : Proper (_ ==> dist_later _ ==> _) f)
end;
try match goal with
| |- @dist_later ?A _ ?n ?x ?y =>
destruct n as [|n]; [exact I|change (@dist A _ n x y)]
end;
try simple apply reflexivity.
| |- dist_later ?n ?x ?y =>
constructor; intros idxName ltName;
repeat match goal with
| H: dist_later n _ _ |- _ => destruct H as [H]; specialize (H idxName ltName) as H
| H: pointwise_relation _ (dist_later n) _ _ |- _ =>
apply (dist_pointwise_lt _ idxName _ _ ltName) in H
end
end.
Tactic Notation "dist_later_intro" :=
let m := fresh "m" in
let Hlt := fresh "Hlt" in
dist_later_intro as m Hlt.
(** We combine [f_contractive_prepare] and [dist_later_intro] into the
[f_contractive] tactic.
For all the goals not solved by [dist_later_intro] (i.e., the ones that are
not [dist_later n x y]), we try reflexivity. Since reflexivity can be very
expensive when unification fails, we use [fast_reflexivity]. *)
Tactic Notation "f_contractive" "as" ident(idxName) ident(ltName) :=
f_contractive_prepare;
try dist_later_intro as idxName ltName;
try fast_reflexivity.
Tactic Notation "f_contractive" :=
let m := fresh "m" in
let Hlt := fresh "Hlt" in
f_contractive as m Hlt.
Ltac solve_contractive :=
solve_proper_core ltac:(fun _ => first [f_contractive | f_equiv]).
(** Limit preserving predicates *)
Class LimitPreserving `{!Cofe A} (P : A Prop) : Prop :=
limit_preserving (c : chain A) : ( n, P (c n)) P (compl c).
Hint Mode LimitPreserving + + ! : typeclass_instances.
(** To perform induction over a fixpoint ([fixpoint_ind]) and to construct the
COFE over a Sigma type ([sig_cofe]) we need the predicate to be limit
preserving: if it holds for every element of a chain, it must hold for the
limit. *)
Class LimitPreserving {SI : sidx} `{!Cofe A} (P : A Prop) : Prop := {
limit_preserving_compl (c : chain A) :
( n, P (c n)) P (compl c);
limit_preserving_lbcompl n Hn (c : bchain A n) :
( m Hm, P (c m Hm)) P (lbcompl Hn c);
}.
Global Hint Mode LimitPreserving - + + ! : typeclass_instances.
Section limit_preserving.
Context `{Cofe A}.
(* These are not instances as they will never fire automatically...
but they can still be helpful in proving things to be limit preserving. *)
Context {SI : sidx} `{!Cofe A}.
Implicit Types P Q : A Prop.
(* These lemmas are not instances as they will never fire automatically...
but they can still be helpful in proving things to be limit preserving. *)
Lemma limit_preserving_sidx_finite `{!SIdxFinite SI} P :
LimitPreserving P c : chain A, ( n, P (c n)) P (compl c).
Proof.
split; [by destruct 1|]. intros Hcompl. split; [done|].
intros n Hn. by destruct (SIdx.limit_finite n).
Qed.
Lemma limit_preserving_ext (P Q : A Prop) :
( x, P x Q x) LimitPreserving P LimitPreserving Q.
Proof. intros HP Hlimit c ?. apply HP, Hlimit=> n; by apply HP. Qed.
Proof.
intros HP [Hcompl Hlbcompl]. split.
- intros c ?. apply HP, Hcompl=> n; by apply HP.
- intros n Hn c HC. apply HP, Hlbcompl=> m Hm. by apply HP.
Qed.
Global Instance limit_preserving_const (P : Prop) : LimitPreserving (λ _ : A, P).
Proof. intros c HP. apply (HP 0). Qed.
Proof.
split.
- intros c HP. apply (HP 0).
- intros n [Hlim Hn] Hc HP. by apply (HP (Sᵢ 0)), Hlim, SIdx.neq_0_lt_0.
Qed.
Lemma limit_preserving_discrete (P : A Prop) :
Proper (dist 0 ==> impl) P LimitPreserving P.
Proof. intros PH c Hc. by rewrite (conv_compl 0). Qed.
Lemma limit_preserving_discrete P :
Proper (dist 0 ==> impl) P LimitPreserving P.
Proof.
intros HP. split.
- intros c Hc. by rewrite (conv_compl 0).
- intros n Hn c HPc.
rewrite (conv_lbcompl _ _ (SIdx.limit_lt_0 _ Hn)). apply HPc.
Qed.
Lemma limit_preserving_and (P1 P2 : A Prop) :
LimitPreserving P1 LimitPreserving P2
Lemma limit_preserving_and P1 P2 :
LimitPreserving P1
LimitPreserving P2
LimitPreserving (λ x, P1 x P2 x).
Proof. intros Hlim1 Hlim2 c Hc. split. apply Hlim1, Hc. apply Hlim2, Hc. Qed.
Proof.
intros [Hcompl1 Hlbcompl1] [Hcompl2 Hlbcompl2]. split.
- intros c Hc. split; [apply Hcompl1|apply Hcompl2]; apply Hc.
- intros n Hn c Hc. split; [apply Hlbcompl1|apply Hlbcompl2]; apply Hc.
Qed.
Lemma limit_preserving_impl (P1 P2 : A Prop) :
Proper (dist 0 ==> impl) P1 LimitPreserving P2
Lemma limit_preserving_impl P1 P2 :
Proper (dist 0 ==> impl) P1
LimitPreserving P2
LimitPreserving (λ x, P1 x P2 x).
Proof.
intros Hlim1 Hlim2 c Hc HP1. apply Hlim2=> n; apply Hc.
eapply Hlim1, HP1. apply dist_le with n; last lia. apply (conv_compl n).
intros HP1 [Hcompl Hlbcompl]. split.
- intros c Hc HP1c. apply Hcompl=> n. eapply Hc, HP1, HP1c.
apply dist_le with n, SIdx.le_0_l. apply conv_compl.
- intros n Hn c Hc HP1c. apply Hlbcompl=> m Hm. eapply Hc, HP1, HP1c.
apply dist_le with m, SIdx.le_0_l. apply conv_lbcompl.
Qed.
(** This is strictly weaker than the [_impl] variant, but sometimes automation
is better at proving [Proper] for [iff] than for [impl]. *)
Lemma limit_preserving_impl' P1 P2 :
Proper (dist 0 ==> iff) P1
LimitPreserving P2
LimitPreserving (λ x, P1 x P2 x).
Proof.
intros HP1. apply limit_preserving_impl. intros ???.
apply iff_impl_subrelation. by apply HP1.
Qed.
Lemma limit_preserving_forall {B} (P : B A Prop) :
( y, LimitPreserving (P y))
LimitPreserving (λ x, y, P y x).
Proof. intros Hlim c Hc y. by apply Hlim. Qed.
Proof.
intros Hlim. split.
- intros c Hc y. by apply Hlim.
- intros n Hn c Hc y. by apply Hlim.
Qed.
(** We need [SIdxFinite] because [compl_bchain_map] does not hold for [≡],
only for a bounded [≡{m}≡]. *)
Lemma limit_preserving_equiv `{!SIdxFinite SI} `{!Cofe B} (f g : A B) :
NonExpansive f NonExpansive g LimitPreserving (λ x, f x g x).
Proof.
intros Hf Hg. apply limit_preserving_sidx_finite=> c Hfg.
apply equiv_dist=> n. by rewrite -!compl_chain_map !conv_compl /= Hfg.
Qed.
End limit_preserving.
(** Fixpoint *)
Program Definition fixpoint_chain {A : ofeT} `{Inhabited A} (f : A A)
`{!Contractive f} : chain A := {| chain_car i := Nat.iter (S i) f inhabitant |}.
Next Obligation.
intros A ? f ? n.
induction n as [|n IH]=> -[|i] //= ?; try lia.
- apply (contractive_0 f).
- apply (contractive_S f), IH; auto with lia.
Qed.
(** A COFE defines a limit operation [lbcompl] for all limit indices. When
defining a fixpoint operator on COFEs, it is convenient
to have a limit operation which can be applied to every index, instead of just
the limit indices. We derive such an operation, called [bcompl] for
"bounded completion". *)
Section bcompl.
Context {SI : sidx} `{!Cofe A, !Inhabited A}.
Definition bcompl : BCompl A := λ n c,
match SIdx.case n with
| inl (inl Hn') => inhabitant
(* If the chain ends in a non-zero non-limit ordinal, we can just take the
last element as the limit. *)
| inl (inr (m Hm)) => c m (SIdx.lt_succ_diag_r' _ _ Hm)
| inr Hlim => lbcompl Hlim c
end.
Program Definition fixpoint_def `{Cofe A, Inhabited A} (f : A A)
`{!Contractive f} : A := compl (fixpoint_chain f).
Definition fixpoint_aux : seal (@fixpoint_def). by eexists. Qed.
Definition fixpoint {A AC AiH} f {Hf} := fixpoint_aux.(unseal) A AC AiH f Hf.
Definition fixpoint_eq : @fixpoint = @fixpoint_def := fixpoint_aux.(seal_eq).
Lemma conv_bcompl {n} (c : bchain A n) m Hm : bcompl n c {m} c m Hm.
Proof.
rewrite /bcompl. destruct (SIdx.case _) as [[->|[m' ->]]|?]; simpl.
- by destruct (SIdx.nlt_0_r m).
- by apply bchain_cauchy, SIdx.lt_succ_r.
- apply conv_lbcompl.
Qed.
Lemma bcompl_ne {n} (c1 c2 : bchain A n) m :
( p (Hp : p < n), c1 p Hp {m} c2 p Hp)
bcompl n c1 {m} bcompl n c2.
Proof.
intros Hc. rewrite /bcompl. destruct (SIdx.case _) as [[?|[m' ->]]|?]; simpl.
- done.
- apply Hc.
- by apply lbcompl_ne.
Qed.
Lemma limit_preserving_bcompl (P : A Prop) n (c : bchain A n) :
n 0 P inhabitant
LimitPreserving P
( m Hm, P (c m Hm)) P (bcompl n c).
Proof.
intros H0 [Hcompl Hlbcompl] HP. rewrite /bcompl.
destruct (SIdx.case _) as [[?|[m' ->]]|?]; naive_solver.
Qed.
End bcompl.
(** We define the fixpoint of a contractive function [f : A → A] for an arbitrary
step-index type [SI]. To explain the fixpoint construction in the general case,
let us first recall the construction in the finite case. To find the fixpoint of
a contractive function [f], we start with some dummy element [x_0] (an arbitrary
inhabitant of [A]) and iterate [f] on it such that [x_1 := f x_0], [x_2 := f x_1],
... (i.e., [x_(n + 1) := f x_n]). We then find the fixpoint as the completion
[compl] of all of these fixpoint approximations (i.e., [x := compl (λ i, x_i)]).
In the general case with ordinals as step-indices, iterating over all natural
numbers is not enough. The way that this is solved is that COFEs provide
completion operations [lbcompl] for all limit ordinals as an additional
component of their definition. Then, conceptually, we can first define our
approximations as before for all natural numbers (i.e., [x_0, x_1, ...]) and
then we can define [x_ω := lbcompl ω (λ n, x_n)] to get the limit of all those
natural number approximations. Once we have [x_ω], we can start our iteration
again (i.e., [x_(ω + 1) := f (x_ω), ...]). We keep repeating this with all
ordinals until we have eventually defined [x_n] for all ordinals [n]. At that
point, we get a fixpoint by using the completion [x := compl (λ n, x_n)]
analogously to the natural number case.
There is one small caveat to this construction. In the definition of the
fixpoint, it is somewhat inconvenient to distinguish between the cases 0,
[Sᵢ n], and limit ordinals explicitly. We can get a very clean definition of the
fixpoint with a trick: we generalize the operation [lbcompl] that only works on
limit ordinals to work on arbitrary ordinals. This is the operation [bcompl]
defined above. *)
Section fixpoint.
Context `{Cofe A, Inhabited A} (f : A A) `{!Contractive f}.
Context {SI : sidx} `{!Cofe A, !Inhabited A} (f : A A) `{!Contractive f}.
(** Getting Coq to agree with the above description of the construction of the
fixpoint takes a little work. To apply the completion operations (i.e.,
[compl] and [bcompl]), we need to know that the fixed point approximations
(i.e., ([x_n] for any [n]) and ([x_m] for any [n < m])) form a "chain".
This is not an issue in the case of natural numbers, where we can
simply proceed in three steps:
1. We first define all [x_n] by recursion on [n].
2. We prove that they form a chain, and
3. We define [x := compl (λ n, x_n)] knowing that [(λ n, x_n)] is a chain.
In the general case of ordinals, our life is harder. We want to use recursion
on ordinals to define [x_n] in terms of its predecessors [x_m] for [m < n].
However, to define [x_n], we need to use the bounded completion operation
[bcompl], which can only be applied to "bounded chains". Thus, while defining
[x_n], we need to know that all the previously defined [x_m] form a chain. In
other words, we need a property about the sequence of elements that we have
just defined. To that end, we define [bfchain], which packages a bounded chain
[bfchain_car] with the property that applying [bcompl] to [bfchain_car] gives
the fixpoint up to index [n]. *)
(** Note that [bfchain] is a private implementation detail, but Coq does not
allow us to make records [Local]. *)
Record bfchain n := {
bfchain_car :> bchain A n;
bfchain_fixpoint p :
p < n f (bcompl n bfchain_car) {p} bcompl n bfchain_car;
}.
Local Lemma bfchain_chain_unique {n m} (c1 : bfchain n) (c2 : bfchain m) p :
p < n p < m bcompl n c1 {p} bcompl m c2.
Proof using Type*.
intros Hn Hm. induction (SIdx.lt_wf p) as [p _ IH].
rewrite -(bfchain_fixpoint _ c2) // -(bfchain_fixpoint _ c1) //.
apply (contractive_dist_later_dist _); split=> p' Hp'.
apply IH; [done|by etrans..].
Qed.
Local Program Definition fixpoint_bchain_go n
(rec : m, m < n bfchain m) : bfchain n :=
{| bfchain_car := {| bchain_car m' Hm' := f (bcompl m' (rec _ Hm')) |} |}.
Next Obligation.
intros n rec m m' Hmn Hmn' Hm; simpl.
apply (contractive_dist_later_dist _); split=> p Hp.
apply bfchain_chain_unique; eauto using SIdx.lt_le_trans.
Qed.
Next Obligation.
intros n rec p Hp; simpl. rewrite (conv_bcompl _ _ Hp) /=.
apply (contractive_dist_later_dist _); split=> p' Hp'. by apply rec.
Qed.
(** We obtain a bounded fixpoint chain for every index [n] by index recursion.
In the recursive case, we construct a new chain up to [n] by taking,
for any [m < n], the limit of the [m]-th chain before applying [f] to it. *)
Local Definition fixpoint_bchain n : bfchain n :=
Fix SIdx.lt_wf bfchain fixpoint_bchain_go n.
Lemma fixpoint_unfold : fixpoint f f (fixpoint f).
(** We obtain a final full chain by repeating this construction for every [n],
using the bounded chains computed before. *)
Local Program Definition fixpoint_chain : chain A :=
{| chain_car n := f (bcompl n (fixpoint_bchain n)) |}.
Next Obligation.
intros n m [Hnm| ->]%SIdx.le_lteq; simpl; [|done].
apply (contractive_dist_later_dist _); split=> p Hp.
apply bfchain_chain_unique; [by etrans|done].
Qed.
Local Definition fixpoint_def : A := compl fixpoint_chain.
Local Definition fixpoint_aux : seal (@fixpoint_def).
Proof using Type. by eexists. Qed.
Definition fixpoint := fixpoint_aux.(unseal).
Local Definition fixpoint_unseal :
@fixpoint = @fixpoint_def := fixpoint_aux.(seal_eq).
(** This lemma does not work well with [rewrite]; we usually define a specific
unfolding lemma for each fixpoint and then [apply fixpoint_unfold] in the
proof of that unfolding lemma. *)
Lemma fixpoint_unfold : fixpoint f (fixpoint).
Proof.
apply equiv_dist=>n.
rewrite fixpoint_eq /fixpoint_def (conv_compl n (fixpoint_chain f)) //.
induction n as [|n IH]; simpl; eauto using contractive_0, contractive_S.
apply equiv_dist=> n. rewrite fixpoint_unseal /fixpoint_def /=.
rewrite !conv_compl /fixpoint_chain /=.
apply (contractive_dist_later_dist _); split=> p Hp.
by rewrite bfchain_fixpoint.
Qed.
End fixpoint.
Section fixpoint.
Context {SI : sidx} `{!Cofe A, !Inhabited A} (f : A A) `{!Contractive f}.
Lemma fixpoint_unique (x : A) : x f x x fixpoint f.
Proof.
rewrite !equiv_dist=> Hx n. induction n as [|n IH]; simpl in *.
- rewrite Hx fixpoint_unfold; eauto using contractive_0.
- rewrite Hx fixpoint_unfold. apply (contractive_S _), IH.
rewrite !equiv_dist=> Hx n. induction (SIdx.lt_wf n) as [n _ IH].
rewrite Hx fixpoint_unfold. f_contractive; eauto.
Qed.
Lemma fixpoint_ne (g : A A) `{!Contractive g} n :
( z, f z {n} g z) fixpoint f {n} fixpoint g.
Proof.
intros Hfg. rewrite fixpoint_eq /fixpoint_def
(conv_compl n (fixpoint_chain f)) (conv_compl n (fixpoint_chain g)) /=.
induction n as [|n IH]; simpl in *; [by rewrite !Hfg|].
rewrite Hfg; apply contractive_S, IH; auto using dist_S.
intros Hfg. induction (SIdx.lt_wf n) as [n _ IH].
rewrite (fixpoint_unfold f) (fixpoint_unfold g) -Hfg.
f_contractive. apply IH; eauto using dist_lt.
Qed.
Lemma fixpoint_proper (g : A A) `{!Contractive g} :
( x, f x g x) fixpoint f fixpoint g.
......@@ -338,32 +747,30 @@ Section fixpoint.
Lemma fixpoint_ind (P : A Prop) :
Proper (() ==> impl) P
( x, P x) ( x, P x P (f x))
( x, P x)
( x, P x P (f x))
LimitPreserving P
P (fixpoint f).
Proof.
intros ? [x Hx] Hincr Hlim. set (chcar i := Nat.iter (S i) f x).
assert (Hcauch : n i : nat, n i chcar i {n} chcar n).
{ intros n. rewrite /chcar. induction n as [|n IH]=> -[|i] //=;
eauto using contractive_0, contractive_S with lia. }
set (fp2 := compl {| chain_cauchy := Hcauch |}).
assert (f fp2 fp2).
{ apply equiv_dist=>n. rewrite /fp2 (conv_compl n) /= /chcar.
induction n as [|n IH]; simpl; eauto using contractive_0, contractive_S. }
rewrite -(fixpoint_unique fp2) //.
apply Hlim=> n /=. by apply Nat_iter_ind.
intros HP [x Hx] Hf Hlim. eapply HP.
{ eapply fixpoint_unique, (@fixpoint_unfold _ _ _ (populate x) f). }
rewrite fixpoint_unseal /fixpoint_def. apply Hlim=> m /=.
apply Hf. rewrite /fixpoint_bchain /Fix.
generalize (SIdx.lt_wf m). revert m. fix IH 2=> m acc.
apply limit_preserving_bcompl; [auto..|].
intros m' Hm'. destruct acc as [acc]; simpl. apply Hf, IH.
Qed.
End fixpoint.
(** Fixpoint of f when f^k is contractive. **)
Definition fixpointK `{Cofe A, Inhabited A} k (f : A A)
(** Fixpoint of [f] when [f^k] is contractive. **)
Definition fixpointK {SI : sidx} `{!Cofe A, !Inhabited A} k (f : A A)
`{!Contractive (Nat.iter k f)} := fixpoint (Nat.iter k f).
Section fixpointK.
Local Set Default Proof Using "Type*".
Context `{Cofe A, Inhabited A} (f : A A) (k : nat).
Context {SI : sidx} {A : ofe} `{!Cofe A, !Inhabited A} (f : A A) (k : nat).
Context {f_contractive : Contractive (Nat.iter k f)} {f_ne : NonExpansive f}.
(* Note than f_ne is crucial here: there are functions f such that f^2 is contractive,
but f is not non-expansive.
Consider for example f: SPred → SPred (where SPred is "downclosed sets of natural numbers").
......@@ -393,7 +800,7 @@ Section fixpointK.
Lemma fixpointK_unfold : fixpointK k f f (fixpointK k f).
Proof.
symmetry. rewrite /fixpointK. apply fixpoint_unique.
by rewrite -Nat_iter_S_r Nat_iter_S -fixpoint_unfold.
by rewrite -Nat.iter_succ_r Nat.iter_succ -fixpoint_unfold.
Qed.
Lemma fixpointK_unique (x : A) : x f x x fixpointK k f.
......@@ -424,19 +831,17 @@ Section fixpointK.
P (fixpointK k f).
Proof.
intros. rewrite /fixpointK. apply fixpoint_ind; eauto.
intros; apply Nat_iter_ind; auto.
intros; apply Nat.iter_ind; auto.
Qed.
End fixpointK.
(** Mutual fixpoints *)
Section fixpointAB.
Local Unset Default Proof Using.
Context `{Cofe A, Cofe B, !Inhabited A, !Inhabited B}.
Context {SI : sidx} {A B : ofe} `{!Cofe A, !Cofe B, !Inhabited A, !Inhabited B}.
Context (fA : A B A).
Context (fB : A B B).
Context `{ n, Proper (dist_later n ==> dist n ==> dist n) fA}.
Context `{ n, Proper (dist_later n ==> dist_later n ==> dist n) fB}.
Context {fA_contractive : n, Proper (dist_later n ==> dist n ==> dist n) fA}.
Context {fB_contractive : n, Proper (dist_later n ==> dist_later n ==> dist n) fB}.
Local Definition fixpoint_AB (x : A) : B := fixpoint (fB x).
Local Instance fixpoint_AB_contractive : Contractive fixpoint_AB.
......@@ -447,7 +852,7 @@ Section fixpointAB.
Local Definition fixpoint_AA (x : A) : A := fA x (fixpoint_AB x).
Local Instance fixpoint_AA_contractive : Contractive fixpoint_AA.
Proof. solve_contractive. Qed.
Proof using fA_contractive. solve_contractive. Qed.
Definition fixpoint_A : A := fixpoint fixpoint_AA.
Definition fixpoint_B : B := fixpoint_AB fixpoint_A.
......@@ -457,13 +862,13 @@ Section fixpointAB.
Lemma fixpoint_B_unfold : fB fixpoint_A fixpoint_B fixpoint_B.
Proof. by rewrite {2}/fixpoint_B /fixpoint_AB (fixpoint_unfold _). Qed.
Instance: Proper (() ==> () ==> ()) fA.
Proof.
apply ne_proper_2=> n x x' ? y y' ?. f_contractive; auto using dist_S.
Local Instance: Proper (() ==> () ==> ()) fA.
Proof using fA_contractive.
apply ne_proper_2=> n x x' ? y y' ?. f_contractive; eauto using dist_lt.
Qed.
Instance: Proper (() ==> () ==> ()) fB.
Proof.
apply ne_proper_2=> n x x' ? y y' ?. f_contractive; auto using dist_S.
Local Instance: Proper (() ==> () ==> ()) fB.
Proof using fB_contractive.
apply ne_proper_2=> n x x' ? y y' ?. f_contractive; eauto using dist_lt.
Qed.
Lemma fixpoint_A_unique p q : fA p q p fB p q q p fixpoint_A.
......@@ -476,7 +881,7 @@ Section fixpointAB.
End fixpointAB.
Section fixpointAB_ne.
Context `{Cofe A, Cofe B, !Inhabited A, !Inhabited B}.
Context {SI : sidx} {A B : ofe} `{!Cofe A, !Cofe B, !Inhabited A, !Inhabited B}.
Context (fA fA' : A B A).
Context (fB fB' : A B B).
Context `{ n, Proper (dist_later n ==> dist n ==> dist n) fA}.
......@@ -496,7 +901,7 @@ Section fixpointAB_ne.
fixpoint_B fA fB {n} fixpoint_B fA' fB'.
Proof.
intros HfA HfB. apply fixpoint_ne=> z. rewrite HfB. f_contractive.
apply fixpoint_A_ne; auto using dist_S.
apply fixpoint_A_ne; eauto using dist_lt.
Qed.
Lemma fixpoint_A_proper :
......@@ -510,24 +915,26 @@ Section fixpointAB_ne.
End fixpointAB_ne.
(** Non-expansive function space *)
Record ofe_mor (A B : ofeT) : Type := OfeMor {
Record ofe_mor {SI : sidx} (A B : ofe) : Type := OfeMor {
ofe_mor_car :> A B;
ofe_mor_ne : NonExpansive ofe_mor_car
}.
Arguments OfeMor {_ _} _ {_}.
Global Arguments OfeMor {_ _ _} _ {_}.
Add Printing Constructor ofe_mor.
Existing Instance ofe_mor_ne.
Global Existing Instance ofe_mor_ne.
Notation "'λne' x .. y , t" :=
(@OfeMor _ _ (λ x, .. (@OfeMor _ _ (λ y, t) _) ..) _)
(@OfeMor _ _ _ (λ x, .. (@OfeMor _ _ _ (λ y, t) _) ..) _)
(at level 200, x binder, y binder, right associativity).
Section ofe_mor.
Context {A B : ofeT}.
Context {SI : sidx} {A B : ofe}.
Global Instance ofe_mor_proper (f : ofe_mor A B) : Proper (() ==> ()) f.
Proof. apply ne_proper, ofe_mor_ne. Qed.
Instance ofe_mor_equiv : Equiv (ofe_mor A B) := λ f g, x, f x g x.
Instance ofe_mor_dist : Dist (ofe_mor A B) := λ n f g, x, f x {n} g x.
Local Instance ofe_mor_equiv : Equiv (ofe_mor A B) := λ f g, x, f x g x.
Local Instance ofe_mor_dist : Dist (ofe_mor A B) := λ n f g, x, f x {n} g x.
Definition ofe_mor_ofe_mixin : OfeMixin (ofe_mor A B).
Proof.
split.
......@@ -537,261 +944,368 @@ Section ofe_mor.
+ by intros f x.
+ by intros f g ? x.
+ by intros f g h ?? x; trans (g x).
- by intros n f g ? x; apply dist_S.
- intros n m f g ? x ?; eauto using dist_le.
Qed.
Canonical Structure ofe_morO := OfeT (ofe_mor A B) ofe_mor_ofe_mixin.
Canonical Structure ofe_morO := Ofe (ofe_mor A B) ofe_mor_ofe_mixin.
Program Definition ofe_mor_chain (c : chain ofe_morO)
(x : A) : chain B := {| chain_car n := c n x |}.
Next Obligation. intros c x n i ?. by apply (chain_cauchy c). Qed.
Program Definition ofe_mor_compl `{Cofe B} : Compl ofe_morO := λ c,
Program Definition ofe_mor_compl `{!Cofe B} : Compl ofe_morO := λ c,
{| ofe_mor_car x := compl (ofe_mor_chain c x) |}.
Next Obligation.
intros ? c n x y Hx. by rewrite (conv_compl n (ofe_mor_chain c x))
(conv_compl n (ofe_mor_chain c y)) /= Hx.
Qed.
Global Program Instance ofe_mor_cofe `{Cofe B} : Cofe ofe_morO :=
{| compl := ofe_mor_compl |}.
Program Definition ofe_mor_bchain {n}
(c : bchain ofe_morO n) (x : A) : bchain B n :=
{| bchain_car n Hn := c n Hn x |}.
Next Obligation. intros n c x m Hm i ??. by apply (bchain_cauchy n c). Qed.
Program Definition ofe_mor_lbcompl `{!Cofe B} : LBCompl ofe_morO := λ n Hn c,
{| ofe_mor_car x := lbcompl Hn (ofe_mor_bchain c x) |}.
Next Obligation.
intros ? n c x; simpl.
by rewrite (conv_compl n (ofe_mor_chain c x)) /=.
intros ? n Hn c m x y Hx. apply lbcompl_ne=> p ?.
rewrite /ofe_mor_bchain /=. by rewrite Hx.
Qed.
Global Instance ofe_mor_car_ne :
NonExpansive2 (@ofe_mor_car A B).
Global Program Instance ofe_mor_cofe `{!Cofe B} : Cofe ofe_morO :=
{| compl := ofe_mor_compl; lbcompl := ofe_mor_lbcompl |}.
Next Obligation. intros ? n c x; simpl. by rewrite conv_compl. Qed.
Next Obligation.
intros ? n Hn m Hm H x; simpl. rewrite (conv_lbcompl Hn) //=.
Qed.
Next Obligation.
intros ? n Hn c1 c2 m Hc x; simpl. apply lbcompl_ne=> p Hp. apply Hc.
Qed.
Global Instance ofe_mor_car_ne : NonExpansive2 (@ofe_mor_car SI A B).
Proof. intros n f g Hfg x y Hx; rewrite Hx; apply Hfg. Qed.
Global Instance ofe_mor_car_proper :
Proper (() ==> () ==> ()) (@ofe_mor_car A B) := ne_proper_2 _.
Proper (() ==> () ==> ()) (@ofe_mor_car SI A B) := ne_proper_2 _.
Lemma ofe_mor_ext (f g : ofe_mor A B) : f g x, f x g x.
Proof. done. Qed.
End ofe_mor.
Arguments ofe_morO : clear implicits.
Global Arguments ofe_morO {_} _ _.
Notation "A -n> B" :=
(ofe_morO A B) (at level 99, B at level 200, right associativity).
Instance ofe_mor_inhabited {A B : ofeT} `{Inhabited B} :
Global Instance ofe_mor_inhabited {SI : sidx} {A B : ofe} `{Inhabited B} :
Inhabited (A -n> B) := populate (λne _, inhabitant).
(** Identity and composition and constant function *)
Definition cid {A} : A -n> A := OfeMor id.
Instance: Params (@cid) 1 := {}.
Definition cconst {A B : ofeT} (x : B) : A -n> B := OfeMor (const x).
Instance: Params (@cconst) 2 := {}.
Definition cid {SI : sidx} {A: ofe} : A -n> A := OfeMor id.
Global Instance: Params (@cid) 2 := {}.
Definition cconst {SI : sidx} {A B : ofe} (x : B) : A -n> B := OfeMor (const x).
Global Instance: Params (@cconst) 3 := {}.
Definition ccompose {A B C}
Definition ccompose {SI : sidx} {A B C: ofe}
(f : B -n> C) (g : A -n> B) : A -n> C := OfeMor (f g).
Instance: Params (@ccompose) 3 := {}.
Global Instance: Params (@ccompose) 4 := {}.
Infix "◎" := ccompose (at level 40, left associativity).
Global Instance ccompose_ne {A B C} :
NonExpansive2 (@ccompose A B C).
Global Instance ccompose_ne {SI : sidx} {A B C: ofe} :
NonExpansive2 (@ccompose SI A B C).
Proof. intros n ?? Hf g1 g2 Hg x. rewrite /= (Hg x) (Hf (g2 x)) //. Qed.
Global Instance ccompose_proper {SI : sidx} {A B C: ofe} :
Proper (() ==> () ==> ()) (@ccompose SI A B C).
Proof. apply ne_proper_2; apply _. Qed.
(* Function space maps *)
Definition ofe_mor_map {A A' B B'} (f : A' -n> A) (g : B -n> B')
(h : A -n> B) : A' -n> B' := g h f.
Instance ofe_mor_map_ne {A A' B B'} n :
Proper (dist n ==> dist n ==> dist n ==> dist n) (@ofe_mor_map A A' B B').
Proof. intros ??? ??? ???. by repeat apply ccompose_ne. Qed.
Definition ofe_morO_map {A A' B B'} (f : A' -n> A) (g : B -n> B') :
(A -n> B) -n> (A' -n> B') := OfeMor (ofe_mor_map f g).
Instance ofe_morO_map_ne {A A' B B'} :
NonExpansive2 (@ofe_morO_map A A' B B').
Definition ofe_mor_map {SI : sidx} {A A' B B': ofe}
(f : A' -n> A) (g : B -n> B') (h : A -n> B) : A' -n> B' :=
g h f.
Global Instance ofe_mor_map_ne {SI : sidx} {A A' B B': ofe} :
NonExpansive3 (@ofe_mor_map SI A A' B B').
Proof. intros n ??? ??? ???. by repeat apply ccompose_ne. Qed.
Definition ofe_morO_map {SI : sidx} {A A' B B': ofe}
(f : A' -n> A) (g : B -n> B') : (A -n> B) -n> (A' -n> B') :=
OfeMor (ofe_mor_map f g).
Global Instance ofe_morO_map_ne {SI : sidx} {A A' B B': ofe} :
NonExpansive2 (@ofe_morO_map SI A A' B B').
Proof.
intros n f f' Hf g g' Hg ?. rewrite /= /ofe_mor_map.
by repeat apply ccompose_ne.
Qed.
(** unit *)
(** * Unit type *)
Section unit.
Instance unit_dist : Dist unit := λ _ _ _, True.
Context {SI : sidx}.
Local Instance unit_dist : Dist unit := λ _ _ _, True.
Definition unit_ofe_mixin : OfeMixin unit.
Proof. by repeat split; try exists 0. Qed.
Canonical Structure unitO : ofeT := OfeT unit unit_ofe_mixin.
Proof. repeat split. Qed.
Canonical Structure unitO : ofe := Ofe unit unit_ofe_mixin.
Global Program Instance unit_cofe : Cofe unitO := { compl x := () }.
Next Obligation. by repeat split; try exists 0. Qed.
Global Program Instance unit_cofe : Cofe unitO :=
{ compl x := (); lbcompl _ _ _ := () }.
Solve All Obligations with by repeat split.
Global Instance unit_ofe_discrete : OfeDiscrete unitO.
Proof. done. Qed.
End unit.
(** empty *)
(** * Empty type *)
Section empty.
Instance Empty_set_dist : Dist Empty_set := λ _ _ _, True.
Context {SI : sidx}.
Local Instance Empty_set_dist : Dist Empty_set := λ _ _ _, True.
Definition Empty_set_ofe_mixin : OfeMixin Empty_set.
Proof. by repeat split; try exists 0. Qed.
Canonical Structure Empty_setO : ofeT := OfeT Empty_set Empty_set_ofe_mixin.
Proof. by repeat split. Qed.
Canonical Structure Empty_setO : ofe := Ofe Empty_set Empty_set_ofe_mixin.
Global Program Instance Empty_set_cofe : Cofe Empty_setO := { compl x := x 0 }.
Next Obligation. by repeat split; try exists 0. Qed.
Global Program Instance Empty_set_cofe : Cofe Empty_setO :=
{| compl c := c 0; lbcompl n Hn c := c _ (SIdx.limit_lt_0 _ Hn) |}.
Solve All Obligations with done.
Global Instance Empty_set_ofe_discrete : OfeDiscrete Empty_setO.
Proof. done. Qed.
End empty.
(** Product *)
(** * Product type *)
Section product.
Context {A B : ofeT}.
Context {SI : sidx} {A B : ofe}.
Instance prod_dist : Dist (A * B) := λ n, prod_relation (dist n) (dist n).
Local Instance prod_dist : Dist (A * B) := λ n, prod_relation (dist n) (dist n).
Global Instance pair_ne :
NonExpansive2 (@pair A B) := _.
Global Instance fst_ne : NonExpansive (@fst A B) := _.
Global Instance snd_ne : NonExpansive (@snd A B) := _.
Definition prod_ofe_mixin : OfeMixin (A * B).
Proof.
split.
- intros x y; unfold dist, prod_dist, equiv, prod_equiv, prod_relation.
rewrite !equiv_dist; naive_solver.
- apply _.
- by intros n [x1 y1] [x2 y2] [??]; split; apply dist_S.
- by intros n m [x1 y1] [x2 y2] [??]; split; eapply dist_le.
Qed.
Canonical Structure prodO : ofeT := OfeT (A * B) prod_ofe_mixin.
Canonical Structure prodO : ofe := Ofe (A * B) prod_ofe_mixin.
Global Program Instance prod_cofe `{Cofe A, Cofe B} : Cofe prodO :=
{ compl c := (compl (chain_map fst c), compl (chain_map snd c)) }.
Next Obligation.
intros ?? n c; split. apply (conv_compl n (chain_map fst c)).
apply (conv_compl n (chain_map snd c)).
Qed.
Global Program Instance prod_cofe `{!Cofe A, !Cofe B} : Cofe prodO := {
compl c := (compl (chain_map fst c), compl (chain_map snd c));
lbcompl n Hn c := (lbcompl Hn (bchain_map fst c), lbcompl Hn (bchain_map snd c))
}.
Next Obligation. split; simpl; rewrite conv_compl //. Qed.
Next Obligation. split; simpl; rewrite conv_lbcompl //=. Qed.
Next Obligation. split; simpl; apply lbcompl_ne=> ?? /=; f_equiv; auto. Qed.
Global Instance prod_discrete (x : A * B) :
Discrete (x.1) Discrete (x.2) Discrete x.
Proof. by intros ???[??]; split; apply (discrete _). Qed.
Proof. by intros ???[??]; split; apply (discrete_0 _). Qed.
Global Instance prod_ofe_discrete :
OfeDiscrete A OfeDiscrete B OfeDiscrete prodO.
Proof. intros ?? [??]; apply _. Qed.
End product.
Arguments prodO : clear implicits.
Typeclasses Opaque prod_dist.
Lemma pair_dist n (a1 a2 : A) (b1 b2 : B) :
(a1, b1) {n} (a2, b2) a1 {n} a2 b1 {n} b2.
Proof. reflexivity. Qed.
End product.
Instance prod_map_ne {A A' B B' : ofeT} n :
Global Arguments prodO {_} _ _.
(** Below we make [prod_dist] type class opaque, so we first lift all
instances *)
Global Instance pair_dist_inj {SI : sidx} {A B : ofe} n :
Inj2 ({n}) ({n}) ({n}) (@pair A B) := _.
Global Instance curry_ne {SI : sidx} {A B C : ofe} n :
Proper ((({n}@{A*B}) ==> ({n}@{C})) ==>
({n}) ==> ({n}) ==> ({n})) curry := _.
Global Instance uncurry_ne {SI : sidx} {A B C : ofe} n :
Proper ((({n}) ==> ({n}) ==> ({n})) ==>
({n}@{A*B}) ==> ({n}@{C})) uncurry := _.
Global Instance curry3_ne {SI : sidx} {A B C D : ofe} n :
Proper ((({n}@{A*B*C}) ==> ({n}@{D})) ==>
({n}) ==> ({n}) ==> ({n}) ==> ({n})) curry3 := _.
Global Instance uncurry3_ne {SI : sidx} {A B C D : ofe} n :
Proper ((({n}) ==> ({n}) ==> ({n}) ==> ({n})) ==>
({n}@{A*B*C}) ==> ({n}@{D})) uncurry3 := _.
Global Instance curry4_ne {SI : sidx} {A B C D E : ofe} n :
Proper ((({n}@{A*B*C*D}) ==> ({n}@{E})) ==>
({n}) ==> ({n}) ==> ({n}) ==> ({n}) ==> ({n})) curry4 := _.
Global Instance uncurry4_ne {SI : sidx} {A B C D E : ofe} n :
Proper ((({n}) ==> ({n}) ==> ({n}) ==> ({n}) ==> ({n})) ==>
({n}@{A*B*C*D}) ==> ({n}@{E})) uncurry4 := _.
Global Typeclasses Opaque prod_dist.
Global Instance prod_map_ne {SI : sidx} {A A' B B' : ofe} n :
Proper ((dist n ==> dist n) ==> (dist n ==> dist n) ==>
dist n ==> dist n) (@prod_map A A' B B').
Proof. by intros f f' Hf g g' Hg ?? [??]; split; [apply Hf|apply Hg]. Qed.
Definition prodO_map {A A' B B'} (f : A -n> A') (g : B -n> B') :
Definition prodO_map {SI : sidx} {A A' B B'} (f : A -n> A') (g : B -n> B') :
prodO A B -n> prodO A' B' := OfeMor (prod_map f g).
Instance prodO_map_ne {A A' B B'} :
NonExpansive2 (@prodO_map A A' B B').
Global Instance prodO_map_ne {SI : sidx} {A A' B B'} :
NonExpansive2 (@prodO_map SI A A' B B').
Proof. intros n f f' Hf g g' Hg [??]; split; [apply Hf|apply Hg]. Qed.
(** Functors *)
Record oFunctor := OFunctor {
oFunctor_car : A `{!Cofe A} B `{!Cofe B}, ofeT;
(** * COFE → OFE Functors *)
Record oFunctor {SI : sidx} := OFunctor {
oFunctor_car : A `{!Cofe A} B `{!Cofe B}, ofe;
oFunctor_map `{!Cofe A1, !Cofe A2, !Cofe B1, !Cofe B2} :
((A2 -n> A1) * (B1 -n> B2)) oFunctor_car A1 B1 -n> oFunctor_car A2 B2;
oFunctor_ne `{!Cofe A1, !Cofe A2, !Cofe B1, !Cofe B2} :
oFunctor_map_ne `{!Cofe A1, !Cofe A2, !Cofe B1, !Cofe B2} :
NonExpansive (@oFunctor_map A1 _ A2 _ B1 _ B2 _);
oFunctor_id `{!Cofe A, !Cofe B} (x : oFunctor_car A B) :
oFunctor_map_id `{!Cofe A, !Cofe B} (x : oFunctor_car A B) :
oFunctor_map (cid,cid) x x;
oFunctor_compose `{!Cofe A1, !Cofe A2, !Cofe A3, !Cofe B1, !Cofe B2, !Cofe B3}
oFunctor_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 :
oFunctor_map (fg, g'f') x oFunctor_map (g,g') (oFunctor_map (f,f') x)
}.
Existing Instance oFunctor_ne.
Instance: Params (@oFunctor_map) 9 := {}.
Global Existing Instance oFunctor_map_ne.
Global Instance: Params (@oFunctor_map) 10 := {}.
Declare Scope oFunctor_scope.
Delimit Scope oFunctor_scope with OF.
Bind Scope oFunctor_scope with oFunctor.
Class oFunctorContractive (F : oFunctor) :=
oFunctor_contractive `{!Cofe A1, !Cofe A2, !Cofe B1, !Cofe B2} :>
Contractive (@oFunctor_map F A1 _ A2 _ B1 _ B2 _).
Hint Mode oFunctorContractive ! : typeclass_instances.
Class oFunctorContractive {SI : sidx} (F : oFunctor) :=
#[global] oFunctor_map_contractive `{!Cofe A1, !Cofe A2, !Cofe B1, !Cofe B2} ::
Contractive (@oFunctor_map SI F A1 _ A2 _ B1 _ B2 _).
Global Hint Mode oFunctorContractive - ! : typeclass_instances.
Definition oFunctor_diag (F: oFunctor) (A: ofeT) `{!Cofe A} : ofeT :=
(** Not a coercion due to the [Cofe] type class argument, and to avoid
ambiguous coercion paths, see https://gitlab.mpi-sws.org/iris/iris/issues/240. *)
Definition oFunctor_apply {SI : sidx} (F : oFunctor) (A : ofe) `{!Cofe A} : ofe :=
oFunctor_car F A A.
(** Note that the implicit argument [Cofe A] is not taken into account when
[oFunctor_diag] is used as a coercion. So, given [F : oFunctor] and [A : ofeT]
one has to write [F A _]. *)
Coercion oFunctor_diag : oFunctor >-> Funclass.
Program Definition constOF (B : ofeT) : oFunctor :=
Program Definition oFunctor_oFunctor_compose {SI : sidx} (F1 F2 : oFunctor)
`{!∀ `{!Cofe A, !Cofe B}, Cofe (oFunctor_car F2 A B)} : oFunctor := {|
oFunctor_car A _ B _ := oFunctor_car F1 (oFunctor_car F2 B A) (oFunctor_car F2 A B);
oFunctor_map A1 _ A2 _ B1 _ B2 _ 'fg :=
oFunctor_map F1 (oFunctor_map F2 (fg.2,fg.1),oFunctor_map F2 fg)
|}.
Next Obligation.
intros SI F1 F2 ? A1 ? A2 ? B1 ? B2 ? n [f1 g1] [f2 g2] [??]; simpl in *.
apply oFunctor_map_ne; split; apply oFunctor_map_ne; by split.
Qed.
Next Obligation.
intros SI F1 F2 ? A ? B ? x; simpl in *. rewrite -{2}(oFunctor_map_id F1 x).
apply equiv_dist=> n. apply oFunctor_map_ne.
split=> y /=; by rewrite !oFunctor_map_id.
Qed.
Next Obligation.
intros SI F1 F2 ? A1 ? A2 ? A3 ? B1 ? B2 ? B3 ? f g f' g' x; simpl in *.
rewrite -oFunctor_map_compose. apply equiv_dist=> n. apply oFunctor_map_ne.
split=> y /=; by rewrite !oFunctor_map_compose.
Qed.
Global Instance oFunctor_oFunctor_compose_contractive_1 {SI : sidx} (F1 F2 : oFunctor)
`{!∀ `{!Cofe A, !Cofe B}, Cofe (oFunctor_car F2 A B)} :
oFunctorContractive F1 oFunctorContractive (oFunctor_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 oFunctor_oFunctor_compose_contractive_2 {SI : sidx} (F1 F2 : oFunctor)
`{!∀ `{!Cofe A, !Cofe B}, Cofe (oFunctor_car F2 A B)} :
oFunctorContractive F2 oFunctorContractive (oFunctor_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 constOF {SI : sidx} (B : ofe) : oFunctor :=
{| oFunctor_car A1 A2 _ _ := B; oFunctor_map A1 _ A2 _ B1 _ B2 _ f := cid |}.
Solve Obligations with done.
Coercion constOF : ofeT >-> oFunctor.
Coercion constOF : ofe >-> oFunctor.
Instance constOF_contractive B : oFunctorContractive (constOF B).
Global Instance constOF_contractive {SI : sidx} B : oFunctorContractive (constOF B).
Proof. rewrite /oFunctorContractive; apply _. Qed.
Program Definition idOF : oFunctor :=
Program Definition idOF {SI : sidx} : oFunctor :=
{| oFunctor_car A1 _ A2 _ := A2; oFunctor_map A1 _ A2 _ B1 _ B2 _ f := f.2 |}.
Solve Obligations with done.
Notation "∙" := idOF : oFunctor_scope.
Program Definition prodOF (F1 F2 : oFunctor) : oFunctor := {|
Program Definition prodOF {SI : sidx} (F1 F2 : oFunctor) : oFunctor := {|
oFunctor_car A _ B _ := prodO (oFunctor_car F1 A B) (oFunctor_car F2 A B);
oFunctor_map A1 _ A2 _ B1 _ B2 _ fg :=
prodO_map (oFunctor_map F1 fg) (oFunctor_map F2 fg)
|}.
Next Obligation.
intros ?? A1 ? A2 ? B1 ? B2 ? n ???; by apply prodO_map_ne; apply oFunctor_ne.
intros ??? A1 ? A2 ? B1 ? B2 ? n ???; by apply prodO_map_ne; apply oFunctor_map_ne.
Qed.
Next Obligation. by intros F1 F2 A ? B ? [??]; rewrite /= !oFunctor_id. Qed.
Next Obligation. by intros ? F1 F2 A ? B ? [??]; rewrite /= !oFunctor_map_id. Qed.
Next Obligation.
intros F1 F2 A1 ? A2 ? A3 ? B1 ? B2 ? B3 ? f g f' g' [??]; simpl.
by rewrite !oFunctor_compose.
intros ? F1 F2 A1 ? A2 ? A3 ? B1 ? B2 ? B3 ? f g f' g' [??]; simpl.
by rewrite !oFunctor_map_compose.
Qed.
Notation "F1 * F2" := (prodOF F1%OF F2%OF) : oFunctor_scope.
Instance prodOF_contractive F1 F2 :
Global Instance prodOF_contractive {SI : sidx} F1 F2 :
oFunctorContractive F1 oFunctorContractive F2
oFunctorContractive (prodOF F1 F2).
Proof.
intros ?? A1 ? A2 ? B1 ? B2 ? n ???;
by apply prodO_map_ne; apply oFunctor_contractive.
by apply prodO_map_ne; apply oFunctor_map_contractive.
Qed.
Program Definition ofe_morOF (F1 F2 : oFunctor) : oFunctor := {|
Program Definition ofe_morOF {SI : sidx} (F1 F2 : oFunctor) : oFunctor := {|
oFunctor_car A _ B _ := oFunctor_car F1 B A -n> oFunctor_car F2 A B;
oFunctor_map A1 _ A2 _ B1 _ B2 _ fg :=
ofe_morO_map (oFunctor_map F1 (fg.2, fg.1)) (oFunctor_map F2 fg)
|}.
Next Obligation.
intros F1 F2 A1 ? A2 ? B1 ? B2 ? n [f g] [f' g'] Hfg; simpl in *.
apply ofe_morO_map_ne; apply oFunctor_ne; split; by apply Hfg.
intros ? F1 F2 A1 ? A2 ? B1 ? B2 ? n [f g] [f' g'] Hfg; simpl in *.
apply ofe_morO_map_ne; apply oFunctor_map_ne; split; by apply Hfg.
Qed.
Next Obligation.
intros F1 F2 A ? B ? [f ?] ?; simpl. rewrite /= !oFunctor_id.
apply (ne_proper f). apply oFunctor_id.
intros ? F1 F2 A ? B ? [f ?] ?; simpl. rewrite /= !oFunctor_map_id.
apply (ne_proper f). apply oFunctor_map_id.
Qed.
Next Obligation.
intros F1 F2 A1 ? A2 ? A3 ? B1 ? B2 ? B3 ? f g f' g' [h ?] ?; simpl in *.
rewrite -!oFunctor_compose. do 2 apply (ne_proper _). apply oFunctor_compose.
intros ? F1 F2 A1 ? A2 ? A3 ? B1 ? B2 ? B3 ? f g f' g' [h ?] ?; simpl in *.
rewrite -!oFunctor_map_compose. do 2 apply (ne_proper _). apply oFunctor_map_compose.
Qed.
Notation "F1 -n> F2" := (ofe_morOF F1%OF F2%OF) : oFunctor_scope.
Instance ofe_morOF_contractive F1 F2 :
Global Instance ofe_morOF_contractive {SI : sidx} F1 F2 :
oFunctorContractive F1 oFunctorContractive F2
oFunctorContractive (ofe_morOF F1 F2).
Proof.
intros ?? A1 ? A2 ? B1 ? B2 ? n [f g] [f' g'] Hfg; simpl in *.
apply ofe_morO_map_ne; apply oFunctor_contractive; destruct n, Hfg; by split.
apply ofe_morO_map_ne; apply oFunctor_map_contractive;
split; intros m Hlt; split; simpl.
all: destruct Hfg as [Hfg]; destruct (Hfg m); auto.
Qed.
(** Sum *)
(** * Sum type *)
Section sum.
Context {A B : ofeT}.
Context {SI : sidx} {A B : ofe}.
Instance sum_dist : Dist (A + B) := λ n, sum_relation (dist n) (dist n).
Local Instance sum_dist : Dist (A + B) := λ n, sum_relation (dist n) (dist n).
Global Instance inl_ne : NonExpansive (@inl A B) := _.
Global Instance inr_ne : NonExpansive (@inr A B) := _.
Global Instance inl_ne_inj : Inj (dist n) (dist n) (@inl A B) := _.
Global Instance inr_ne_inj : Inj (dist n) (dist n) (@inr A B) := _.
Global Instance inl_ne_inj n : Inj (dist n) (dist n) (@inl A B) := _.
Global Instance inr_ne_inj n : Inj (dist n) (dist n) (@inr A B) := _.
Definition sum_ofe_mixin : OfeMixin (A + B).
Proof.
split.
- intros x y; split=> Hx.
+ destruct Hx=> n; constructor; by apply equiv_dist.
+ destruct (Hx 0); constructor; apply equiv_dist=> n; by apply (inj _).
+ destruct (Hx 0); constructor; apply equiv_dist=> n; by apply (inj _).
- apply _.
- destruct 1; constructor; by apply dist_S.
- destruct 1; constructor; eapply dist_le; eauto.
Qed.
Canonical Structure sumO : ofeT := OfeT (A + B) sum_ofe_mixin.
Canonical Structure sumO : ofe := Ofe (A + B) sum_ofe_mixin.
(** The [compl] operator for sums is tricky. We are are given a [chain (A + B)]
and need to turn it into a [chain A] or [chain B] to use the [compl] operation
on [A] and [B], respectively. A priori it seems that a [chain (A + B)] might
contain [inl] and [inr] elements in a mixed fashion, but this is not the case.
Once we have inspected the first element [c 0ᵢ] and established it is an [inl]
or [inr], we know the whole chain contains [inl] or [inr] elements.
Encoding this knowledge using dependent types is annoying, hence the functions
[inl_chain] and [inr_chain] take a default value [a] and [b] to handle the
impossible [inr] and [inl] cases. The function [sum_compl] only calls
[inl_chain] if the first element is [inl a], so it can easily supply such a
default (and similar for [inr]). In the proofs of [sum_cofe] we show that
these cases cannot occur. *)
Program Definition inl_chain (c : chain sumO) (a : A) : chain A :=
{| chain_car n := match c n return _ with inl a' => a' | _ => a end |}.
Next Obligation. intros c a n i ?; simpl. by destruct (chain_cauchy c n i). Qed.
......@@ -799,93 +1313,168 @@ Section sum.
{| chain_car n := match c n return _ with inr b' => b' | _ => b end |}.
Next Obligation. intros c b n i ?; simpl. by destruct (chain_cauchy c n i). Qed.
Definition sum_compl `{Cofe A, Cofe B} : Compl sumO := λ c,
match c 0 with
Definition sum_compl `{!Cofe A, !Cofe B} : Compl sumO := λ c,
match c 0 with
| inl a => inl (compl (inl_chain c a))
| inr b => inr (compl (inr_chain c b))
end.
Global Program Instance sum_cofe `{Cofe A, Cofe B} : Cofe sumO :=
{ compl := sum_compl }.
(** The definition of [bcompl] follows the same pattern as [compl] with the
caveat that we need to show that [0ᵢ] is below [n] (for which we use
[SIdx.limit_lt_0]). *)
Program Definition inl_bchain {n} (c : bchain sumO n) (a : A) : bchain A n :=
{| bchain_car n Hn := match c n Hn return _ with inl a' => a' | _ => a end |}.
Next Obligation.
intros n c a m p Hm Hp Hmp; simpl.
by destruct (bchain_cauchy n c m p Hm Hp Hmp).
Qed.
Program Definition inr_bchain {n} (c : bchain sumO n) (b : B) : bchain B n :=
{| bchain_car n Hn := match c n Hn return _ with inr b' => b' | _ => b end |}.
Next Obligation.
intros n c b m p Hm Hp Hmp; simpl.
by destruct (bchain_cauchy n c m p Hm Hp Hmp).
Qed.
Definition sum_lbcompl `{!Cofe A, !Cofe B} : LBCompl sumO := λ n Hn c,
match c 0 (SIdx.limit_lt_0 _ Hn) with
| inl a => inl (lbcompl Hn (inl_bchain c a))
| inr b => inr (lbcompl Hn (inr_bchain c b))
end.
Global Program Instance sum_cofe `{!Cofe A, !Cofe B} : Cofe sumO :=
{ compl := sum_compl; lbcompl := sum_lbcompl }.
Next Obligation.
intros ?? n c; rewrite /compl /sum_compl.
feed inversion (chain_cauchy c 0 n); first by auto with lia; constructor.
oinversion (chain_cauchy c 0 n); first apply SIdx.le_0_l.
- rewrite (conv_compl n (inl_chain c _)) /=. destruct (c n); naive_solver.
- rewrite (conv_compl n (inr_chain c _)) /=. destruct (c n); naive_solver.
Qed.
Next Obligation.
intros ?? n Hn c m Hm; rewrite /sum_lbcompl.
oinversion (bchain_cauchy n c 0 m (SIdx.limit_lt_0 _ Hn) Hm);
first apply SIdx.le_0_l.
- rewrite (conv_lbcompl Hn _ Hm) /=. destruct (c m _); naive_solver.
- rewrite (conv_lbcompl Hn _ Hm) /=. destruct (c m _); naive_solver.
Qed.
Next Obligation.
intros ?? n Hn c1 c2 m Hc. rewrite /sum_lbcompl.
destruct (Hc 0 (SIdx.limit_lt_0 _ Hn));
rewrite /= lbcompl_ne //=; intros p Hp; by destruct (Hc p Hp).
Qed.
Global Instance inl_discrete (x : A) : Discrete x Discrete (inl x).
Proof. inversion_clear 2; constructor; by apply (discrete _). Qed.
Proof. inversion_clear 2; constructor; by apply (discrete_0 _). Qed.
Global Instance inr_discrete (y : B) : Discrete y Discrete (inr y).
Proof. inversion_clear 2; constructor; by apply (discrete _). Qed.
Proof. inversion_clear 2; constructor; by apply (discrete_0 _). Qed.
Global Instance sum_ofe_discrete :
OfeDiscrete A OfeDiscrete B OfeDiscrete sumO.
Proof. intros ?? [?|?]; apply _. Qed.
End sum.
Arguments sumO : clear implicits.
Typeclasses Opaque sum_dist.
Global Arguments sumO {_} _ _.
Global Typeclasses Opaque sum_dist.
Instance sum_map_ne {A A' B B' : ofeT} n :
Global Instance sum_map_ne {SI : sidx} {A A' B B' : ofe} n :
Proper ((dist n ==> dist n) ==> (dist n ==> dist n) ==>
dist n ==> dist n) (@sum_map A A' B B').
Proof.
intros f f' Hf g g' Hg ??; destruct 1; constructor; [by apply Hf|by apply Hg].
Qed.
Definition sumO_map {A A' B B'} (f : A -n> A') (g : B -n> B') :
Definition sumO_map {SI : sidx} {A A' B B'} (f : A -n> A') (g : B -n> B') :
sumO A B -n> sumO A' B' := OfeMor (sum_map f g).
Instance sumO_map_ne {A A' B B'} :
NonExpansive2 (@sumO_map A A' B B').
Global Instance sumO_map_ne {SI : sidx} {A A' B B'} :
NonExpansive2 (@sumO_map SI A A' B B').
Proof. intros n f f' Hf g g' Hg [?|?]; constructor; [apply Hf|apply Hg]. Qed.
Program Definition sumOF (F1 F2 : oFunctor) : oFunctor := {|
Program Definition sumOF {SI : sidx} (F1 F2 : oFunctor) : oFunctor := {|
oFunctor_car A _ B _ := sumO (oFunctor_car F1 A B) (oFunctor_car F2 A B);
oFunctor_map A1 _ A2 _ B1 _ B2 _ fg :=
sumO_map (oFunctor_map F1 fg) (oFunctor_map F2 fg)
|}.
Next Obligation.
intros ?? A1 ? A2 ? B1 ? B2 ? n ???; by apply sumO_map_ne; apply oFunctor_ne.
intros ??? A1 ? A2 ? B1 ? B2 ? n ???;
by apply sumO_map_ne; apply oFunctor_map_ne.
Qed.
Next Obligation.
by intros ? F1 F2 A ? B ? [?|?]; rewrite /= !oFunctor_map_id.
Qed.
Next Obligation. by intros F1 F2 A ? B ? [?|?]; rewrite /= !oFunctor_id. Qed.
Next Obligation.
intros F1 F2 A1 ? A2 ? A3 ? B1 ? B2 ? B3 ? f g f' g' [?|?]; simpl;
by rewrite !oFunctor_compose.
intros ? F1 F2 A1 ? A2 ? A3 ? B1 ? B2 ? B3 ? f g f' g' [?|?]; simpl;
by rewrite !oFunctor_map_compose.
Qed.
Notation "F1 + F2" := (sumOF F1%OF F2%OF) : oFunctor_scope.
Instance sumOF_contractive F1 F2 :
Global Instance sumOF_contractive {SI : sidx} F1 F2 :
oFunctorContractive F1 oFunctorContractive F2
oFunctorContractive (sumOF F1 F2).
Proof.
intros ?? A1 ? A2 ? B1 ? B2 ? n ???;
by apply sumO_map_ne; apply oFunctor_contractive.
by apply sumO_map_ne; apply oFunctor_map_contractive.
Qed.
(** Discrete OFE *)
(** * Discrete OFEs *)
Section discrete_ofe.
Context `{Equiv A} (Heq : @Equivalence A ()).
Context {SI : sidx} `{Equiv A} (Heq : @Equivalence A ()).
Instance discrete_dist : Dist A := λ n x y, x y.
Local Instance discrete_dist : Dist A := λ n x y, x y.
Definition discrete_ofe_mixin : OfeMixin A.
Proof using Type*.
split.
- intros x y; split; [done|intros Hn; apply (Hn 0)].
- intros x y; split; [done|intros Hn; apply (Hn 0)].
- done.
- done.
Qed.
Global Instance discrete_ofe_discrete : OfeDiscrete (OfeT A discrete_ofe_mixin).
Global Instance discrete_ofe_discrete : OfeDiscrete (Ofe A discrete_ofe_mixin).
Proof. by intros x y. Qed.
Global Program Instance discrete_cofe : Cofe (OfeT A discrete_ofe_mixin) :=
{ compl c := c 0 }.
Global Program Instance discrete_cofe : Cofe (Ofe A discrete_ofe_mixin) :=
{ compl c := c 0; lbcompl n Hn c := c _ (SIdx.limit_lt_0 _ Hn) }.
Next Obligation.
intros n c. rewrite /compl /=;
symmetry; apply (chain_cauchy c 0 n). lia.
intros n c; simpl. symmetry. apply (chain_cauchy c 0 n), SIdx.le_0_l.
Qed.
Next Obligation.
intros n Hn c m Hm; simpl.
symmetry; apply (bchain_cauchy n c 0 m), SIdx.le_0_l.
Qed.
Next Obligation. simpl; eauto. Qed.
End discrete_ofe.
Notation discreteO A := (OfeT A (discrete_ofe_mixin _)).
Notation leibnizO A := (OfeT A (@discrete_ofe_mixin _ equivL _)).
(** The combinators [discreteO] and [leibnizO] should be used with care. There
are two ways in which they can be used:
1. To define an OFE on a ground type, such as [nat], [expr], etc. The OFE
instance should be defined as [Canonical Structure tyO := leibnizO ty] or
[Canonical Structure tyO := discreteO ty], so not using [Definition]. See
[natO] below for an example. Make sure to avoid overlapping instances, so
always check if no instance has already been defined. For most of the types
from Coq, std++, and Iris, instances are present in Iris. The convention is
to use the name [tyO] for the OFE instance of a type [ty].
2. As part of abstractions that are parametrized with a [Type], but where an
[ofe] is needed to use (camera) combinators. See [ghost_var] as an example.
In this case, the public API of the abstraction should exclusively use
[Type], i.e., the use of [leibnizO] or [discreteO] should not leak. Otherwise
client code can end up with overlapping instances, and thus experience odd
unification failures.
You should *never* use [leibnizO] or [discreteO] on compound types such as
[list nat]. That creates overlapping canonical instances for the head symbol
(e.g., [listO] and [leibnizO (list nat)]) and confuses unification. Instead, you
have two options:
- declare/use a canonical instance for the ground type, e.g., [listO natO].
- declare a newtype, e.g., [Record ty := Ty { ty_car : list nat }], and then
declare a canonical instance for that type, e.g.,
[Canonical Structure tyO := leibnizO ty]. *)
(** The combinator [discreteO A] lifts an existing [Equiv A] instance into a
discrete OFE. *)
Notation discreteO A := (Ofe A (discrete_ofe_mixin _)).
(** The combinator [leibnizO A] lifts Leibniz equality [=] into a discrete OFE.
The implementation forces the [Equivalence] proof to be [eq_equivalence] so that
Coq does not accidentally use another one, like [ofe_equivalence], in the case of
aliases. See also https://gitlab.mpi-sws.org/iris/iris/issues/299 *)
Notation leibnizO A := (Ofe A (@discrete_ofe_mixin _ _ equivL eq_equivalence)).
(** In order to define a discrete CMRA with carrier [A] (in the file [cmra.v])
we need to determine the [Equivalence A] proof that was used to construct the
......@@ -900,68 +1489,103 @@ Notation discrete_ofe_equivalence_of A := ltac:(
| discrete_ofe_mixin ?H => exact H
end) (only parsing).
Instance leibnizO_leibniz A : LeibnizEquiv (leibnizO A).
Global Instance leibnizO_leibniz {SI : sidx} A : LeibnizEquiv (leibnizO A).
Proof. by intros x y. Qed.
Canonical Structure boolO := leibnizO bool.
Canonical Structure natO := leibnizO nat.
Canonical Structure positiveO := leibnizO positive.
Canonical Structure NO := leibnizO N.
Canonical Structure ZO := leibnizO Z.
(** * Basic Coq types *)
Canonical Structure boolO {SI : sidx} : ofe := leibnizO bool.
Canonical Structure natO {SI : sidx} : ofe := leibnizO nat.
Canonical Structure positiveO {SI : sidx} : ofe := leibnizO positive.
Canonical Structure NO {SI : sidx} : ofe := leibnizO N.
Canonical Structure ZO {SI : sidx} : ofe := leibnizO Z.
Section prop.
Context {SI : sidx}.
Local Instance Prop_equiv : Equiv Prop := iff.
Local Instance Prop_equivalence : Equivalence (≡@{Prop}) := _.
Canonical Structure PropO := discreteO Prop.
End prop.
(* Option *)
(** * Option type *)
Section option.
Context {A : ofeT}.
Context {SI : sidx} {A : ofe}.
Instance option_dist : Dist (option A) := λ n, option_Forall2 (dist n).
Lemma dist_option_Forall2 n mx my : mx {n} my option_Forall2 (dist n) mx my.
Local Instance option_dist : Dist (option A) := λ n, option_Forall2 (dist n).
Lemma option_dist_Forall2 n mx my : mx {n} my option_Forall2 (dist n) mx my.
Proof. done. Qed.
Definition option_ofe_mixin : OfeMixin (option A).
Proof.
split.
- intros mx my; split; [by destruct 1; constructor; apply equiv_dist|].
intros Hxy; destruct (Hxy 0); constructor; apply equiv_dist.
by intros n; feed inversion (Hxy n).
intros Hxy; destruct (Hxy 0); constructor; apply equiv_dist.
by intros n; oinversion (Hxy n).
- apply _.
- destruct 1; constructor; by apply dist_S.
- destruct 1; constructor; by eapply dist_le.
Qed.
Canonical Structure optionO := OfeT (option A) option_ofe_mixin.
Canonical Structure optionO := Ofe (option A) option_ofe_mixin.
Global Instance Some_ne : NonExpansive (@Some A).
Proof. intros ????. by econstructor. Qed.
Global Instance Some_dist_inj n : Inj (dist n) (dist n) (@Some A).
Proof. by inversion_clear 1. Qed.
Program Definition option_chain (c : chain optionO) (x : A) : chain A :=
{| chain_car n := default x (c n) |}.
Next Obligation. intros c x n i ?; simpl. by destruct (chain_cauchy c n i). Qed.
Definition option_compl `{Cofe A} : Compl optionO := λ c,
match c 0 with Some x => Some (compl (option_chain c x)) | None => None end.
Global Program Instance option_cofe `{Cofe A} : Cofe optionO :=
{ compl := option_compl }.
Definition option_compl `{!Cofe A} : Compl optionO := λ c,
match c 0 with Some x => Some (compl (option_chain c x)) | None => None end.
Program Definition option_bchain n (c : bchain optionO n) (x : A) : bchain A n :=
{| bchain_car n Hn := default x (c n Hn) |}.
Next Obligation.
intros ? n c; rewrite /compl /option_compl.
feed inversion (chain_cauchy c 0 n); auto with lia; [].
constructor. rewrite (conv_compl n (option_chain c _)) /=.
destruct (c n); naive_solver.
intros n c x m p Hm Hp Hmp; simpl.
by destruct (bchain_cauchy n c m p Hm Hp Hmp).
Qed.
Definition option_lbcompl `{!Cofe A} : LBCompl optionO := λ n Hn c,
match c _ (SIdx.limit_lt_0 _ Hn) with
| Some x => Some (lbcompl Hn (option_bchain n c x))
| None => None
end.
Global Program Instance option_cofe `{!Cofe A} : Cofe optionO :=
{ compl := option_compl; lbcompl := option_lbcompl }.
Next Obligation.
intros ? n c. rewrite /compl /option_compl.
oinversion (chain_cauchy c 0 n); [by apply SIdx.le_0_l|..]; f_equiv.
rewrite (conv_compl n (option_chain c _)) /=. destruct (c n); naive_solver.
Qed.
Next Obligation.
intros ? n Hn c m Hm. rewrite /lbcompl /option_lbcompl.
oinversion (bchain_cauchy n c 0 m (SIdx.limit_lt_0 _ Hn) Hm);
[by apply SIdx.le_0_l|..]; f_equiv.
rewrite (conv_lbcompl _ _ Hm) /=. destruct (c m Hm); naive_solver.
Qed.
Next Obligation.
intros ? n Hn c1 c2 m Hc. rewrite /lbcompl /option_lbcompl.
destruct (Hc 0 (SIdx.limit_lt_0 _ Hn)); f_equiv.
apply lbcompl_ne=> p Hp /=. by destruct (Hc p Hp).
Qed.
Global Instance option_ofe_discrete : OfeDiscrete A OfeDiscrete optionO.
Proof. destruct 2; constructor; by apply (discrete _). Qed.
Proof. destruct 2; constructor; by apply (discrete_0 _). Qed.
Global Instance Some_ne : NonExpansive (@Some A).
Proof. by constructor. Qed.
Global Instance is_Some_ne n : Proper (dist n ==> iff) (@is_Some A).
Proof. destruct 1; split; eauto. Qed.
Global Instance Some_dist_inj : Inj (dist n) (dist n) (@Some A).
Proof. by inversion_clear 1. Qed.
Global Instance from_option_ne {B} (R : relation B) (f : A B) n :
Proper (dist n ==> R) f Proper (R ==> dist n ==> R) (from_option f).
Global Instance from_option_ne {B} (R : relation B) n :
Proper ((dist (A:=A) n ==> R) ==> R ==> dist n ==> R) from_option.
Proof. destruct 3; simpl; auto. Qed.
Global Instance None_discrete : Discrete (@None A).
Proof. inversion_clear 1; constructor. Qed.
Global Instance Some_discrete x : Discrete x Discrete (Some x).
Proof. by intros ?; inversion_clear 1; constructor; apply discrete. Qed.
Proof. by intros ?; inversion_clear 1; constructor; apply discrete_0. Qed.
Lemma dist_None n mx : mx {n} None mx = None.
Proof. split; [by inversion_clear 1|by intros ->]. Qed.
Lemma dist_Some n x y : Some x {n} Some y x {n} y.
Proof. split; [by inversion_clear 1 | by intros ->]. Qed.
Lemma dist_Some_inv_l n mx my x :
mx {n} my mx = Some x y, my = Some y x {n} y.
Proof. destruct 1; naive_solver. Qed.
......@@ -974,97 +1598,143 @@ Section option.
Proof. intros ?%(dist_Some_inv_r _ _ _ y); naive_solver. Qed.
End option.
Typeclasses Opaque option_dist.
Arguments optionO : clear implicits.
Global Typeclasses Opaque option_dist.
Global Arguments optionO {_} _.
Instance option_fmap_ne {A B : ofeT} n:
Proper ((dist n ==> dist n) ==> dist n ==> dist n) (@fmap option _ A B).
Global Instance option_fmap_ne {SI : sidx} {A B : ofe} n :
Proper ((dist n ==> dist n) ==> ({n}@{option A}) ==> ({n}@{option B})) fmap.
Proof. intros f f' Hf ?? []; constructor; auto. Qed.
Instance option_mbind_ne {A B : ofeT} n:
Proper ((dist n ==> dist n) ==> dist n ==> dist n) (@mbind option _ A B).
Global Instance option_mbind_ne {SI : sidx} {A B : ofe} n :
Proper ((dist n ==> dist n) ==> ({n}@{option A}) ==> ({n}@{option B})) mbind.
Proof. destruct 2; simpl; auto. Qed.
Instance option_mjoin_ne {A : ofeT} n:
Proper (dist n ==> dist n) (@mjoin option _ A).
Global Instance option_mjoin_ne {SI : sidx} {A : ofe} n :
Proper (dist n ==> ({n}@{option A})) mjoin.
Proof. destruct 1 as [?? []|]; simpl; by constructor. Qed.
Definition optionO_map {A B} (f : A -n> B) : optionO A -n> optionO B :=
Global Instance option_fmap_dist_inj {SI : sidx} {A B : ofe} (f : A B) n :
Inj ({n}) ({n}) f Inj ({n}@{option A}) ({n}@{option B}) (fmap f).
Proof. apply option_fmap_inj. Qed.
Lemma fmap_Some_dist {SI : sidx}
{A B : ofe} (f : A B) (mx : option A) (y : B) n :
f <$> mx {n} Some y x : A, mx = Some x y {n} f x.
Proof.
split; [|by intros (x&->&->)].
intros (?&?%fmap_Some&?)%dist_Some_inv_r'; naive_solver.
Qed.
Definition optionO_map {SI : sidx} {A B: ofe} (f : A -n> B) :
optionO A -n> optionO B :=
OfeMor (fmap f : optionO A optionO B).
Instance optionO_map_ne A B : NonExpansive (@optionO_map A B).
Global Instance optionO_map_ne {SI : sidx} (A B: ofe) :
NonExpansive (@optionO_map SI A B).
Proof. by intros n f f' Hf []; constructor; apply Hf. Qed.
Program Definition optionOF (F : oFunctor) : oFunctor := {|
Program Definition optionOF {SI : sidx} (F : oFunctor) : oFunctor := {|
oFunctor_car A _ B _ := optionO (oFunctor_car F A B);
oFunctor_map A1 _ A2 _ B1 _ B2 _ fg := optionO_map (oFunctor_map F fg)
|}.
Next Obligation.
by intros F A1 ? A2 ? B1 ? B2 ? n f g Hfg; apply optionO_map_ne, oFunctor_ne.
intros ? F A1 ? A2 ? B1 ? B2 ? n f g Hfg;
by apply optionO_map_ne, oFunctor_map_ne.
Qed.
Next Obligation.
intros F A ? B ? x. rewrite /= -{2}(option_fmap_id x).
apply option_fmap_equiv_ext=>y; apply oFunctor_id.
intros ? F A ? B ? x. rewrite /= -{2}(option_fmap_id x).
apply option_fmap_equiv_ext=> y; apply oFunctor_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 oFunctor_compose.
intros ? F A1 ? A2 ? A3 ? B1 ? B2 ? B3 ? f g f' g' x.
rewrite /= -option_fmap_compose.
apply option_fmap_equiv_ext=> y; apply oFunctor_map_compose.
Qed.
Instance optionOF_contractive F :
Global Instance optionOF_contractive {SI : sidx} F :
oFunctorContractive F oFunctorContractive (optionOF F).
Proof.
by intros ? A1 ? A2 ? B1 ? B2 ? n f g Hfg; apply optionO_map_ne, oFunctor_contractive.
by intros ? A1 ? A2 ? B1 ? B2 ? n f g Hfg;
apply optionO_map_ne, oFunctor_map_contractive.
Qed.
(** Later *)
(** * Later type *)
(** Note that the projection [later_car] is not non-expansive (see also the
lemma [later_car_anti_contractive] below), so it cannot be used in the logic.
If you need to get a witness out, you should use the lemma [Next_uninj]
instead. *)
Record later (A : Type) : Type := Next { later_car : A }.
Add Printing Constructor later.
Arguments Next {_} _.
Arguments later_car {_} _.
Instance: Params (@Next) 1 := {}.
Global Arguments Next {_} _.
Global Arguments later_car {_} _.
Global Instance: Params (@Next) 1 := {}.
Section later.
Context {A : ofeT}.
Instance later_equiv : Equiv (later A) := λ x y, later_car x later_car y.
Instance later_dist : Dist (later A) := λ n x y,
Context {SI : sidx} {A : ofe}.
Local Instance later_equiv : Equiv (later A) := λ x y, later_car x later_car y.
Local Instance later_dist : Dist (later A) := λ n x y,
dist_later n (later_car x) (later_car y).
Definition later_ofe_mixin : OfeMixin (later A).
Proof.
split.
- intros x y; unfold equiv, later_equiv; rewrite !equiv_dist.
split. intros Hxy [|n]; [done|apply Hxy]. intros Hxy n; apply (Hxy (S n)).
- intros [x] [y]; unfold equiv, later_equiv; rewrite !equiv_dist.
split; first by intros Hxy n; split; intros m Hm.
intros H n. eapply (H (Sᵢ n)), SIdx.lt_succ_diag_r.
- split; rewrite /dist /later_dist.
+ by intros [x].
+ by intros [x] [y].
+ by intros [x] [y] [z] ??; trans y.
- intros [|n] [x] [y] ?; [done|]; rewrite /dist /later_dist; by apply dist_S.
Qed.
Canonical Structure laterO : ofeT := OfeT (later A) later_ofe_mixin.
Program Definition later_chain (c : chain laterO) : chain A :=
{| chain_car n := later_car (c (S n)) |}.
Next Obligation. intros c n i ?; apply (chain_cauchy c (S n)); lia. Qed.
Global Program Instance later_cofe `{Cofe A} : Cofe laterO :=
{ compl c := Next (compl (later_chain c)) }.
Next Obligation.
intros ? [|n] c; [done|by apply (conv_compl n (later_chain c))].
- intros n m [x] [y] Hxy ?; split; intros p Hp.
by eapply Hxy, SIdx.lt_le_trans.
Qed.
Canonical Structure laterO : ofe := Ofe (later A) later_ofe_mixin.
Global Instance Next_contractive : Contractive (@Next A).
Proof. by intros [|n] x y. Qed.
Global Instance Later_inj n : Inj (dist n) (dist (S n)) (@Next A).
Proof. by intros x y. Qed.
Proof. by intros n x y. Qed.
Global Instance Next_inj n : Inj (dist_later n) (dist n) (@Next A).
Proof. by intros x y H. Qed.
Lemma Next_uninj x : a, x Next a.
Proof. by exists (later_car x). Qed.
Instance later_car_anti_contractive n :
Local Instance later_car_anti_contractive n :
Proper (dist n ==> dist_later n) later_car.
Proof. move=> [x] [y] /= Hxy. done. Qed.
(* f is contractive iff it can factor into `Next` and a non-expansive function. *)
Lemma contractive_alt {B : ofeT} (f : A B) :
Program Definition later_chain (c : chain laterO) : chain A :=
{| chain_car n := later_car (c (Sᵢ n)) |}.
Next Obligation.
intros c n i ?%SIdx.succ_le_mono.
apply (chain_cauchy c (Sᵢ n)); eauto using SIdx.lt_succ_diag_r.
Qed.
Program Definition later_limit_bchain {n} (c : bchain laterO n)
(Hn : SIdx.limit n) : bchain A n :=
{| bchain_car m Hm := later_car (c (Sᵢ m) _) |}.
Next Obligation. intros n _ Hn. apply Hn. Qed.
Next Obligation.
intros n c Hn m p ???%SIdx.succ_le_mono; simpl.
apply (bchain_cauchy n c (Sᵢ m) (Sᵢ p)); eauto using SIdx.lt_succ_diag_r.
Qed.
Global Program Instance later_cofe `{!Cofe A} : Cofe laterO := {
compl c := Next (compl (later_chain c));
lbcompl n Hn c := Next (lbcompl Hn (later_limit_bchain c Hn))
}.
Next Obligation.
intros ? n c; split; intros m Hm%SIdx.le_succ_l; simpl. rewrite conv_compl /=.
symmetry; apply (chain_cauchy c (Sᵢ m) n); eauto using SIdx.lt_succ_diag_r.
Qed.
Next Obligation.
intros ? n ? c m Hm; simpl; split; intros p Hp; simpl.
rewrite (conv_lbcompl _ _ (transitivity Hp Hm)) /=.
symmetry; eapply (bchain_cauchy n c (Sᵢ p)); eauto using SIdx.lt_succ_diag_r.
by apply SIdx.le_succ_l.
Qed.
Next Obligation.
intros ? n Hn c1 c2 m Hc; split; intros p Hp; simpl.
apply lbcompl_ne=> q Hq; by apply Hc.
Qed.
(** [f] is contractive iff it can factor into [Next] and a non-expansive
function. *)
Lemma contractive_alt {B : ofe} (f : A B) :
Contractive f g : later A B, NonExpansive g x, f x g (Next x).
Proof.
split.
......@@ -1073,56 +1743,74 @@ Section later.
Qed.
End later.
Arguments laterO : clear implicits.
Global Arguments laterO {_} _.
Definition later_map {A B} (f : A B) (x : later A) : later B :=
Next (f (later_car x)).
Instance later_map_ne {A B : ofeT} (f : A B) n :
Proper (dist (pred n) ==> dist (pred n)) f
Global Instance later_map_ne {SI : sidx} {A B : ofe} (f : A B) n :
Proper (dist_later n ==> dist_later n) f
Proper (dist n ==> dist n) (later_map f) | 0.
Proof. destruct n as [|n]; intros Hf [x] [y] ?; do 2 red; simpl; auto. Qed.
Instance later_map_proper {A B : ofeT} (f : A B) :
Proof.
intros P [x] [y] H; rewrite /later_map //=.
split; intros m Hm; apply P, Hm. apply H.
Qed.
Global Instance later_map_ne' {SI : sidx} {A B : ofe} (f : A B) :
NonExpansive f NonExpansive (later_map f).
Proof.
intros ?? [x] [y] H. unfold later_map; simpl.
split; intros ??; simpl. f_equiv. by apply H.
Qed.
Global Instance later_map_proper {SI : sidx} {A B : ofe} (f : A B) :
Proper (() ==> ()) f
Proper (() ==> ()) (later_map f).
Proof. solve_proper. Qed.
Lemma later_map_id {A} (x : later A) : later_map id x = x.
Lemma later_map_Next {SI : sidx} {A B : ofe} (f : A B) x :
later_map f (Next x) = Next (f x).
Proof. done. Qed.
Lemma later_map_id {SI : sidx} {A} (x : later A) : later_map id x = x.
Proof. by destruct x. Qed.
Lemma later_map_compose {A B C} (f : A B) (g : B C) (x : later A) :
Lemma later_map_compose {SI : sidx}
{A B C} (f : A B) (g : B C) (x : later A) :
later_map (g f) x = later_map g (later_map f x).
Proof. by destruct x. Qed.
Lemma later_map_ext {A B : ofeT} (f g : A B) x :
Lemma later_map_ext {SI : sidx} {A B : ofe} (f g : A B) x :
( x, f x g x) later_map f x later_map g x.
Proof. destruct x; intros Hf; apply Hf. Qed.
Definition laterO_map {A B} (f : A -n> B) : laterO A -n> laterO B :=
Definition laterO_map {SI : sidx}
{A B: ofe} (f : A -n> B) : laterO A -n> laterO B :=
OfeMor (later_map f).
Instance laterO_map_contractive (A B : ofeT) : Contractive (@laterO_map A B).
Proof. intros [|n] f g Hf n'; [done|]; apply Hf; lia. Qed.
Global Instance laterO_map_contractive {SI : sidx} (A B : ofe) :
Contractive (@laterO_map SI A B).
Proof. intros n f g Hlater [x]; split; intros ??; simpl. by apply Hlater. Qed.
Program Definition laterOF (F : oFunctor) : oFunctor := {|
Program Definition laterOF {SI : sidx} (F : oFunctor) : oFunctor := {|
oFunctor_car A _ B _ := laterO (oFunctor_car F A B);
oFunctor_map A1 _ A2 _ B1 _ B2 _ fg := laterO_map (oFunctor_map F fg)
|}.
Next Obligation.
intros F A1 ? A2 ? B1 ? B2 ? n fg fg' ?.
by apply (contractive_ne laterO_map), oFunctor_ne.
intros ? F A1 ? A2 ? B1 ? B2 ? n fg fg' ?.
by apply (contractive_ne laterO_map), oFunctor_map_ne.
Qed.
Next Obligation.
intros F A ? B ? x; simpl. rewrite -{2}(later_map_id x).
apply later_map_ext=>y. by rewrite oFunctor_id.
intros ? F A ? B ? x; simpl. rewrite -{2}(later_map_id x).
apply later_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 -later_map_compose.
apply later_map_ext=>y; apply oFunctor_compose.
intros ? F A1 ? A2 ? A3 ? B1 ? B2 ? B3 ? f g f' g' x; simpl. rewrite -later_map_compose.
apply later_map_ext=>y; apply oFunctor_map_compose.
Qed.
Notation "▶ F" := (laterOF F%OF) (at level 20, right associativity) : oFunctor_scope.
Instance laterOF_contractive F : oFunctorContractive (laterOF F).
Global Instance laterOF_contractive {SI : sidx} F : oFunctorContractive (laterOF F).
Proof.
intros A1 ? A2 ? B1 ? B2 ? n fg fg' Hfg. apply laterO_map_contractive.
destruct n as [|n]; simpl in *; first done. apply oFunctor_ne, Hfg.
split; intros ???; simpl. by eapply oFunctor_map_ne, Hfg.
Qed.
(** Dependently-typed functions over a discrete domain *)
(** * Dependently-typed functions over a discrete domain *)
(** This separate notion is useful whenever we need dependent functions, and
whenever we want to avoid the hassle of the bundled non-expansive function type.
......@@ -1136,14 +1824,14 @@ We make [discrete_fun] a definition so that we can register it as a canonical
structure. We do not bundle the [Proper] proof to keep [discrete_fun] easier to
use. It turns out all the desired OFE and functorial properties do not rely on
this [Proper] instance. *)
Definition discrete_fun {A} (B : A ofeT) := x : A, B x.
Definition discrete_fun {SI : sidx} {A} (B : A ofe) := x : A, B x.
Section discrete_fun.
Context {A : Type} {B : A ofeT}.
Context {SI : sidx} {A : Type} {B : A ofe}.
Implicit Types f g : discrete_fun B.
Instance discrete_fun_equiv : Equiv (discrete_fun B) := λ f g, x, f x g x.
Instance discrete_fun_dist : Dist (discrete_fun B) := λ n f g, x, f x {n} g x.
Local Instance discrete_fun_equiv : Equiv (discrete_fun B) := λ f g, x, f x g x.
Local Instance discrete_fun_dist : Dist (discrete_fun B) := λ n f g, x, f x {n} g x.
Definition discrete_fun_ofe_mixin : OfeMixin (discrete_fun B).
Proof.
split.
......@@ -1153,19 +1841,30 @@ Section discrete_fun.
+ by intros f x.
+ by intros f g ? x.
+ by intros f g h ?? x; trans (g x).
- by intros n f g ? x; apply dist_S.
- intros n m f g ? H x. by eapply dist_le.
Qed.
Canonical Structure discrete_funO := OfeT (discrete_fun B) discrete_fun_ofe_mixin.
Canonical Structure discrete_funO := Ofe (discrete_fun B) discrete_fun_ofe_mixin.
Program Definition discrete_fun_chain `(c : chain discrete_funO)
Program Definition discrete_fun_chain (c : chain discrete_funO)
(x : A) : chain (B x) := {| chain_car n := c n x |}.
Next Obligation. intros c x n i ?. by apply (chain_cauchy c). Qed.
Global Program Instance discrete_fun_cofe `{ x, Cofe (B x)} : Cofe discrete_funO :=
{ compl c x := compl (discrete_fun_chain c x) }.
Next Obligation. intros ? n c x. apply (conv_compl n (discrete_fun_chain c x)). Qed.
Program Definition discrete_fun_bchain {n} (c : bchain discrete_funO n)
(x : A) : bchain (B x) n := {| bchain_car n Hn := c n Hn x |}.
Next Obligation. intros n c x m p Hm Hp Hmp. by apply (bchain_cauchy n c). Qed.
Global Program Instance discrete_fun_cofe `{!∀ x, Cofe (B x)} :
Cofe discrete_funO := {
compl c x := compl (discrete_fun_chain c x);
lbcompl n Hn c x := lbcompl Hn (discrete_fun_bchain c x)
}.
Next Obligation. intros ? n c x. by apply conv_compl. Qed.
Next Obligation. intros ? n Hn c m Hm x. by rewrite (conv_lbcompl _ _ Hm). Qed.
Next Obligation.
intros ? n Hn c1 c2 m Hc x. apply lbcompl_ne=> ?? /=. by apply Hc.
Qed.
Global Instance discrete_fun_inhabited `{ x, Inhabited (B x)} : Inhabited discrete_funO :=
populate (λ _, inhabitant).
Global Instance discrete_fun_inhabited `{ x, Inhabited (B x)} :
Inhabited discrete_funO := populate (λ _, inhabitant).
Global Instance discrete_fun_lookup_discrete `{EqDecision A} f x :
Discrete f Discrete (f x).
Proof.
......@@ -1178,64 +1877,68 @@ Section discrete_fun.
Qed.
End discrete_fun.
Arguments discrete_funO {_} _.
Global Arguments discrete_funO {_ _} _.
Notation "A -d> B" :=
(@discrete_funO A (λ _, B)) (at level 99, B at level 200, right associativity).
(@discrete_funO _ A (λ _, B)) (at level 99, B at level 200, right associativity).
Definition discrete_fun_map {A} {B1 B2 : A ofeT} (f : x, B1 x B2 x)
Definition discrete_fun_map {SI : sidx} {A} {B1 B2 : A ofe} (f : x, B1 x B2 x)
(g : discrete_fun B1) : discrete_fun B2 := λ x, f _ (g x).
Lemma discrete_fun_map_ext {A} {B1 B2 : A ofeT} (f1 f2 : x, B1 x B2 x)
Lemma discrete_fun_map_ext {SI : sidx} {A} {B1 B2 : A ofe} (f1 f2 : x, B1 x B2 x)
(g : discrete_fun B1) :
( x, f1 x (g x) f2 x (g x)) discrete_fun_map f1 g discrete_fun_map f2 g.
Proof. done. Qed.
Lemma discrete_fun_map_id {A} {B : A ofeT} (g : discrete_fun B) :
Lemma discrete_fun_map_id {SI : sidx} {A} {B : A ofe} (g : discrete_fun B) :
discrete_fun_map (λ _, id) g = g.
Proof. done. Qed.
Lemma discrete_fun_map_compose {A} {B1 B2 B3 : A ofeT}
Lemma discrete_fun_map_compose {SI : sidx} {A} {B1 B2 B3 : A ofe}
(f1 : x, B1 x B2 x) (f2 : x, B2 x B3 x) (g : discrete_fun B1) :
discrete_fun_map (λ x, f2 x f1 x) g = discrete_fun_map f2 (discrete_fun_map f1 g).
Proof. done. Qed.
Instance discrete_fun_map_ne {A} {B1 B2 : A ofeT} (f : x, B1 x B2 x) n :
Global Instance discrete_fun_map_ne {SI : sidx} {A} {B1 B2 : A ofe}
(f : x, B1 x B2 x) n :
( x, Proper (dist n ==> dist n) (f x))
Proper (dist n ==> dist n) (discrete_fun_map f).
Proof. by intros ? y1 y2 Hy x; rewrite /discrete_fun_map (Hy x). Qed.
Definition discrete_funO_map {A} {B1 B2 : A ofeT} (f : discrete_fun (λ x, B1 x -n> B2 x)) :
Definition discrete_funO_map {SI : sidx} {A} {B1 B2 : A ofe}
(f : discrete_fun (λ x, B1 x -n> B2 x)) :
discrete_funO B1 -n> discrete_funO B2 := OfeMor (discrete_fun_map f).
Instance discrete_funO_map_ne {A} {B1 B2 : A ofeT} :
NonExpansive (@discrete_funO_map A B1 B2).
Global Instance discrete_funO_map_ne {SI : sidx} {A} {B1 B2 : A ofe} :
NonExpansive (@discrete_funO_map SI A B1 B2).
Proof. intros n f1 f2 Hf g x; apply Hf. Qed.
Program Definition discrete_funOF {C} (F : C oFunctor) : oFunctor := {|
Program Definition discrete_funOF {SI : sidx} {C} (F : C oFunctor) : oFunctor := {|
oFunctor_car A _ B _ := discrete_funO (λ c, oFunctor_car (F c) A B);
oFunctor_map A1 _ A2 _ B1 _ B2 _ fg := discrete_funO_map (λ c, oFunctor_map (F c) fg)
|}.
Next Obligation.
intros C F A1 ? A2 ? B1 ? B2 ? n ?? g. by apply discrete_funO_map_ne=>?; apply oFunctor_ne.
intros ? C F A1 ? A2 ? B1 ? B2 ? n ?? g.
by apply discrete_funO_map_ne=>?; apply oFunctor_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 oFunctor_id.
intros ? C F A ? B ? g; simpl. rewrite -{2}(discrete_fun_map_id g).
apply discrete_fun_map_ext=> y; apply oFunctor_map_id.
Qed.
Next Obligation.
intros C F A1 ? A2 ? A3 ? B1 ? B2 ? B3 ? f1 f2 f1' f2' g.
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 oFunctor_compose.
apply discrete_fun_map_ext=>y; apply oFunctor_map_compose.
Qed.
Notation "T -d> F" := (@discrete_funOF T%type (λ _, F%OF)) : oFunctor_scope.
Notation "T -d> F" := (@discrete_funOF _ T%type (λ _, F%OF)) : oFunctor_scope.
Instance discrete_funOF_contractive {C} (F : C oFunctor) :
Global Instance discrete_funOF_contractive {SI : sidx} {C} (F : C oFunctor) :
( c, oFunctorContractive (F c)) oFunctorContractive (discrete_funOF F).
Proof.
intros ? A1 ? A2 ? B1 ? B2 ? n ?? g.
by apply discrete_funO_map_ne=>c; apply oFunctor_contractive.
by apply discrete_funO_map_ne=>c; apply oFunctor_map_contractive.
Qed.
(** Constructing isomorphic OFEs *)
Lemma iso_ofe_mixin {A : ofeT} `{Equiv B, Dist B} (g : B A)
(** * Constructing isomorphic OFEs *)
Lemma iso_ofe_mixin {SI : sidx}
{A : ofe} {B : Type} `{!Equiv B, !Dist B} (g : B A)
(g_equiv : y1 y2, y1 y2 g y1 g y2)
(g_dist : n y1 y2, y1 {n} y2 g y1 {n} g y2) : OfeMixin B.
Proof.
......@@ -1245,50 +1948,63 @@ Proof.
+ intros y. by apply g_dist.
+ intros y1 y2. by rewrite !g_dist.
+ intros y1 y2 y3. rewrite !g_dist. intros ??; etrans; eauto.
- intros n y1 y2. rewrite !g_dist. apply dist_S.
- intros n m y1 y2. rewrite !g_dist. by eapply dist_le.
Qed.
Section iso_cofe_subtype.
Context {A B : ofeT} `{Cofe A} (P : A Prop) (f : x, P x B) (g : B A).
Context {SI : sidx} {A B : ofe} `{!Cofe A}.
Context (P : A Prop) (f : x, P x B) (g : B A).
Context (g_dist : n y1 y2, y1 {n} y2 g y1 {n} g y2).
Let Hgne : NonExpansive g.
Proof. intros n y1 y2. apply g_dist. Qed.
Existing Instance Hgne.
Proof. intros n y1 y2. apply g_dist. Defined.
Local Existing Instance Hgne.
Context (gf : x Hx, g (f x Hx) x).
Context (Hlimit : c : chain B, P (compl (chain_map g c))).
Program Definition iso_cofe_subtype : Cofe B :=
{| compl c := f (compl (chain_map g c)) _ |}.
Context (Hblimit : n (Hn : SIdx.limit n) c, P (lbcompl Hn (bchain_map g c))).
Program Definition iso_cofe_subtype : Cofe B := {|
compl c := f (compl (chain_map g c)) _;
lbcompl n Hn c := f (lbcompl Hn (bchain_map g c)) _
|}.
Next Obligation. apply Hlimit. Qed.
Next Obligation. apply Hblimit. Qed.
Next Obligation. intros n c; simpl. apply g_dist. by rewrite gf conv_compl. Qed.
Next Obligation.
intros n Hn c m Hm; simpl. apply g_dist. by rewrite gf (conv_lbcompl _ _ Hm).
Qed.
Next Obligation.
intros n c; simpl. apply g_dist. by rewrite gf conv_compl.
intros n Hn c1 c2 m Hc; simpl. apply g_dist. rewrite !gf.
apply lbcompl_ne=> ?? /=. by rewrite Hc.
Qed.
End iso_cofe_subtype.
Lemma iso_cofe_subtype' {A B : ofeT} `{Cofe A}
Lemma iso_cofe_subtype' {SI : sidx} {A B : ofe} `{!Cofe A}
(P : A Prop) (f : x, P x B) (g : B A)
(Pg : y, P (g y))
(g_dist : n y1 y2, y1 {n} y2 g y1 {n} g y2)
(gf : x Hx, g (f x Hx) x)
(Hlimit : LimitPreserving P) : Cofe B.
Proof. apply: (iso_cofe_subtype P f g)=> // c. apply Hlimit=> ?; apply Pg. Qed.
Proof. destruct Hlimit. apply: (iso_cofe_subtype P f g); eauto. Qed.
Definition iso_cofe {A B : ofeT} `{Cofe A} (f : A B) (g : B A)
Definition iso_cofe {SI : sidx} {A B : ofe} `{!Cofe A} (f : A B) (g : B A)
(g_dist : n y1 y2, y1 {n} y2 g y1 {n} g y2)
(gf : x, g (f x) x) : Cofe B.
Proof. by apply (iso_cofe_subtype (λ _, True) (λ x _, f x) g). Qed.
(** Sigma *)
(** * Sigma type *)
Section sigma.
Context {A : ofeT} {P : A Prop}.
Context {SI : sidx} {A : ofe} {P : A Prop}.
Implicit Types x : sig P.
(* TODO: Find a better place for this Equiv instance. It also
should not depend on A being an OFE. *)
Instance sig_equiv : Equiv (sig P) := λ x1 x2, `x1 `x2.
Instance sig_dist : Dist (sig P) := λ n x1 x2, `x1 {n} `x2.
Local Instance sig_equiv : Equiv (sig P) := λ x1 x2, `x1 `x2.
Local Instance sig_dist : Dist (sig P) := λ n x1 x2, `x1 {n} `x2.
Definition sig_equiv_alt x y : x y `x `y := reflexivity _.
Definition sig_dist_alt n x y : x {n} y `x {n} `y := reflexivity _.
Definition sig_equiv_def x y : (x y) = (`x `y) := reflexivity _.
Definition sig_dist_def n x y : (x {n} y) = (`x {n} `y) := reflexivity _.
Lemma exist_ne n a1 a2 (H1 : P a1) (H2 : P a2) :
a1 {n} a2 a1 H1 {n} a2 H2.
......@@ -1298,25 +2014,26 @@ Section sigma.
Proof. by intros n [a Ha] [b Hb] ?. Qed.
Definition sig_ofe_mixin : OfeMixin (sig P).
Proof. by apply (iso_ofe_mixin proj1_sig). Qed.
Canonical Structure sigO : ofeT := OfeT (sig P) sig_ofe_mixin.
Canonical Structure sigO : ofe := Ofe (sig P) sig_ofe_mixin.
Global Instance sig_cofe `{Cofe A, !LimitPreserving P} : Cofe sigO.
Global Instance sig_cofe `{!Cofe A, !LimitPreserving P} : Cofe sigO.
Proof. apply (iso_cofe_subtype' P (exist P) proj1_sig)=> //. by intros []. Qed.
Global Instance sig_discrete (x : sig P) : Discrete (`x) Discrete x.
Proof. intros ? y. rewrite sig_dist_alt sig_equiv_alt. apply (discrete _). Qed.
Proof. intros ? y. rewrite sig_dist_def sig_equiv_def. apply (discrete_0 _). Qed.
Global Instance sig_ofe_discrete : OfeDiscrete A OfeDiscrete sigO.
Proof. intros ??. apply _. Qed.
End sigma.
Arguments sigO {_} _.
Global Arguments sigO {_ _} _.
(** Ofe for [sigT]. The first component must be discrete
and use Leibniz equality, while the second component might be any OFE. *)
(** * SigmaT type *)
(** Ofe for [sigT]. The first component must be discrete and use Leibniz
equality, while the second component might be any OFE. *)
Section sigT.
Import EqNotations.
Context {A : Type} {P : A ofeT}.
Context {SI : sidx} {A : Type} {P : A ofe}.
Implicit Types x : sigT P.
(**
......@@ -1326,7 +2043,7 @@ Section sigT.
Unlike in the topos of trees, with (C)OFEs we cannot use step-indexed equality
on the first component.
*)
Instance sigT_dist : Dist (sigT P) := λ n x1 x2,
Local Instance sigT_dist : Dist (sigT P) := λ n x1 x2,
Heq : projT1 x1 = projT1 x2, rew Heq in projT2 x1 {n} projT2 x2.
(**
......@@ -1336,7 +2053,7 @@ Section sigT.
By defining [equiv] in terms of [dist], we can define an OFE
without assuming UIP, at the cost of complex reasoning on [equiv].
*)
Instance sigT_equiv : Equiv (sigT P) := λ x1 x2,
Local Instance sigT_equiv : Equiv (sigT P) := λ x1 x2,
n, x1 {n} x2.
(** Unfolding lemmas.
......@@ -1348,8 +2065,10 @@ Section sigT.
Heq : projT1 x1 = projT1 x2, (rew Heq in projT2 x1) {n} projT2 x2 :=
reflexivity _.
Definition sigT_dist_proj1 n {x y} : x {n} y projT1 x = projT1 y := proj1_ex.
Definition sigT_equiv_proj1 {x y} : x y projT1 x = projT1 y := λ H, proj1_ex (H 0).
Definition sigT_dist_proj1 n {x y} :
x {n} y projT1 x = projT1 y := proj1_ex.
Definition sigT_equiv_proj1 {x y} :
x y projT1 x = projT1 y := λ H, proj1_ex (H 0).
Definition sigT_ofe_mixin : OfeMixin (sigT P).
Proof.
......@@ -1362,18 +2081,18 @@ Section sigT.
destruct 1 as [-> Heq1].
destruct 1 as [-> Heq2]. exists eq_refl => /=. by trans y.
- setoid_rewrite sigT_dist_eq.
move => [xa x] [ya y] /=. destruct 1 as [-> Heq].
exists eq_refl. exact: dist_S.
move => m [xa x] [ya y] /=. destruct 1 as [-> Heq].
exists eq_refl. eauto using dist_le.
Qed.
Canonical Structure sigTO : ofeT := OfeT (sigT P) sigT_ofe_mixin.
Canonical Structure sigTO : ofe := Ofe (sigT P) sigT_ofe_mixin.
Lemma sigT_equiv_eq_alt `{!∀ a b : A, ProofIrrel (a = b)} x1 x2 :
x1 x2
Heq : projT1 x1 = projT1 x2, rew Heq in projT2 x1 projT2 x2.
Proof.
setoid_rewrite equiv_dist; setoid_rewrite sigT_dist_eq; split => Heq.
- move: (Heq 0) => [H0eq1 _].
setoid_rewrite equiv_dist. setoid_rewrite sigT_dist_eq. split => Heq.
- move: (Heq 0) => [H0eq1 _].
exists H0eq1 => n. move: (Heq n) => [] Hneq1.
by rewrite (proof_irrel H0eq1 Hneq1).
- move: Heq => [Heq1 Heqn2] n. by exists Heq1.
......@@ -1417,63 +2136,122 @@ Section sigT.
Global Instance existT_proper_2 a : Proper (() ==> ()) (@existT A P a).
Proof. apply ne_proper, _. Qed.
Implicit Types (c : chain sigTO).
Implicit Types c : chain sigTO.
Global Instance sigT_discrete x : Discrete (projT2 x) Discrete x.
Proof.
move: x => [xa x] ? [ya y] [] /=; intros -> => /= Hxy n.
exists eq_refl => /=. apply equiv_dist, (discrete _), Hxy.
exists eq_refl => /=. apply equiv_dist, (discrete_0 _), Hxy.
Qed.
Global Instance sigT_ofe_discrete : ( a, OfeDiscrete (P a)) OfeDiscrete sigTO.
Proof. intros ??. apply _. Qed.
Lemma sigT_chain_const_proj1 c n : projT1 (c n) = projT1 (c 0).
Proof. refine (sigT_dist_proj1 _ (chain_cauchy c 0 n _)). lia. Qed.
Lemma sigT_chain_const_proj1 c n : projT1 (c n) = projT1 (c 0).
Proof.
refine (sigT_dist_proj1 _ (chain_cauchy c 0 n _)). apply SIdx.le_0_l.
Qed.
(* For this COFE construction we need UIP (Uniqueness of Identity Proofs)
on [A] (i.e. [∀ x y : A, ProofIrrel (x = y)]. UIP is most commonly obtained
from decidable equality (by Hedberg’s theorem, see
[stdpp.proof_irrel.eq_pi]). *)
Lemma sigT_bchain_const_proj1 n Hn (c: bchain sigTO n) m Hm :
projT1 (c m Hm) = projT1 (c 0 Hn).
Proof.
refine (sigT_dist_proj1 _ (bchain_cauchy n c 0 m _ _ _)). apply SIdx.le_0_l.
Qed.
(** For this COFE construction we need UIP (Uniqueness of Identity Proofs)
on [A] (i.e. [∀ x y : A, ProofIrrel (x = y)]. UIP is most commonly obtained
from decidable equality (by Hedberg’s theorem, see
[stdpp.proof_irrel.eq_pi]). *)
Section cofe.
Context `{!∀ a b : A, ProofIrrel (a = b)} `{!∀ a, Cofe (P a)}.
Program Definition chain_map_snd c : chain (P (projT1 (c 0))) :=
(** The definitions [sigT_compl] and [sigT_lbcompl] are similar to
[sum_compl] and [sum_bcompl]. We are given a [chain (sigT P)], i.e., a
sequence of dependent pairs [existT x_i y_i], but know that all [x_i]s are
the same. So, given a [chain (sigT P)], we are able to construct a
[chain (P (projT1 (c 0ᵢ)))] (which involves some programming with dependent
types), which we can then give to the [compl] operation on [P]. *)
Program Definition chain_map_snd c : chain (P (projT1 (c 0))) :=
{| chain_car n := rew (sigT_chain_const_proj1 c n) in projT2 (c n) |}.
Next Obligation.
move => c n i Hle /=.
(* [Hgoal] is our thesis, up to casts: *)
case: (chain_cauchy c n i Hle) => [Heqin Hgoal] /=.
(* Pretty delicate. We have two casts to [projT1 (c 0)].
(* Pretty delicate. We have two casts to [projT1 (c 0)].
We replace those by one cast. *)
move: (sigT_chain_const_proj1 c i) (sigT_chain_const_proj1 c n)
=> Heqi0 Heqn0.
(* Rewrite [projT1 (c 0)] to [projT1 (c n)] in goal and [Heqi0]: *)
(* Rewrite [projT1 (c 0)] to [projT1 (c n)] in goal and [Heqi0]: *)
destruct Heqn0.
by rewrite /= (proof_irrel Heqi0 Heqin).
Qed.
Definition sigT_compl : Compl sigTO :=
λ c, existT (projT1 (chain_car c 0)) (compl (chain_map_snd c)).
Definition sigT_compl : Compl sigTO := λ c,
existT (projT1 (c 0)) (compl (chain_map_snd c)).
Global Program Instance sigT_cofe : Cofe sigTO := { compl := sigT_compl }.
Program Definition bchain_map_snd n Hn (c : bchain sigTO n) :
bchain (P (projT1 (c 0 Hn))) n :=
{| bchain_car m Hm :=
rew (sigT_bchain_const_proj1 n Hn c m Hm) in projT2 (c m Hm) |}.
Next Obligation.
move => n Hn c m p Hm Hp Hle /=.
(* [Hgoal] is our thesis, up to casts: *)
case: (bchain_cauchy n c m p Hm Hp Hle) => [Heqin Hgoal] /=.
(* Pretty delicate. We have two casts to [projT1 (c 0ᵢ)].
We replace those by one cast. *)
move: (sigT_bchain_const_proj1 n Hn c m Hm)
(sigT_bchain_const_proj1 n Hn c p Hp)=> Heqm0 Heqp0.
(* Rewrite [projT1 (c 0ᵢ)] to [projT1 (c n)] in goal and [Heqi0]: *)
destruct Heqm0.
by rewrite /= (proof_irrel Heqp0 Heqin).
Qed.
Definition sigT_lbcompl : LBCompl sigTO := λ n Hn (c : bchain sigTO n),
existT (projT1 (c _ (SIdx.limit_lt_0 _ Hn)))
(lbcompl Hn (bchain_map_snd n (SIdx.limit_lt_0 _ Hn) c)).
Global Program Instance sigT_cofe : Cofe sigTO :=
{ compl := sigT_compl; lbcompl := sigT_lbcompl }.
Next Obligation.
intros n c. rewrite /sigT_compl sigT_dist_eq /=.
exists (symmetry (sigT_chain_const_proj1 c n)).
(* Our thesis, up to casts: *)
pose proof (conv_compl n (chain_map_snd c)) as Hgoal.
move: (compl (chain_map_snd c)) Hgoal => pc0 /=.
destruct (sigT_chain_const_proj1 c n); simpl. done.
by destruct (sigT_chain_const_proj1 c n).
Qed.
Next Obligation.
intros n Hn c m Hm. rewrite /sigT_lbcompl sigT_dist_eq /=.
exists (symmetry (sigT_bchain_const_proj1 n (SIdx.limit_lt_0 _ Hn) c m Hm)).
(* Our thesis, up to casts: *)
pose proof (conv_lbcompl Hn
(bchain_map_snd n (SIdx.limit_lt_0 _ Hn) c) Hm) as Hgoal.
move: (lbcompl Hn (bchain_map_snd n _ c)) Hgoal => pc0 /=.
by destruct (sigT_bchain_const_proj1 n _ c m Hm).
Qed.
Next Obligation.
intros n Hn c1 c2 m Hc. rewrite /sigT_lbcompl sigT_dist_eq /=.
destruct (Hc 0 (SIdx.limit_lt_0 _ Hn)) as [eq Ht]. exists eq; simpl.
enough (lbcompl Hn (rew [λ x : A, bchain (P x) n] eq in
bchain_map_snd n (SIdx.limit_lt_0 _ Hn) c1)
{m} lbcompl Hn (bchain_map_snd n (SIdx.limit_lt_0 _ Hn) c2)) as Hlbcompl.
{ rewrite -Hlbcompl. clear Ht Hlbcompl. by destruct eq. }
apply lbcompl_ne=> p Hp /=. destruct (Hc p Hp) as [eq' H'].
rewrite -(@map_subst _ (λ y, bchain (P y) n) P (λ y d, d p Hp) _ _ eq) /=.
rewrite rew_compose. revert H'.
move: (sigT_bchain_const_proj1 n _ c2 p Hp)
(eq_trans (sigT_bchain_const_proj1 n _ c1 p Hp) eq)=> e1 e2.
destruct e1; simpl. intros <-. by rewrite /= (proof_irrel e2 eq').
Qed.
End cofe.
End sigT.
Arguments sigTO {_} _.
Global Arguments sigTO {_ _} _.
Section sigTOF.
Context {A : Type}.
Context {SI : sidx} {A : Type}.
Program Definition sigT_map {P1 P2 : A ofeT} :
Program Definition sigT_map {P1 P2 : A ofe} :
discrete_funO (λ a, P1 a -n> P2 a) -n>
sigTO P1 -n> sigTO P2 :=
λne f xpx, existT _ (f _ (projT2 xpx)).
......@@ -1486,26 +2264,114 @@ Section sigTOF.
Qed.
Program Definition sigTOF (F : A oFunctor) : oFunctor := {|
oFunctor_car A CA B CB := sigTO (λ a, oFunctor_car (F a) A _ B CB);
oFunctor_car A CA B CB := sigTO (λ a, oFunctor_car (F a) A B);
oFunctor_map A1 _ A2 _ B1 _ B2 _ fg := sigT_map (λ a, oFunctor_map (F a) fg)
|}.
Next Obligation.
repeat intro. exists eq_refl => /=. solve_proper.
Qed.
Next Obligation.
simpl; intros. apply (existT_proper eq_refl), oFunctor_id.
simpl; intros. apply (existT_proper eq_refl), oFunctor_map_id.
Qed.
Next Obligation.
simpl; intros. apply (existT_proper eq_refl), oFunctor_compose.
simpl; intros. apply (existT_proper eq_refl), oFunctor_map_compose.
Qed.
Global Instance sigTOF_contractive {F} :
( a, oFunctorContractive (F a)) oFunctorContractive (sigTOF F).
Proof.
repeat intro. apply sigT_map => a. exact: oFunctor_contractive.
repeat intro. apply sigT_map => a. exact: oFunctor_map_contractive.
Qed.
End sigTOF.
Arguments sigTOF {_} _%OF.
Global Arguments sigTOF {_ _} _%_OF.
Notation "{ x & P }" := (sigTOF (λ x, P%OF)) : oFunctor_scope.
Notation "{ x : A & P }" := (@sigTOF A%type (λ x, P%OF)) : oFunctor_scope.
Notation "{ x : A & P }" := (@sigTOF _ A%type (λ x, P%OF)) : oFunctor_scope.
(** * Isomorphisms between OFEs *)
Record ofe_iso {SI : sidx} (A B : ofe) := OfeIso {
ofe_iso_1 : A -n> B;
ofe_iso_2 : B -n> A;
ofe_iso_12 y : ofe_iso_1 (ofe_iso_2 y) y;
ofe_iso_21 x : ofe_iso_2 (ofe_iso_1 x) x;
}.
Global Arguments OfeIso {_ _ _} _ _ _ _.
Global Arguments ofe_iso_1 {_ _ _} _.
Global Arguments ofe_iso_2 {_ _ _} _.
Global Arguments ofe_iso_12 {_ _ _} _ _.
Global Arguments ofe_iso_21 {_ _ _} _ _.
Section ofe_iso.
Context {SI : sidx} {A B : ofe}.
Local Instance ofe_iso_equiv : Equiv (ofe_iso A B) := λ I1 I2,
ofe_iso_1 I1 ofe_iso_1 I2 ofe_iso_2 I1 ofe_iso_2 I2.
Local Instance ofe_iso_dist : Dist (ofe_iso A B) := λ n I1 I2,
ofe_iso_1 I1 {n} ofe_iso_1 I2 ofe_iso_2 I1 {n} ofe_iso_2 I2.
Global Instance ofe_iso_1_ne : NonExpansive (ofe_iso_1 (A:=A) (B:=B)).
Proof. by destruct 1. Qed.
Global Instance ofe_iso_2_ne : NonExpansive (ofe_iso_2 (A:=A) (B:=B)).
Proof. by destruct 1. Qed.
Lemma ofe_iso_ofe_mixin : OfeMixin (ofe_iso A B).
Proof. by apply (iso_ofe_mixin (λ I, (ofe_iso_1 I, ofe_iso_2 I))). Qed.
Canonical Structure ofe_isoO : ofe := Ofe (ofe_iso A B) ofe_iso_ofe_mixin.
Global Instance ofe_iso_cofe `{!SIdxFinite SI}
`{!Cofe A, !Cofe B} : Cofe ofe_isoO.
Proof.
apply (iso_cofe_subtype'
(λ I : prodO (A -n> B) (B -n> A),
( y, I.1 (I.2 y) y) ( x, I.2 (I.1 x) x))
(λ I HI, OfeIso (I.1) (I.2) (proj1 HI) (proj2 HI))
(λ I, (ofe_iso_1 I, ofe_iso_2 I))); [by intros []|done..|].
apply limit_preserving_and; apply limit_preserving_forall=> ?;
apply limit_preserving_equiv; solve_proper.
Qed.
End ofe_iso.
Global Arguments ofe_isoO {_} _ _.
Program Definition iso_ofe_refl {SI : sidx} {A: ofe} : ofe_iso A A :=
OfeIso cid cid _ _.
Solve Obligations with done.
Definition iso_ofe_sym {SI : sidx} {A B : ofe} (I : ofe_iso A B) : ofe_iso B A :=
OfeIso (ofe_iso_2 I) (ofe_iso_1 I) (ofe_iso_21 I) (ofe_iso_12 I).
Global Instance iso_ofe_sym_ne {SI : sidx} {A B : ofe} :
NonExpansive (iso_ofe_sym (A:=A) (B:=B)).
Proof. intros n I1 I2 []; split; simpl; by f_equiv. Qed.
Program Definition iso_ofe_trans {SI : sidx} {A B C: ofe}
(I : ofe_iso A B) (J : ofe_iso B C) : ofe_iso A C :=
OfeIso (ofe_iso_1 J ofe_iso_1 I) (ofe_iso_2 I ofe_iso_2 J) _ _.
Next Obligation. intros ? A B C I J z; simpl. by rewrite !ofe_iso_12. Qed.
Next Obligation. intros ? A B C I J z; simpl. by rewrite !ofe_iso_21. Qed.
Global Instance iso_ofe_trans_ne {SI : sidx} {A B C} :
NonExpansive2 (iso_ofe_trans (A:=A) (B:=B) (C:=C)).
Proof. intros n I1 I2 [] J1 J2 []; split; simpl; by f_equiv. Qed.
Program Definition iso_ofe_cong {SI : sidx} (F : oFunctor) `{!Cofe A, !Cofe B}
(I : ofe_iso A B) : ofe_iso (oFunctor_apply F A) (oFunctor_apply F B) :=
OfeIso (oFunctor_map F (ofe_iso_2 I, ofe_iso_1 I))
(oFunctor_map F (ofe_iso_1 I, ofe_iso_2 I)) _ _.
Next Obligation.
intros ? F A ? B ? I x. rewrite -oFunctor_map_compose -{2}(oFunctor_map_id F x).
apply equiv_dist=> n.
apply oFunctor_map_ne; split=> ? /=; by rewrite ?ofe_iso_12 ?ofe_iso_21.
Qed.
Next Obligation.
intros ? F A ? B ? I y. rewrite -oFunctor_map_compose -{2}(oFunctor_map_id F y).
apply equiv_dist=> n.
apply oFunctor_map_ne; split=> ? /=; by rewrite ?ofe_iso_12 ?ofe_iso_21.
Qed.
Global Instance iso_ofe_cong_ne {SI : sidx} (F : oFunctor) `{!Cofe A, !Cofe B} :
NonExpansive (iso_ofe_cong F (A:=A) (B:=B)).
Proof. intros n I1 I2 []; split; simpl; by f_equiv. Qed.
Global Instance iso_ofe_cong_contractive {SI : sidx}
(F : oFunctor) `{!Cofe A, !Cofe B} :
oFunctorContractive F Contractive (iso_ofe_cong F (A:=A) (B:=B)).
Proof. intros ? n I1 I2 HI; split; simpl; f_contractive; by destruct HI. Qed.
From iris.algebra Require Export cmra.
From iris.prelude Require Import options.
(* The [IsOp a b1 b2] class is used in two directions: to "split" input [a] into
outputs [b1] and [b2], and to "merge" inputs [b1] and [b2] into output [a],
where in both cases we have [a ≡ b1 ⋅ b2].
Since the [IsOp a b1 b2] class is used in two directions, there are some
subtleties we need to account for:
- If we want to "merge", we want the "op" instance to be used *last*. That is,
before using [IsOp (b1 ⋅ b2) b1 b2], we want to traverse the structure of the
term to merge constructors, and we want it to combine terms like [q/2] and
[q/2] into [q] instead of [q/2 ⋅ q/2].
- If we want to "split", we want the "op" instance to be used *first*. That is,
we want to use [IsOp (b1 ⋅ b2) b1 b2] eagerly, so that for instance, a term
like [q1 ⋅ q2] is turned into [q1] and [q2] and not two times [(q1 ⋅ q2) / 2].
To achieve this, there are various classes with different modes:
- [IsOp a b1 b2]. This class has no mode, so it can be used even to
combine/merge evars. This class has only one direct instance
[IsOp (a ⋅ b) a b] with cost 100 (so it is used last), ensuring that the
"op" rule is used last when merging.
- [IsOp' a b1 b2]. This class requires either [a] OR both [b1] and [b2] to be inputs.
All usual instances should be of this class to avoid loops.
- [IsOp'LR a b1 b2]. This class requires [a] to be an input and has just one
instance [IsOp'LR (a ⋅ b) a b] with cost 0. This ensures that the "op"
rule is used first when splitting.
*)
Class IsOp {SI : sidx} {A : cmra} (a b1 b2 : A) := is_op : a b1 b2.
Global Arguments is_op {_ _} _ _ _ {_}.
Global Hint Mode IsOp - + - - - : typeclass_instances.
Global Instance is_op_op {SI : sidx} {A : cmra} (a b : A) : IsOp (a b) a b | 100.
Proof. by rewrite /IsOp. Qed.
Class IsOp' {SI : sidx} {A : cmra} (a b1 b2 : A) :=
#[global] is_op' :: IsOp a b1 b2.
Global Hint Mode IsOp' - + ! - - : typeclass_instances.
Global Hint Mode IsOp' - + - ! ! : typeclass_instances.
Class IsOp'LR {SI : sidx} {A : cmra} (a b1 b2 : A) := is_op_lr : IsOp a b1 b2.
Global Existing Instance is_op_lr | 0.
Global Hint Mode IsOp'LR - + ! - - : typeclass_instances.
Global Instance is_op_lr_op {SI : sidx} {A : cmra} (a b : A) : IsOp'LR (a b) a b | 0.
Proof. by rewrite /IsOp'LR /IsOp. Qed.
(* FromOp *)
(* TODO: Worst case there could be a lot of backtracking on these instances,
try to refactor. *)
Global Instance is_op_pair {SI : sidx} {A B : cmra} (a b1 b2 : A) (a' b1' b2' : B) :
IsOp a b1 b2 IsOp a' b1' b2' IsOp' (a,a') (b1,b1') (b2,b2').
Proof. by constructor. Qed.
Global Instance is_op_pair_core_id_l {SI : sidx} {A B : cmra} (a : A) (a' b1' b2' : B) :
CoreId a IsOp a' b1' b2' IsOp' (a,a') (a,b1') (a,b2').
Proof. constructor=> //=. by rewrite -core_id_dup. Qed.
Global Instance is_op_pair_core_id_r {SI : sidx} {A B : cmra} (a b1 b2 : A) (a' : B) :
CoreId a' IsOp a b1 b2 IsOp' (a,a') (b1,a') (b2,a').
Proof. constructor=> //=. by rewrite -core_id_dup. Qed.
Global Instance is_op_Some {SI : sidx} {A : cmra} (a : A) b1 b2 :
IsOp a b1 b2 IsOp' (Some a) (Some b1) (Some b2).
Proof. by constructor. Qed.
From iris.algebra Require Export gmap coPset local_updates.
From iris.algebra Require Import updates proofmode_classes.
From iris.prelude Require Import options.
(** The camera [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]. The key connectives are [reservation_map_data k a] (the "points-to"
assertion of this map), which associates data [a : A] with a key [k : positive],
and [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 [reservation_map_token_union] enables one to split [reservation_map_token]
w.r.t. disjoint union. That is, if we have [E1 ## E2], then we get
[reservation_map_token (E1 ∪ E2) = reservation_map_token E1 ⋅ reservation_map_token E2].
- The lemma [reservation_map_alloc] provides a frame preserving update to
associate data to a key: [reservation_map_token E ~~> 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 reservation_map (A : Type) := ReservationMap {
reservation_map_data_proj : gmap positive A;
reservation_map_token_proj : coPset_disj
}.
Add Printing Constructor reservation_map.
Global Arguments ReservationMap {_} _ _.
Global Arguments reservation_map_data_proj {_} _.
Global Arguments reservation_map_token_proj {_} _.
Global Instance: Params (@ReservationMap) 1 := {}.
Global Instance: Params (@reservation_map_data_proj) 1 := {}.
Global Instance: Params (@reservation_map_token_proj) 1 := {}.
Definition reservation_map_data {SI : sidx} {A : cmra}
(k : positive) (a : A) : reservation_map A :=
ReservationMap {[ k := a ]} ε.
Definition reservation_map_token {SI : sidx} {A : cmra}
(E : coPset) : reservation_map A :=
ReservationMap (CoPset E).
Global Instance: Params (@reservation_map_data) 3 := {}.
(* Ofe *)
Section ofe.
Context {SI : sidx} {A : ofe}.
Implicit Types x y : reservation_map A.
Local Instance reservation_map_equiv : Equiv (reservation_map A) := λ x y,
reservation_map_data_proj x reservation_map_data_proj y
reservation_map_token_proj x = reservation_map_token_proj y.
Local Instance reservation_map_dist : Dist (reservation_map A) := λ n x y,
reservation_map_data_proj x {n} reservation_map_data_proj y
reservation_map_token_proj x = reservation_map_token_proj y.
Global Instance ReservationMap_ne :
NonExpansive2 (@ReservationMap A).
Proof. by split. Qed.
Global Instance ReservationMap_proper :
Proper (() ==> (=) ==> ()) (@ReservationMap A).
Proof. by split. Qed.
Global Instance reservation_map_data_proj_ne :
NonExpansive (@reservation_map_data_proj A).
Proof. by destruct 1. Qed.
Global Instance reservation_map_data_proj_proper :
Proper (() ==> ()) (@reservation_map_data_proj A).
Proof. by destruct 1. Qed.
Definition reservation_map_ofe_mixin : OfeMixin (reservation_map A).
Proof.
by apply (iso_ofe_mixin
(λ x, (reservation_map_data_proj x, reservation_map_token_proj x))).
Qed.
Canonical Structure reservation_mapO :=
Ofe (reservation_map A) reservation_map_ofe_mixin.
Global Instance ReservationMap_discrete a b :
Discrete a Discrete b Discrete (ReservationMap a b).
Proof. intros ?? [??] [??]; split; unfold_leibniz; by eapply discrete_0. Qed.
Global Instance reservation_map_ofe_discrete :
OfeDiscrete A OfeDiscrete reservation_mapO.
Proof. intros ? [??]; apply _. Qed.
End ofe.
Global Arguments reservation_mapO {_} _.
(* Camera *)
Section cmra.
Context {SI : sidx} {A : cmra}.
Implicit Types a b : A.
Implicit Types x y : reservation_map A.
Implicit Types k : positive.
Global Instance reservation_map_data_ne i :
NonExpansive (@reservation_map_data SI A i).
Proof. solve_proper. Qed.
Global Instance reservation_map_data_proper k :
Proper (() ==> ()) (@reservation_map_data SI A k).
Proof. solve_proper. Qed.
Global Instance reservation_map_data_discrete k a :
Discrete a Discrete (reservation_map_data k a).
Proof. intros. apply ReservationMap_discrete; apply _. Qed.
Global Instance reservation_map_token_discrete E :
Discrete (@reservation_map_token SI A E).
Proof. intros. apply ReservationMap_discrete; apply _. Qed.
Local Instance reservation_map_valid_instance : Valid (reservation_map A) := λ x,
match reservation_map_token_proj x with
| CoPset E =>
(reservation_map_data_proj x)
(* dom (reservation_map_data_proj x) ⊥ E *)
i, reservation_map_data_proj x !! i = None i E
| CoPsetInvalid => False
end.
Global Arguments reservation_map_valid_instance !_ /.
Local Instance reservation_map_validN_instance : ValidN (reservation_map A) := λ n x,
match reservation_map_token_proj x with
| CoPset E =>
{n} (reservation_map_data_proj x)
(* dom (reservation_map_data_proj x) ⊥ E *)
i, reservation_map_data_proj x !! i = None i E
| CoPsetInvalid => False
end.
Global Arguments reservation_map_validN_instance !_ /.
Local Instance reservation_map_pcore_instance : PCore (reservation_map A) := λ x,
Some (ReservationMap (core (reservation_map_data_proj x)) ε).
Local Instance reservation_map_op_instance : Op (reservation_map A) := λ x y,
ReservationMap (reservation_map_data_proj x reservation_map_data_proj y)
(reservation_map_token_proj x reservation_map_token_proj y).
Definition reservation_map_valid_eq :
valid = λ x, match reservation_map_token_proj x with
| CoPset E =>
(reservation_map_data_proj x)
(* dom (reservation_map_data_proj x) ⊥ E *)
i, reservation_map_data_proj x !! i = None i E
| CoPsetInvalid => False
end := eq_refl _.
Definition reservation_map_validN_eq :
validN = λ n x, match reservation_map_token_proj x with
| CoPset E =>
{n} (reservation_map_data_proj x)
(* dom (reservation_map_data_proj x) ⊥ E *)
i, reservation_map_data_proj x !! i = None i E
| CoPsetInvalid => False
end := eq_refl _.
Lemma reservation_map_included x y :
x y
reservation_map_data_proj x reservation_map_data_proj y
reservation_map_token_proj x reservation_map_token_proj y.
Proof.
split; [intros [[z1 z2] Hz]; split; [exists z1|exists z2]; apply Hz|].
intros [[z1 Hz1] [z2 Hz2]]; exists (ReservationMap z1 z2); split; auto.
Qed.
Lemma reservation_map_data_proj_validN n x : {n} x {n} reservation_map_data_proj x.
Proof. by destruct x as [? [?|]]=> // -[??]. Qed.
Lemma reservation_map_token_proj_validN n x : {n} x {n} reservation_map_token_proj x.
Proof. by destruct x as [? [?|]]=> // -[??]. Qed.
Lemma reservation_map_cmra_mixin : CmraMixin (reservation_map A).
Proof.
apply cmra_total_mixin.
- eauto.
- by intros n x y1 y2 [Hy Hy']; split; simpl; rewrite ?Hy ?Hy'.
- solve_proper.
- intros n [m1 [E1|]] [m2 [E2|]] [Hm ?]=> // -[??]; split; simplify_eq/=.
+ by rewrite -Hm.
+ intros i. by rewrite -(dist_None n) -Hm dist_None.
- pose 0.
intros [m [E|]]; rewrite reservation_map_valid_eq
reservation_map_validN_eq /= ?cmra_valid_validN; naive_solver.
- intros n m [k [E|]]; rewrite reservation_map_validN_eq /=;
naive_solver eauto using cmra_validN_le.
- split; simpl; [by rewrite assoc|by rewrite assoc_L].
- split; simpl; [by rewrite comm|by rewrite comm_L].
- split; simpl; [by rewrite cmra_core_l|by rewrite left_id_L].
- split; simpl; [by rewrite cmra_core_idemp|done].
- intros ??; rewrite! reservation_map_included; intros [??].
by split; simpl; apply: cmra_core_mono. (* FIXME: FIXME(Coq #6294): needs new unification *)
- intros n [m1 [E1|]] [m2 [E2|]]=> //=; rewrite reservation_map_validN_eq /=.
rewrite {1}/op /cmra_op /=. case_decide; last done.
intros [Hm Hdisj]; split; first by eauto using cmra_validN_op_l.
intros i. move: (Hdisj i). rewrite lookup_op.
case: (m1 !! i); case: (m2 !! i); set_solver.
- intros n x y1 y2 ? [??]; simpl in *.
destruct (cmra_extend n (reservation_map_data_proj x)
(reservation_map_data_proj y1) (reservation_map_data_proj y2))
as (m1&m2&?&?&?); auto using reservation_map_data_proj_validN.
destruct (cmra_extend n (reservation_map_token_proj x)
(reservation_map_token_proj y1) (reservation_map_token_proj y2))
as (E1&E2&?&?&?); auto using reservation_map_token_proj_validN.
by exists (ReservationMap m1 E1), (ReservationMap m2 E2).
Qed.
Canonical Structure reservation_mapR :=
Cmra (reservation_map A) reservation_map_cmra_mixin.
Global Instance reservation_map_cmra_discrete :
CmraDiscrete A CmraDiscrete reservation_mapR.
Proof.
split; first apply _.
intros [m [E|]]; rewrite reservation_map_validN_eq reservation_map_valid_eq //=.
by intros [?%cmra_discrete_valid ?].
Qed.
Local Instance reservation_map_empty_instance : Unit (reservation_map A) := ReservationMap ε ε.
Lemma reservation_map_ucmra_mixin : UcmraMixin (reservation_map A).
Proof.
split; simpl.
- rewrite reservation_map_valid_eq /=. split; [apply ucmra_unit_valid|]. set_solver.
- split; simpl; [by rewrite left_id|by rewrite left_id_L].
- do 2 constructor; [apply (core_id_core _)|done].
Qed.
Canonical Structure reservation_mapUR :=
Ucmra (reservation_map A) reservation_map_ucmra_mixin.
Global Instance reservation_map_data_core_id k a :
CoreId a CoreId (reservation_map_data k a).
Proof. do 2 constructor; simpl; auto. apply core_id_core, _. Qed.
Lemma reservation_map_data_valid k a : (reservation_map_data k a) a.
Proof. rewrite reservation_map_valid_eq /= singleton_valid. set_solver. Qed.
Lemma reservation_map_token_valid E : (reservation_map_token E).
Proof. rewrite reservation_map_valid_eq /=. split; first done. by left. Qed.
Lemma reservation_map_data_op k a b :
reservation_map_data k (a b) = reservation_map_data k a reservation_map_data k b.
Proof.
by rewrite {2}/op /reservation_map_op_instance /reservation_map_data /= singleton_op left_id_L.
Qed.
Lemma reservation_map_data_mono k a b :
a b reservation_map_data k a reservation_map_data k b.
Proof. intros [c ->]. by rewrite reservation_map_data_op. Qed.
Global Instance reservation_map_data_is_op k a b1 b2 :
IsOp a b1 b2
IsOp' (reservation_map_data k a) (reservation_map_data k b1) (reservation_map_data k b2).
Proof. rewrite /IsOp' /IsOp=> ->. by rewrite reservation_map_data_op. Qed.
Lemma reservation_map_token_union E1 E2 :
E1 ## E2
reservation_map_token (E1 E2) = reservation_map_token E1 reservation_map_token E2.
Proof.
intros. by rewrite /op /reservation_map_op_instance
/reservation_map_token /= coPset_disj_union // left_id_L.
Qed.
Lemma reservation_map_token_difference E1 E2 :
E1 E2
reservation_map_token E2 = reservation_map_token E1 reservation_map_token (E2 E1).
Proof.
intros. rewrite -reservation_map_token_union; last set_solver.
by rewrite -union_difference_L.
Qed.
Lemma reservation_map_token_valid_op E1 E2 :
(reservation_map_token E1 reservation_map_token E2) E1 ## E2.
Proof.
rewrite reservation_map_valid_eq /= {1}/op /cmra_op /=. case_decide; last done.
split; [done|]; intros _. split.
- by rewrite left_id.
- intros i. rewrite lookup_op lookup_empty. auto.
Qed.
Lemma reservation_map_alloc E k a :
k E a reservation_map_token E ~~> reservation_map_data k a.
Proof.
intros ??. apply cmra_total_update=> n [mf [Ef|]] //.
rewrite reservation_map_validN_eq /= {1}/op {1}/cmra_op /=.
case_decide; last done.
rewrite !left_id_L. intros [Hmf Hdisj]; 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.
- 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 reservation_map_updateP P (Q : reservation_map A Prop) k a :
a ~~>: P
( a', P a' Q (reservation_map_data k a'))
reservation_map_data k a ~~>: Q.
Proof.
intros Hup HP. apply cmra_total_updateP=> n [mf [Ef|]] //.
rewrite reservation_map_validN_eq /= left_id_L. intros [Hmf Hdisj].
destruct (Hup n (mf !! k)) as (a'&?&?).
{ move: (Hmf (k)).
by rewrite lookup_op lookup_singleton Some_op_opM. }
exists (reservation_map_data k a'); split; first by eauto.
rewrite /= left_id_L. 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.
- intros j. move: (Hdisj j).
rewrite !lookup_op !op_None !lookup_singleton_None. naive_solver.
Qed.
Lemma reservation_map_update k a b :
a ~~> b
reservation_map_data k a ~~> reservation_map_data k b.
Proof.
rewrite !cmra_update_updateP. eauto using reservation_map_updateP with subst.
Qed.
End cmra.
Global Arguments reservation_mapR {_} _.
Global Arguments reservation_mapUR {_} _.
From iris.prelude Require Export prelude.
From iris.prelude Require Import options.
Set Primitive Projections.
(** Step-Indices
In this file, we declare an interface for step-indices (see [sidx]).
Step-indexing in Iris enables all kinds of recursive constructions such
as mixed-variance recursive definitions and mixed-variance recursive types.
For more information on step-indexing and how it is used in Iris, see the
journal paper "Iris from the Ground Up" (which can be found at
https://doi.org/10.1017/S0956796818000151) or the Iris appendix.
Traditionally, step-indexing is done by using natural numbers as indices.
However, it can also be generalized to richer sets of indices such as ω^2 (i.e.,
pairs of natural numbers with lexicographic ordering), ω^ω (i.e., lists of
natural numbers with lexicographic ordering), countable ordinals, and even
uncountable ordinals. Large parts of Iris are agnostic of the step-indexing type
that is used, which allows us to factor out the notion of a step-index via the
interface of an "index type" (see [sidx] and [SIdxMixin]) and define them
generically in the index type (e.g., the [algebra] folder).
Note that transfinite step-indexing is not a strict improvement on finite
step-indexing (i.e., step-indexing with natural numbers)—they are incomparable
(see the paper on "Transfinite Iris" for an explanation of the trade-offs,
which can be found at https://doi.org/10.1145/3453483.3454031). Thus, to retain
the benefits of both, we make definitions parametric instead of specialized to
natural numbers or a particular type of ordinals. *)
(** An index type [I] comes with
- a well-founded strict linear order [sidx_lt] (written [i < j]),
- a total linear order [sidx_le] (written [i ≤ j]),
- a zero index [sidx_zero] (written [0ᵢ]),
- and a successor operation on indices [sidx_succ] (written [Sᵢ]).
The less-or-equal relation [≤] must be compatible with the less-than relation
[<] ([sidx_mixin_le_lteq]). That is, [i ≤ j] iff [i < j] or [i = j]. The [0ᵢ]
index has to be the smallest step-index (i.e., [¬ i ≺ zero] for all [i]) and the
successor operation [Sᵢ] should yield for a step-index [i] the least index
greater than [i]. The strict order [<] should be _computationally_ decidable,
and, finally, it should be computationally decidable whether a given step-index
is a limit index (i.e., it cannot be reached from any predecessor by applying
the successor operation).
The less-or-equal operation [sidx_le] is not strictly necessary since [i ≤ j]
is equivalent to [i < j] or [i = j]. We add it as an additional parameter to
this interface for engineering reasons: the addition of [sidx_le] simplifies
backwards compatibly to developments that assume step-indices are always natural
numbers and (implicitly) that [(i ≤ j)%sidx] unifies with [(i ≤ j)%nat] for
natural numbers. *)
Record SIdxMixin {I} (ilt ile : relation I) (zero : I) (succ : I I) := {
sidx_mixin_lt_trans : Transitive ilt;
sidx_mixin_lt_wf : well_founded ilt;
sidx_mixin_lt_trichotomy : TrichotomyT ilt;
sidx_mixin_le_lteq n m : ile n m ilt n m n = m;
sidx_mixin_nlt_0_r n : ¬ ilt n zero;
sidx_mixin_lt_succ_diag_r n : ilt n (succ n);
sidx_mixin_le_succ_l_2 n m : ilt n m ile (succ n) m;
(** An index is either a successor or a limit index. The second disjunct in
[weak_case] considers [0ᵢ] to be a limit index, which is often not desired. So
typically, instead of [weak_case] you want to use [Index.case] which has an
explicit disjunct for [0ᵢ] and uses [Index.limit] to define limit indices,
excluding [0ᵢ]. *)
sidx_mixin_weak_case n : {m | n = succ m} + ( m, ilt m n ilt (succ m) n);
}.
(** [sidx] is both a canonical structure and a typeclass.
When working with a concrete index type, we usually fix it globally. Declaring
an appropriate instance and structure enables Coq to infer the globally fixed
index type automatically in almost all places. For an example of how to
instantiate [sidx], see [stepindex_finite.v].
Using both a canonical structure and a type class is not very common. In this
case, it maximizes the inference that is done by Coq when we use the index type.
The fact that the step-index is a type class makes sure in, for example, lemma
statements that it will be inferred from the context automatically. It also
means that if we work with finite step-indices then the finite step-index
instance will be used by default as soon as we (or any of our dependencies)
import [stepindex_finite.v]. The canonical structure, on the other hand, makes
sure that the index type is inferred automatically as soon as we use either
natural numbers or ordinals as the step-index in a concrete lemma. *)
Structure sidx := SIdx {
sidx_car :> Type;
sidx_lt : relation sidx_car;
sidx_le : relation sidx_car;
sidx_zero : sidx_car;
sidx_succ : sidx_car sidx_car;
sidx_mixin : SIdxMixin sidx_lt sidx_le sidx_zero sidx_succ;
}.
Existing Class sidx.
Global Arguments SIdx _ {_ _ _ _} _.
Global Arguments sidx_car {_}.
Global Arguments sidx_lt {_}.
Global Arguments sidx_le {_}.
Global Arguments sidx_zero {_}.
Global Arguments sidx_succ {_} _.
Declare Scope sidx_scope.
Delimit Scope sidx_scope with sidx.
Bind Scope sidx_scope with sidx_car.
Local Open Scope sidx_scope.
(** We cannot overload [0] and [S], so we define custom notations with an [i]
subscript. Overloading [0] is disallowed, one should use [Number Notation], but
that only works for "concrete" number types defined as an [Inductive] (not
"abstract" number types represented by a class). One also cannot overload [S]
since it is a constructor (of [nat]) and not a [Notation]. *)
Notation "0ᵢ" := sidx_zero (at level 0).
Notation "'Sᵢ'" := sidx_succ (at level 0).
(** The remaining operations have notations in the `%sidx` scope. *)
Notation "(<)" := sidx_lt (only parsing) : sidx_scope.
Notation "n < m" := (sidx_lt n m) : sidx_scope.
Notation "(≤)" := sidx_le (only parsing) : sidx_scope.
Notation "n ≤ m" := (sidx_le n m) : sidx_scope.
Global Hint Extern 0 (_ _)%sidx => reflexivity : core.
(** Finite indices are exactly the natural numbers *)
Class SIdxFinite (SI : sidx) :=
finite_index n : n = 0 m, n = Sᵢ m.
Module SIdx.
Section sidx.
Context {SI : sidx}.
Implicit Type n m p : SI.
(** * Lifting of the mixin properties *)
(** Not an [Instance], we get that via [StrictOrder]. *)
Lemma lt_trans : Transitive (<).
Proof. eapply sidx_mixin_lt_trans, sidx_mixin. Qed.
Lemma lt_wf : well_founded (<).
Proof. eapply sidx_mixin_lt_wf, sidx_mixin. Qed.
Global Instance lt_trichotomy : TrichotomyT (<).
Proof. eapply sidx_mixin_lt_trichotomy, sidx_mixin. Qed.
Lemma le_lteq n m : n m n < m n = m.
Proof. eapply sidx_mixin_le_lteq, sidx_mixin. Qed.
Lemma nlt_0_r n : ¬ n < 0.
Proof. eapply sidx_mixin_nlt_0_r, sidx_mixin. Qed.
Lemma lt_succ_diag_r n : n < Sᵢ n.
Proof. eapply sidx_mixin_lt_succ_diag_r, sidx_mixin. Qed.
Lemma le_succ_l_2 n m : n < m Sᵢ n m.
Proof. eapply sidx_mixin_le_succ_l_2, sidx_mixin. Qed.
Lemma weak_case n : { m | n = Sᵢ m } + ( m, m < n Sᵢ m < n).
Proof. eapply sidx_mixin_weak_case, sidx_mixin. Qed.
(** * Limit indices *)
(** [weak_case] allows us to decide whether a number is a limit index.
However, according to [m < n → Sᵢ m < n] (right disjunct), 0 is a limit
index. In many definitions and proofs (e.g., [case] and [rec]), it is helpful
to think of 0 as a special case rather than a degenerate limit index. We
therefore exclude 0 from the notion of limit indices. *)
Record limit (n : SI) := Limit {
limit_gt_S m : m < n Sᵢ m < n;
limit_neq_0 : n 0;
}.
(** * Derived properties *)
Global Instance inhabited : Inhabited SI := populate 0.
Global Instance lt_strict : StrictOrder (<).
Proof.
split; [|apply lt_trans].
intros n Hn. induction (lt_wf n) as [n _ IH].
by apply IH in Hn.
Qed.
Global Instance le_po : PartialOrder ().
Proof.
split; [split|].
- intros n. apply le_lteq. auto.
- intros n m p. rewrite /sidx_le. rewrite !le_lteq.
intros [Hnm|Hnm] [Hmp|Hmp]; subst; [|by eauto..].
left. by etrans.
- intros n m [H1| ->]%le_lteq [H2|H2]%le_lteq; [|by eauto..].
destruct (irreflexivity (<) n). by trans m.
Qed.
Lemma lt_le_incl n m : n < m n m.
Proof. intros. apply le_lteq. auto. Qed.
Local Hint Resolve lt_le_incl : core.
Global Instance le_total : Total ().
Proof. intros n m. destruct (trichotomy (<) n m); naive_solver. Qed.
Lemma lt_le_trans n m p : n < m m p n < p.
Proof. intros ? [| ->]%le_lteq; [|done]. by trans m. Qed.
Lemma le_lt_trans n m p : n m m < p n < p.
Proof. intros [| ->]%le_lteq ?; [|done]. by trans m. Qed.
Lemma le_succ_l n m : Sᵢ n m n < m.
Proof.
split; [|apply le_succ_l_2].
intros. eapply lt_le_trans; [|done]. apply lt_succ_diag_r.
Qed.
Lemma lt_succ_r n m : n < Sᵢ m n m.
Proof.
split.
- intros. destruct (total () n m) as [|[H1| ->]%le_lteq]; eauto.
destruct (irreflexivity (<) n). eapply lt_le_trans; [done|].
by apply le_succ_l.
- intros. eapply le_lt_trans; [done|]. apply lt_succ_diag_r.
Qed.
Lemma succ_le_mono n m : n m Sᵢ n Sᵢ m.
Proof. by rewrite le_succ_l lt_succ_r. Qed.
Lemma succ_lt_mono n m : n < m Sᵢ n < Sᵢ m.
Proof. by rewrite lt_succ_r le_succ_l. Qed.
Global Instance succ_inj : Inj (=) (=) Sᵢ.
Proof.
intros n m Heq.
apply (anti_symm ()); apply succ_le_mono; by rewrite Heq.
Qed.
Lemma le_succ_diag_r n : n Sᵢ n.
Proof. apply lt_le_incl, lt_succ_diag_r. Qed.
Lemma neq_0_lt_0 n : n 0 0 < n.
Proof.
split.
- destruct (trichotomy (<) 0 n) as [|[|?%nlt_0_r]]; naive_solver.
- intros ? ->. by destruct (nlt_0_r 0).
Qed.
Lemma lt_ge_cases n m : n < m m n.
Proof.
destruct (trichotomyT (<) n m) as [[| ->]|]; auto using lt_le_incl.
Qed.
Lemma le_gt_cases n m : n m m < n.
Proof. destruct (lt_ge_cases m n); auto. Qed.
Lemma le_ngt n m : n m ¬ m < n.
Proof.
split.
- intros ??. apply (irreflexivity (<) n); eauto using le_lt_trans.
- intros ?. by destruct (lt_ge_cases m n).
Qed.
Lemma lt_nge n m : n < m ¬ m n.
Proof.
split.
- intros ??. apply (irreflexivity (<) n); eauto using lt_le_trans.
- intros ?. by destruct (le_gt_cases m n).
Qed.
Lemma le_neq n m : n < m n m n m.
Proof.
split.
- intros. split; [eauto using lt_le_incl|].
intros ->. by apply (irreflexivity (<) m).
- intros [? Heq]. apply lt_nge=> ?. by apply Heq, (anti_symm ()).
Qed.
Lemma nlt_succ_r n m : ¬ m < Sᵢ n n < m.
Proof. by rewrite -le_ngt le_succ_l. Qed.
(** These instances have a very high cost, so that the default instances for
[nat] are prefered over these ones in case of finite step indexing. *)
Global Instance eq_dec : EqDecision SI | 1000.
Proof.
intros n m. destruct (trichotomyT (<) n m) as [[H| ->]|H].
- right. intros ->. by apply (irreflexivity (<) m).
- by left.
- right. intros ->. by apply (irreflexivity (<) m).
Qed.
Global Instance lt_dec : RelDecision (<) | 1000.
Proof. apply _. Defined.
Global Instance le_dec : RelDecision () | 1000.
Proof.
intros n m. destruct (trichotomyT (<) n m) as [[| ->]|]; [left; by auto..|].
right. intro. destruct (irreflexivity (<) n); eauto using le_lt_trans.
Qed.
Lemma le_0_l n : 0 n.
Proof. apply le_ngt, nlt_0_r. Qed.
Lemma le_0_r n : n 0 n = 0.
Proof. split; [|by intros ->]. intros. by apply (anti_symm ()), le_0_l. Qed.
Lemma neq_succ_0 n : Sᵢ n 0.
Proof. apply neq_0_lt_0, lt_succ_r, le_0_l. Qed.
Lemma succ_neq n : n Sᵢ n.
Proof.
intros Hn. destruct (irreflexivity (<) n). rewrite {2}Hn.
apply lt_succ_diag_r.
Qed.
Lemma limit_0 : ¬limit 0.
Proof. by intros [? []]. Qed.
Lemma limit_lt_0 n : limit n 0 < n.
Proof. intros. apply neq_0_lt_0; intros ->. by apply limit_0. Qed.
Lemma limit_S n : ¬limit (Sᵢ n).
Proof.
intros [Hlim _]. apply (irreflexivity (<) (Sᵢ n)), Hlim, lt_succ_diag_r.
Qed.
Lemma limit_finite `{!SIdxFinite SI} n : ¬limit n.
Proof.
destruct (finite_index n) as [->|[m ->]]; auto using limit_0, limit_S.
Qed.
Lemma case n : (n = 0) + { m | n = Sᵢ m } + limit n.
Proof.
destruct (decide (n = 0)); [by auto|].
destruct (weak_case n); by auto using limit.
Qed.
(** * Ordinal recursion *)
(** We define the standard recursion scheme for ordinal induction (i.e.,
transfinite induction) [rec] with the right "unfolding" lemmas:
P 0 →
(∀ i, P i → P (Sᵢ i)) →
(∀ i, limit i, (∀ j, j < i → P j) → P i) →
∀ i, P i *)
Lemma lt_succ_diag_r' n m : n = Sᵢ m m < n.
Proof. intros ->. apply lt_succ_diag_r. Qed.
Definition rec (P : SI Type)
(s : P 0)
(f : n, P n P (Sᵢ n))
(lim : n, limit n ( m, m < n P m) P n) :
n, P n :=
Fix lt_wf _ $ λ n IH,
match case n with
| inl (inl EQ) => eq_rect_r P s EQ
| inl (inr (m EQ)) => eq_rect_r P (f m (IH m (lt_succ_diag_r' _ _ EQ))) EQ
| inr Hlim => lim n Hlim IH
end.
Definition rec_lim_ext {P : SI Type}
(lim : n, limit n ( m, m < n P m) P n) :=
n Hn Hn' f g,
( m Hm, f m Hm = g m Hm)
lim n Hn f = lim n Hn' g.
Lemma rec_unfold P s f lim n :
rec_lim_ext lim
rec P s f lim n =
match case n with
| inl (inl EQ) => eq_rect_r P s EQ
| inl (inr (m EQ)) => eq_rect_r P (f m (rec P s f lim m)) EQ
| inr Hlim => lim n Hlim (λ m _, rec P s f lim m)
end.
Proof.
intros Hext. unfold rec at 1. apply Fix_eq=> m g h EQ.
destruct (case m) as [[?|[m' ->]]|?]; [done| |by auto].
rewrite /eq_rect_r /= EQ //.
Qed.
Lemma rec_zero P s f lim : rec P s f lim 0 = s.
Proof.
rewrite /rec /Fix. destruct (lt_wf 0); simpl.
destruct (case 0) as [[H0|[m ?]]|?].
- by rewrite (proof_irrel H0 (eq_refl _)).
- by destruct (neq_succ_0 m).
- by destruct limit_0.
Qed.
Lemma rec_succ P s f lim n :
rec_lim_ext lim
rec P s f lim (Sᵢ n) = f n (rec P s f lim n).
Proof.
intros Hext. rewrite rec_unfold //.
destruct (case _) as [[?|[m Hm]]|?].
- by destruct (neq_succ_0 n).
- apply (inj _) in Hm as Hm'; subst. by rewrite (proof_irrel Hm (eq_refl _)).
- by destruct (limit_S n).
Qed.
Lemma rec_lim P s f lim n Hn :
rec_lim_ext lim
rec P s f lim n = lim n Hn (λ m _, rec P s f lim m).
Proof.
intros Hext. rewrite rec_unfold //.
destruct (case _) as [[->|[m ->]]|[Hlim H0]].
- by destruct limit_0.
- by destruct (limit_S m).
- by apply Hext.
Qed.
End sidx.
End SIdx.
From iris.algebra Require Import stepindex ofe cmra.
From iris.prelude Require Import options.
(** * [sidx] instance for [nat] *)
(** This file provides an instantiation of the [sidx] stepindex type for [nat],
which is the stepindex type traditionally used by Iris.
Side-effect: every development importing this file will automatically use finite
indices due to the declared instances and canonical structures for [sidx]. *)
Lemma nat_sidx_mixin : SIdxMixin lt le 0 S.
Proof.
constructor; try lia.
- apply _.
- exact lt_wf.
- intros m n. destruct (lt_eq_lt_dec m n); naive_solver.
- intros [|n]; eauto with lia.
Qed.
(** This declares the globally used index type to be [natSI]. All following
lemmas are implicitly specialized to [natSI]! *)
Canonical Structure natSI : sidx := SIdx nat nat_sidx_mixin.
Global Existing Instance natSI | 0.
Global Instance nat_sidx_finite : SIdxFinite natSI.
Proof. intros [|n]; eauto. Qed.
(** We define a notion of finite OFEs and COFEs that matches Iris's traditional
definitions, and makes it easier to define OFEs and COFEs specialized to the
[natSI] index type. *)
Section finite.
Local Set Default Proof Using "Type*".
(** Variants of lemmas with [S] and [≤] that use the definitions on [nat]
directly. These are convenient in combination with [lia], which does not
unfold the projections of the [sidx] class. *)
Lemma dist_later_S {A : ofe} (n : nat) (a b : A) :
a {n} b dist_later (S n) a b.
Proof. apply dist_later_S. Qed.
Lemma dist_S {A : ofe} n (x y : A) : x {S n} y x {n} y.
Proof. apply dist_S. Qed.
Lemma dist_le {A : ofe} (n m : nat) (x y : A) :
x {n} y m n x {m} y.
Proof. by apply dist_le. Qed.
Lemma contractive_S {A B : ofe} (f : A B) `{!Contractive f} (n : nat) x y :
x {n} y f x {S n} f y.
Proof. by apply contractive_S. Qed.
Lemma conv_compl_S `{!Cofe A} n (c : chain A) : compl c {n} c (S n).
Proof. apply conv_compl_S. Qed.
Lemma cmra_validN_S {A : cmra} n (x : A) : {S n} x {n} x.
Proof. intros ?. by eapply cmra_validN_lt, SIdx.lt_succ_diag_r. Qed.
Lemma cmra_includedN_S {A : cmra} n (x y : A) : x {S n} y x {n} y.
Proof. apply cmra_includedN_S. Qed.
(** We define [dist_later_fin], an equivalent (see dist_later_fin_iff) version
of [dist_later] that uses a [match] on the step-index instead of the
quantification over smaller step-indicies. The definition of [dist_later_fin]
matches how [dist_later] used to be defined (i.e., with a [match] on the
step-index), so [dist_later_fin] simplifies adapting existing Iris
developments that used to rely on the reduction behavior of [dist_later].
The "fin" indicates that when, in the future, the step-index is abstracted away,
this equivalence will only hold for finite step-indices (as in, ordinals without
"limit" steps such as natural numbers). *)
Definition dist_later_fin {A : ofe} (n : nat) (x y : A) :=
match n with 0 => True | S n => x {n} y end.
Lemma dist_later_fin_iff {A : ofe} (n : nat) (x y : A):
dist_later n x y dist_later_fin n x y.
Proof.
destruct n; unfold dist_later_fin; first by split; eauto using dist_later_0.
by rewrite dist_later_S.
Qed.
(** Shorthand for defining OFEs that only work for [natSI] *)
Lemma ofe_mixin_finite A `{!Equiv A, !Dist A} :
( x y : A, x y n, x {n} y)
( n, Equivalence (@dist natSI A _ n))
( n (x y : A), x {S n} y x {n} y) (* [S] instead of [<] *)
OfeMixin A.
Proof.
intros; split; [done..|].
intros n m x y Heq Hlt. induction Hlt; eauto.
Qed.
(** Shorthand for defining COFEs that only work for [natSI] *)
Program Definition cofe_finite {A} (compl : Compl A)
(conv_compl: n c, compl c {n} c n) : Cofe A :=
{| compl := compl; lbcompl n Hn := False_rect _ (SIdx.limit_finite _ Hn) |}.
Next Obligation. auto. Qed.
Next Obligation. intros. simpl. by destruct (SIdx.limit_finite _ _). Qed.
Next Obligation. intros. simpl. by destruct (SIdx.limit_finite _ _). Qed.
End finite.
(** For backwards compatibility, we define the tactic [f_contractive_fin] that
works with [dist_later_fin] and crucially relies on the step-indices being [nat]
and the reduction behavior of [dist_later_fin].
The tactic [f_contractive_fin] is useful in Iris developments that define custom
notions of [dist] and [dist_later] (e.g., RustBelt), but should be avoided if
possible. *)
(** The tactic [dist_later_fin_intro] is a special case of [dist_later_intro],
which only works for natural numbers as step-indices. It changes [dist_later] to
[dist_later_fin], which only makes sense on natural numbers. We keep
[dist_later_fin_intro] around for backwards compatibility. *)
(** For the goal [dist_later n x y], the tactic [dist_later_fin_intro] changes
the goal to [dist_later_fin] and takes care of the case where [n=0], such
that we are only left with the case where [n = S n'] for some [n']. Changing
[dist_later] to [dist_later_fin] enables reduction and thus works better with
custom versions of [dist] as used e.g. by LambdaRust. *)
Ltac dist_later_fin_intro :=
match goal with
| |- @dist_later _ ?A _ ?n ?x ?y =>
apply dist_later_fin_iff;
destruct n as [|n]; [exact I|change (@dist natSI A _ n x y)]
end.
Tactic Notation "f_contractive_fin" :=
f_contractive_prepare;
try dist_later_fin_intro;
try fast_reflexivity.
(**
This file formalizes the STS construction from the original Iris paper (POPL15).
DISCLAIMER: The definition of STSs is included in the Iris development for
historical purposes. If you plan to mechanize an Iris proof in Coq, it is
usually better to use a more direct encoding of the ghost state you need as a
resource algebra (camera). STSs are very painful to use in Coq, and they are
therefore barely used in practice.
The type [stsT] describes state-transition systems: a type of states, a type of
tokens, a step relation between states, and a token assignment function. Then
[sts_resR sts], for [sts: stsT], is the resource algebra of "STS resources",
which can be fragments ("we are in one of these states", where the set of states
needs to be closed under transitions performed without the locally owned
tokens), or authoritative ("we are exactly in this state").
The construction is performed via an intermediate internal type, [sts.car]. The
reason for this intermediate step is that composition of two STS resources is
defined only if their token sets are disjoint and the state sets are not
disjoint (i.e., they have at least one element in common). This condition is not
decidable, so we cannot use the usual approach (used e.g. in [gset_disj]) of
just composing those pairs to a dedicated "invalid" element. Instead, [sts_res]
consists of an [sts.car] element (fragment or authoritative), together with a
[Prop] defining whether this element is valid. That way we can "defer" the
validity check from composition to RA validity.
*)
From stdpp Require Export propset.
From iris.algebra Require Export cmra.
From iris.algebra Require Import dra.
Set Default Proof Using "Type".
From iris.algebra Require Export cmra updates.
From iris.prelude Require Import options.
Local Arguments valid _ _ !_ /.
Local Arguments op _ _ !_ !_ /.
Local Arguments core _ _ !_ /.
......@@ -14,9 +40,9 @@ Structure stsT := Sts {
prim_step : relation state;
tok : state propset token;
}.
Arguments Sts {_ _} _ _.
Arguments prim_step {_} _ _.
Arguments tok {_} _.
Global Arguments Sts {_ _} _ _.
Global Arguments prim_step {_} _ _.
Global Arguments tok {_} _.
Notation states sts := (propset (state sts)).
Notation tokens sts := (propset (token sts)).
......@@ -48,22 +74,22 @@ Definition up_set (S : states sts) (T : tokens sts) : states sts :=
S ≫= λ s, up s T.
(** Tactic setup *)
Hint Resolve Step : core.
Hint Extern 50 (equiv (A:=propset _) _ _) => set_solver : sts.
Hint Extern 50 (¬equiv (A:=propset _) _ _) => set_solver : sts.
Hint Extern 50 (_ _) => set_solver : sts.
Hint Extern 50 (_ _) => set_solver : sts.
Hint Extern 50 (_ ## _) => set_solver : sts.
Local Hint Resolve Step : core.
Local Hint Extern 50 (equiv (A:=propset _) _ _) => set_solver : sts.
Local Hint Extern 50 (¬equiv (A:=propset _) _ _) => set_solver : sts.
Local Hint Extern 50 (_ _) => set_solver : sts.
Local Hint Extern 50 (_ _) => set_solver : sts.
Local Hint Extern 50 (_ ## _) => set_solver : sts.
(** ** Setoids *)
Instance frame_step_mono : Proper (flip () ==> (=) ==> (=) ==> impl) frame_step.
Local Instance frame_step_mono : Proper (flip () ==> (=) ==> (=) ==> impl) frame_step.
Proof.
intros ?? HT ?? <- ?? <-; destruct 1; econstructor;
eauto with sts; set_solver.
Qed.
Global Instance frame_step_proper : Proper (() ==> (=) ==> (=) ==> iff) frame_step.
Proof. move=> ?? /set_equiv_spec [??]; split; by apply frame_step_mono. Qed.
Instance closed_proper' : Proper (() ==> () ==> impl) closed.
Proof. move=> ?? /set_equiv_subseteq [??]; split; by apply frame_step_mono. Qed.
Local Instance closed_proper' : Proper (() ==> () ==> impl) closed.
Proof. destruct 3; constructor; intros; setoid_subst; eauto. Qed.
Global Instance closed_proper : Proper (() ==> () ==> iff) closed.
Proof. by split; apply closed_proper'. Qed.
......@@ -75,7 +101,7 @@ Proof.
Qed.
Global Instance up_proper : Proper ((=) ==> () ==> ()) up.
Proof.
by move=> ??? ?? /set_equiv_spec [??]; split; apply up_preserving.
by move=> ??? ?? /set_equiv_subseteq [??]; split; apply up_preserving.
Qed.
Global Instance up_set_preserving : Proper (() ==> flip () ==> ()) up_set.
Proof.
......@@ -84,7 +110,7 @@ Proof.
Qed.
Global Instance up_set_proper : Proper (() ==> () ==> ()) up_set.
Proof.
move=> S1 S2 /set_equiv_spec [??] T1 T2 /set_equiv_spec [??];
move=> S1 S2 /set_equiv_subseteq [??] T1 T2 /set_equiv_subseteq [??];
split; by apply up_set_preserving.
Qed.
......@@ -149,7 +175,7 @@ Lemma closed_up_empty s : closed (up s ∅) ∅.
Proof. eauto using closed_up with sts. Qed.
Lemma up_closed S T : closed S T up_set S T S.
Proof.
intros ?; apply set_equiv_spec; split; auto using subseteq_up_set.
intros ?; apply set_equiv_subseteq; split; auto using subseteq_up_set.
intros s; unfold up_set; rewrite elem_of_bind; intros (s'&Hstep&?).
induction Hstep; eauto using closed_step.
Qed.
......@@ -168,138 +194,276 @@ End sts.
Notation steps := (rtc step).
Notation frame_steps T := (rtc (frame_step T)).
(* The type of bounds we can give to the state of an STS. This is the type
that we equip with an RA structure. *)
(* The type of bounds we can give to the state of an STS. On paper, this is the
type that we equip with an RA structure. In Coq we have to do some work to
model composition only being defined under some non-computable conditions. *)
Inductive car (sts : stsT) :=
| auth : state sts propset (token sts) car sts
| frag : propset (state sts) propset (token sts) car sts.
Arguments auth {_} _ _.
Arguments frag {_} _ _.
Global Arguments auth {_} _ _.
Global Arguments frag {_} _ _.
End sts.
Notation stsT := sts.stsT.
Notation Sts := sts.Sts.
(** * STSs form a disjoint RA *)
Section sts_dra.
Context (sts : stsT).
(** * STSs form an RA *)
Section sts_res.
Context {SI : sidx} {sts : stsT}.
Import sts.
Implicit Types S : states sts.
Implicit Types T : tokens sts.
Inductive sts_equiv : Equiv (car sts) :=
Inductive sts_car_equiv : Equiv (car sts) :=
| auth_equiv s T1 T2 : T1 T2 auth s T1 auth s T2
| frag_equiv S1 S2 T1 T2 : T1 T2 S1 S2 frag S1 T1 frag S2 T2.
Existing Instance sts_equiv.
Instance sts_valid : Valid (car sts) := λ x,
Local Existing Instance sts_car_equiv.
Local Instance sts_car_valid_instance : Valid (car sts) := λ x,
match x with
| auth s T => tok s ## T
| frag S' T => closed S' T s, s S'
end.
Instance sts_core : Core (car sts) := λ x,
match x with
Local Instance sts_car_core_instance : PCore (car sts) := λ x,
Some match x with
| frag S' _ => frag (up_set S' )
| auth s _ => frag (up s )
end.
Inductive sts_disjoint : Disjoint (car sts) :=
Inductive sts_car_disjoint_instance : Disjoint (car sts) :=
| frag_frag_disjoint S1 S2 T1 T2 :
( s, s S1 S2) T1 ## T2 frag S1 T1 ## frag S2 T2
| auth_frag_disjoint s S T1 T2 : s S T1 ## T2 auth s T1 ## frag S T2
| frag_auth_disjoint s S T1 T2 : s S T1 ## T2 frag S T1 ## auth s T2.
Existing Instance sts_disjoint.
Instance sts_op : Op (car sts) := λ x1 x2,
Local Existing Instance sts_car_disjoint_instance.
Local Instance sts_op_instance : Op (car sts) := λ x1 x2,
match x1, x2 with
| frag S1 T1, frag S2 T2 => frag (S1 S2) (T1 T2)
| auth s T1, frag _ T2 => auth s (T1 T2)
| frag _ T1, auth s T2 => auth s (T1 T2)
| auth s T1, auth _ T2 => auth s (T1 T2)(* never happens *)
| auth s T1, auth _ T2 => auth s (T1 T2) (* never happens *)
end.
Hint Extern 50 (equiv (A:=propset _) _ _) => set_solver : sts.
Hint Extern 50 ( s : state sts, _) => set_solver : sts.
Hint Extern 50 (_ _) => set_solver : sts.
Hint Extern 50 (_ _) => set_solver : sts.
Hint Extern 50 (_ ## _) => set_solver : sts.
Local Hint Extern 50 (equiv (A:=propset _) _ _) => set_solver : sts.
Local Hint Extern 50 ( s : state sts, _) => set_solver : sts.
Local Hint Extern 50 (_ _) => set_solver : sts.
Local Hint Extern 50 (_ _) => set_solver : sts.
Local Hint Extern 50 (_ ## _) => set_solver : sts.
Global Instance auth_proper s : Proper (() ==> ()) (@auth sts s).
Proof. by constructor. Qed.
Global Instance frag_proper : Proper (() ==> () ==> ()) (@frag sts).
Proof. by constructor. Qed.
Instance sts_equivalence: Equivalence (() : relation (car sts)).
Local Instance sts_car_equivalence: Equivalence (() : relation (car sts)).
Proof.
split.
- by intros []; constructor.
- by destruct 1; constructor.
- destruct 1; inversion_clear 1; constructor; etrans; eauto.
Qed.
Lemma sts_dra_mixin : DraMixin (car sts).
Local Instance sts_car_op_proper : Proper ((≡@{car sts}) ==> () ==> ()) ().
Proof. by do 2 destruct 1; constructor; setoid_subst. Qed.
Local Instance sts_car_core_proper : Proper ((≡@{car sts}) ==> ()) core.
Proof. by destruct 1; constructor; setoid_subst. Qed.
Local Instance sts_car_valid_proper : Proper ((≡@{car sts}) ==> impl) valid.
Proof. by destruct 1; simpl; intros ?; setoid_subst. Qed.
Local Instance sts_car_valid_proper' : Proper ((≡@{car sts}) ==> iff) valid.
Proof. by split; apply: sts_car_valid_proper. Qed.
Local Instance sts_car_disjoint_proper (x : car sts) :
Proper (() ==> impl) (disjoint x).
Proof.
split.
- apply _.
- by do 2 destruct 1; constructor; setoid_subst.
- by destruct 1; constructor; setoid_subst.
- by destruct 1; simpl; intros ?; setoid_subst.
- by intros ? [|]; destruct 1; inversion_clear 1; econstructor; setoid_subst.
- destruct 3; simpl in *; destruct_and?; eauto using closed_op;
match goal with H : closed _ _ |- _ => destruct H end; set_solver.
- intros []; naive_solver eauto using closed_up, closed_up_set,
by intros ? [|]; destruct 1; inversion_clear 1; econstructor; setoid_subst.
Qed.
Local Instance sts_car_disjoint_symmetric : Symmetric (@disjoint (car sts) _).
Proof. destruct 1; constructor; auto with sts. Qed.
Local Instance sts_car_disjoint_proper' :
Proper ((≡@{car sts}) ==> () ==> iff) disjoint.
Proof.
intros x1 x2 Hx y1 y2 Hy; split.
- by rewrite Hy (symmetry_iff (##) x1) (symmetry_iff (##) x2) Hx.
- by rewrite -Hy (symmetry_iff (##) x2) (symmetry_iff (##) x1) -Hx.
Qed.
Local Lemma sts_car_core_valid (x : car sts) :
x core x.
Proof.
destruct x; naive_solver eauto using closed_up, closed_up_set,
elem_of_up, elem_of_up_set with sts.
- intros [] [] [] _ _ _ _ _; constructor; rewrite ?assoc; auto with sts.
- destruct 4; inversion_clear 1; constructor; auto with sts.
- destruct 4; inversion_clear 1; constructor; auto with sts.
- destruct 1; constructor; auto with sts.
- destruct 3; constructor; auto with sts.
- intros []; constructor; eauto with sts.
- intros []; constructor; auto with sts.
- intros [s T|S T]; constructor; auto with sts.
+ rewrite (up_closed (up _ _)); auto using closed_up with sts.
+ rewrite (up_closed (up_set _ _)); eauto using closed_up_set with sts.
- intros x y. exists (core (x y))=> ?? Hxy; split_and?.
+ destruct Hxy; constructor; unfold up_set; set_solver.
+ destruct Hxy; simpl;
eauto using closed_up_set_empty, closed_up_empty with sts.
+ destruct Hxy; econstructor;
repeat match goal with
| |- context [ up_set ?S ?T ] =>
unless (S up_set S T) by done; pose proof (subseteq_up_set S T)
| |- context [ up ?s ?T ] =>
unless (s up s T) by done; pose proof (elem_of_up s T)
end; auto with sts.
Qed.
Canonical Structure stsDR : draT := DraT (car sts) sts_dra_mixin.
End sts_dra.
(** * The STS Resource Algebra *)
Qed.
Local Lemma sts_car_op_valid (x y : car sts) :
x y x ## y (x y).
Proof.
destruct 3; simpl in *; destruct_and?; eauto using closed_op;
select (closed _ _) (fun H => destruct H); set_solver.
Qed.
Local Lemma sts_car_op_assoc (x y z : car sts) :
x y z x ## y x y ## z x (y z) (x y) z.
Proof.
destruct x, y, z; intros _ _ _ _ _; constructor; rewrite ?assoc; auto with sts.
Qed.
Local Lemma sts_car_op_comm (x y : car sts) :
x y x ## y x y y x.
Proof. destruct 3; constructor; auto with sts. Qed.
Local Lemma sts_car_disjoint_ll (x y z : car sts) :
x y z x ## y x y ## z x ## z.
Proof.
destruct 4; inversion_clear 1; constructor; auto with sts.
Qed.
Local Lemma sts_car_disjoint_rl (x y z : car sts) :
x y z y ## z x ## y z x ## y.
Proof. intros ???. rewrite !(symmetry_iff _ x). by apply sts_car_disjoint_ll. Qed.
Local Lemma sts_car_disjoint_lr (x y z : car sts) :
x y z x ## y x y ## z y ## z.
Proof. intros ????. rewrite sts_car_op_comm //. by apply sts_car_disjoint_ll. Qed.
Local Lemma sts_car_disjoint_move_l (x y z : car sts) :
x y z x ## y x y ## z x ## y z.
Proof. destruct 4; inversion_clear 1; constructor; auto with sts. Qed.
Local Lemma sts_car_disjoint_move_r (a b c : car sts) :
a b c b ## c a ## b c a b ## c.
Proof.
intros; symmetry; rewrite sts_car_op_comm; eauto using sts_car_disjoint_rl.
apply sts_car_disjoint_move_l; auto; by rewrite sts_car_op_comm.
Qed.
Local Hint Immediate sts_car_disjoint_move_l sts_car_disjoint_move_r : core.
Local Lemma sts_car_core_disjoint_l (x : car sts) : x core x ## x.
Proof. destruct x; constructor; eauto with sts. Qed.
Local Lemma sts_car_core_l (x : car sts) : x core x x x.
Proof. destruct x; constructor; eauto with sts. Qed.
Local Lemma sts_car_core_idemp (x : car sts) : x core (core x) core x.
Proof.
destruct x as [s T|S T]; constructor; auto with sts.
+ rewrite (up_closed (up _ _)); auto using closed_up with sts.
+ rewrite (up_closed (up_set _ _)); eauto using closed_up_set with sts.
Qed.
Local Lemma sts_car_core_mono (x y : car sts) :
z, x y x ## y core (x y) core x z z core x ## z.
Proof.
exists (core (x y))=> ?? Hxy; split_and?.
+ destruct Hxy; constructor; unfold up_set; set_solver.
+ destruct Hxy; simpl;
eauto using closed_up_set_empty, closed_up_empty with sts.
+ destruct Hxy; econstructor;
repeat match goal with
| |- context [ up_set ?S ?T ] =>
unless (S up_set S T) by done; pose proof (subseteq_up_set S T)
| |- context [ up ?s ?T ] =>
unless (s up s T) by done; pose proof (elem_of_up s T)
end; auto with sts.
Qed.
(** The resource type for [sts]. *)
Record sts_res := StsRes {
(** The underlying STS carrier element, storing the actual data. *)
sts_car : car sts;
(** Defines whether this element is valid. *)
sts_valid : Prop;
(** Valid elements must have a valid carrier element. *)
sts_valid_prf : sts_valid sts_car
}.
Add Printing Constructor sts_res.
Global Arguments StsRes _ _ {_}.
(** Setoid and OFE for [sts_res]. *)
Local Instance sts_equiv : Equiv sts_res := λ x y,
(sts_valid x sts_valid y) (sts_valid x sts_car x sts_car y).
Local Instance sts_equivalence : Equivalence (@equiv sts_res _).
Proof.
split; unfold equiv, sts_equiv.
- by intros [x px ?]; simpl.
- intros [x px ?] [y py ?]; naive_solver.
- intros [x px ?] [y py ?] [z pz ?] [? Hxy] [? Hyz]; simpl in *.
split; [|intros; trans y]; tauto.
Qed.
Canonical Structure sts_resO {SI : sidx} : ofe := discreteO sts_res.
(** RA for [sts_res]. *)
Local Instance sts_res_valid_instance : Valid sts_res := sts_valid.
Local Program Instance sts_res_pcore_instance : PCore sts_res := λ x,
Some (StsRes (core (sts_car x)) ( x)).
Next Obligation.
intros []; naive_solver eauto using sts_car_core_valid.
Qed.
Local Program Instance sts_res_op_instance : Op sts_res := λ x y,
StsRes (sts_car x sts_car y)
( x y sts_car x ## sts_car y).
Next Obligation.
intros [] []; naive_solver eauto using sts_car_op_valid.
Qed.
Definition sts_res_ra_mixin : RAMixin sts_res.
Proof.
apply ra_total_mixin; first eauto.
- intros ??? [? Heq]; split; simpl; [|intros (?&?&?); by rewrite Heq].
split; intros (?&?&?); split_and!;
first [rewrite ?Heq; tauto|rewrite -?Heq; tauto|tauto].
- by intros ?? [? Heq]; split; [done|]; simpl; intros ?; rewrite Heq.
- intros ?? [??]; naive_solver.
- intros [x px ?] [y py ?] [z pz ?]; split; simpl;
[intuition eauto 2 using sts_car_disjoint_lr, sts_car_disjoint_rl
|intuition eauto using sts_car_op_assoc, sts_car_disjoint_rl].
- intros [x px ?] [y py ?]; split; naive_solver eauto using sts_car_op_comm.
- intros [x px ?]; split;
naive_solver eauto using sts_car_core_l, sts_car_core_disjoint_l.
- intros [x px ?]; split; naive_solver eauto using sts_car_core_idemp.
- intros [x px ?] [y py ?] [[z pz ?] [? Hy]]; simpl in *.
destruct (sts_car_core_mono x z) as (z'&Hz').
unshelve eexists (StsRes z' (px py pz)).
{ intros (?&?&?); apply Hz'; tauto. }
split; simpl; first tauto.
intros. rewrite Hy //. tauto.
- by intros [x px ?] [y py ?] (?&?&?).
Qed.
Canonical Structure sts_resR : cmra :=
discreteR sts_res sts_res_ra_mixin.
Global Instance sts_res_disrete_cmra : CmraDiscrete sts_resR.
Proof. apply discrete_cmra_discrete. Qed.
Global Instance sts_res_cmra_total : CmraTotal sts_resR.
Proof. rewrite /CmraTotal; eauto. Qed.
Local Definition to_sts_res (x : car sts) : sts_res :=
@StsRes x (valid x) id.
Global Instance to_sts_res_proper : Proper (() ==> ()) to_sts_res.
Proof. by intros x1 x2 Hx; split; rewrite /= Hx. Qed.
Lemma to_sts_res_op a b :
( (a b) a b a ## b)
to_sts_res (a b) to_sts_res a to_sts_res b.
Proof. split; naive_solver eauto using sts_car_op_valid. Qed.
End sts_res.
Global Arguments sts_resR {_} _.
(** Finally, the general theory of STS that should be used by users *)
Notation stsC sts := (validityO (stsDR sts)).
Notation stsR sts := (validityR (stsDR sts)).
Section sts_definitions.
Context {sts : stsT}.
Definition sts_auth (s : sts.state sts) (T : sts.tokens sts) : stsR sts :=
to_validity (sts.auth s T).
Definition sts_frag (S : sts.states sts) (T : sts.tokens sts) : stsR sts :=
to_validity (sts.frag S T).
Definition sts_frag_up (s : sts.state sts) (T : sts.tokens sts) : stsR sts :=
Context {SI : sidx} {sts : stsT}.
Definition sts_auth (s : sts.state sts) (T : sts.tokens sts) : sts_resR sts :=
to_sts_res (sts.auth s T).
Definition sts_frag (S : sts.states sts) (T : sts.tokens sts) : sts_resR sts :=
to_sts_res (sts.frag S T).
Definition sts_frag_up (s : sts.state sts) (T : sts.tokens sts) : sts_resR sts :=
sts_frag (sts.up s T) T.
End sts_definitions.
Instance: Params (@sts_auth) 2 := {}.
Instance: Params (@sts_frag) 1 := {}.
Instance: Params (@sts_frag_up) 2 := {}.
Global Instance: Params (@sts_auth) 3 := {}.
Global Instance: Params (@sts_frag) 2 := {}.
Global Instance: Params (@sts_frag_up) 3 := {}.
Section stsRA.
Import sts.
Context {sts : stsT}.
Context {SI : sidx} {sts : stsT}.
Implicit Types s : state sts.
Implicit Types S : states sts.
Implicit Types T : tokens sts.
Arguments dra_valid _ !_/.
Local Arguments cmra_valid _ !_/.
(** Setoids *)
Global Instance sts_auth_proper s : Proper (() ==> ()) (sts_auth s).
Proof. solve_proper. Qed.
Global Instance sts_frag_proper : Proper (() ==> () ==> ()) (@sts_frag sts).
Global Instance sts_frag_proper : Proper (() ==> () ==> ()) (@sts_frag SI sts).
Proof. solve_proper. Qed.
Global Instance sts_frag_up_proper s : Proper (() ==> ()) (sts_frag_up s).
Proof. solve_proper. Qed.
......@@ -313,7 +477,9 @@ Lemma sts_frag_up_valid s T : ✓ sts_frag_up s T ↔ tok s ## T.
Proof.
split.
- move=>/sts_frag_valid [H _]. apply H, elem_of_up.
- intros. apply sts_frag_valid; split. by apply closed_up. set_solver.
- intros. apply sts_frag_valid; split.
+ by apply closed_up.
+ set_solver+.
Qed.
Lemma sts_auth_frag_valid_inv s S T1 T2 :
......@@ -321,14 +487,14 @@ Lemma sts_auth_frag_valid_inv s S T1 T2 :
Proof. by intros (?&?&Hdisj); inversion Hdisj. Qed.
(** Op *)
Lemma sts_op_auth_frag s S T :
Lemma sts_auth_frag_op s S T :
s S closed S T sts_auth s sts_frag S T sts_auth s T.
Proof.
intros; split; [split|constructor; set_solver]; simpl.
- intros (?&?&?); by apply closed_disjoint with S.
- intros; split_and?; last constructor; set_solver.
Qed.
Lemma sts_op_auth_frag_up s T :
Lemma sts_auth_frag_up_op s T :
sts_auth s sts_frag_up s T sts_auth s T.
Proof.
intros; split; [split|constructor; set_solver]; simpl.
......@@ -340,18 +506,18 @@ Proof.
+ constructor; last set_solver. apply elem_of_up.
Qed.
Lemma sts_op_frag S1 S2 T1 T2 :
Lemma sts_frag_op S1 S2 T1 T2 :
T1 ## T2 sts.closed S1 T1 sts.closed S2 T2
sts_frag (S1 S2) (T1 T2) sts_frag S1 T1 sts_frag S2 T2.
Proof.
intros HT HS1 HS2. rewrite /sts_frag -to_validity_op //.
intros HT HS1 HS2. rewrite /sts_frag -to_sts_res_op //.
move=>/=[?[? ?]]. split_and!; [set_solver..|constructor; set_solver].
Qed.
(* Notice that the following does *not* hold -- the composition of the
two closures is weaker than the closure with the itnersected token
set. Also see up_op.
Lemma sts_op_frag_up s T1 T2 :
Lemma sts_frag_up_op s T1 T2 :
T1 ## T2 → sts_frag_up s (T1 ∪ T2) ≡ sts_frag_up s T1 ⋅ sts_frag_up s T2.
*)
......@@ -359,8 +525,10 @@ Lemma sts_op_frag_up s T1 T2 :
Lemma sts_update_auth s1 s2 T1 T2 :
steps (s1,T1) (s2,T2) sts_auth s1 T1 ~~> sts_auth s2 T2.
Proof.
intros ?; apply validity_update.
inversion 3 as [|? S ? Tf|]; simplify_eq/=; destruct_and?.
intros ?. apply cmra_discrete_total_update.
intros [x x_val Hx_val]; simpl. intros (Htok & Hval & Hdisj).
specialize (Hx_val Hval).
inversion Hdisj as [|? S ? Tf|]; simplify_eq/=; destruct_and?.
destruct (steps_closed s1 s2 T1 T2 S Tf) as (?&?&?); auto; [].
repeat (done || constructor).
Qed.
......@@ -368,11 +536,19 @@ Qed.
Lemma sts_update_frag S1 S2 T1 T2 :
(closed S1 T1 closed S2 T2 S1 S2 T2 T1) sts_frag S1 T1 ~~> sts_frag S2 T2.
Proof.
rewrite /sts_frag=> HC HS HT. apply validity_update.
inversion 3 as [|? S ? Tf|]; simplify_eq/=;
rewrite /sts_frag=> HC HS HT. apply cmra_discrete_total_update.
intros [x x_val Hx_val]; simpl. intros (Htok & Hval & Hdisj).
specialize (Hx_val Hval).
inversion Hdisj as [|? S ? Tf|]; simplify_eq/=;
(destruct HC as (? & ? & ?); first by destruct_and?).
- split_and!. done. set_solver. constructor; set_solver.
- split_and!. done. set_solver. constructor; set_solver.
- split_and!; first done.
+ set_solver.
+ done.
+ constructor; set_solver.
- split_and!; first done.
+ set_solver.
+ done.
+ constructor; set_solver.
Qed.
Lemma sts_update_frag_up s1 S2 T1 T2 :
......@@ -387,14 +563,18 @@ Lemma sts_up_set_intersection S1 Sf Tf :
closed Sf Tf S1 Sf S1 up_set (S1 Sf) Tf.
Proof.
intros Hclf. apply (anti_symm ()).
- move=>s [HS1 HSf]. split. by apply HS1. by apply subseteq_up_set.
- move=>s [HS1 HSf]. split.
+ by apply HS1.
+ by apply subseteq_up_set.
- move=>s [HS1 [s' [/elem_of_PropSet Hsup Hs']]]. split; first done.
eapply closed_steps, Hsup; first done. set_solver +Hs'.
Qed.
Global Instance sts_frag_core_id S : CoreId (sts_frag S ).
Proof.
constructor; split=> //= [[??]]. by rewrite /dra.dra_core /= sts.up_closed.
constructor; split=> //= [[??]].
(* FIXME: rewriting with [sts.up_closed] for some reason fails here. *)
f_equiv. by rewrite sts.up_closed.
Qed.
(** Inclusion *)
......@@ -447,8 +627,8 @@ Structure stsT := Sts {
state : Type;
prim_step : relation state;
}.
Arguments Sts {_} _.
Arguments prim_step {_} _ _.
Global Arguments Sts {_} _.
Global Arguments prim_step {_} _ _.
Notation states sts := (propset (state sts)).
Definition stsT_token := Empty_set.
......@@ -490,7 +670,7 @@ Notation sts_notokT := sts_notok.stsT.
Notation Sts_NoTok := sts_notok.Sts.
Section sts_notokRA.
Context {sts : sts_notokT}.
Context {SI : sidx} {sts : sts_notokT}.
Import sts_notok.
Implicit Types s : state sts.
Implicit Types S : states sts.
......
(** This file provides an "unbounded" version of the fractional camera whose
elements are in the interval (0,..) instead of (0,1]. *)
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 [frac] is used more often, we
define [ufrac] through a [Definition] instead of a [Notation]. That way, Coq
infers the [frac] camera by default when using the [Qp] type. *)
Definition ufrac := Qp.
Section ufrac.
Context {SI : sidx}.
Implicit Types p q : ufrac.
Canonical Structure ufracO := leibnizO ufrac.
Local Instance ufrac_valid_instance : Valid ufrac := λ x, True.
Local Instance ufrac_pcore_instance : PCore ufrac := λ _, None.
Local Instance ufrac_op_instance : Op ufrac := λ x y, (x + y)%Qp.
Lemma ufrac_op p q : p q = (p + q)%Qp.
Proof. done. Qed.
Lemma ufrac_included p q : p q (p < q)%Qp.
Proof. by rewrite Qp.lt_sum. Qed.
Corollary ufrac_included_weak p q : p q (p q)%Qp.
Proof. rewrite ufrac_included. apply Qp.lt_le_incl. Qed.
Definition ufrac_ra_mixin : RAMixin ufrac.
Proof. split; try apply _; try done. Qed.
Canonical Structure ufracR := discreteR ufrac ufrac_ra_mixin.
Global Instance ufrac_cmra_discrete : CmraDiscrete ufracR.
Proof. apply discrete_cmra_discrete. Qed.
Global Instance ufrac_cancelable q : Cancelable q.
Proof. intros n p1 p2 _. apply (inj (Qp.add q)). Qed.
Global Instance ufrac_id_free q : IdFree q.
Proof. intros p _. apply Qp.add_id_free. Qed.
Global Instance is_op_ufrac q : IsOp' q (q/2)%Qp (q/2)%Qp.
Proof. by rewrite /IsOp' /IsOp ufrac_op Qp.div_2. Qed.
End ufrac.
From iris.algebra Require Export cmra.
Set Default Proof Using "Type".
From iris.prelude Require Import options.
(** * Frame preserving updates *)
(* This quantifies over [option A] for the frame. That is necessary to
make the following hold:
x ~~> P → Some c ~~> Some P
*)
Definition cmra_updateP {A : cmraT} (x : A) (P : A Prop) := n mz,
Definition cmra_updateP {SI : sidx} {A : cmra} (x : A) (P : A Prop) := n mz,
{n} (x ? mz) y, P y {n} (y ? mz).
Instance: Params (@cmra_updateP) 1 := {}.
Global Instance: Params (@cmra_updateP) 1 := {}.
Infix "~~>:" := cmra_updateP (at level 70).
Definition cmra_update {A : cmraT} (x y : A) := n mz,
Definition cmra_update {SI : sidx} {A : cmra} (x y : A) := n mz,
{n} (x ? mz) {n} (y ? mz).
Infix "~~>" := cmra_update (at level 70).
Instance: Params (@cmra_update) 1 := {}.
Global Instance: Params (@cmra_update) 1 := {}.
Section updates.
Context {A : cmraT}.
Context {SI : sidx} {A : cmra}.
Implicit Types x y : A.
Global Instance cmra_updateP_proper :
Proper (() ==> pointwise_relation _ iff ==> iff) (@cmra_updateP A).
Proper (() ==> pointwise_relation _ iff ==> iff) (@cmra_updateP SI A).
Proof.
rewrite /pointwise_relation /cmra_updateP=> x x' Hx P P' HP;
split=> ? n mz; setoid_subst; naive_solver.
Qed.
Global Instance cmra_update_proper :
Proper (() ==> () ==> iff) (@cmra_update A).
Proper (() ==> () ==> iff) (@cmra_update SI A).
Proof.
rewrite /cmra_update=> x x' Hx y y' Hy; split=> ? n mz ?; setoid_subst; auto.
Qed.
......@@ -52,13 +52,28 @@ Lemma cmra_update_exclusive `{!Exclusive x} y:
Proof. move=>??[z|]=>[/exclusiveN_l[]|_]. by apply cmra_valid_validN. Qed.
(** Updates form a preorder. *)
Global Instance cmra_update_preorder : PreOrder (@cmra_update A).
(** We set this rewrite relation's cost above the stdlib's
([impl], [iff], [eq], ...) and [≡] but below [⊑].
[eq] (at 100) < [≡] (at 150) < [cmra_update] (at 170) < [⊑] (at 200) *)
Global Instance cmra_update_rewrite_relation :
RewriteRelation (@cmra_update SI A) | 170 := {}.
Global Instance cmra_update_preorder : PreOrder (@cmra_update SI A).
Proof.
split.
- intros x. by apply cmra_update_updateP, cmra_updateP_id.
- intros x y z. rewrite !cmra_update_updateP.
eauto using cmra_updateP_compose with subst.
Qed.
Global Instance cmra_update_proper_update :
Proper (flip cmra_update ==> cmra_update ==> impl) (@cmra_update SI A).
Proof.
intros x1 x2 Hx y1 y2 Hy ?. etrans; [apply Hx|]. by etrans; [|apply Hy].
Qed.
Global Instance cmra_update_flip_proper_update :
Proper (cmra_update ==> flip cmra_update ==> flip impl) (@cmra_update SI A).
Proof.
intros x1 x2 Hx y1 y2 Hy ?. etrans; [apply Hx|]. by etrans; [|apply Hy].
Qed.
Lemma cmra_updateP_op (P1 P2 Q : A Prop) x1 x2 :
x1 ~~>: P1 x2 ~~>: P2 ( y1 y2, P1 y1 P2 y2 Q (y1 y2))
......@@ -80,54 +95,74 @@ Proof.
rewrite !cmra_update_updateP; eauto using cmra_updateP_op with congruence.
Qed.
Global Instance cmra_update_op_proper :
Proper (cmra_update ==> cmra_update ==> cmra_update) (op (A:=A)).
Proof. intros x1 x2 Hx y1 y2 Hy. by apply cmra_update_op. Qed.
Global Instance cmra_update_op_flip_proper :
Proper (flip cmra_update ==> flip cmra_update ==> flip cmra_update) (op (A:=A)).
Proof. intros x1 x2 Hx y1 y2 Hy. by apply cmra_update_op. Qed.
Lemma cmra_update_op_l x y : x y ~~> x.
Proof. intros n mz. rewrite comm cmra_op_opM_assoc. apply cmra_validN_op_r. Qed.
Lemma cmra_update_op_r x y : x y ~~> y.
Proof. rewrite comm. apply cmra_update_op_l. Qed.
Lemma cmra_update_valid0 x y : ({0} x x ~~> y) x ~~> y.
Lemma cmra_update_included x y : x y y ~~> x.
Proof. intros [z ->]. apply cmra_update_op_l. Qed.
Lemma cmra_update_valid0 x y : ({0} x x ~~> y) x ~~> y.
Proof.
intros H n mz Hmz. apply H, Hmz.
apply (cmra_validN_le n); last lia.
destruct mz. eapply cmra_validN_op_l, Hmz. apply Hmz.
apply (cmra_validN_le n), SIdx.le_0_l.
destruct mz.
- eapply cmra_validN_op_l, Hmz.
- apply Hmz.
Qed.
(** ** Frame preserving updates for total CMRAs *)
Section total_updates.
Local Set Default Proof Using "Type*".
Context `{CmraTotal A}.
Lemma cmra_total_updateP x (P : A Prop) :
x ~~>: P n z, {n} (x z) y, P y {n} (y z).
Proof.
split=> Hup; [intros n z; apply (Hup n (Some z))|].
intros n [z|] ?; simpl; [by apply Hup|].
destruct (Hup n (core x)) as (y&?&?); first by rewrite cmra_core_r.
eauto using cmra_validN_op_l.
Qed.
Lemma cmra_total_update x y : x ~~> y n z, {n} (x z) {n} (y z).
Proof. rewrite cmra_update_updateP cmra_total_updateP. naive_solver. Qed.
(** ** Frame preserving updates for total and discete CMRAs *)
Lemma cmra_total_updateP `{!CmraTotal A} x (P : A Prop) :
x ~~>: P n z, {n} (x z) y, P y {n} (y z).
Proof.
split=> Hup; [intros n z; apply (Hup n (Some z))|].
intros n [z|] ?; simpl; [by apply Hup|].
destruct (Hup n (core x)) as (y&?&?); first by rewrite cmra_core_r.
eauto using cmra_validN_op_l.
Qed.
Lemma cmra_total_update `{!CmraTotal A} x y :
x ~~> y n z, {n} (x z) {n} (y z).
Proof. rewrite cmra_update_updateP cmra_total_updateP. naive_solver. Qed.
Context `{CmraDiscrete A}.
Lemma cmra_discrete_updateP `{!CmraDiscrete A} (x : A) (P : A Prop) :
x ~~>: P mz, (x ? mz) y, P y (y ? mz).
Proof.
unfold cmra_updateP. setoid_rewrite <-cmra_discrete_valid_iff.
pose 0. naive_solver.
Qed.
Lemma cmra_discrete_update `{!CmraDiscrete A} (x y : A) :
x ~~> y mz, (x ? mz) (y ? mz).
Proof.
unfold cmra_update. setoid_rewrite <-cmra_discrete_valid_iff.
pose 0. naive_solver.
Qed.
Lemma cmra_discrete_updateP (x : A) (P : A Prop) :
x ~~>: P z, (x z) y, P y (y z).
Proof.
rewrite cmra_total_updateP; setoid_rewrite <-cmra_discrete_valid_iff.
naive_solver eauto using 0.
Qed.
Lemma cmra_discrete_update (x y : A) :
x ~~> y z, (x z) (y z).
Proof.
rewrite cmra_total_update; setoid_rewrite <-cmra_discrete_valid_iff.
naive_solver eauto using 0.
Qed.
End total_updates.
Lemma cmra_discrete_total_updateP `{!CmraDiscrete A, !CmraTotal A}
(x : A) (P : A Prop) :
x ~~>: P z, (x z) y, P y (y z).
Proof.
rewrite cmra_total_updateP; setoid_rewrite <-cmra_discrete_valid_iff.
pose 0. naive_solver.
Qed.
Lemma cmra_discrete_total_update `{!CmraDiscrete A, !CmraTotal A} (x y : A) :
x ~~> y z, (x z) (y z).
Proof.
rewrite cmra_total_update; setoid_rewrite <-cmra_discrete_valid_iff.
pose 0. naive_solver.
Qed.
End updates.
(** * Transport *)
Section cmra_transport.
Context {A B : cmraT} (H : A = B).
Context {SI : sidx} {A B : cmra} (H : A = B).
Notation T := (cmra_transport H).
Lemma cmra_transport_updateP (P : A Prop) (Q : B Prop) x :
x ~~>: P ( y, P y Q (T y)) T x ~~>: Q.
......@@ -137,9 +172,60 @@ Section cmra_transport.
Proof. eauto using cmra_transport_updateP. Qed.
End cmra_transport.
(** * Isomorphism *)
Section iso_cmra.
Context {SI : sidx} {A B : cmra} (f : A B) (g : B A).
Lemma iso_cmra_updateP (P : B Prop) (Q : A Prop) y
(gf : x, g (f x) x)
(g_op : y1 y2, g (y1 y2) g y1 g y2)
(g_validN : n y, {n} (g y) {n} y) :
y ~~>: P
( y', P y' Q (g y'))
g y ~~>: Q.
Proof.
intros Hup Hx n mz Hmz.
destruct (Hup n (f <$> mz)) as (y'&HPy'&Hy'%g_validN).
{ apply g_validN. destruct mz as [z|]; simpl in *; [|done].
by rewrite g_op gf. }
exists (g y'); split; [by eauto|].
destruct mz as [z|]; simpl in *; [|done].
revert Hy'. by rewrite g_op gf.
Qed.
Lemma iso_cmra_updateP' (P : B Prop) y
(gf : x, g (f x) x)
(g_op : y1 y2, g (y1 y2) g y1 g y2)
(g_validN : n y, {n} (g y) {n} y) :
y ~~>: P
g y ~~>: λ x, y, x = g y P y.
Proof. eauto using iso_cmra_updateP. Qed.
End iso_cmra.
Section update_lift_cmra.
Context {SI : sidx} {A B : cmra}.
Implicit Types a : A.
Implicit Types b : B.
(** This lemma shows that if [f] maps non-deterministic updates from [B] to [A]
(i.e., [cmra_updateP] / [~~>:]), then [f] also maps deterministic updates from
[B] to [A] (i.e., [cmra_update] / [~~>]) *)
Lemma cmra_update_lift_updateP (f : B A) b b' :
( P, b ~~>: P f b ~~>: λ a', b', a' = f b' P b')
b ~~> b'
f b ~~> f b'.
Proof.
intros Hgen Hupd.
eapply cmra_update_updateP, cmra_updateP_weaken.
{ eapply Hgen, cmra_update_updateP, Hupd. }
naive_solver.
Qed.
End update_lift_cmra.
(** * Product *)
Section prod.
Context {A B : cmraT}.
Context {SI : sidx} {A B : cmra}.
Implicit Types x : A * B.
Lemma prod_updateP P1 P2 (Q : A * B Prop) x :
......@@ -162,7 +248,7 @@ End prod.
(** * Option *)
Section option.
Context {A : cmraT}.
Context {SI : sidx} {A : cmra}.
Implicit Types x y : A.
Lemma option_updateP (P : A Prop) (Q : option A Prop) x :
......