From 111386885885145a9552c5eabcaacd8b606f44a4 Mon Sep 17 00:00:00 2001
From: Dmitry Khalanskiy <Dmitry.Khalanskiy@jetbrains.com>
Date: Tue, 21 Jan 2020 16:43:51 +0300
Subject: [PATCH] Add `list_singletonM_included`

A lemma that allows to relate a singleton with another list.
---
 theories/algebra/list.v | 14 ++++++++++++++
 1 file changed, 14 insertions(+)

diff --git a/theories/algebra/list.v b/theories/algebra/list.v
index b1685f2eb..437f24baf 100644
--- a/theories/algebra/list.v
+++ b/theories/algebra/list.v
@@ -413,6 +413,20 @@ Section properties.
   Lemma list_singleton_snoc l x:
     {[length l := x]} ⋅ l ≡ l ++ [x].
   Proof. elim: l => //= ?? <-. by rewrite left_id. Qed.
+  Lemma list_singletonM_included i x l:
+    {[i := x]} ≼ l ↔ (∃ x', l !! i = Some x' ∧ x ≼ x').
+  Proof.
+    rewrite list_lookup_included. split.
+    { move /(_ i). rewrite list_lookup_singletonM option_included_total.
+      naive_solver. }
+    intros (y&Hi&?) j. destruct (Nat.lt_total j i) as [?|[->|?]].
+    - rewrite list_lookup_singletonM_lt //.
+      destruct (lookup_lt_is_Some_2 l j) as [z Hz].
+      { trans i; eauto using lookup_lt_Some. }
+      rewrite Hz. by apply Some_included_2, ucmra_unit_least.
+    - rewrite list_lookup_singletonM Hi. by apply Some_included_2.
+    - rewrite list_lookup_singletonM_gt //. apply: ucmra_unit_least.
+  Qed.
 
   (* Update *)
   Lemma list_singleton_updateP (P : A → Prop) (Q : list A → Prop) x :
-- 
GitLab