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

Merge branch 'ralf/cmra-op-discrete' into 'master'

strengthen cmra_op_discrete

See merge request iris/iris!618
parents 6ed0574d c9d861f5
No related branches found
No related tags found
No related merge requests found
......@@ -64,6 +64,8 @@ HeapLang, which is now in a separate package `coq-iris-heap-lang`.
inequality instead of RA composition and validity (also in `base_logic` and
the higher layers).
* Move `algebra.base` module to `prelude.prelude`.
* Strengthen `cmra_op_discrete` to assume only `✓{0} (x1 ⋅ x2)` instead of `✓
(x1 ⋅ x2)`.
**Changes in `bi`:**
......
......@@ -549,11 +549,11 @@ Qed.
Lemma cmra_discrete_included_r x y : Discrete y x {0} y x y.
Proof. intros ? [x' ?]. exists x'. by apply (discrete y). Qed.
Lemma cmra_op_discrete x1 x2 :
(x1 x2) Discrete x1 Discrete x2 Discrete (x1 x2).
{0} (x1 x2) Discrete x1 Discrete x2 Discrete (x1 x2).
Proof.
intros ??? z Hz.
destruct (cmra_extend 0 z x1 x2) as (y1&y2&Hz'&?&?); auto; simpl in *.
{ rewrite -?Hz. by apply cmra_valid_validN. }
{ rewrite -?Hz. done. }
by rewrite Hz' (discrete x1 y1) // (discrete x2 y2).
Qed.
......
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