Skip to content
Snippets Groups Projects
Commit f66d2618 authored by Ralf Jung's avatar Ralf Jung
Browse files

ghost-var lib: more lemmas for when you own the two halves

parent 892dfbb6
No related branches found
No related tags found
No related merge requests found
...@@ -25,6 +25,10 @@ Section lemmas. ...@@ -25,6 +25,10 @@ Section lemmas.
Global Instance to_frac_inj : Inj2 () () () (@to_frac_agree A). Global Instance to_frac_inj : Inj2 () () () (@to_frac_agree A).
Proof. by intros q1 a1 q2 a2 [??%(inj to_agree)]. Qed. Proof. by intros q1 a1 q2 a2 [??%(inj to_agree)]. Qed.
Lemma frac_agree_op q1 q2 a :
to_frac_agree (q1 + q2) a to_frac_agree q1 a to_frac_agree q2 a.
Proof. rewrite /to_frac_agree -pair_op agree_idemp //. Qed.
Lemma frac_agree_op_valid q1 a1 q2 a2 : Lemma frac_agree_op_valid q1 a1 q2 a2 :
(to_frac_agree q1 a1 to_frac_agree q2 a2) (to_frac_agree q1 a1 to_frac_agree q2 a2)
(q1 + q2)%Qp a1 a2. (q1 + q2)%Qp a1 a2.
......
...@@ -30,7 +30,7 @@ Section lemmas. ...@@ -30,7 +30,7 @@ Section lemmas.
Proof. apply _. Qed. Proof. apply _. Qed.
Global Instance ghost_var_fractional γ a : Fractional (λ q, ghost_var γ q a). Global Instance ghost_var_fractional γ a : Fractional (λ q, ghost_var γ q a).
Proof. intros q1 q2. rewrite /ghost_var -own_op -pair_op agree_idemp //. Qed. Proof. intros q1 q2. rewrite /ghost_var -own_op -frac_agree_op //. Qed.
Global Instance ghost_var_as_fractional γ a q : Global Instance ghost_var_as_fractional γ a q :
AsFractional (ghost_var γ q a) (λ q, ghost_var γ q a) q. AsFractional (ghost_var γ q a) (λ q, ghost_var γ q a) q.
Proof. split. done. apply _. Qed. Proof. split. done. apply _. Qed.
...@@ -43,7 +43,7 @@ Section lemmas. ...@@ -43,7 +43,7 @@ Section lemmas.
|==> γ, ghost_var γ 1 a. |==> γ, ghost_var γ 1 a.
Proof. iApply own_alloc. done. Qed. Proof. iApply own_alloc. done. Qed.
Lemma ghost_var_op_valid γ a1 q1 a2 q2 : Lemma ghost_var_valid_2 γ a1 q1 a2 q2 :
ghost_var γ q1 a1 -∗ ghost_var γ q2 a2 -∗ ⌜✓ (q1 + q2)%Qp a1 = a2⌝. ghost_var γ q1 a1 -∗ ghost_var γ q2 a2 -∗ ⌜✓ (q1 + q2)%Qp a1 = a2⌝.
Proof. Proof.
iIntros "Hvar1 Hvar2". iIntros "Hvar1 Hvar2".
...@@ -61,7 +61,20 @@ Section lemmas. ...@@ -61,7 +61,20 @@ Section lemmas.
ghost_var γ 1 a ==∗ ghost_var γ 1 b. ghost_var γ 1 a ==∗ ghost_var γ 1 b.
Proof. Proof.
iApply own_update. apply cmra_update_exclusive. done. iApply own_update. apply cmra_update_exclusive. done.
Qed. Qed.
Lemma ghost_var_update_2 b γ a1 q1 a2 q2 :
(q1 + q2 = 1)%Qp
ghost_var γ q1 a1 -∗ ghost_var γ q2 a2 ==∗ ghost_var γ q1 b ghost_var γ q2 b.
Proof.
iIntros (Hq) "[H1 H2]".
iDestruct (ghost_var_valid_2 with "H1 H2") as %[_ ->].
iCombine "H1 H2" as "H". rewrite -frac_agree_op Hq.
iMod (ghost_var_update with "H") as "H".
rewrite -Hq. iApply ghost_var_split. done.
Qed.
Lemma ghost_var_update_halves b γ a1 a2 :
ghost_var γ (1/2) a1 -∗ ghost_var γ (1/2) a2 ==∗ ghost_var γ (1/2) b ghost_var γ (1/2) b.
Proof. iApply ghost_var_update_2. apply Qp_half_half. Qed.
End lemmas. End lemmas.
......
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