From 45954c6f960790761ae4ab9afcbd72d8b6448aa6 Mon Sep 17 00:00:00 2001 From: Robbert Krebbers <mail@robbertkrebbers.nl> Date: Thu, 6 Oct 2016 10:58:27 +0200 Subject: [PATCH] =?UTF-8?q?More=20properties=20about=20Some=20=5F=20?= =?UTF-8?q?=E2=89=BC=20Some=20=5F.?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit These are very useful when dealing with the authoritative CMRA. --- algebra/cmra.v | 21 ++++++++++++++++++++- algebra/gmap.v | 38 +++++++++++++++----------------------- 2 files changed, 35 insertions(+), 24 deletions(-) diff --git a/algebra/cmra.v b/algebra/cmra.v index 79a64561a..85e0e2f7c 100644 --- a/algebra/cmra.v +++ b/algebra/cmra.v @@ -1208,8 +1208,14 @@ Section option. Lemma Some_included x y : Some x ≼ Some y ↔ x ≡ y ∨ x ≼ y. Proof. rewrite option_included; naive_solver. Qed. - Lemma Some_included' `{CMRATotal A} x y : Some x ≼ Some y ↔ x ≼ y. + Lemma Some_included_2 x y : x ≼ y → Some x ≼ Some y. + Proof. rewrite Some_included; eauto. Qed. + Lemma Some_included_total `{CMRATotal A} x y : Some x ≼ Some y ↔ x ≼ y. Proof. rewrite Some_included. split. by intros [->|?]. eauto. Qed. + Lemma Some_included_exclusive x `{!Exclusive x} y : + Some x ≼ Some y → ✓ y → x ≡ y. + Proof. move=> /Some_included [//|/exclusive_included]; tauto. Qed. + Lemma is_Some_included mx my : mx ≼ my → is_Some mx → is_Some my. Proof. rewrite -!not_eq_None_Some option_included. naive_solver. Qed. End option. @@ -1217,6 +1223,19 @@ End option. Arguments optionR : clear implicits. Arguments optionUR : clear implicits. +Section option_prod. + Context {A B : cmraT}. + Lemma Some_pair_included (x1 x2 : A) (y1 y2 : B) : + Some (x1,y1) ≼ Some (x2,y2) → Some x1 ≼ Some x2 ∧ Some y1 ≼ Some y2. + Proof. rewrite !Some_included. intros [[??]|[??]%prod_included]; eauto. Qed. + Lemma Some_pair_included_total_1 `{CMRATotal A} (x1 x2 : A) (y1 y2 : B) : + Some (x1,y1) ≼ Some (x2,y2) → x1 ≼ x2 ∧ Some y1 ≼ Some y2. + Proof. intros ?%Some_pair_included. by rewrite -(Some_included_total x1). Qed. + Lemma Some_pair_included_total_2 `{CMRATotal B} (x1 x2 : A) (y1 y2 : B) : + Some (x1,y1) ≼ Some (x2,y2) → Some x1 ≼ Some x2 ∧ y1 ≼ y2. + Proof. intros ?%Some_pair_included. by rewrite -(Some_included_total y1). Qed. +End option_prod. + Instance option_fmap_cmra_monotone {A B : cmraT} (f: A → B) `{!CMRAMonotone f} : CMRAMonotone (fmap f : option A → option B). Proof. diff --git a/algebra/gmap.v b/algebra/gmap.v index aa08442b1..c679e9781 100644 --- a/algebra/gmap.v +++ b/algebra/gmap.v @@ -256,37 +256,29 @@ Global Instance gmap_singleton_persistent i (x : A) : Proof. intros. by apply persistent_total, core_singleton'. Qed. Lemma singleton_includedN n m i x : - {[ i := x ]} ≼{n} m ↔ ∃ y, m !! i ≡{n}≡ Some y ∧ (x ≼{n} y ∨ x ≡{n}≡ y). + {[ i := x ]} ≼{n} m ↔ ∃ y, m !! i ≡{n}≡ Some y ∧ Some x ≼{n} Some y. Proof. split. - - move=> [m' /(_ i)]; rewrite lookup_op lookup_singleton. - case (m' !! i)=> [y|]=> Hm. - + exists (x ⋅ y); eauto using cmra_includedN_l. - + exists x; eauto. - - intros (y&Hi&[[z ?]| ->]). - + exists (<[i:=z]>m)=> j; destruct (decide (i = j)) as [->|]. - * rewrite Hi lookup_op lookup_singleton lookup_insert. by constructor. - * by rewrite lookup_op lookup_singleton_ne // lookup_insert_ne // left_id. - + exists (delete i m)=> j; destruct (decide (i = j)) as [->|]. - * by rewrite Hi lookup_op lookup_singleton lookup_delete. - * by rewrite lookup_op lookup_singleton_ne // lookup_delete_ne // left_id. + - move=> [m' /(_ i)]; rewrite lookup_op lookup_singleton=> Hi. + exists (x ⋅? m' !! i). rewrite -Some_op_opM. + split. done. apply cmra_includedN_l. + - intros (y&Hi&[mz Hy]). exists (partial_alter (λ _, mz) i m). + intros j; destruct (decide (i = j)) as [->|]. + + by rewrite lookup_op lookup_singleton lookup_partial_alter Hi. + + by rewrite lookup_op lookup_singleton_ne// lookup_partial_alter_ne// left_id. Qed. (* We do not have [x ≼ y ↔ ∀ n, x ≼{n} y], so we cannot use the previous lemma *) Lemma singleton_included m i x : - {[ i := x ]} ≼ m ↔ ∃ y, m !! i ≡ Some y ∧ (x ≼ y ∨ x ≡ y). + {[ i := x ]} ≼ m ↔ ∃ y, m !! i ≡ Some y ∧ Some x ≼ Some y. Proof. split. - move=> [m' /(_ i)]; rewrite lookup_op lookup_singleton. - case (m' !! i)=> [y|]=> Hm. - + exists (x ⋅ y); eauto using cmra_included_l. - + exists x; eauto. - - intros (y&Hi&[[z ?]| ->]). - + exists (<[i:=z]>m)=> j; destruct (decide (i = j)) as [->|]. - * rewrite Hi lookup_op lookup_singleton lookup_insert. by constructor. - * by rewrite lookup_op lookup_singleton_ne // lookup_insert_ne // left_id. - + exists (delete i m)=> j; destruct (decide (i = j)) as [->|]. - * by rewrite Hi lookup_op lookup_singleton lookup_delete. - * by rewrite lookup_op lookup_singleton_ne // lookup_delete_ne // left_id. + exists (x ⋅? m' !! i). rewrite -Some_op_opM. + split. done. apply cmra_included_l. + - intros (y&Hi&[mz Hy]). exists (partial_alter (λ _, mz) i m). + intros j; destruct (decide (i = j)) as [->|]. + + by rewrite lookup_op lookup_singleton lookup_partial_alter Hi. + + by rewrite lookup_op lookup_singleton_ne// lookup_partial_alter_ne// left_id. Qed. Lemma insert_updateP (P : A → Prop) (Q : gmap K A → Prop) m i x : -- GitLab