diff --git a/theories/algebra/gmap.v b/theories/algebra/gmap.v index 3a5cbe52a1b8384ff7293c7c8a28de9911b37939..51ae2084dc7c7148caac26bf9b7e991719c6253c 100644 --- a/theories/algebra/gmap.v +++ b/theories/algebra/gmap.v @@ -308,7 +308,7 @@ Global Instance gmap_singleton_core_id i (x : A) : CoreId x → CoreId {[ i := x ]}. Proof. intros. by apply core_id_total, core_singleton'. Qed. -Lemma singleton_includedN n m i x : +Lemma singleton_includedN_l n m i x : {[ i := x ]} ≼{n} m ↔ ∃ y, m !! i ≡{n}≡ Some y ∧ Some x ≼{n} Some y. Proof. split. @@ -321,7 +321,7 @@ Proof. + 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 : +Lemma singleton_included_l m i x : {[ i := x ]} ≼ m ↔ ∃ y, m !! i ≡ Some y ∧ Some x ≼ Some y. Proof. split. @@ -333,13 +333,21 @@ Proof. + 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 singleton_included_exclusive m i x : +Lemma singleton_included_exclusive_l m i x : Exclusive x → ✓ m → {[ i := x ]} ≼ m ↔ m !! i ≡ Some x. Proof. - intros ? Hm. rewrite singleton_included. split; last by eauto. + intros ? Hm. rewrite singleton_included_l. split; last by eauto. intros (y&?&->%(Some_included_exclusive _)); eauto using lookup_valid_Some. Qed. +Lemma singleton_included i x y : + {[ i := x ]} ≼ ({[ i := y ]} : gmap K A) ↔ x ≡ y ∨ x ≼ y. +Proof. + rewrite singleton_included_l. split. + - intros (y'&Hi&?). rewrite lookup_insert in Hi. + apply Some_included. by rewrite Hi. + - intros ?. exists y. by rewrite lookup_insert Some_included. +Qed. Global Instance singleton_cancelable i x : Cancelable (Some x) → Cancelable {[ i := x ]}. diff --git a/theories/base_logic/lib/gen_heap.v b/theories/base_logic/lib/gen_heap.v index 4a3065576a14f7566d73898207a35d592e19f688..1ee1d569d94ecb82507f40f9002b0b1b1bea40bd 100644 --- a/theories/base_logic/lib/gen_heap.v +++ b/theories/base_logic/lib/gen_heap.v @@ -146,7 +146,7 @@ Section to_gen_heap. Lemma gen_heap_singleton_included σ l q v : {[l := (q, to_agree v)]} ≼ to_gen_heap σ → σ !! l = Some v. Proof. - rewrite singleton_included=> -[[q' av] []]. + rewrite singleton_included_l=> -[[q' av] []]. rewrite /to_gen_heap lookup_fmap fmap_Some_equiv => -[v' [Hl [/= -> ->]]]. move=> /Some_pair_included_total_2 [_] /to_agree_included /leibniz_equiv_iff -> //. Qed. diff --git a/theories/base_logic/lib/proph_map.v b/theories/base_logic/lib/proph_map.v index 4b4c1c11710fcf98e1d87ed6fe454ada2016c4e1..6793cd7a74904761aeb3c9097029d59f1abd468f 100644 --- a/theories/base_logic/lib/proph_map.v +++ b/theories/base_logic/lib/proph_map.v @@ -102,7 +102,7 @@ Section to_proph_map. Lemma proph_map_singleton_included R p vs : {[p := Excl vs]} ≼ to_proph_map R → R !! p = Some vs. Proof. - rewrite singleton_included_exclusive; last by apply to_proph_map_valid. + rewrite singleton_included_exclusive_l; last by apply to_proph_map_valid. by rewrite leibniz_equiv_iff /to_proph_map lookup_fmap fmap_Some=> -[v' [-> [->]]]. Qed. End to_proph_map.