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

Basic properties of frame preserving updates and those for products.

parent 4882ecf8
No related branches found
No related tags found
No related merge requests found
...@@ -318,6 +318,15 @@ Proof. ...@@ -318,6 +318,15 @@ Proof.
* by intros Hx z ?; exists y; split; [done|apply (Hx z)]. * by intros Hx z ?; exists y; split; [done|apply (Hx z)].
* by intros Hx z n ?; destruct (Hx z n) as (?&<-&?). * by intros Hx z n ?; destruct (Hx z n) as (?&<-&?).
Qed. Qed.
Lemma ra_updateP_id (P : A Prop) x : P x x ⇝: P.
Proof. by intros ? z n ?; exists x. Qed.
Lemma ra_updateP_compose (P Q : A Prop) x :
x ⇝: P ( y, P y y ⇝: Q) x ⇝: Q.
Proof.
intros Hx Hy z n ?. destruct (Hx z n) as (y&?&?); auto. by apply (Hy y).
Qed.
Lemma ra_updateP_weaken (P Q : A Prop) x : x ⇝: P ( y, P y Q y) x ⇝: Q.
Proof. eauto using ra_updateP_compose, ra_updateP_id. Qed.
End cmra. End cmra.
Hint Extern 0 (_ {0} _) => apply cmra_includedN_0. Hint Extern 0 (_ {0} _) => apply cmra_includedN_0.
...@@ -384,14 +393,14 @@ Section discrete. ...@@ -384,14 +393,14 @@ Section discrete.
Qed. Qed.
Definition discreteRA : cmraT := Definition discreteRA : cmraT :=
CMRAT (cofe_mixin A) discrete_cmra_mixin discrete_extend_mixin. CMRAT (cofe_mixin A) discrete_cmra_mixin discrete_extend_mixin.
Lemma discrete_updateP (x : A) (P : A Prop) `{!Inhabited (sig P)} : Lemma discrete_updateP (x : discreteRA) (P : A Prop) `{!Inhabited (sig P)} :
( z, (x z) y, P y (y z)) (x : discreteRA) ⇝: P. ( z, (x z) y, P y (y z)) x ⇝: P.
Proof. Proof.
intros Hvalid z [|n]; [|apply Hvalid]. intros Hvalid z [|n]; [|apply Hvalid].
by destruct (_ : Inhabited (sig P)) as [[y ?]]; exists y. by destruct (_ : Inhabited (sig P)) as [[y ?]]; exists y.
Qed. Qed.
Lemma discrete_update (x y : A) : Lemma discrete_update (x y : discreteRA) :
( z, (x z) (y z)) (x : discreteRA) y. ( z, (x z) (y z)) x y.
Proof. intros Hvalid z [|n]; [done|apply Hvalid]. Qed. Proof. intros Hvalid z [|n]; [done|apply Hvalid]. Qed.
End discrete. End discrete.
...@@ -465,6 +474,15 @@ Section prod. ...@@ -465,6 +474,15 @@ Section prod.
* by split; rewrite /=left_id. * by split; rewrite /=left_id.
* by intros ? [??]; split; apply (timeless _). * by intros ? [??]; split; apply (timeless _).
Qed. Qed.
Lemma prod_update x y : x.1 y.1 x.2 y.2 x y.
Proof. intros ?? z n [??]; split; simpl in *; auto. Qed.
Lemma prod_updateP (P : A * B Prop) P1 P2 x :
x.1 ⇝: P1 x.2 ⇝: P2 ( y, P1 (y.1) P2 (y.2) P y) x ⇝: P.
Proof.
intros Hx1 Hx2 HP z n [??]; simpl in *.
destruct (Hx1 (z.1) n) as (a&?&?), (Hx2 (z.2) n) as (b&?&?); auto.
exists (a,b); repeat split; auto.
Qed.
End prod. End prod.
Arguments prodRA : clear implicits. Arguments prodRA : clear implicits.
......
0% Loading or .
You are about to add 0 people to the discussion. Proceed with caution.
Finish editing this message first!
Please register or to comment