Newer
Older
Robbert Krebbers
committed
(** This file collects general purpose definitions and theorems on lists that
are not in the Coq standard library. *)
From Coq Require Export Permutation.
From stdpp Require Export numbers base option.
(** FIXME: Workaround for https://github.com/coq/coq/issues/14571 *)
(** Remove the instances [Permutation_cons] and [Permutation_app'] since their
priorities are 10, which is above the priority 5 of [proper_relation], and add
them back with the right priority (default = 0, since these instances have no
premises). *)
Global Remove Hints Permutation_cons Permutation_app' : typeclass_instances.
Global Existing Instances Permutation_cons Permutation_app'.
Global Arguments length {_} _ : assert.
Global Arguments cons {_} _ _ : assert.
Global Arguments app {_} _ _ : assert.
Global Instance: Params (@length) 1 := {}.
Global Instance: Params (@cons) 1 := {}.
Global Instance: Params (@app) 1 := {}.
(** [head] and [tail] are defined as [parsing only] for [hd_error] and [tl] in
the Coq standard library. We redefine these notations to make sure they also
pretty print properly. *)
Robbert Krebbers
committed
Notation head := hd_error.
Notation tail := tl.
Notation take := firstn.
Notation drop := skipn.
Robbert Krebbers
committed
Global Arguments head {_} _ : assert.
Global Arguments tail {_} _ : assert.
Global Arguments take {_} !_ !_ / : assert.
Global Arguments drop {_} !_ !_ / : assert.
Global Instance: Params (@head) 1 := {}.
Global Instance: Params (@tail) 1 := {}.
Global Instance: Params (@take) 1 := {}.
Global Instance: Params (@drop) 1 := {}.
Global Arguments Permutation {_} _ _ : assert.
Global Arguments Forall_cons {_} _ _ _ _ _ : assert.
Notation "(::)" := cons (only parsing) : list_scope.
Notation "( x ::.)" := (cons x) (only parsing) : list_scope.
Notation "(.:: l )" := (λ x, cons x l) (only parsing) : list_scope.
Notation "(++)" := app (only parsing) : list_scope.
Notation "( l ++.)" := (app l) (only parsing) : list_scope.
Notation "(.++ k )" := (λ l, app l k) (only parsing) : list_scope.
Infix "≡ₚ" := Permutation (at level 70, no associativity) : stdpp_scope.
Notation "(≡ₚ)" := Permutation (only parsing) : stdpp_scope.
Notation "( x ≡ₚ.)" := (Permutation x) (only parsing) : stdpp_scope.
Notation "(.≡ₚ x )" := (λ y, y ≡ₚ x) (only parsing) : stdpp_scope.
Notation "(≢ₚ)" := (λ x y, ¬x ≡ₚ y) (only parsing) : stdpp_scope.
Notation "x ≢ₚ y":= (¬x ≡ₚ y) (at level 70, no associativity) : stdpp_scope.
Notation "( x ≢ₚ.)" := (λ y, x ≢ₚ y) (only parsing) : stdpp_scope.
Notation "(.≢ₚ x )" := (λ y, y ≢ₚ x) (only parsing) : stdpp_scope.
Robbert Krebbers
committed
Infix "≡ₚ@{ A }" :=
(@Permutation A) (at level 70, no associativity, only parsing) : stdpp_scope.
Notation "(≡ₚ@{ A } )" := (@Permutation A) (only parsing) : stdpp_scope.
Global Instance maybe_cons {A} : Maybe2 (@cons A) := λ l,
match l with x :: l => Some (x,l) | _ => None end.
(** Setoid equality lifted to lists *)
Inductive list_equiv `{Equiv A} : Equiv (list A) :=
| nil_equiv : [] ≡ []
| cons_equiv x y l k : x ≡ y → l ≡ k → x :: l ≡ y :: k.
Global Existing Instance list_equiv.
(** The operation [l !! i] gives the [i]th element of the list [l], or [None]
in case [i] is out of bounds. *)
Global Instance list_lookup {A} : Lookup nat A (list A) :=
fix go i l {struct l} : option A := let _ : Lookup _ _ _ := @go in
| [] => None | x :: l => match i with 0 => Some x | S i => l !! i end
(** The operation [l !!! i] is a total version of the lookup operation
[l !! i]. *)
Global Instance list_lookup_total `{!Inhabited A} : LookupTotal nat A (list A) :=
fix go i l {struct l} : A := let _ : LookupTotal _ _ _ := @go in
match l with
| [] => inhabitant
| x :: l => match i with 0 => x | S i => l !!! i end
end.
(** The operation [alter f i l] applies the function [f] to the [i]th element
of [l]. In case [i] is out of bounds, the list is returned unchanged. *)
Global Instance list_alter {A} : Alter nat A (list A) := λ f,
fix go i l {struct l} :=
| x :: l => match i with 0 => f x :: l | S i => x :: go i l end
(** The operation [<[i:=x]> l] overwrites the element at position [i] with the
value [x]. In case [i] is out of bounds, the list is returned unchanged. *)
Global Instance list_insert {A} : Insert nat A (list A) :=
fix go i y l {struct l} := let _ : Insert _ _ _ := @go in
match l with
| [] => []
| x :: l => match i with 0 => y :: l | S i => x :: <[i:=y]>l end
end.
Fixpoint list_inserts {A} (i : nat) (k l : list A) : list A :=
match k with
| [] => l
| y :: k => <[i:=y]>(list_inserts (S i) k l)
end.
Global Instance: Params (@list_inserts) 1 := {}.
(** The operation [delete i l] removes the [i]th element of [l] and moves
all consecutive elements one position ahead. In case [i] is out of bounds,
the list is returned unchanged. *)
Global Instance list_delete {A} : Delete nat (list A) :=
fix go (i : nat) (l : list A) {struct l} : list A :=
match l with
| [] => []
| x :: l => match i with 0 => l | S i => x :: @delete _ _ go i l end
Robbert Krebbers
committed
end.
(** The function [option_list o] converts an element [Some x] into the
singleton list [[x]], and [None] into the empty list [[]]. *)
Definition option_list {A} : option A → list A := option_rect _ (λ x, [x]) [].
Global Instance: Params (@option_list) 1 := {}.
Global Instance maybe_list_singleton {A} : Maybe (λ x : A, [x]) := λ l,
match l with [x] => Some x | _ => None end.
(** The function [filter P l] returns the list of elements of [l] that
satisfies [P]. The order remains unchanged. *)
Global Instance list_filter {A} : Filter A (list A) :=
fix go P _ l := let _ : Filter _ _ := @go in
| x :: l => if decide (P x) then x :: filter P l else filter P l
end.
(** The function [list_find P l] returns the first index [i] whose element
satisfies the predicate [P]. *)
Definition list_find {A} P `{∀ x, Decision (P x)} : list A → option (nat * A) :=
| [] => None
| x :: l => if decide (P x) then Some (0,x) else prod_map S id <$> go l
Global Instance: Params (@list_find) 3 := {}.
(** The function [replicate n x] generates a list with length [n] of elements
with value [x]. *)
Fixpoint replicate {A} (n : nat) (x : A) : list A :=
Robbert Krebbers
committed
match n with 0 => [] | S n => x :: replicate n x end.
Global Instance: Params (@replicate) 2 := {}.
(** The function [rotate n l] rotates the list [l] by [n], e.g., [rotate 1
[x0; x1; ...; xm]] becomes [x1; ...; xm; x0]. Rotating by a multiple of
[length l] is the identity function. **)
Definition rotate {A} (n : nat) (l : list A) : list A :=
drop (n `mod` length l) l ++ take (n `mod` length l) l.
Global Instance: Params (@rotate) 2 := {}.
(** The function [rotate_take s e l] returns the range between the
indices [s] (inclusive) and [e] (exclusive) of [l]. If [e ≤ s], all
elements after [s] and before [e] are returned. *)
Definition rotate_take {A} (s e : nat) (l : list A) : list A :=
take (rotate_nat_sub s e (length l)) (rotate s l).
Global Instance: Params (@rotate_take) 3 := {}.
(** The function [reverse l] returns the elements of [l] in reverse order. *)
Definition reverse {A} (l : list A) : list A := rev_append l [].
Global Instance: Params (@reverse) 1 := {}.
(** The function [last l] returns the last element of the list [l], or [None]
if the list [l] is empty. *)
Fixpoint last {A} (l : list A) : option A :=
match l with [] => None | [x] => Some x | _ :: l => last l end.
Global Instance: Params (@last) 1 := {}.
Global Arguments last : simpl nomatch.
(** The function [resize n y l] takes the first [n] elements of [l] in case
[length l ≤ n], and otherwise appends elements with value [x] to [l] to obtain
a list of length [n]. *)
Fixpoint resize {A} (n : nat) (y : A) (l : list A) : list A :=
match l with
| [] => replicate n y
Robbert Krebbers
committed
| x :: l => match n with 0 => [] | S n => x :: resize n y l end
Global Arguments resize {_} !_ _ !_ : assert.
Global Instance: Params (@resize) 2 := {}.
(** The function [reshape k l] transforms [l] into a list of lists whose sizes
are specified by [k]. In case [l] is too short, the resulting list will be
padded with empty lists. In case [l] is too long, it will be truncated. *)
Fixpoint reshape {A} (szs : list nat) (l : list A) : list (list A) :=
match szs with
| [] => [] | sz :: szs => take sz l :: reshape szs (drop sz l)
Global Instance: Params (@reshape) 2 := {}.
Definition sublist_lookup {A} (i n : nat) (l : list A) : option (list A) :=
guard (i + n ≤ length l); Some (take n (drop i l)).
Definition sublist_alter {A} (f : list A → list A)
(i n : nat) (l : list A) : list A :=
take i l ++ f (take n (drop i l)) ++ drop (i + n) l.
(** Functions to fold over a list. We redefine [foldl] with the arguments in
the same order as in Haskell. *)
Notation foldr := fold_right.
Definition foldl {A B} (f : A → B → A) : A → list B → A :=
fix go a l := match l with [] => a | x :: l => go (f a x) l end.
(** The monadic operations. *)
Global Instance list_ret: MRet list := λ A x, x :: @nil A.
Global Instance list_fmap : FMap list := λ A B f,
fix go (l : list A) := match l with [] => [] | x :: l => f x :: go l end.
Global Instance list_omap : OMap list := λ A B f,
fix go (l : list A) :=
match l with
| [] => []
| x :: l => match f x with Some y => y :: go l | None => go l end
end.
Global Instance list_bind : MBind list := λ A B f,
fix go (l : list A) := match l with [] => [] | x :: l => f x ++ go l end.
Global Instance list_join: MJoin list :=
fix go A (ls : list (list A)) : list A :=
Robbert Krebbers
committed
match ls with [] => [] | l :: ls => l ++ @mjoin _ go _ ls end.
Definition mapM `{MBind M, MRet M} {A B} (f : A → M B) : list A → M (list B) :=
match l with [] => mret [] | x :: l => y ← f x; k ← go l; mret (y :: k) end.
(** We define stronger variants of the map function that allow the mapped
function to use the index of the elements. *)
Fixpoint imap {A B} (f : nat → A → B) (l : list A) : list B :=
match l with
| [] => []
| x :: l => f 0 x :: imap (f ∘ S) l
end.
Robbert Krebbers
committed
Definition zipped_map {A B} (f : list A → list A → A → B) :
list A → list A → list B := fix go l k :=
match k with
| [] => []
| x :: k => f l k x :: go (x :: l) k
end.
Robbert Krebbers
committed
Fixpoint imap2 {A B C} (f : nat → A → B → C) (l : list A) (k : list B) : list C :=
| [], _ | _, [] => []
| x :: l, y :: k => f 0 x y :: imap2 (f ∘ S) l k
Robbert Krebbers
committed
Inductive zipped_Forall {A} (P : list A → list A → A → Prop) :
list A → list A → Prop :=
| zipped_Forall_nil l : zipped_Forall P l []
| zipped_Forall_cons l k x :
P l k x → zipped_Forall P (x :: l) k → zipped_Forall P l (x :: k).
Global Arguments zipped_Forall_nil {_ _} _ : assert.
Global Arguments zipped_Forall_cons {_ _} _ _ _ _ _ : assert.
(** The function [mask f βs l] applies the function [f] to elements in [l] at
positions that are [true] in [βs]. *)
Fixpoint mask {A} (f : A → A) (βs : list bool) (l : list A) : list A :=
match βs, l with
| β :: βs, x :: l => (if β then f x else x) :: mask f βs l
| _, _ => l
end.
(** The function [permutations l] yields all permutations of [l]. *)
Fixpoint interleave {A} (x : A) (l : list A) : list (list A) :=
match l with
| [] => [[x]]| y :: l => (x :: y :: l) :: ((y ::.) <$> interleave x l)
end.
Fixpoint permutations {A} (l : list A) : list (list A) :=
match l with [] => [[]] | x :: l => permutations l ≫= interleave x end.
(** The predicate [suffix] holds if the first list is a suffix of the second.
The predicate [prefix] holds if the first list is a prefix of the second. *)
Definition suffix {A} : relation (list A) := λ l1 l2, ∃ k, l2 = k ++ l1.
Definition prefix {A} : relation (list A) := λ l1 l2, ∃ k, l2 = l1 ++ k.
Infix "`suffix_of`" := suffix (at level 70) : stdpp_scope.
Infix "`prefix_of`" := prefix (at level 70) : stdpp_scope.
Global Hint Extern 0 (_ `prefix_of` _) => reflexivity : core.
Global Hint Extern 0 (_ `suffix_of` _) => reflexivity : core.
Section prefix_suffix_ops.
Context `{EqDecision A}.
Definition max_prefix : list A → list A → list A * list A * list A :=
fix go l1 l2 :=
match l1, l2 with
| [], l2 => ([], l2, [])
| l1, [] => (l1, [], [])
| x1 :: l1, x2 :: l2 =>
Robbert Krebbers
committed
if decide_rel (=) x1 x2
then prod_map id (x1 ::.) (go l1 l2) else (x1 :: l1, x2 :: l2, [])
Definition max_suffix (l1 l2 : list A) : list A * list A * list A :=
match max_prefix (reverse l1) (reverse l2) with
| (k1, k2, k3) => (reverse k1, reverse k2, reverse k3)
end.
Definition strip_prefix (l1 l2 : list A) := (max_prefix l1 l2).1.2.
Definition strip_suffix (l1 l2 : list A) := (max_suffix l1 l2).1.2.
End prefix_suffix_ops.
(** A list [l1] is a sublist of [l2] if [l2] is obtained by removing elements
from [l1] without changing the order. *)
Inductive sublist {A} : relation (list A) :=
| sublist_nil : sublist [] []
Robbert Krebbers
committed
| sublist_skip x l1 l2 : sublist l1 l2 → sublist (x :: l1) (x :: l2)
| sublist_cons x l1 l2 : sublist l1 l2 → sublist l1 (x :: l2).
Infix "`sublist_of`" := sublist (at level 70) : stdpp_scope.
Global Hint Extern 0 (_ `sublist_of` _) => reflexivity : core.
Robbert Krebbers
committed
(** A list [l2] submseteq a list [l1] if [l2] is obtained by removing elements
from [l1] while possiblity changing the order. *)
Inductive submseteq {A} : relation (list A) :=
| submseteq_nil : submseteq [] []
| submseteq_skip x l1 l2 : submseteq l1 l2 → submseteq (x :: l1) (x :: l2)
| submseteq_swap x y l : submseteq (y :: x :: l) (x :: y :: l)
| submseteq_cons x l1 l2 : submseteq l1 l2 → submseteq l1 (x :: l2)
| submseteq_trans l1 l2 l3 : submseteq l1 l2 → submseteq l2 l3 → submseteq l1 l3.
Infix "⊆+" := submseteq (at level 70) : stdpp_scope.
Global Hint Extern 0 (_ ⊆+ _) => reflexivity : core.
Robbert Krebbers
committed
(** Removes [x] from the list [l]. The function returns a [Some] when the
removal succeeds and [None] when [x] is not in [l]. *)
Fixpoint list_remove `{EqDecision A} (x : A) (l : list A) : option (list A) :=
match l with
| [] => None
| y :: l => if decide (x = y) then Some l else (y ::.) <$> list_remove x l
end.
(** Removes all elements in the list [k] from the list [l]. The function returns
a [Some] when the removal succeeds and [None] some element of [k] is not in [l]. *)
Fixpoint list_remove_list `{EqDecision A} (k : list A) (l : list A) : option (list A) :=
match k with
| [] => Some l | x :: k => list_remove x l ≫= list_remove_list k
end.
Inductive Forall3 {A B C} (P : A → B → C → Prop) :
list A → list B → list C → Prop :=
| Forall3_nil : Forall3 P [] [] []
| Forall3_cons x y z l k k' :
P x y z → Forall3 P l k k' → Forall3 P (x :: l) (y :: k) (z :: k').
Robbert Krebbers
committed
(** Set operations on lists *)
Global Instance list_subseteq {A} : SubsetEq (list A) := λ l1 l2, ∀ x, x ∈ l1 → x ∈ l2.
Robbert Krebbers
committed
Section list_set.
Context `{dec : EqDecision A}.
Global Instance elem_of_list_dec : RelDecision (∈@{list A}).
Robbert Krebbers
committed
refine (
Robbert Krebbers
committed
fix go x l :=
Robbert Krebbers
committed
match l return Decision (x ∈ l) with
| [] => right _
Robbert Krebbers
committed
| y :: l => cast_if_or (decide (x = y)) (go x l)
Robbert Krebbers
committed
end); clear go dec; subst; try (by constructor); abstract by inversion 1.
Defined.
Fixpoint remove_dups (l : list A) : list A :=
match l with
| [] => []
| x :: l =>
if decide_rel (∈) x l then remove_dups l else x :: remove_dups l
end.
Fixpoint list_difference (l k : list A) : list A :=
match l with
| [] => []
| x :: l =>
if decide_rel (∈) x k
then list_difference l k else x :: list_difference l k
Robbert Krebbers
committed
end.
Definition list_union (l k : list A) : list A := list_difference l k ++ k.
Robbert Krebbers
committed
Fixpoint list_intersection (l k : list A) : list A :=
match l with
| [] => []
| x :: l =>
if decide_rel (∈) x k
then x :: list_intersection l k else list_intersection l k
Robbert Krebbers
committed
end.
Definition list_intersection_with (f : A → A → option A) :
list A → list A → list A := fix go l k :=
match l with
| [] => []
| x :: l => foldr (λ y,
match f x y with None => id | Some z => (z ::.) end) (go l k) k
Robbert Krebbers
committed
end.
End list_set.
Robbert Krebbers
committed
(** These next functions allow to efficiently encode lists of positives (bit
strings) into a single positive and go in the other direction as well. This is
for example used for the countable instance of lists and in namespaces.
The main functions are [positives_flatten] and [positives_unflatten]. *)
Fixpoint positives_flatten_go (xs : list positive) (acc : positive) : positive :=
match xs with
| [] => acc
| x :: xs => positives_flatten_go xs (acc~1~0 ++ Preverse (Pdup x))
end.
Robbert Krebbers
committed
(** Flatten a list of positives into a single positive by duplicating the bits
of each element, so that:
- [0 -> 00]
- [1 -> 11]
and then separating each element with [10]. *)
Definition positives_flatten (xs : list positive) : positive :=
positives_flatten_go xs 1.
Fixpoint positives_unflatten_go
(p : positive)
(acc_xs : list positive)
(acc_elm : positive)
: option (list positive) :=
match p with
| 1 => Some acc_xs
| p'~0~0 => positives_unflatten_go p' acc_xs (acc_elm~0)
| p'~1~1 => positives_unflatten_go p' acc_xs (acc_elm~1)
| p'~1~0 => positives_unflatten_go p' (acc_elm :: acc_xs) 1
| _ => None
end%positive.
(** Unflatten a positive into a list of positives, assuming the encoding
Robbert Krebbers
committed
used by [positives_flatten]. *)
Definition positives_unflatten (p : positive) : option (list positive) :=
positives_unflatten_go p [] 1.
(** * Basic tactics on lists *)
(** The tactic [discriminate_list] discharges a goal if it submseteq
a list equality involving [(::)] and [(++)] of two lists that have a different
length as one of its hypotheses. *)
Tactic Notation "discriminate_list" hyp(H) :=
apply (f_equal length) in H;
repeat (csimpl in H || rewrite app_length in H); exfalso; lia.
Tactic Notation "discriminate_list" :=
Robbert Krebbers
committed
match goal with H : _ =@{list _} _ |- _ => discriminate_list H end.
(** The tactic [simplify_list_eq] simplifies hypotheses involving
equalities on lists using injectivity of [(::)] and [(++)]. Also, it simplifies
lookups in singleton lists. *)
Lemma app_inj_1 {A} (l1 k1 l2 k2 : list A) :
length l1 = length k1 → l1 ++ l2 = k1 ++ k2 → l1 = k1 ∧ l2 = k2.
Proof. revert k1. induction l1; intros [|??]; naive_solver. Qed.
Lemma app_inj_2 {A} (l1 k1 l2 k2 : list A) :
length l2 = length k2 → l1 ++ l2 = k1 ++ k2 → l1 = k1 ∧ l2 = k2.
Proof.
intros ? Hl. apply app_inj_1; auto.
apply (f_equal length) in Hl. rewrite !app_length in Hl. lia.
Qed.
Robbert Krebbers
committed
repeat match goal with
| H : _ ++ _ = _ ++ _ |- _ => first
[ apply app_inv_head in H | apply app_inv_tail in H
| apply app_inj_1 in H; [destruct H|done]
| apply app_inj_2 in H; [destruct H|done] ]
Robbert Krebbers
committed
destruct i; [change (Some x = Some y) in H | discriminate]
Robbert Krebbers
committed
(** * General theorems *)
Section general_properties.
Robbert Krebbers
committed
Implicit Types x y z : A.
Implicit Types l k : list A.
Global Instance: Inj2 (=) (=) (=) (@cons A).
Robbert Krebbers
committed
Proof. by injection 1. Qed.
Global Instance: ∀ k, Inj (=) (=) (k ++.).
Proof. intros ???. apply app_inv_head. Qed.
Global Instance: ∀ k, Inj (=) (=) (.++ k).
Proof. intros ???. apply app_inv_tail. Qed.
Global Instance: Assoc (=) (@app A).
Proof. intros ???. apply app_assoc. Qed.
Global Instance: LeftId (=) [] (@app A).
Proof. done. Qed.
Global Instance: RightId (=) [] (@app A).
Proof. intro. apply app_nil_r. Qed.
Robbert Krebbers
committed
Lemma app_nil l1 l2 : l1 ++ l2 = [] ↔ l1 = [] ∧ l2 = [].
Proof. split; [apply app_eq_nil|]. by intros [-> ->]. Qed.
Robbert Krebbers
committed
Lemma app_singleton l1 l2 x :
l1 ++ l2 = [x] ↔ l1 = [] ∧ l2 = [x] ∨ l1 = [x] ∧ l2 = [].
Proof. split; [apply app_eq_unit|]. by intros [[-> ->]|[-> ->]]. Qed.
Robbert Krebbers
committed
Lemma cons_middle x l1 l2 : l1 ++ x :: l2 = l1 ++ [x] ++ l2.
Proof. done. Qed.
Lemma list_eq l1 l2 : (∀ i, l1 !! i = l2 !! i) → l1 = l2.
revert l2. induction l1 as [|x l1 IH]; intros [|y l2] H.
- done.
- discriminate (H 0).
- discriminate (H 0).
- f_equal; [by injection (H 0)|]. apply (IH _ $ λ i, H (S i)).
Global Instance list_eq_dec {dec : EqDecision A} : EqDecision (list A) :=
list_eq_dec dec.
Global Instance list_eq_nil_dec l : Decision (l = []).
Proof. by refine match l with [] => left _ | _ => right _ end. Defined.
Lemma list_singleton_reflect l :
option_reflect (λ x, l = [x]) (length l ≠ 1) (maybe (λ x, [x]) l).
Proof. by destruct l as [|? []]; constructor. Defined.
Definition nil_length : length (@nil A) = 0 := eq_refl.
Definition cons_length x l : length (x :: l) = S (length l) := eq_refl.
Robbert Krebbers
committed
Lemma nil_or_length_pos l : l = [] ∨ length l ≠ 0.
Proof. destruct l; simpl; auto with lia. Qed.
Lemma nil_length_inv l : length l = 0 → l = [].
Lemma lookup_cons_ne_0 l x i : i ≠ 0 → (x :: l) !! i = l !! pred i.
Proof. by destruct i. Qed.
Lemma lookup_total_cons_ne_0 `{!Inhabited A} l x i :
i ≠ 0 → (x :: l) !!! i = l !!! pred i.
Proof. by destruct i. Qed.
Lemma lookup_nil i : @nil A !! i = None.
Proof. by destruct i. Qed.
Lemma lookup_total_nil `{!Inhabited A} i : @nil A !!! i = inhabitant.
Proof. by destruct i. Qed.
Robbert Krebbers
committed
Lemma lookup_tail l i : tail l !! i = l !! S i.
Proof. by destruct l. Qed.
Lemma lookup_total_tail `{!Inhabited A} l i : tail l !!! i = l !!! S i.
Proof. by destruct l. Qed.
Lemma lookup_lt_Some l i x : l !! i = Some x → i < length l.
Proof. revert i. induction l; intros [|?] ?; naive_solver auto with arith. Qed.
Lemma lookup_lt_is_Some_1 l i : is_Some (l !! i) → i < length l.
Proof. intros [??]; eauto using lookup_lt_Some. Qed.
Lemma lookup_lt_is_Some_2 l i : i < length l → is_Some (l !! i).
Proof. revert i. induction l; intros [|?] ?; naive_solver auto with lia. Qed.
Lemma lookup_lt_is_Some l i : is_Some (l !! i) ↔ i < length l.
Proof. split; auto using lookup_lt_is_Some_1, lookup_lt_is_Some_2. Qed.
Lemma lookup_ge_None l i : l !! i = None ↔ length l ≤ i.
Proof. rewrite eq_None_not_Some, lookup_lt_is_Some. lia. Qed.
Lemma lookup_ge_None_1 l i : l !! i = None → length l ≤ i.
Proof. by rewrite lookup_ge_None. Qed.
Lemma lookup_ge_None_2 l i : length l ≤ i → l !! i = None.
Proof. by rewrite lookup_ge_None. Qed.
Lemma list_eq_same_length l1 l2 n :
length l2 = n → length l1 = n →
(∀ i x y, i < n → l1 !! i = Some x → l2 !! i = Some y → x = y) → l1 = l2.
intros <- Hlen Hl; apply list_eq; intros i. destruct (l2 !! i) as [x|] eqn:Hx.
- destruct (lookup_lt_is_Some_2 l1 i) as [y Hy].
{ rewrite Hlen; eauto using lookup_lt_Some. }
rewrite Hy; f_equal; apply (Hl i); eauto using lookup_lt_Some.
- by rewrite lookup_ge_None, Hlen, <-lookup_ge_None.
Lemma nth_lookup l i d : nth i l d = default d (l !! i).
Proof. revert i. induction l as [|x l IH]; intros [|i]; simpl; auto. Qed.
Lemma nth_lookup_Some l i d x : l !! i = Some x → nth i l d = x.
Proof. rewrite nth_lookup. by intros ->. Qed.
Lemma nth_lookup_or_length l i d : {l !! i = Some (nth i l d)} + {length l ≤ i}.
rewrite nth_lookup. destruct (l !! i) eqn:?; eauto using lookup_ge_None_1.
Lemma list_lookup_total_alt `{!Inhabited A} l i :
l !!! i = default inhabitant (l !! i).
Proof. revert i. induction l; intros []; naive_solver. Qed.
Lemma list_lookup_total_correct `{!Inhabited A} l i x :
l !! i = Some x → l !!! i = x.
Proof. rewrite list_lookup_total_alt. by intros ->. Qed.
Lemma list_lookup_lookup_total `{!Inhabited A} l i :
is_Some (l !! i) → l !! i = Some (l !!! i).
Proof. rewrite list_lookup_total_alt; by intros [x ->]. Qed.
Lemma list_lookup_lookup_total_lt `{!Inhabited A} l i :
i < length l → l !! i = Some (l !!! i).
Proof. intros ?. by apply list_lookup_lookup_total, lookup_lt_is_Some_2. Qed.
Lemma list_lookup_alt `{!Inhabited A} l i x :
l !! i = Some x ↔ i < length l ∧ l !!! i = x.
Proof.
naive_solver eauto using list_lookup_lookup_total_lt,
list_lookup_total_correct, lookup_lt_Some.
Qed.
Lemma lookup_app l1 l2 i :
(l1 ++ l2) !! i =
match l1 !! i with Some x => Some x | None => l2 !! (i - length l1) end.
Proof. revert i. induction l1 as [|x l1 IH]; intros [|i]; naive_solver. Qed.
Lemma lookup_app_l l1 l2 i : i < length l1 → (l1 ++ l2) !! i = l1 !! i.
Proof. rewrite lookup_app. by intros [? ->]%lookup_lt_is_Some. Qed.
Lemma lookup_total_app_l `{!Inhabited A} l1 l2 i :
i < length l1 → (l1 ++ l2) !!! i = l1 !!! i.
Proof. intros. by rewrite !list_lookup_total_alt, lookup_app_l. Qed.
Lemma lookup_app_l_Some l1 l2 i x : l1 !! i = Some x → (l1 ++ l2) !! i = Some x.
Proof. rewrite lookup_app. by intros ->. Qed.
Lemma lookup_app_r l1 l2 i :
length l1 ≤ i → (l1 ++ l2) !! i = l2 !! (i - length l1).
Proof. rewrite lookup_app. by intros ->%lookup_ge_None. Qed.
Lemma lookup_total_app_r `{!Inhabited A} l1 l2 i :
length l1 ≤ i → (l1 ++ l2) !!! i = l2 !!! (i - length l1).
Proof. intros. by rewrite !list_lookup_total_alt, lookup_app_r. Qed.
Lemma lookup_app_Some l1 l2 i x :
(l1 ++ l2) !! i = Some x ↔
l1 !! i = Some x ∨ length l1 ≤ i ∧ l2 !! (i - length l1) = Some x.
Proof.
rewrite lookup_app. destruct (l1 !! i) eqn:Hi.
- apply lookup_lt_Some in Hi. naive_solver lia.
- apply lookup_ge_None in Hi. naive_solver lia.
Lemma lookup_cons x l i :
(x :: l) !! i =
match i with 0 => Some x | S i => l !! i end.
Proof. reflexivity. Qed.
Lemma lookup_cons_Some x l i y :
(x :: l) !! i = Some y ↔
(i = 0 ∧ x = y) ∨ (1 ≤ i ∧ l !! (i - 1) = Some y).
Proof.
rewrite lookup_cons. destruct i as [|i].
- naive_solver lia.
- replace (S i - 1) with i by lia. naive_solver lia.
Qed.
Lemma list_lookup_singleton x i :
[x] !! i =
match i with 0 => Some x | S _ => None end.
Proof. reflexivity. Qed.
Lemma list_lookup_singleton_Some x i y :
[x] !! i = Some y ↔ i = 0 ∧ x = y.
Proof. rewrite lookup_cons_Some. naive_solver. Qed.
Lemma list_lookup_middle l1 l2 x n :
n = length l1 → (l1 ++ x :: l2) !! n = Some x.
Proof. intros ->. by induction l1. Qed.
Lemma list_lookup_total_middle `{!Inhabited A} l1 l2 x n :
n = length l1 → (l1 ++ x :: l2) !!! n = x.
Proof. intros. by rewrite !list_lookup_total_alt, list_lookup_middle. Qed.
Lemma list_insert_alter l i x : <[i:=x]>l = alter (λ _, x) i l.
Proof. by revert i; induction l; intros []; intros; f_equal/=. Qed.
Robbert Krebbers
committed
Lemma alter_length f l i : length (alter f i l) = length l.
Proof. revert i. by induction l; intros [|?]; f_equal/=. Qed.
Robbert Krebbers
committed
Lemma insert_length l i x : length (<[i:=x]>l) = length l.
Proof. revert i. by induction l; intros [|?]; f_equal/=. Qed.
Robbert Krebbers
committed
Lemma list_lookup_alter f l i : alter f i l !! i = f <$> l !! i.
Proof.
revert i.
induction l as [|?? IHl]; [done|].
intros [|i]; [done|]. apply (IHl i).
Qed.
Lemma list_lookup_total_alter `{!Inhabited A} f l i :
i < length l → alter f i l !!! i = f (l !!! i).
Proof.
intros [x Hx]%lookup_lt_is_Some_2.
by rewrite !list_lookup_total_alt, list_lookup_alter, Hx.
Qed.
Lemma list_lookup_alter_ne f l i j : i ≠ j → alter f i l !! j = l !! j.
Proof. revert i j. induction l; [done|]. intros [] []; naive_solver. Qed.
Lemma list_lookup_total_alter_ne `{!Inhabited A} f l i j :
i ≠ j → alter f i l !!! j = l !!! j.
Proof. intros. by rewrite !list_lookup_total_alt, list_lookup_alter_ne. Qed.
Robbert Krebbers
committed
Lemma list_lookup_insert l i x : i < length l → <[i:=x]>l !! i = Some x.
Proof. revert i. induction l; intros [|?] ?; f_equal/=; auto with lia. Qed.
Lemma list_lookup_total_insert `{!Inhabited A} l i x :
i < length l → <[i:=x]>l !!! i = x.
Proof. intros. by rewrite !list_lookup_total_alt, list_lookup_insert. Qed.
Lemma list_lookup_insert_ne l i j x : i ≠ j → <[i:=x]>l !! j = l !! j.
Proof. revert i j. induction l; [done|]. intros [] []; naive_solver. Qed.
Lemma list_lookup_total_insert_ne `{!Inhabited A} l i j x :
i ≠ j → <[i:=x]>l !!! j = l !!! j.
Proof. intros. by rewrite !list_lookup_total_alt, list_lookup_insert_ne. Qed.
Lemma list_lookup_insert_Some l i x j y :
<[i:=x]>l !! j = Some y ↔
i = j ∧ x = y ∧ j < length l ∨ i ≠ j ∧ l !! j = Some y.
Proof.
destruct (decide (i = j)) as [->|];
[split|rewrite list_lookup_insert_ne by done; tauto].
- intros Hy. assert (j < length l).
{ rewrite <-(insert_length l j x); eauto using lookup_lt_Some. }
rewrite list_lookup_insert in Hy by done; naive_solver.
- intros [(?&?&?)|[??]]; rewrite ?list_lookup_insert; naive_solver.
Qed.
Lemma list_insert_commute l i j x y :
i ≠ j → <[i:=x]>(<[j:=y]>l) = <[j:=y]>(<[i:=x]>l).
Proof. revert i j. by induction l; intros [|?] [|?] ?; f_equal/=; auto. Qed.
Lemma list_insert_id' l i x : (i < length l → l !! i = Some x) → <[i:=x]>l = l.
Proof. revert i. induction l; intros [|i] ?; f_equal/=; naive_solver lia. Qed.
Lemma list_insert_id l i x : l !! i = Some x → <[i:=x]>l = l.
Proof. intros ?. by apply list_insert_id'. Qed.
Lemma list_insert_ge l i x : length l ≤ i → <[i:=x]>l = l.
Proof. revert i. induction l; intros [|i] ?; f_equal/=; auto with lia. Qed.
Lemma list_insert_insert l i x y : <[i:=x]> (<[i:=y]> l) = <[i:=x]> l.
Proof. revert i. induction l; intros [|i]; f_equal/=; auto. Qed.
Robbert Krebbers
committed
Lemma list_lookup_other l i x :
length l ≠ 1 → l !! i = Some x → ∃ j y, j ≠ i ∧ l !! j = Some y.
intros. destruct i, l as [|x0 [|x1 l]]; simplify_eq/=.
- by exists 1, x1.
- by exists 0, x0.
Robbert Krebbers
committed
Lemma alter_app_l f l1 l2 i :
i < length l1 → alter f i (l1 ++ l2) = alter f i l1 ++ l2.
Proof. revert i. induction l1; intros [|?] ?; f_equal/=; auto with lia. Qed.
Robbert Krebbers
committed
Lemma alter_app_r f l1 l2 i :
alter f (length l1 + i) (l1 ++ l2) = l1 ++ alter f i l2.
Proof. revert i. induction l1; intros [|?]; f_equal/=; auto. Qed.
Robbert Krebbers
committed
Lemma alter_app_r_alt f l1 l2 i :
length l1 ≤ i → alter f i (l1 ++ l2) = l1 ++ alter f (i - length l1) l2.
Proof.
intros. assert (i = length l1 + (i - length l1)) as Hi by lia.
rewrite Hi at 1. by apply alter_app_r.
Qed.
Lemma list_alter_id f l i : (∀ x, f x = x) → alter f i l = l.
Proof. intros ?. revert i. induction l; intros [|?]; f_equal/=; auto. Qed.
Lemma list_alter_ext f g l k i :
(∀ x, l !! i = Some x → f x = g x) → l = k → alter f i l = alter g i k.
Proof. intros H ->. revert i H. induction k; intros [|?] ?; f_equal/=; auto. Qed.
Lemma list_alter_compose f g l i :
alter (f ∘ g) i l = alter f i (alter g i l).
Proof. revert i. induction l; intros [|?]; f_equal/=; auto. Qed.
Lemma list_alter_commute f g l i j :
i ≠ j → alter f i (alter g j l) = alter g j (alter f i l).
Proof. revert i j. induction l; intros [|?][|?] ?; f_equal/=; auto with lia. Qed.
Robbert Krebbers
committed
Lemma insert_app_l l1 l2 i x :
i < length l1 → <[i:=x]>(l1 ++ l2) = <[i:=x]>l1 ++ l2.
Proof. revert i. induction l1; intros [|?] ?; f_equal/=; auto with lia. Qed.
Robbert Krebbers
committed
Lemma insert_app_r l1 l2 i x : <[length l1+i:=x]>(l1 ++ l2) = l1 ++ <[i:=x]>l2.
Proof. revert i. induction l1; intros [|?]; f_equal/=; auto. Qed.
Robbert Krebbers
committed
Lemma insert_app_r_alt l1 l2 i x :
length l1 ≤ i → <[i:=x]>(l1 ++ l2) = l1 ++ <[i - length l1:=x]>l2.
Proof.
intros. assert (i = length l1 + (i - length l1)) as Hi by lia.
rewrite Hi at 1. by apply insert_app_r.
Qed.
Robbert Krebbers
committed
Lemma delete_middle l1 l2 x : delete (length l1) (l1 ++ x :: l2) = l1 ++ l2.
Proof. induction l1; f_equal/=; auto. Qed.
Lemma length_delete l i :
is_Some (l !! i) → length (delete i l) = length l - 1.
Proof.
rewrite lookup_lt_is_Some. revert i.
induction l as [|x l IH]; intros [|i] ?; simpl in *; [lia..|].
rewrite IH by lia. lia.
Qed.
Lemma lookup_delete_lt l i j : j < i → delete i l !! j = l !! j.
Proof. revert i j; induction l; intros [] []; naive_solver eauto with lia. Qed.
Lemma lookup_total_delete_lt `{!Inhabited A} l i j :
j < i → delete i l !!! j = l !!! j.
Proof. intros. by rewrite !list_lookup_total_alt, lookup_delete_lt. Qed.
Lemma lookup_delete_ge l i j : i ≤ j → delete i l !! j = l !! S j.
Proof. revert i j; induction l; intros [] []; naive_solver eauto with lia. Qed.
Lemma lookup_total_delete_ge `{!Inhabited A} l i j :
i ≤ j → delete i l !!! j = l !!! S j.
Proof. intros. by rewrite !list_lookup_total_alt, lookup_delete_ge. Qed.
Lemma inserts_length l i k : length (list_inserts i k l) = length l.
Proof.
revert i. induction k; intros ?; csimpl; rewrite ?insert_length; auto.
Qed.
Lemma list_lookup_inserts l i k j :
i ≤ j < i + length k → j < length l →
list_inserts i k l !! j = k !! (j - i).
Proof.
revert i j. induction k as [|y k IH]; csimpl; intros i j ??; [lia|].
destruct (decide (i = j)) as [->|].
{ by rewrite list_lookup_insert, Nat.sub_diag
by (rewrite inserts_length; lia). }
rewrite list_lookup_insert_ne, IH by lia.
by replace (j - i) with (S (j - S i)) by lia.
Qed.
Lemma list_lookup_total_inserts `{!Inhabited A} l i k j :
i ≤ j < i + length k → j < length l →
list_inserts i k l !!! j = k !!! (j - i).
Proof. intros. by rewrite !list_lookup_total_alt, list_lookup_inserts. Qed.
Lemma list_lookup_inserts_lt l i k j :
j < i → list_inserts i k l !! j = l !! j.
Proof.
revert i j. induction k; intros i j ?; csimpl;
rewrite ?list_lookup_insert_ne by lia; auto with lia.
Qed.
Lemma list_lookup_total_inserts_lt `{!Inhabited A}l i k j :
j < i → list_inserts i k l !!! j = l !!! j.
Proof. intros. by rewrite !list_lookup_total_alt, list_lookup_inserts_lt. Qed.
Lemma list_lookup_inserts_ge l i k j :
i + length k ≤ j → list_inserts i k l !! j = l !! j.
Proof.
revert i j. induction k; csimpl; intros i j ?;
rewrite ?list_lookup_insert_ne by lia; auto with lia.
Qed.
Lemma list_lookup_total_inserts_ge `{!Inhabited A} l i k j :
i + length k ≤ j → list_inserts i k l !!! j = l !!! j.
Proof. intros. by rewrite !list_lookup_total_alt, list_lookup_inserts_ge. Qed.
Lemma list_lookup_inserts_Some l i k j y :
list_inserts i k l !! j = Some y ↔
(j < i ∨ i + length k ≤ j) ∧ l !! j = Some y ∨
i ≤ j < i + length k ∧ j < length l ∧ k !! (j - i) = Some y.
Proof.
destruct (decide (j < i)).
{ rewrite list_lookup_inserts_lt by done; intuition lia. }
destruct (decide (i + length k ≤ j)).
{ rewrite list_lookup_inserts_ge by done; intuition lia. }
split.
- intros Hy. assert (j < length l).
{ rewrite <-(inserts_length l i k); eauto using lookup_lt_Some. }
rewrite list_lookup_inserts in Hy by lia. intuition lia.
- intuition. by rewrite list_lookup_inserts by lia.
Qed.
Lemma list_insert_inserts_lt l i j x k :
i < j → <[i:=x]>(list_inserts j k l) = list_inserts j k (<[i:=x]>l).
Proof.
revert i j. induction k; intros i j ?; simpl;
rewrite 1?list_insert_commute by lia; auto with f_equal.
Qed.
Lemma list_inserts_app_l l1 l2 l3 i :
list_inserts i (l1 ++ l2) l3 = list_inserts (length l1 + i) l2 (list_inserts i l1 l3).
Proof.
revert i; induction l1 as [|x l1 IH]; [done|].
intro i. simpl. rewrite IH, Nat.add_succ_r. apply list_insert_inserts_lt. lia.
Qed.
Lemma list_inserts_app_r l1 l2 l3 i :
list_inserts (length l2 + i) l1 (l2 ++ l3) = l2 ++ list_inserts i l1 l3.
Proof.
revert i; induction l1 as [|x l1 IH]; [done|].
822
823
824
825
826
827
828
829
830
831
832
833
834
835
836
837
838
839
840
841
842
843
844
845
846
847
intros i. simpl. by rewrite plus_n_Sm, IH, insert_app_r.
Qed.
Lemma list_inserts_nil l1 i : list_inserts i l1 [] = [].
Proof.
revert i; induction l1 as [|x l1 IH]; [done|].
intro i. simpl. by rewrite IH.
Qed.
Lemma list_inserts_cons l1 l2 i x :
list_inserts (S i) l1 (x :: l2) = x :: list_inserts i l1 l2.
Proof.
revert i; induction l1 as [|y l1 IH]; [done|].
intro i. simpl. by rewrite IH.
Qed.
Lemma list_inserts_0_r l1 l2 l3 :
length l1 = length l2 → list_inserts 0 l1 (l2 ++ l3) = l1 ++ l3.
Proof.
revert l2. induction l1 as [|x l1 IH]; intros [|y l2] ?; simplify_eq/=; [done|].
rewrite list_inserts_cons. simpl. by rewrite IH.
Qed.
Lemma list_inserts_0_l l1 l2 l3 :
length l1 = length l3 → list_inserts 0 (l1 ++ l2) l3 = l1.
Proof.
revert l3. induction l1 as [|x l1 IH]; intros [|z l3] ?; simplify_eq/=.
{ by rewrite list_inserts_nil. }
rewrite list_inserts_cons. simpl. by rewrite IH.
Qed.
849
850
851
852
853
854
855
856
857
858
859
860
861
862
863
864
865
866
867
868
869
870
871
872
873
874
875
876
877
878
879
880
881
882
883
884
885
886
887
888
889
(** ** Properties of the [reverse] function *)
Lemma reverse_nil : reverse [] =@{list A} [].
Proof. done. Qed.
Lemma reverse_singleton x : reverse [x] = [x].
Proof. done. Qed.
Lemma reverse_cons l x : reverse (x :: l) = reverse l ++ [x].
Proof. unfold reverse. by rewrite <-!rev_alt. Qed.
Lemma reverse_snoc l x : reverse (l ++ [x]) = x :: reverse l.
Proof. unfold reverse. by rewrite <-!rev_alt, rev_unit. Qed.
Lemma reverse_app l1 l2 : reverse (l1 ++ l2) = reverse l2 ++ reverse l1.
Proof. unfold reverse. rewrite <-!rev_alt. apply rev_app_distr. Qed.
Lemma reverse_length l : length (reverse l) = length l.
Proof. unfold reverse. rewrite <-!rev_alt. apply rev_length. Qed.
Lemma reverse_involutive l : reverse (reverse l) = l.
Proof. unfold reverse. rewrite <-!rev_alt. apply rev_involutive. Qed.
Lemma reverse_lookup l i :
i < length l →
reverse l !! i = l !! (length l - S i).
Proof.
revert i. induction l as [|x l IH]; simpl; intros i Hi; [done|].
rewrite reverse_cons.
destruct (decide (i = length l)); subst.
+ by rewrite list_lookup_middle, Nat.sub_diag by by rewrite reverse_length.
+ rewrite lookup_app_l by (rewrite reverse_length; lia).
rewrite IH by lia.
by assert (length l - i = S (length l - S i)) as -> by lia.
Qed.
Lemma reverse_lookup_Some l i x :
reverse l !! i = Some x ↔ l !! (length l - S i) = Some x ∧ i < length l.
Proof.
split.
- destruct (decide (i < length l)); [ by rewrite reverse_lookup|].
rewrite lookup_ge_None_2; [done|]. rewrite reverse_length. lia.
- intros [??]. by rewrite reverse_lookup.
Qed.
Global Instance: Inj (=) (=) (@reverse A).
Proof.
intros l1 l2 Hl.
by rewrite <-(reverse_involutive l1), <-(reverse_involutive l2), Hl.
Qed.
(** ** Properties of the [elem_of] predicate *)
Robbert Krebbers
committed
Lemma not_elem_of_nil x : x ∉ [].
Robbert Krebbers
committed
Lemma elem_of_nil x : x ∈ [] ↔ False.
Proof. intuition. by destruct (not_elem_of_nil x). Qed.
Robbert Krebbers
committed
Lemma elem_of_nil_inv l : (∀ x, x ∉ l) → l = [].
Proof. destruct l; [done|]. by edestruct 1; constructor. Qed.
Lemma elem_of_not_nil x l : x ∈ l → l ≠ [].
Proof. intros ? ->. by apply (elem_of_nil x). Qed.
Robbert Krebbers
committed
Lemma elem_of_cons l x y : x ∈ y :: l ↔ x = y ∨ x ∈ l.
Proof. by split; [inversion 1; subst|intros [->|?]]; constructor. Qed.
Robbert Krebbers
committed
Lemma not_elem_of_cons l x y : x ∉ y :: l ↔ x ≠ y ∧ x ∉ l.
Robbert Krebbers
committed
Lemma elem_of_app l1 l2 x : x ∈ l1 ++ l2 ↔ x ∈ l1 ∨ x ∈ l2.
induction l1 as [|y l1 IH]; simpl.
- rewrite elem_of_nil. naive_solver.
- rewrite !elem_of_cons, IH. naive_solver.
Robbert Krebbers
committed
Lemma not_elem_of_app l1 l2 x : x ∉ l1 ++ l2 ↔ x ∉ l1 ∧ x ∉ l2.
Robbert Krebbers
committed
Lemma elem_of_list_singleton x y : x ∈ [y] ↔ x = y.
Proof. rewrite elem_of_cons, elem_of_nil. tauto. Qed.
Lemma elem_of_reverse_2 x l : x ∈ l → x ∈ reverse l.
Proof.
induction 1; rewrite reverse_cons, elem_of_app,
?elem_of_list_singleton; intuition.
Qed.
Lemma elem_of_reverse x l : x ∈ reverse l ↔ x ∈ l.
Proof.
split; auto using elem_of_reverse_2.
intros. rewrite <-(reverse_involutive l). by apply elem_of_reverse_2.
Qed.
Lemma elem_of_list_lookup_1 l x : x ∈ l → ∃ i, l !! i = Some x.
Proof.
induction 1 as [|???? IH]; [by exists 0 |].
destruct IH as [i ?]; auto. by exists (S i).
Qed.
Lemma elem_of_list_lookup_total_1 `{!Inhabited A} l x :
x ∈ l → ∃ i, i < length l ∧ l !!! i = x.
Proof.
intros [i Hi]%elem_of_list_lookup_1.
eauto using lookup_lt_Some, list_lookup_total_correct.
Qed.
Lemma elem_of_list_lookup_2 l i x : l !! i = Some x → x ∈ l.
Proof.
revert i. induction l; intros [|i] ?; simplify_eq/=; constructor; eauto.
Qed.
Lemma elem_of_list_lookup_total_2 `{!Inhabited A} l i :
i < length l → l !!! i ∈ l.
Proof. intros. by eapply elem_of_list_lookup_2, list_lookup_lookup_total_lt. Qed.
Lemma elem_of_list_lookup l x : x ∈ l ↔ ∃ i, l !! i = Some x.
Proof. firstorder eauto using elem_of_list_lookup_1, elem_of_list_lookup_2. Qed.
Lemma elem_of_list_lookup_total `{!Inhabited A} l x :
x ∈ l ↔ ∃ i, i < length l ∧ l !!! i = x.
Proof.
naive_solver eauto using elem_of_list_lookup_total_1, elem_of_list_lookup_total_2.
Qed.
Lemma elem_of_list_split_length l i x :
l !! i = Some x → ∃ l1 l2, l = l1 ++ x :: l2 ∧ i = length l1.
Proof.
revert i; induction l as [|y l IH]; intros [|i] Hl; simplify_eq/=.
- exists []; eauto.
- destruct (IH _ Hl) as (?&?&?&?); simplify_eq/=.
eexists (y :: _); eauto.
Qed.
Robbert Krebbers
committed
Lemma elem_of_list_split l x : x ∈ l → ∃ l1 l2, l = l1 ++ x :: l2.
intros [? (?&?&?&_)%elem_of_list_split_length]%elem_of_list_lookup_1; eauto.
Lemma elem_of_list_split_l `{EqDecision A} l x :
x ∈ l → ∃ l1 l2, l = l1 ++ x :: l2 ∧ x ∉ l1.
Proof.
induction 1 as [x l|x y l ? IH].
{ exists [], l. rewrite elem_of_nil. naive_solver. }
destruct (decide (x = y)) as [->|?].
- exists [], l. rewrite elem_of_nil. naive_solver.
- destruct IH as (l1 & l2 & -> & ?).
exists (y :: l1), l2. rewrite elem_of_cons. naive_solver.
Qed.
Lemma elem_of_list_split_r `{EqDecision A} l x :
x ∈ l → ∃ l1 l2, l = l1 ++ x :: l2 ∧ x ∉ l2.
Proof.
induction l as [|y l IH] using rev_ind.
{ by rewrite elem_of_nil. }
destruct (decide (x = y)) as [->|].
- exists l, []. rewrite elem_of_nil. naive_solver.
- rewrite elem_of_app, elem_of_list_singleton. intros [?| ->]; try done.
destruct IH as (l1 & l2 & -> & ?); auto.
exists l1, (l2 ++ [y]).
rewrite elem_of_app, elem_of_list_singleton, <-(assoc_L (++)). naive_solver.
Qed.
Lemma list_elem_of_insert l i x : i < length l → x ∈ <[i:=x]>l.
Proof. intros. by eapply elem_of_list_lookup_2, list_lookup_insert. Qed.
Lemma nth_elem_of l i d : i < length l → nth i l d ∈ l.
Proof.
intros; eapply elem_of_list_lookup_2.
destruct (nth_lookup_or_length l i d); [done | by lia].
Qed.
Robbert Krebbers
committed
Jonas Kastberg
committed
Lemma not_elem_of_app_cons_inv_l x y l1 l2 k1 k2 :
x ∉ k1 → y ∉ l1 →
l1 ++ x :: l2 = k1 ++ y :: k2 →
l1 = k1 ∧ x = y ∧ l2 = k2.
Proof.
revert k1. induction l1 as [|x' l1 IH]; intros [|y' k1] Hx Hy ?; simplify_eq/=;
try apply not_elem_of_cons in Hx as [??];
try apply not_elem_of_cons in Hy as [??]; naive_solver.
Qed.
Jonas Kastberg
committed
Lemma not_elem_of_app_cons_inv_r x y l1 l2 k1 k2 :