Skip to content
Snippets Groups Projects
list.v 140 KiB
Newer Older
(* Copyright (c) 2012-2014, Robbert Krebbers. *)
(* 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. *)
Require Export numbers base decidable option.
Robbert Krebbers's avatar
Robbert Krebbers committed

Arguments length {_} _.
Robbert Krebbers's avatar
Robbert Krebbers committed
Arguments cons {_} _ _.
Arguments app {_} _ _.
Arguments Permutation {_} _ _.
Arguments Forall_cons {_} _ _ _ _ _.
Robbert Krebbers's avatar
Robbert Krebbers committed

Notation tail := tl.
Notation take := firstn.
Notation drop := skipn.
Arguments take {_} !_ !_ /.
Arguments drop {_} !_ !_ /.

Robbert Krebbers's avatar
Robbert Krebbers committed
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.

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.

(** * Definitions *)
(** The operation [l !! i] gives the [i]th element of the list [l], or [None]
in case [i] is out of bounds. *)
Robbert Krebbers's avatar
Robbert Krebbers committed
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. *)
Robbert Krebbers's avatar
Robbert Krebbers committed
Instance list_alter {A} (f : A  A) : AlterD nat A (list A) f :=
  fix go i l {struct l} := let _ : AlterD _ _ _ f := @go in
  match l with
  | [] => []
  | x :: l => match i with 0 => f x :: l | S i => x :: alter f 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.
(** 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 :=
  | x :: l => match i with 0 => l | S i => x :: @delete _ _ go i l end

(** The function [option_list o] converts an element [Some x] into the
singleton list [[x]], and [None] into the empty list [[]]. *)
Robbert Krebbers's avatar
Robbert Krebbers committed
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.
Robbert Krebbers's avatar
Robbert Krebbers committed

(** 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
Robbert Krebbers's avatar
Robbert Krebbers committed
  match l with
  | [] => []
  | 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 :=
  fix go l :=
  match l with
  | [] => None | x :: l => if decide (P x) then Some 0 else S <$> go l
Robbert Krebbers's avatar
Robbert Krebbers committed

(** 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 :=
  match n with 0 => [] | S n => x :: replicate n x end.
Robbert Krebbers's avatar
Robbert Krebbers committed

(** 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.
Robbert Krebbers's avatar
Robbert Krebbers committed
(** 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
  | x :: l => match n with 0 => [] | S n => x :: resize n y l end
Robbert Krebbers's avatar
Robbert Krebbers committed
  end.
Arguments resize {_} !_ _ !_.

(** 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 {A B} (f : A  B) : FMapD list f :=
  fix go (l : list A) :=
  match l with [] => [] | x :: l => f x :: @fmap _ _ _ f go l end.
Instance list_bind {A B} (f : A  list B) : MBindD list f :=
  fix go (l : list A) :=
  match l with [] => [] | x :: l => f x ++ @mbind _ _ _ f go l end.
Instance list_join: MJoin list :=
  fix go A (ls : list (list A)) : list A :=
  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) :=
  fix go l :=
  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) :=
  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.
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.

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.
Infix "`suffix_of`" := suffix_of (at level 70) : C_scope.
Infix "`prefix_of`" := prefix_of (at level 70) : C_scope.
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 =>
      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.
(** 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 [] []
  | 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`" := sublist (at level 70) : C_scope.

(** A list [l2] contains a list [l1] if [l2] is obtained by removing elements
from [l1] while possiblity changing the order. *)
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)
  | contains_trans l1 l2 l3 : contains l1 l2  contains l2 l3  contains l1 l3.
Infix "`contains`" := contains (at level 70) : C_scope.

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
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').

(** 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
  Definition list_union (l k : list A) : list A := list_difference l k ++ k.
  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
    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 (simpl in H || rewrite app_length in H); exfalso; lia.
Tactic Notation "discriminate_list_equality" :=
  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_injective_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_injective_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_injective_1; auto.
  apply (f_equal length) in Hl. rewrite !app_length in Hl. lia.
Qed.
Ltac simplify_list_equality :=
  repeat match goal with
  | _ => progress simplify_equality
  | H : _ ++ _ = _ ++ _ |- _ => first
    [ apply app_inv_head in H | apply app_inv_tail in H
    | apply app_injective_1 in H; [destruct H|done]
    | apply app_injective_2 in H; [destruct H|done] ]
Robbert Krebbers's avatar
Robbert Krebbers committed
  | H : [?x] !! ?i = Some ?y |- _ =>
    destruct i; [change (Some x = Some y) in H | discriminate]
  end;
  try discriminate_list_equality.
Ltac simplify_list_equality' :=
  repeat (progress simpl in * || simplify_list_equality).
(** * General theorems *)
Section general_properties.
Robbert Krebbers's avatar
Robbert Krebbers committed
Context {A : Type}.
Implicit Types x y z : A.
Implicit Types l k : list A.
Robbert Krebbers's avatar
Robbert Krebbers committed

Global Instance: Injective2 (=) (=) (=) (@cons A).
Proof. by injection 1. Qed.
Global Instance:  k, Injective (=) (=) (k ++).
Proof. intros ???. apply app_inv_head. Qed.
Global Instance:  k, Injective (=) (=) (++ k).
Proof. intros ???. apply app_inv_tail. Qed.
Global Instance: Associative (=) (@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.
Lemma app_nil l1 l2 : l1 ++ l2 = []  l1 = []  l2 = [].
Proof. split. apply app_eq_nil. by intros [-> ->]. Qed.
Lemma app_singleton l1 l2 x :
  l1 ++ l2 = [x]  l1 = []  l2 = [x]  l1 = [x]  l2 = [].
Proof. split. apply app_eq_unit. by intros [[-> ->]|[-> ->]]. Qed.
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)).
Global Instance list_eq_dec {dec :  x y, Decision (x = y)} :  l k,
  Decision (l = k) := 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) (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.
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.
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_length l1 l2 :
Robbert Krebbers's avatar
Robbert Krebbers committed
  length l2 = length l1 
  ( i x y, l1 !! i = Some x  l2 !! i = Some y  x = y)  l1 = l2.
Robbert Krebbers's avatar
Robbert Krebbers committed
Proof.
  intros Hl ?; apply list_eq; intros i. destruct (l2 !! i) as [x|] eqn:Hx.
  * destruct (lookup_lt_is_Some_2 l1 i) as [y ?]; [|naive_solver].
    rewrite <-Hl. eauto using lookup_lt_Some.
  * by rewrite lookup_ge_None, <-Hl, <-lookup_ge_None.
Robbert Krebbers's avatar
Robbert Krebbers committed
Qed.
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 : (l1 ++ l2) !! (length l1 + i) = l2 !! i.
Proof. revert i. induction l1; intros [|i]; simplify_equality'; auto. Qed.
Lemma lookup_app_r_alt l1 l2 i j :
  j = length l1  (l1 ++ l2) !! (j + i) = l2 !! i.
Proof. intros ->. by apply lookup_app_r. Qed.
Lemma lookup_app_r_Some l1 l2 i x :
  l2 !! i = Some x  (l1 ++ l2) !! (length l1 + i) = Some x.
Proof. by rewrite lookup_app_r. Qed.
Lemma lookup_app_minus_r l1 l2 i :
  length l1  i  (l1 ++ l2) !! i = l2 !! (i - length l1).
Proof. intros. rewrite <-(lookup_app_r l1 l2). f_equal. lia. Qed.
Lemma lookup_app_inv l1 l2 i x :
  (l1 ++ l2) !! i = Some x  l1 !! i = Some x  l2 !! (i - length l1) = Some x.
Proof. revert i. induction l1; intros [|i] ?; simplify_equality'; auto. 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 alter_length f l i : length (alter f i l) = length l.
Proof. revert i. by induction l; intros [|?]; f_equal'. Qed.
Lemma insert_length l i x : length (<[i:=x]>l) = length l.
Proof. revert i. by induction l; intros [|?]; f_equal'. Qed.
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.
Robbert Krebbers's avatar
Robbert Krebbers committed
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 [] [] ?; simpl; auto with congruence.
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_other l i x :
  length l  1  l !! i = Some x   j y, j  i  l !! j = Some y.
Robbert Krebbers's avatar
Robbert Krebbers committed
Proof.
  intros. destruct i, l as [|x0 [|x1 l]]; simplify_equality'.
Robbert Krebbers's avatar
Robbert Krebbers committed
  * by exists 1 x1.
  * by exists 0 x0.
Qed.
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.
  alter f (length l1 + i) (l1 ++ l2) = l1 ++ alter f i l2.
Proof. revert i. induction l1; intros [|?]; f_equal'; auto. Qed.
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_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.
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.
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.
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.
Lemma delete_middle l1 l2 x : delete (length l1) (l1 ++ x :: l2) = l1 ++ l2.
Proof. induction l1; f_equal'; auto. Qed.
(** ** Properties of the [elem_of] predicate *)
Proof. by inversion 1. Qed.
Lemma elem_of_nil x : x  []  False.
Proof. intuition. by destruct (not_elem_of_nil x). Qed.
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.
Lemma elem_of_cons l x y : x  y :: l  x = y  x  l.
Proof. split; [inversion 1; subst|intros [->|?]]; constructor (done). Qed.
Lemma not_elem_of_cons l x y : x  y :: l  x  y  x  l.
Robbert Krebbers's avatar
Robbert Krebbers committed
Proof. rewrite elem_of_cons. tauto. Qed.
Lemma elem_of_app l1 l2 x : x  l1 ++ l2  x  l1  x  l2.
  induction l1.
  * split; [by right|]. intros [Hx|]; [|done]. by destruct (elem_of_nil x).
  * simpl. rewrite !elem_of_cons, IHl1. tauto.
Lemma not_elem_of_app l1 l2 x : x  l1 ++ l2  x  l1  x  l2.
Robbert Krebbers's avatar
Robbert Krebbers committed
Proof. rewrite elem_of_app. tauto. Qed.
Lemma elem_of_list_singleton x y : x  [y]  x = y.
Proof. rewrite elem_of_cons, elem_of_nil. tauto. Qed.
Robbert Krebbers's avatar
Robbert Krebbers committed
Global Instance elem_of_list_permutation_proper x : Proper (() ==> iff) (x ).
Proof. induction 1; rewrite ?elem_of_nil, ?elem_of_cons; intuition. Qed.
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|].
  by exists (y :: l1) l2.
Lemma elem_of_list_lookup_1 l x : x  l   i, l !! i = Some x.
  induction 1 as [|???? IH]; [by exists 0 |].
  destruct IH as [i ?]; auto. by exists (S i).
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.
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.

(** ** Properties of the [NoDup] predicate *)
Lemma NoDup_nil : NoDup (@nil A)  True.
Proof. split; constructor. Qed.
Lemma NoDup_cons x l : NoDup (x :: l)  x  l  NoDup l.
Proof. split. by inversion 1. intros [??]. by constructor. Qed.
Lemma NoDup_cons_11 x l : NoDup (x :: l)  x  l.
Proof. rewrite NoDup_cons. by intros [??]. Qed.
Lemma NoDup_cons_12 x l : NoDup (x :: l)  NoDup l.
Proof. rewrite NoDup_cons. by intros [??]. Qed.
Proof. constructor. apply not_elem_of_nil. constructor. Qed.
Lemma NoDup_app l k : NoDup (l ++ k)  NoDup l  ( x, x  l  x  k)  NoDup k.
Robbert Krebbers's avatar
Robbert Krebbers committed
Proof.
  induction l; simpl.
  * rewrite NoDup_nil. setoid_rewrite elem_of_nil. naive_solver.
  * rewrite !NoDup_cons.
Robbert Krebbers's avatar
Robbert Krebbers committed
    setoid_rewrite elem_of_cons. setoid_rewrite elem_of_app. naive_solver.
Robbert Krebbers's avatar
Robbert Krebbers committed
Qed.
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 (injective S), (Hl (S i) (S j) x').
Robbert Krebbers's avatar
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
    | x :: l =>
      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
  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.
(** ** 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 *)
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).
  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 :
    list_find P l = Some i   x, l !! i = Some x  P x.
    revert i. induction l; intros [] ?; simplify_option_equality; eauto.
  Qed.
  Lemma list_find_elem_of l x : x  l  P x   i, list_find P l = Some i.
  Proof.
    induction 1 as [|x y l ? IH]; intros; simplify_option_equality; eauto.
    by destruct IH as [i ->]; [|exists (S i)].
  Qed.
End find.

Section find_eq.
  Context `{ x y, Decision (x = y)}.
  Lemma list_find_eq_Some l i x : list_find (x =) l = Some i  l !! i = Some x.
  Proof.
    intros.
    destruct (list_find_Some (x =) l i) as (?&?&?); auto with congruence.
  Qed.
  Lemma list_find_eq_elem_of l x : x  l   i, list_find (x=) l = Some i.
  Proof. eauto using list_find_elem_of. Qed.
End find_eq.

(** ** Properties of the [reverse] function *)
Lemma reverse_nil : reverse [] = @nil A.
Proof. done. Qed.
Lemma reverse_singleton x : reverse [x] = [x].
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 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: Injective (=) (=) (@reverse A).
Proof.
  intros l1 l2 Hl.
  by rewrite <-(reverse_involutive l1), <-(reverse_involutive l2), Hl.
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.
(** ** 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.
Lemma take_nil n : take n (@nil A) = [].
Robbert Krebbers's avatar
Robbert Krebbers committed
Proof. by destruct n. Qed.
Lemma take_app l k : take (length l) (l ++ k) = l.
Proof. induction l; f_equal'; auto. Qed.
Lemma take_app_alt l k n : n = length l  take n (l ++ k) = l.
Robbert Krebbers's avatar
Robbert Krebbers committed
Proof. intros Hn. by rewrite Hn, take_app. Qed.
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.
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.
Lemma take_ge l n : length l  n  take n l = l.
Proof. revert n. induction l; intros [|?] ?; f_equal'; auto with lia. Qed.
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_idempotent l n : take n (take n l) = take n l.
Robbert Krebbers's avatar
Robbert Krebbers committed
Proof. by rewrite take_take, Min.min_idempotent. Qed.
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.
Robbert Krebbers's avatar
Robbert Krebbers committed
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).
Robbert Krebbers's avatar
Robbert Krebbers committed
Proof.
  revert n m. induction l; intros [|?][|?]; simpl; auto using take_nil with lia.
Robbert Krebbers's avatar
Robbert Krebbers committed
Qed.
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.
Lemma lookup_take_ge l n i : n  i  take n l !! i = None.
Robbert Krebbers's avatar
Robbert Krebbers committed
Proof. revert n i. induction l; intros [|?] [|?] ?; simpl; auto with lia. Qed.
Lemma take_alter f l n i : n  i  take n (alter f i l) = take n l.
Robbert Krebbers's avatar
Robbert Krebbers committed
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.
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.
Lemma drop_nil n : drop n (@nil A) = [].
Robbert Krebbers's avatar
Robbert Krebbers committed
Proof. by destruct n. Qed.
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.
Lemma drop_all l : drop (length l) l = [].
Proof. by apply drop_ge. Qed.
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_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 lookup_drop l n i : drop n l !! i = l !! (n + i).
Robbert Krebbers's avatar
Robbert Krebbers committed
Proof. revert n i. induction l; intros [|i] ?; simpl; auto. Qed.
Lemma drop_alter f l n i : i < n  drop n (alter f i l) = drop n l.
Robbert Krebbers's avatar
Robbert Krebbers committed
Proof.
  intros. apply list_eq. intros j.
  by rewrite !lookup_drop, !list_lookup_alter_ne by lia.
Qed.
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.
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.

(** ** Properties of the [replicate] function *)
Lemma replicate_length n x : length (replicate n x) = n.
Proof. induction n; simpl; auto. Qed.
Lemma lookup_replicate n x i : i < n  replicate n x !! i = Some x.
Proof. revert i. induction n; intros [|?]; naive_solver auto with lia. Qed.
Lemma lookup_replicate_inv n x y i :
  replicate n x !! i = Some y  y = x  i < n.
Proof. revert i. induction n; intros [|?]; naive_solver auto with lia. Qed.
Lemma lookup_replicate_None n x i : n  i  replicate n x !! i = None.
Proof.
  rewrite eq_None_not_Some, Nat.le_ngt. split.
  * intros Hin [x' Hx']; destruct Hin.
    by destruct (lookup_replicate_inv n x x' i).
  * intros Hx ?. destruct Hx. exists x; auto using lookup_replicate.
Qed.
Lemma elem_of_replicate_inv x n y : x  replicate n y  x = y.
Proof. induction n; simpl; rewrite ?elem_of_nil, ?elem_of_cons; intuition. Qed.
Lemma replicate_S n x : replicate (S n) x = x :: replicate  n x.
Robbert Krebbers's avatar
Robbert Krebbers committed
  replicate (n + m) x = replicate n x ++ replicate m x.
Proof. induction n; f_equal'; auto. Qed.
Lemma take_replicate n m x : take n (replicate m x) = replicate (min n m) x.
Proof. revert m. by induction n; intros [|?]; f_equal'. Qed.
Lemma take_replicate_plus n m x : take n (replicate (n + m) x) = replicate n x.
Robbert Krebbers's avatar
Robbert Krebbers committed
Proof. by rewrite take_replicate, min_l by lia. Qed.
Lemma drop_replicate n m x : drop n (replicate m x) = replicate (m - n) x.
Proof. revert m. by induction n; intros [|?]; f_equal'. Qed.
Lemma drop_replicate_plus n m x : drop n (replicate (n + m) x) = replicate m x.
Robbert Krebbers's avatar
Robbert Krebbers committed
Proof. rewrite drop_replicate. f_equal. lia. Qed.
Lemma replicate_as_elem_of x n l :
  replicate n x = l  length l = n   y, y  l  y = x.
  split; [intros <-; eauto using elem_of_replicate_inv, replicate_length|].
  intros [<- Hl]. symmetry. induction l as [|y l IH]; f_equal'.
  * apply Hl. by left.
  * apply IH. intros ??. apply Hl. by right.
Lemma reverse_replicate n x : reverse (replicate n x) = replicate n x.
  symmetry. apply replicate_as_elem_of.
  rewrite reverse_length, replicate_length. split; auto.
  intros y. rewrite elem_of_reverse. by apply elem_of_replicate_inv.
Lemma replicate_false βs n : length βs = n  replicate n false =.>* βs.
Proof. intros <-. by induction βs; simpl; constructor. Qed.
(** ** Properties of the [resize] function *)
Lemma resize_spec l n x : resize n x l = take n l ++ replicate (n - length l) x.
Proof. revert n. induction l; intros [|?]; f_equal'; auto. Qed.
Lemma resize_0 l x : resize 0 x l = [].
Robbert Krebbers's avatar
Robbert Krebbers committed
Proof. by destruct l. Qed.
Lemma resize_nil n x : resize n x [] = replicate n x.
Proof. rewrite resize_spec. rewrite take_nil. f_equal'. lia. Qed.
Lemma resize_ge l n x :
  length l  n  resize n x l = l ++ replicate (n - length l) x.
Robbert Krebbers's avatar
Robbert Krebbers committed
Proof. intros. by rewrite resize_spec, take_ge. Qed.
Lemma resize_le l n x : n  length l  resize n x l = take n l.
Robbert Krebbers's avatar
Robbert Krebbers committed
Proof.
  intros. rewrite resize_spec, (proj2 (Nat.sub_0_le _ _)) by done.
  simpl. by rewrite (right_id_L [] (++)).
Robbert Krebbers's avatar
Robbert Krebbers committed
Qed.
Lemma resize_all l x : resize (length l) x l = l.
Robbert Krebbers's avatar
Robbert Krebbers committed
Proof. intros. by rewrite resize_le, take_ge. Qed.
Lemma resize_all_alt l n x : n = length l  resize n x l = l.
Proof. intros ->. by rewrite resize_all. Qed.
Robbert Krebbers's avatar
Robbert Krebbers committed
  resize (n + m) x l = resize n x l ++ resize m x (drop n l).
Proof.
  revert n m. induction l; intros [|?] [|?]; f_equal'; auto.
  * by rewrite Nat.add_0_r, (right_id_L [] (++)).
Robbert Krebbers's avatar
Robbert Krebbers committed
  * by rewrite replicate_plus.
Qed.
Lemma resize_plus_eq l n m x :
  length l = n  resize (n + m) x l = l ++ replicate m x.
Proof. intros <-. by rewrite resize_plus, resize_all, drop_all, resize_nil. Qed.
Lemma resize_app_le l1 l2 n x :
  n  length l1  resize n x (l1 ++ l2) = resize n x l1.
Robbert Krebbers's avatar
Robbert Krebbers committed
Proof.
  intros. by rewrite !resize_le, take_app_le by (rewrite ?app_length; lia).
Robbert Krebbers's avatar
Robbert Krebbers committed
Qed.
Lemma resize_app l1 l2 n x : n = length l1  resize n x (l1 ++ l2) = l1.
Proof. intros ->. by rewrite resize_app_le, resize_all. Qed.
Lemma resize_app_ge l1 l2 n x :
  length l1  n  resize n x (l1 ++ l2) = l1 ++ resize (n - length l1) x l2.
Robbert Krebbers's avatar
Robbert Krebbers committed
Proof.
  intros. rewrite !resize_spec, take_app_ge, (associative_L (++)) by done.
Robbert Krebbers's avatar
Robbert Krebbers committed
  do 2 f_equal. rewrite app_length. lia.
Qed.
Lemma resize_length l n x : length (resize n x l) = n.
Proof. rewrite resize_spec, app_length, replicate_length, take_length. lia. Qed.
Lemma resize_replicate x n m : resize n x (replicate m x) = replicate n x.
Proof. revert m. induction n; intros [|?]; f_equal'; auto. Qed.
Lemma resize_resize l n m x : n  m  resize n x (resize m x l) = resize n x l.
Robbert Krebbers's avatar
Robbert Krebbers committed
Proof.
  revert n m. induction l; simpl.
  * intros. by rewrite !resize_nil, resize_replicate.
  * intros [|?] [|?] ?; f_equal'; auto with lia.
Robbert Krebbers's avatar
Robbert Krebbers committed
Qed.
Lemma resize_idempotent l n x : resize n x (resize n x l) = resize n x l.
Robbert Krebbers's avatar
Robbert Krebbers committed
Proof. by rewrite resize_resize. Qed.
Lemma resize_take_le l n m x : n  m  resize n x (take m l) = resize n x l.
Proof. revert n m. induction l; intros [|?][|?] ?; f_equal'; auto with lia. Qed.
Lemma resize_take_eq l n x : resize n x (take n l) = resize n x l.
Robbert Krebbers's avatar
Robbert Krebbers committed
Proof. by rewrite resize_take_le. Qed.
Lemma take_resize l n m x : take n (resize m x l) = resize (min n m) x l.
Robbert Krebbers's avatar
Robbert Krebbers committed
Proof.
  revert n m. induction l; intros [|?][|?]; f_equal'; auto using take_replicate.
Robbert Krebbers's avatar
Robbert Krebbers committed
Qed.
Lemma take_resize_le l n m x : n  m  take n (resize m x l) = resize n x l.
Robbert Krebbers's avatar
Robbert Krebbers committed
Proof. intros. by rewrite take_resize, Min.min_l. Qed.
Lemma take_resize_eq l n x : take n (resize n x l) = resize n x l.
Robbert Krebbers's avatar
Robbert Krebbers committed
Proof. intros. by rewrite take_resize, Min.min_l. Qed.
Lemma take_resize_plus l n m x : take n (resize (n + m) x l) = resize n x l.
Robbert Krebbers's avatar
Robbert Krebbers committed
Proof. by rewrite take_resize, min_l by lia. Qed.
Lemma drop_resize_le l n m x :
  n  m  drop n (resize m x l) = resize (m - n) x (drop n l).
Robbert Krebbers's avatar
Robbert Krebbers committed
Proof.
  revert n m. induction l; simpl.
  * intros. by rewrite drop_nil, !resize_nil, drop_replicate.
  * intros [|?] [|?] ?; simpl; try case_match; auto with lia.
Qed.
Robbert Krebbers's avatar
Robbert Krebbers committed
  drop n (resize (n + m) x l) = resize m x (drop n l).
Proof. rewrite drop_resize_le by lia. f_equal. lia. Qed.
Lemma lookup_resize l n x i : i < n  i < length l  resize n x l !! i = l !! i.
Proof.
  intros ??. destruct (decide (n < length l)).
  * by rewrite resize_le, lookup_take by lia.
  * by rewrite resize_ge, lookup_app_l by lia.
Qed.
Lemma lookup_resize_new l n x i :
  length l  i  i < n  resize n x l !! i = Some x.
Proof.
  intros ??. rewrite resize_ge by lia.
  replace i with (length l + (i - length l)) by lia.
  by rewrite lookup_app_r, lookup_replicate by lia.
Qed.
Lemma lookup_resize_old l n x i : n  i  resize n x l !! i = None.
Proof. intros ?. apply lookup_ge_None_2. by rewrite resize_length. Qed.
End general_properties.
Section more_general_properties.
Context {A : Type}.
Implicit Types x y z : A.
Implicit Types l k : list A.

(** ** Properties of the [reshape] function *)
Lemma reshape_length szs l : length (reshape szs l) = length szs.
Proof. revert l. by induction szs; intros; f_equal'. Qed.
Lemma join_reshape szs l :
  sum_list szs = length l  mjoin (reshape szs l) = l.
Proof.
  revert l. induction szs as [|sz szs IH]; simpl; intros l Hl; [by destruct l|].
  by rewrite IH, take_drop by (rewrite drop_length; lia).
Qed.
Lemma sum_list_replicate n m : sum_list (replicate m n) = m * n.