diff --git a/theories/algebra/lib/frac_agree.v b/theories/algebra/lib/frac_agree.v index b096e871ff6055db9289a1491f75ef826275ed0a..d2e805d776cede6ee63061770c508771e4d72848 100644 --- a/theories/algebra/lib/frac_agree.v +++ b/theories/algebra/lib/frac_agree.v @@ -25,6 +25,10 @@ Section lemmas. Global Instance to_frac_inj : Inj2 (≡) (≡) (≡) (@to_frac_agree A). 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 : ✓ (to_frac_agree q1 a1 â‹… to_frac_agree q2 a2) → ✓ (q1 + q2)%Qp ∧ a1 ≡ a2. diff --git a/theories/base_logic/lib/ghost_var.v b/theories/base_logic/lib/ghost_var.v index 20257482fae231b37683e5ceab0f2dc1374d9679..66af6d3f71a098df63530a1eb4e5f0ddf13e34d8 100644 --- a/theories/base_logic/lib/ghost_var.v +++ b/theories/base_logic/lib/ghost_var.v @@ -30,7 +30,7 @@ Section lemmas. Proof. apply _. Qed. 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 : AsFractional (ghost_var γ q a) (λ q, ghost_var γ q a) q. Proof. split. done. apply _. Qed. @@ -43,7 +43,7 @@ Section lemmas. ⊢ |==> ∃ γ, ghost_var γ 1 a. 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âŒ. Proof. iIntros "Hvar1 Hvar2". @@ -61,7 +61,20 @@ Section lemmas. ghost_var γ 1 a ==∗ ghost_var γ 1 b. Proof. 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.