Newer
Older
(* Copyright (c) 2012-2015, Robbert Krebbers. *)
Robbert Krebbers
committed
(* This file is distributed under the terms of the BSD license. *)
(** This file collects general purpose definitions and theorems on lists that
are not in the Coq standard library. *)
Robbert Krebbers
committed
Require Export Permutation.
Require Export prelude.numbers prelude.base prelude.decidable prelude.option.
Arguments length {_} _.
Arguments cons {_} _ _.
Arguments app {_} _ _.
Arguments Permutation {_} _ _.
Notation tail := tl.
Notation take := firstn.
Notation drop := skipn.
Robbert Krebbers
committed
Arguments take {_} !_ !_ /.
Arguments drop {_} !_ !_ /.
Notation "(::)" := cons (only parsing) : C_scope.
Notation "( x ::)" := (cons x) (only parsing) : C_scope.
Notation "(:: l )" := (λ x, cons x l) (only parsing) : C_scope.
Notation "(++)" := app (only parsing) : C_scope.
Notation "( l ++)" := (app l) (only parsing) : C_scope.
Notation "(++ k )" := (λ l, app l k) (only parsing) : C_scope.
Robbert Krebbers
committed
Infix "≡ₚ" := Permutation (at level 70, no associativity) : C_scope.
Notation "(≡ₚ)" := Permutation (only parsing) : C_scope.
Notation "( x ≡ₚ)" := (Permutation x) (only parsing) : C_scope.
Notation "(≡ₚ x )" := (λ y, y ≡ₚ x) (only parsing) : C_scope.
Notation "(≢ₚ)" := (λ x y, ¬x ≡ₚ y) (only parsing) : C_scope.
Notation "x ≢ₚ y":= (¬x ≡ₚ y) (at level 70, no associativity) : C_scope.
Notation "( x ≢ₚ)" := (λ y, x ≢ₚ y) (only parsing) : C_scope.
Notation "(≢ₚ x )" := (λ y, y ≢ₚ x) (only parsing) : C_scope.
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.
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. *)
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 [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. *)
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. *)
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.
(** 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. *)
Instance list_delete {A} : Delete nat (list A) :=
fix go (i : nat) (l : list A) {struct l} : list A :=
match l with
| [] => []
Robbert Krebbers
committed
| 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]) [].
Definition list_singleton {A} (l : list A) : option A :=
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. *)
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
(** 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.
(** The function [reverse l] returns the elements of [l] in reverse order. *)
Definition reverse {A} (l : list A) : list A := rev_append l [].
(** 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.
(** 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
(** 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)
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. *)
Instance list_ret: MRet list := λ A x, x :: @nil A.
Instance list_fmap : FMap list := λ A B f,
fix go (l : list A) := match l with [] => [] | x :: l => f x :: go l end.
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.
Instance list_bind : MBind list := λ A B f,
fix go (l : list A) := match l with [] => [] | x :: l => f x ++ go l end.
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 map and fold that allow the mapped
function to use the index of the elements. *)
Definition imap_go {A B} (f : nat → A → B) : nat → list A → list B :=
fix go (n : nat) (l : list A) :=
Robbert Krebbers
committed
match l with [] => [] | x :: l => f n x :: go (S n) l end.
Definition imap {A B} (f : nat → A → B) : list A → list B := imap_go f 0.
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.
Definition imap2_go {A B C} (f : nat → A → B → C) :
nat → list A → list B → list C:=
fix go (n : nat) (l : list A) (k : list B) :=
match l, k with
| [], _ |_, [] => [] | x :: l, y :: k => f n x y :: go (S n) l k
end.
Definition imap2 {A B C} (f : nat → A → B → C) :
list A → list B → list C := imap2_go f 0.
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).
Arguments zipped_Forall_nil {_ _} _.
Arguments zipped_Forall_cons {_ _} _ _ _ _ _.
(** 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_of] holds if the first list is a suffix of the second.
The predicate [prefix_of] holds if the first list is a prefix of the second. *)
Definition suffix_of {A} : relation (list A) := λ l1 l2, ∃ k, l2 = k ++ l1.
Definition prefix_of {A} : relation (list A) := λ l1 l2, ∃ k, l2 = l1 ++ k.
Robbert Krebbers
committed
Infix "`suffix_of`" := suffix_of (at level 70) : C_scope.
Infix "`prefix_of`" := prefix_of (at level 70) : C_scope.
Hint Extern 0 (?x `prefix_of` ?y) => reflexivity.
Hint Extern 0 (?x `suffix_of` ?y) => reflexivity.
Section prefix_suffix_ops.
Context `{∀ x y : A, Decision (x = y)}.
Definition max_prefix_of : 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, [])
end.
Definition max_suffix_of (l1 l2 : list A) : list A * list A * list A :=
match max_prefix_of (reverse l1) (reverse l2) with
| (k1, k2, k3) => (reverse k1, reverse k2, reverse k3)
end.
Definition strip_prefix (l1 l2 : list A) := (max_prefix_of l1 l2).1.2.
Definition strip_suffix (l1 l2 : list A) := (max_suffix_of 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).
Robbert Krebbers
committed
Infix "`sublist`" := sublist (at level 70) : C_scope.
Hint Extern 0 (?x `sublist` ?y) => reflexivity.
Robbert Krebbers
committed
(** A list [l2] contains a list [l1] if [l2] is obtained by removing elements
from [l1] while possiblity changing the order. *)
Robbert Krebbers
committed
Inductive contains {A} : relation (list A) :=
| contains_nil : contains [] []
| contains_skip x l1 l2 : contains l1 l2 → contains (x :: l1) (x :: l2)
| contains_swap x y l : contains (y :: x :: l) (x :: y :: l)
| contains_cons x l1 l2 : contains l1 l2 → contains l1 (x :: l2)
Robbert Krebbers
committed
| contains_trans l1 l2 l3 : contains l1 l2 → contains l2 l3 → contains l1 l3.
Infix "`contains`" := contains (at level 70) : C_scope.
Hint Extern 0 (?x `contains` ?y) => reflexivity.
Robbert Krebbers
committed
Section contains_dec_help.
Context {A} {dec : ∀ x y : A, Decision (x = y)}.
Fixpoint list_remove (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.
Fixpoint list_remove_list (k : list A) (l : list A) : option (list A) :=
match k with
| [] => Some l | x :: k => list_remove x l ≫= list_remove_list k
Robbert Krebbers
committed
end.
End contains_dec_help.
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 *)
Section list_set.
Context {A} {dec : ∀ x y : A, Decision (x = y)}.
Global Instance elem_of_list_dec {dec : ∀ x y : A, Decision (x = y)}
(x : A) : ∀ l, Decision (x ∈ l).
Proof.
refine (
fix go l :=
match l return Decision (x ∈ l) with
| [] => right _
| y :: l => cast_if_or (decide (x = y)) (go l)
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
end.
End list_set.
(** * Basic tactics on lists *)
(** The tactic [discriminate_list_equality] discharges a goal if it contains
a list equality involving [(::)] and [(++)] of two lists that have a different
length as one of its hypotheses. *)
Tactic Notation "discriminate_list_equality" hyp(H) :=
apply (f_equal length) in H;
repeat (csimpl in H || rewrite app_length in H); exfalso; lia.
Tactic Notation "discriminate_list_equality" :=
Robbert Krebbers
committed
match goal with
| H : @eq (list _) _ _ |- _ => discriminate_list_equality H
end.
(** The tactic [simplify_list_equality] 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.
Ltac simplify_list_equality :=
Robbert Krebbers
committed
repeat match goal with
| _ => progress simplify_equality'
| 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.
Section setoid.
Context `{Equiv A} `{!Equivalence ((≡) : relation A)}.
Global Instance map_equivalence : Equivalence ((≡) : relation (list A)).
Proof.
split.
* intros l; induction l; constructor; auto.
* induction 1; constructor; auto.
* intros l1 l2 l3 Hl; revert l3.
induction Hl; inversion_clear 1; constructor; try etransitivity; eauto.
Qed.
Global Instance cons_proper : Proper ((≡) ==> (≡) ==> (≡)) (@cons A).
Proof. by constructor. Qed.
Global Instance app_proper : Proper ((≡) ==> (≡) ==> (≡)) (@app A).
Proof.
induction 1 as [|x y l k ?? IH]; intros ?? Htl; simpl; auto.
by apply cons_equiv, IH.
Qed.
Global Instance list_leibniz `{!LeibnizEquiv A} : LeibnizEquiv (list A).
Proof. induction 1; f_equal; fold_leibniz; auto. Qed.
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.
Proof.
revert l2. induction l1; intros [|??] H.
* discriminate (H 0).
* discriminate (H 0).
* f_equal; [by injection (H 0)|]. apply (IHl1 _ $ λ i, H (S i)).
Robbert Krebbers
committed
Global Instance list_eq_dec {dec : ∀ x y, Decision (x = y)} : ∀ l k,
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) (list_singleton 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 = [].
Proof. by destruct l. Qed.
Lemma lookup_nil i : @nil A !! i = None.
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_lt_Some l i x : l !! i = Some x → i < length l.
Proof.
revert i. induction l; intros [|?] ?; simplify_equality'; 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 [|?] ?; simplify_equality'; eauto 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 lookup_app_l l1 l2 i : i < length l1 → (l1 ++ l2) !! i = l1 !! i.
Proof. revert i. induction l1; intros [|?]; simpl; auto with lia. Qed.
Lemma lookup_app_l_Some l1 l2 i x : l1 !! i = Some x → (l1 ++ l2) !! i = Some x.
Proof. intros. rewrite lookup_app_l; eauto using lookup_lt_Some. Qed.
Lemma lookup_app_r l1 l2 i :
length l1 ≤ i → (l1 ++ l2) !! i = l2 !! (i - length l1).
Proof. revert i. induction l1; intros [|?]; simpl; auto with lia. 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.
split.
* revert i. induction l1 as [|y l1 IH]; intros [|i] ?;
simplify_equality'; auto with lia.
destruct (IH i) as [?|[??]]; auto with lia.
* intros [?|[??]]; auto using lookup_app_l_Some. by rewrite lookup_app_r.
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_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. done. intros [|i]. done. apply (IHl i). Qed.
Lemma list_lookup_alter_ne f l i j : i ≠ j → alter f i l !! j = l !! j.
revert i j. induction l; [done|]. intros [][] ?; csimpl; auto with congruence.
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_insert_ne l i j x : i ≠ j → <[i:=x]>l !! j = l !! j.
revert i j. induction l; [done|]. intros [] [] ?; simpl; auto with congruence.
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.
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_equality'.
* 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.
560
561
562
563
564
565
566
567
568
569
570
571
572
573
574
575
576
577
578
579
580
581
582
583
584
585
586
587
588
589
590
591
592
593
594
595
596
597
598
599
600
601
602
603
604
605
606
607
608
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_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_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_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.
(** ** 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.
Robbert Krebbers
committed
* split; [by right|]. intros [Hx|]; [|done]. by destruct (elem_of_nil x).
* simpl. rewrite !elem_of_cons, IHl1. tauto.
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.
Global Instance elem_of_list_permutation_proper x : Proper ((≡ₚ) ==> iff) (x ∈).
Proof. induction 1; rewrite ?elem_of_nil, ?elem_of_cons; intuition. Qed.
Robbert Krebbers
committed
Lemma elem_of_list_split l x : x ∈ l → ∃ l1 l2, l = l1 ++ x :: l2.
induction 1 as [x l|x y l ? [l1 [l2 ->]]]; [by eexists [], l|].
Robbert Krebbers
committed
Lemma elem_of_list_lookup_1 l x : x ∈ l → ∃ i, l !! i = Some x.
Robbert Krebbers
committed
induction 1 as [|???? IH]; [by exists 0 |].
destruct IH as [i ?]; auto. by exists (S i).
Robbert Krebbers
committed
Lemma elem_of_list_lookup_2 l i x : l !! i = Some x → x ∈ l.
revert i. induction l; intros [|i] ?; simplify_equality'; constructor; eauto.
Robbert Krebbers
committed
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_omap {B} (f : A → option B) l (y : B) :
y ∈ omap f l ↔ ∃ x, x ∈ l ∧ f x = Some y.
Proof.
split.
* induction l as [|x l]; csimpl; repeat case_match; inversion 1; subst;
setoid_rewrite elem_of_cons; naive_solver.
* intros (x&Hx&?). by induction Hx; csimpl; repeat case_match;
simplify_equality; try constructor; auto.
Robbert Krebbers
committed
(** ** Properties of the [NoDup] predicate *)
Lemma NoDup_nil : NoDup (@nil A) ↔ True.
Proof. split; constructor. Qed.
Robbert Krebbers
committed
Lemma NoDup_cons x l : NoDup (x :: l) ↔ x ∉ l ∧ NoDup l.
Proof. split. by inversion 1. intros [??]. by constructor. Qed.
Robbert Krebbers
committed
Lemma NoDup_cons_11 x l : NoDup (x :: l) → x ∉ l.
Proof. rewrite NoDup_cons. by intros [??]. Qed.
Robbert Krebbers
committed
Lemma NoDup_cons_12 x l : NoDup (x :: l) → NoDup l.
Proof. rewrite NoDup_cons. by intros [??]. Qed.
Robbert Krebbers
committed
Lemma NoDup_singleton x : NoDup [x].
Proof. constructor. apply not_elem_of_nil. constructor. Qed.
Robbert Krebbers
committed
Lemma NoDup_app l k : NoDup (l ++ k) ↔ NoDup l ∧ (∀ x, x ∈ l → x ∉ k) ∧ NoDup k.
Robbert Krebbers
committed
* rewrite NoDup_nil. setoid_rewrite elem_of_nil. naive_solver.
setoid_rewrite elem_of_cons. setoid_rewrite elem_of_app. naive_solver.
Robbert Krebbers
committed
Global Instance NoDup_proper: Proper ((≡ₚ) ==> iff) (@NoDup A).
Proof.
induction 1 as [|x l k Hlk IH | |].
* by rewrite !NoDup_nil.
* by rewrite !NoDup_cons, IH, Hlk.
* rewrite !NoDup_cons, !elem_of_cons. intuition.
* intuition.
Qed.
Lemma NoDup_lookup l i j x :
NoDup l → l !! i = Some x → l !! j = Some x → i = j.
Proof.
intros Hl. revert i j. induction Hl as [|x' l Hx Hl IH].
{ intros; simplify_equality. }
intros [|i] [|j] ??; simplify_equality'; eauto with f_equal;
exfalso; eauto using elem_of_list_lookup_2.
Qed.
Lemma NoDup_alt l :
NoDup l ↔ ∀ i j x, l !! i = Some x → l !! j = Some x → i = j.
split; eauto using NoDup_lookup.
induction l as [|x l IH]; intros Hl; constructor.
* rewrite elem_of_list_lookup. intros [i ?].
by feed pose proof (Hl (S i) 0 x); auto.
* apply IH. intros i j x' ??. by apply (inj S), (Hl (S i) (S j) x').
Robbert Krebbers
committed
Section no_dup_dec.
Context `{!∀ x y, Decision (x = y)}.
Global Instance NoDup_dec: ∀ l, Decision (NoDup l) :=
fix NoDup_dec l :=
match l return Decision (NoDup l) with
| [] => left NoDup_nil_2
Robbert Krebbers
committed
match decide_rel (∈) x l with
| left Hin => right (λ H, NoDup_cons_11 _ _ H Hin)
| right Hin =>
match NoDup_dec l with
| left H => left (NoDup_cons_2 _ _ Hin H)
| right H => right (H ∘ NoDup_cons_12 _ _)
end
end
Robbert Krebbers
committed
Lemma elem_of_remove_dups l x : x ∈ remove_dups l ↔ x ∈ l.
Proof.
split; induction l; simpl; repeat case_decide;
rewrite ?elem_of_cons; intuition (simplify_equality; auto).
Qed.
Lemma NoDup_remove_dups l : NoDup (remove_dups l).
Proof.
induction l; simpl; repeat case_decide; try constructor; auto.
by rewrite elem_of_remove_dups.
Qed.
Robbert Krebbers
committed
End no_dup_dec.
732
733
734
735
736
737
738
739
740
741
742
743
744
745
746
747
748
749
750
751
752
753
754
755
756
757
758
759
760
761
762
763
764
765
766
767
768
769
770
771
772
773
774
775
776
777
778
779
780
781
782
783
784
785
786
787
788
789
790
791
792
793
794
(** ** Set operations on lists *)
Section list_set.
Context {dec : ∀ x y, Decision (x = y)}.
Lemma elem_of_list_difference l k x : x ∈ list_difference l k ↔ x ∈ l ∧ x ∉ k.
Proof.
split; induction l; simpl; try case_decide;
rewrite ?elem_of_nil, ?elem_of_cons; intuition congruence.
Qed.
Lemma NoDup_list_difference l k : NoDup l → NoDup (list_difference l k).
Proof.
induction 1; simpl; try case_decide.
* constructor.
* done.
* constructor. rewrite elem_of_list_difference; intuition. done.
Qed.
Lemma elem_of_list_union l k x : x ∈ list_union l k ↔ x ∈ l ∨ x ∈ k.
Proof.
unfold list_union. rewrite elem_of_app, elem_of_list_difference.
intuition. case (decide (x ∈ k)); intuition.
Qed.
Lemma NoDup_list_union l k : NoDup l → NoDup k → NoDup (list_union l k).
Proof.
intros. apply NoDup_app. repeat split.
* by apply NoDup_list_difference.
* intro. rewrite elem_of_list_difference. intuition.
* done.
Qed.
Lemma elem_of_list_intersection l k x :
x ∈ list_intersection l k ↔ x ∈ l ∧ x ∈ k.
Proof.
split; induction l; simpl; repeat case_decide;
rewrite ?elem_of_nil, ?elem_of_cons; intuition congruence.
Qed.
Lemma NoDup_list_intersection l k : NoDup l → NoDup (list_intersection l k).
Proof.
induction 1; simpl; try case_decide.
* constructor.
* constructor. rewrite elem_of_list_intersection; intuition. done.
* done.
Qed.
Lemma elem_of_list_intersection_with f l k x :
x ∈ list_intersection_with f l k ↔ ∃ x1 x2,
x1 ∈ l ∧ x2 ∈ k ∧ f x1 x2 = Some x.
Proof.
split.
* induction l as [|x1 l IH]; simpl; [by rewrite elem_of_nil|].
intros Hx. setoid_rewrite elem_of_cons.
cut ((∃ x2, x2 ∈ k ∧ f x1 x2 = Some x)
∨ x ∈ list_intersection_with f l k); [naive_solver|].
clear IH. revert Hx. generalize (list_intersection_with f l k).
induction k; simpl; [by auto|].
case_match; setoid_rewrite elem_of_cons; naive_solver.
* intros (x1&x2&Hx1&Hx2&Hx). induction Hx1 as [x1|x1 ? l ? IH]; simpl.
+ generalize (list_intersection_with f l k).
induction Hx2; simpl; [by rewrite Hx; left |].
case_match; simpl; try setoid_rewrite elem_of_cons; auto.
+ generalize (IH Hx). clear Hx IH Hx2.
generalize (list_intersection_with f l k).
induction k; simpl; intros; [done|].
case_match; simpl; rewrite ?elem_of_cons; auto.
Qed.
End list_set.
(** ** Properties of the [filter] function *)
Robbert Krebbers
committed
Section filter.
Context (P : A → Prop) `{∀ x, Decision (P x)}.
Lemma elem_of_list_filter l x : x ∈ filter P l ↔ P x ∧ x ∈ l.
Proof.
unfold filter. induction l; simpl; repeat case_decide;
rewrite ?elem_of_nil, ?elem_of_cons; naive_solver.
Qed.
Lemma NoDup_filter l : NoDup l → NoDup (filter P l).
Robbert Krebbers
committed
Proof.
unfold filter. induction 1; simpl; repeat case_decide;
rewrite ?NoDup_nil, ?NoDup_cons, ?elem_of_list_filter; tauto.
Qed.
End filter.
(** ** Properties of the [find] function *)
Section find.
Context (P : A → Prop) `{∀ x, Decision (P x)}.
Lemma list_find_Some l i x :
list_find P l = Some (i,x) → l !! i = Some x ∧ P x.
revert i; induction l; intros [] ?;
repeat (match goal with x : prod _ _ |- _ => destruct x end
|| simplify_option_equality); eauto.
Lemma list_find_elem_of l x : x ∈ l → P x → is_Some (list_find P l).
induction 1 as [|x y l ? IH]; intros; simplify_option_equality; eauto.
by destruct IH as [[i x'] ->]; [|exists (S i, x')].
(** ** Properties of the [reverse] function *)
Lemma reverse_nil : reverse [] = @nil A.
Proof. done. Qed.
Robbert Krebbers
committed
Lemma reverse_singleton x : reverse [x] = [x].
Proof. done. Qed.
Robbert Krebbers
committed
Lemma reverse_cons l x : reverse (x :: l) = reverse l ++ [x].
Proof. unfold reverse. by rewrite <-!rev_alt. Qed.
Robbert Krebbers
committed
Lemma reverse_snoc l x : reverse (l ++ [x]) = x :: reverse l.
Proof. unfold reverse. by rewrite <-!rev_alt, rev_unit. Qed.
Robbert Krebbers
committed
Lemma reverse_app l1 l2 : reverse (l1 ++ l2) = reverse l2 ++ reverse l1.
Proof. unfold reverse. rewrite <-!rev_alt. apply rev_app_distr. Qed.
Robbert Krebbers
committed
Lemma reverse_length l : length (reverse l) = length l.
Proof. unfold reverse. rewrite <-!rev_alt. apply rev_length. Qed.
Robbert Krebbers
committed
Lemma reverse_involutive l : reverse (reverse l) = l.
Proof. unfold reverse. rewrite <-!rev_alt. apply rev_involutive. 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.
Global Instance: Inj (=) (=) (@reverse A).
Proof.
intros l1 l2 Hl.
by rewrite <-(reverse_involutive l1), <-(reverse_involutive l2), Hl.
Qed.
Lemma sum_list_with_app (f : A → nat) l k :
sum_list_with f (l ++ k) = sum_list_with f l + sum_list_with f k.
Proof. induction l; simpl; lia. Qed.
Lemma sum_list_with_reverse (f : A → nat) l :
sum_list_with f (reverse l) = sum_list_with f l.
Proof.
induction l; simpl; rewrite ?reverse_cons, ?sum_list_with_app; simpl; lia.
Qed.
(** ** Properties of the [last] function *)
Lemma last_snoc x l : last (l ++ [x]) = Some x.
Proof. induction l as [|? []]; simpl; auto. Qed.
Lemma last_reverse l : last (reverse l) = head l.
Proof. by destruct l as [|x l]; rewrite ?reverse_cons, ?last_snoc. Qed.
Lemma head_reverse l : head (reverse l) = last l.
Proof. by rewrite <-last_reverse, reverse_involutive. Qed.
Robbert Krebbers
committed
(** ** Properties of the [take] function *)
Definition take_drop i l : take i l ++ drop i l = l := firstn_skipn i l.
Lemma take_drop_middle l i x :
l !! i = Some x → take i l ++ x :: drop (S i) l = l.
Proof.
revert i x. induction l; intros [|?] ??; simplify_equality'; f_equal; auto.
Qed.
Robbert Krebbers
committed
Lemma take_nil n : take n (@nil A) = [].
Robbert Krebbers
committed
Lemma take_app l k : take (length l) (l ++ k) = l.
Proof. induction l; f_equal'; auto. Qed.
Robbert Krebbers
committed
Lemma take_app_alt l k n : n = length l → take n (l ++ k) = l.
Proof. intros ->. by apply take_app. Qed.
Lemma take_app3_alt l1 l2 l3 n : n = length l1 → take n ((l1 ++ l2) ++ l3) = l1.
Proof. intros ->. by rewrite <-(assoc_L (++)), take_app. Qed.
Robbert Krebbers
committed
Lemma take_app_le l k n : n ≤ length l → take n (l ++ k) = take n l.
Proof. revert n. induction l; intros [|?] ?; f_equal'; auto with lia. Qed.
Lemma take_plus_app l k n m :
length l = n → take (n + m) (l ++ k) = l ++ take m k.
Proof. intros <-. induction l; f_equal'; auto. Qed.
Robbert Krebbers
committed
Lemma take_app_ge l k n :
length l ≤ n → take n (l ++ k) = l ++ take (n - length l) k.
Proof. revert n. induction l; intros [|?] ?; f_equal'; auto with lia. Qed.
Robbert Krebbers
committed
Lemma take_ge l n : length l ≤ n → take n l = l.
Proof. revert n. induction l; intros [|?] ?; f_equal'; auto with lia. Qed.
Robbert Krebbers
committed
Lemma take_take l n m : take n (take m l) = take (min n m) l.
Proof. revert n m. induction l; intros [|?] [|?]; f_equal'; auto. Qed.
Lemma take_idemp l n : take n (take n l) = take n l.
Proof. by rewrite take_take, Min.min_idempotent. Qed.
Robbert Krebbers
committed
Lemma take_length l n : length (take n l) = min n (length l).
Proof. revert n. induction l; intros [|?]; f_equal'; done. Qed.
Lemma take_length_le l n : n ≤ length l → length (take n l) = n.
Proof. rewrite take_length. apply Min.min_l. Qed.
Lemma take_length_ge l n : length l ≤ n → length (take n l) = length l.
Proof. rewrite take_length. apply Min.min_r. Qed.
Lemma take_drop_commute l n m : take n (drop m l) = drop m (take (m + n) l).
revert n m. induction l; intros [|?][|?]; simpl; auto using take_nil with lia.
Lemma lookup_take l n i : i < n → take n l !! i = l !! i.
Proof. revert n i. induction l; intros [|n] [|i] ?; simpl; auto with lia. Qed.
Robbert Krebbers
committed
Lemma lookup_take_ge l n i : n ≤ i → take n l !! i = None.
Proof. revert n i. induction l; intros [|?] [|?] ?; simpl; auto with lia. Qed.
Robbert Krebbers
committed
Lemma take_alter f l n i : n ≤ i → take n (alter f i l) = take n l.
Proof.
intros. apply list_eq. intros j. destruct (le_lt_dec n j).
* by rewrite !lookup_take_ge.
* by rewrite !lookup_take, !list_lookup_alter_ne by lia.
Qed.
Robbert Krebbers
committed
Lemma take_insert l n i x : n ≤ i → take n (<[i:=x]>l) = take n l.
Proof.
intros. apply list_eq. intros j. destruct (le_lt_dec n j).
* by rewrite !lookup_take_ge.
* by rewrite !lookup_take, !list_lookup_insert_ne by lia.
Qed.
(** ** Properties of the [drop] function *)
Lemma drop_0 l : drop 0 l = l.
Proof. done. Qed.
Robbert Krebbers
committed
Lemma drop_nil n : drop n (@nil A) = [].
Robbert Krebbers
committed
Lemma drop_length l n : length (drop n l) = length l - n.
Proof. revert n. by induction l; intros [|i]; f_equal'. Qed.
Lemma drop_ge l n : length l ≤ n → drop n l = [].
Proof. revert n. induction l; intros [|??]; simpl in *; auto with lia. Qed.
Robbert Krebbers
committed
Lemma drop_all l : drop (length l) l = [].
Lemma drop_drop l n1 n2 : drop n1 (drop n2 l) = drop (n2 + n1) l.
Proof. revert n2. induction l; intros [|?]; simpl; rewrite ?drop_nil; auto. Qed.
Lemma drop_app_le l k n :
n ≤ length l → drop n (l ++ k) = drop n l ++ k.
Proof. revert n. induction l; intros [|?]; simpl; auto with lia. Qed.
Lemma drop_app l k : drop (length l) (l ++ k) = k.
Proof. by rewrite drop_app_le, drop_all. Qed.
Lemma drop_app_alt l k n : n = length l → drop n (l ++ k) = k.
Proof. intros ->. by apply drop_app. Qed.
Lemma drop_app3_alt l1 l2 l3 n :
n = length l1 → drop n ((l1 ++ l2) ++ l3) = l2 ++ l3.
Proof. intros ->. by rewrite <-(assoc_L (++)), drop_app. Qed.
Lemma drop_app_ge l k n :
length l ≤ n → drop n (l ++ k) = drop (n - length l) k.
Proof.
intros. rewrite <-(Nat.sub_add (length l) n) at 1 by done.
by rewrite Nat.add_comm, <-drop_drop, drop_app.
Qed.
Lemma drop_plus_app l k n m :
length l = n → drop (n + m) (l ++ k) = drop m k.
Proof. intros <-. by rewrite <-drop_drop, drop_app. Qed.
Robbert Krebbers
committed
Lemma lookup_drop l n i : drop n l !! i = l !! (n + i).
Proof. revert n i. induction l; intros [|i] ?; simpl; auto. Qed.
Robbert Krebbers
committed
Lemma drop_alter f l n i : i < n → drop n (alter f i l) = drop n l.
Proof.
intros. apply list_eq. intros j.
by rewrite !lookup_drop, !list_lookup_alter_ne by lia.
Qed.
Robbert Krebbers
committed
Lemma drop_insert l n i x : i < n → drop n (<[i:=x]>l) = drop n l.
Proof.
intros. apply list_eq. intros j.
by rewrite !lookup_drop, !list_lookup_insert_ne by lia.
Qed.
Robbert Krebbers
committed
Lemma delete_take_drop l i : delete i l = take i l ++ drop (S i) l.
Proof. revert i. induction l; intros [|?]; f_equal'; auto. Qed.
Lemma take_take_drop l n m : take n l ++ take m (drop n l) = take (n + m) l.
Proof. revert n m. induction l; intros [|?] [|?]; f_equal'; auto. Qed.
Lemma drop_take_drop l n m : n ≤ m → drop n (take m l) ++ drop m l = drop n l.
Proof.
revert n m. induction l; intros [|?] [|?] ?;
f_equal'; auto using take_drop with lia.
Qed.
(** ** Properties of the [replicate] function *)
Robbert Krebbers
committed
Lemma replicate_length n x : length (replicate n x) = n.
Lemma lookup_replicate n x y i :
replicate n x !! i = Some y ↔ y = x ∧ i < n.
Proof.
split.
* revert i. induction n; intros [|?]; naive_solver auto with lia.
* intros [-> Hi]. revert i Hi.
induction n; intros [|?]; naive_solver auto with lia.
Qed.
Lemma lookup_replicate_1 n x y i :
replicate n x !! i = Some y → y = x ∧ i < n.
Proof. by rewrite lookup_replicate. Qed.
Lemma lookup_replicate_2 n x i : i < n → replicate n x !! i = Some x.
Proof. by rewrite lookup_replicate. Qed.
Lemma lookup_replicate_None n x i : n ≤ i ↔ replicate n x !! i = None.