Newer
Older
(* Copyright (c) 2012-2013, 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 numbers base decidable 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.
(** * Definitions *)
(** 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 : nat) (l : list A) {struct l} : option A :=
Robbert Krebbers
committed
| x :: l => match i with 0 => Some x | S i => @lookup _ _ _ go i l 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} (f : A → A) : AlterD nat A (list A) f :=
fix go (i : nat) (l : list A) {struct l} :=
Robbert Krebbers
committed
| x :: l => match i with 0 => f x :: l | S i => x :: @alter _ _ _ f go i 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 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. *)
Robbert Krebbers
committed
Instance list_insert {A} : Insert nat A (list A) := λ i x, alter (λ _, x) i.
Robbert Krebbers
committed
(** 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]) [].
(** 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 :=
match l with
| [] => []
| x :: l =>
Robbert Krebbers
committed
if decide (P x)
then x :: @filter _ _ (@go) _ _ l
else @filter _ _ (@go) _ _ l
end.
(** 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 [].
Fixpoint last' {A} (x : A) (l : list A) : A :=
Robbert Krebbers
committed
match l with [] => x | x :: l => last' x l end.
Definition last {A} (l : list A) : option A :=
Robbert Krebbers
committed
match l with [] => None | x :: l => Some (last' x 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
(** 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) :=
Robbert Krebbers
committed
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) :=
Robbert Krebbers
committed
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 :=
Robbert Krebbers
committed
match ls with [] => [] | l :: ls => l ++ @mjoin _ go _ ls 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 ifoldr {A B} (f : nat → B → A → A) (a : nat → A) :
nat → list B → A := fix go n l :=
match l with [] => a n | b :: l => f n b (go (S n) l) end.
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 {_ _} _ _ _ _ _.
(** Zipping lists. *)
Definition zip_with {A B C} (f : A → B → C) : list A → list B → list C :=
fix go l1 l2 :=
Robbert Krebbers
committed
match l1, l2 with x1 :: l1, x2 :: l2 => f x1 x2 :: go l1 l2 | _ , _ => [] end.
Notation zip := (zip_with pair).
(** 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.
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 snd_map (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) := snd $ fst $ max_prefix_of l1 l2.
Definition strip_suffix (l1 l2 : list A) := snd $ fst $ max_suffix_of l1 l2.
End prefix_suffix_ops.
(** A list [l1] is a sub list 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
206
207
208
209
210
211
212
213
214
215
216
217
218
219
220
221
222
223
224
225
226
227
228
229
230
231
232
233
| sublist_skip x l1 l2 : sublist l1 l2 → sublist (x :: l1) (x :: l2)
| sublist_insert 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] without 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_insert 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
end.
End contains_dec_help.
(** The [same_length] view allows convenient induction over two lists with the
same length. *)
Inductive same_length {A B} : list A → list B → Prop :=
| same_length_nil : same_length [] []
Robbert Krebbers
committed
239
240
241
242
243
244
245
246
247
248
249
250
251
252
253
254
255
256
257
258
259
260
261
262
263
264
265
266
267
268
269
270
271
272
273
274
275
276
277
278
279
280
281
282
283
284
285
286
287
288
| same_length_cons x1 x2 l1 l2 :
same_length l1 l2 → same_length (x1 :: l1) (x2 :: l2).
Infix "`same_length`" := same_length (at level 70) : C_scope.
(** 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
end.
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;
Robbert Krebbers
committed
repeat (simpl 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. *)
Robbert Krebbers
committed
Ltac simplify_list_equality :=
repeat match goal with
| _ => progress simplify_equality
| H : _ ++ _ = _ ++ _ |- _ => first
Robbert Krebbers
committed
[ apply app_inj_tail in H; destruct H
| apply app_inv_head in H | apply app_inv_tail in H ]
Robbert Krebbers
committed
destruct i; [change (Some x = Some y) in H | discriminate]
end;
try discriminate_list_equality.
Robbert Krebbers
committed
(** * General theorems *)
Section general_properties.
Robbert Krebbers
committed
Implicit Types x y z : A.
Implicit Types l k : list A.
Robbert Krebbers
committed
Global Instance: Injective2 (=) (=) (=) (@cons A).
Proof. by injection 1. Qed.
Global Instance: ∀ x, Injective (=) (=) (x ::).
Robbert Krebbers
committed
Global Instance: ∀ l, Injective (=) (=) (:: l).
Robbert Krebbers
committed
Global Instance: ∀ k, Injective (=) (=) (k ++).
Proof. intros ???. apply app_inv_head. Qed.
Robbert Krebbers
committed
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.
Robbert Krebbers
committed
Lemma app_nil l1 l2 : l1 ++ l2 = [] ↔ l1 = [] ∧ l2 = [].
Proof. split. apply app_eq_nil. by intros [??]; subst. Qed.
Lemma app_singleton l1 l2 x :
l1 ++ l2 = [x] ↔ l1 = [] ∧ l2 = [x] ∨ l1 = [x] ∧ l2 = [].
Proof. split. apply app_eq_unit. by intros [[??]|[??]]; subst. Qed.
Lemma cons_middle x l1 l2 : l1 ++ x :: l2 = l1 ++ [x] ++ l2.
Proof. done. Qed.
Lemma app_inj l1 k1 l2 k2 :
length l1 = length k1 → l1 ++ l2 = k1 ++ k2 → l1 = k1 ∧ l2 = k2.
Proof. revert k1. induction l1; intros [|??]; naive_solver. Qed.
Robbert Krebbers
committed
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).
Robbert Krebbers
committed
* f_equal; [by injection (H 0) |]. apply IHl1. intro. apply (H (S _)).
Robbert Krebbers
committed
Lemma list_eq_nil l : (∀ i, l !! i = None) → l = nil.
Robbert Krebbers
committed
Global Instance list_eq_dec {dec : ∀ x y, Decision (x = y)} : ∀ l k,
Robbert Krebbers
committed
Definition list_singleton_dec l : { x | l = [x] } + { length l ≠ 1 }.
Proof. by refine match l with [x] => inleft (x↾_) | _ => inright _ end. Defined.
Robbert Krebbers
committed
Lemma nil_or_length_pos l : l = [] ∨ length l ≠ 0.
Proof. destruct l; simpl; auto with lia. Qed.
Robbert Krebbers
committed
Lemma nil_length 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.
Robbert Krebbers
committed
Lemma lookup_lt_length l i : is_Some (l !! i) ↔ i < length l.
revert i. induction l.
* split; by inversion 1.
* intros [|?]; simpl.
+ split; eauto with arith.
+ by rewrite <-NPeano.Nat.succ_lt_mono.
Robbert Krebbers
committed
Lemma lookup_lt_length_1 l i : is_Some (l !! i) → i < length l.
Robbert Krebbers
committed
Lemma lookup_lt_length_alt l i x : l !! i = Some x → i < length l.
Proof. intros Hl. by rewrite <-lookup_lt_length, Hl. Qed.
Robbert Krebbers
committed
Lemma lookup_lt_length_2 l i : i < length l → is_Some (l !! i).
Proof. apply lookup_lt_length. Qed.
Robbert Krebbers
committed
Lemma lookup_ge_length l i : l !! i = None ↔ length l ≤ i.
Proof. rewrite eq_None_not_Some, lookup_lt_length. lia. Qed.
Robbert Krebbers
committed
Lemma lookup_ge_length_1 l i : l !! i = None → length l ≤ i.
Proof. by rewrite lookup_ge_length. Qed.
Robbert Krebbers
committed
Lemma lookup_ge_length_2 l i : length l ≤ i → l !! i = None.
Proof. by rewrite lookup_ge_length. Qed.
Robbert Krebbers
committed
Lemma list_eq_length_eq l1 l2 :
Robbert Krebbers
committed
(∀ i x y, l1 !! i = Some x → l2 !! i = Some y → x = y) → l1 = l2.
Proof.
intros Hlength Hlookup. apply list_eq. intros i.
destruct (l2 !! i) as [x|] eqn:E.
Robbert Krebbers
committed
* feed inversion (lookup_lt_length_2 l1 i) as [y]; [|eauto with f_equal].
pose proof (lookup_lt_length_alt l2 i x E). lia.
* rewrite lookup_ge_length in E |- *. lia.
Qed.
Robbert Krebbers
committed
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.
Robbert Krebbers
committed
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_length_alt. Qed.
Robbert Krebbers
committed
Lemma lookup_app_r l1 l2 i : (l1 ++ l2) !! (length l1 + i) = l2 !! i.
Robbert Krebbers
committed
revert i. induction l1; intros [|i]; simpl in *; simplify_equality; auto.
Robbert Krebbers
committed
Lemma lookup_app_r_alt l1 l2 i :
length l1 ≤ i → (l1 ++ l2) !! i = l2 !! (i - length l1).
Proof.
intros. assert (i = length l1 + (i - length l1)) as Hi by lia.
rewrite Hi at 1. by apply lookup_app_r.
Qed.
Robbert Krebbers
committed
Lemma lookup_app_r_Some l1 l2 i x :
l2 !! i = Some x → (l1 ++ l2) !! (length l1 + i) = Some x.
Robbert Krebbers
committed
Lemma lookup_app_r_Some_alt l1 l2 i x :
length l1 ≤ i → l2 !! (i - length l1) = Some x → (l1 ++ l2) !! i = Some x.
Proof. intro. by rewrite lookup_app_r_alt. Qed.
Robbert Krebbers
committed
Lemma lookup_app_inv l1 l2 i x :
(l1 ++ l2) !! i = Some x → l1 !! i = Some x ∨ l2 !! (i - length l1) = Some x.
Robbert Krebbers
committed
revert i. induction l1; intros [|i] ?; simpl in *; simplify_equality; auto.
Robbert Krebbers
committed
Lemma list_lookup_middle l1 l2 x : (l1 ++ x :: l2) !! length l1 = Some x.
Proof. by induction l1; simpl. Qed.
Robbert Krebbers
committed
Lemma alter_length f l i : length (alter f i l) = length l.
Proof. revert i. induction l; intros [|?]; simpl; auto with lia. Qed.
Robbert Krebbers
committed
Lemma insert_length l i x : length (<[i:=x]>l) = length l.
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.
Robbert Krebbers
committed
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 [|i] [|j] ?; try done. apply (IHl i). congruence.
Qed.
Robbert Krebbers
committed
Lemma list_lookup_insert l i x : i < length l → <[i:=x]>l !! i = Some x.
Robbert Krebbers
committed
intros Hi. unfold insert, list_insert. rewrite list_lookup_alter.
by feed inversion (lookup_lt_length_2 l i).
Qed.
Robbert Krebbers
committed
Lemma list_lookup_insert_ne l i j x :
i ≠ j → <[i:=x]>l !! j = l !! j.
Proof. apply list_lookup_alter_ne. 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.
Proof.
intros Hl Hi.
destruct i; destruct l as [|x0 [|x1 l]]; simpl in *; simplify_equality.
* by exists 1 x1.
* by exists 0 x0.
Qed.
Robbert Krebbers
committed
Lemma alter_app_l f l1 l2 i :
i < length l1 → alter f i (l1 ++ l2) = alter f i l1 ++ l2.
Robbert Krebbers
committed
revert i. induction l1; intros [|?] ?; simpl in *; f_equal; auto with lia.
Robbert Krebbers
committed
Lemma alter_app_r f l1 l2 i :
alter f (length l1 + i) (l1 ++ l2) = l1 ++ alter f i l2.
Robbert Krebbers
committed
Proof. revert i. induction l1; intros [|?]; simpl in *; 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.
Robbert Krebbers
committed
Lemma insert_app_l l1 l2 i x :
i < length l1 → <[i:=x]>(l1 ++ l2) = <[i:=x]>l1 ++ l2.
Robbert Krebbers
committed
Lemma insert_app_r l1 l2 i x : <[length l1+i:=x]>(l1 ++ l2) = l1 ++ <[i:=x]>l2.
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. apply alter_app_r_alt. Qed.
Robbert Krebbers
committed
Lemma insert_consecutive_length l i k :
length (insert_consecutive i k l) = length l.
Proof. revert i. by induction k; intros; simpl; rewrite ?insert_length. Qed.
Robbert Krebbers
committed
Robbert Krebbers
committed
Lemma delete_middle l1 l2 x : delete (length l1) (l1 ++ x :: l2) = l1 ++ l2.
Proof. induction l1; simpl; f_equal; auto. 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.
Robbert Krebbers
committed
Lemma elem_of_cons l x y : x ∈ y :: l ↔ x = y ∨ x ∈ l.
Proof.
split.
* inversion 1; subst. by left. by right.
* intros [?|?]; subst. by left. by right.
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.
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.
Robbert Krebbers
committed
Lemma elem_of_list_split l x : x ∈ l → ∃ l1 l2, l = l1 ++ x :: l2.
Proof.
induction 1 as [x l|x y l ? [l1 [l2 ?]]].
* by eexists [], l.
* subst. by exists (y :: l1) l2.
Qed.
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.
Proof.
revert i. induction l; intros [|i] ?;
simpl; simplify_equality; constructor; eauto.
Qed.
Robbert Krebbers
committed
548
549
550
551
552
553
554
555
556
557
558
559
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
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.
(** ** 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 list_difference_nodup 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_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 list_intersection_nodup 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 [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.
Robbert Krebbers
committed
Lemma NoDup_Permutation l k : NoDup l → NoDup k → (∀ x, x ∈ l ↔ x ∈ k) → l ≡ₚ k.
Proof.
intros Hl. revert k. induction Hl as [|x l Hin ? IH].
Robbert Krebbers
committed
* intros k _ Hk. rewrite (elem_of_nil_inv k); [done |].
intros x. rewrite <-Hk, elem_of_nil. intros [].
Robbert Krebbers
committed
* intros k Hk Hlk. destruct (elem_of_list_split k x) as [l1 [l2 ?]]; subst.
{ rewrite <-Hlk. by constructor. }
rewrite <-Permutation_middle, NoDup_cons in Hk.
Robbert Krebbers
committed
destruct Hk as [??]. apply Permutation_cons_app, IH; [done |].
Robbert Krebbers
committed
rewrite <-Permutation_middle, !elem_of_cons in Hlk. naive_solver.
Robbert Krebbers
committed
Section no_dup_dec.
Context `{!∀ x y, Decision (x = y)}.
Robbert Krebbers
committed
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 remove_dups_nodup 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.
(** ** 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 filter_nodup 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 [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.
(** ** Properties of the [take] function *)
Robbert Krebbers
committed
Definition take_drop := @firstn_skipn A.
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; simpl; f_equal; auto. Qed.
Robbert Krebbers
committed
Lemma take_app_alt l k n : n = length l → take n (l ++ k) = l.
Proof. intros Hn. by rewrite Hn, take_app. Qed.
Robbert Krebbers
committed
Lemma take_app_le l k n : n ≤ length l → take n (l ++ k) = take n l.
Robbert Krebbers
committed
revert n. induction l; intros [|?] ?; simpl in *; f_equal; auto with lia.
Robbert Krebbers
committed
Lemma take_app_ge l k n :
length l ≤ n → take n (l ++ k) = l ++ take (n - length l) k.
Robbert Krebbers
committed
revert n. induction l; intros [|?] ?; simpl in *; f_equal; auto with lia.
Robbert Krebbers
committed
Lemma take_ge l n : length l ≤ n → take n l = l.
Robbert Krebbers
committed
revert n. induction l; intros [|?] ?; simpl in *; f_equal; auto with lia.
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 [|?] [|?]; simpl; f_equal; auto. Qed.
Robbert Krebbers
committed
Lemma take_idempotent 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 [|?]; simpl; f_equal; done. Qed.
Robbert Krebbers
committed
Lemma take_length_alt l n : n ≤ length l → length (take n l) = n.
Proof. rewrite take_length. apply Min.min_l. Qed.
Robbert Krebbers
committed
Lemma lookup_take l n i : i < n → take n l !! i = l !! i.
Proof.
revert n i. induction l; intros [|n] i ?; trivial.
* auto with lia.
* destruct i; simpl; auto with arith.
Qed.
Robbert Krebbers
committed
Lemma lookup_take_ge l n i : n ≤ i → take n l !! i = None.
Robbert Krebbers
committed
revert n i. induction l; intros [|?] [|?] ?; simpl; auto with lia.
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. apply take_alter. Qed.
(** ** Properties of the [drop] function *)
Robbert Krebbers
committed
Lemma drop_nil n : drop n (@nil A) = [].
Robbert Krebbers
committed
Lemma drop_app l k : drop (length l) (l ++ k) = k.
Proof. induction l; simpl; f_equal; auto. Qed.
Robbert Krebbers
committed
Lemma drop_app_alt l k n : n = length l → drop n (l ++ k) = k.
Proof. intros Hn. by rewrite Hn, drop_app. Qed.
Robbert Krebbers
committed
Lemma drop_length l n : length (drop n l) = length l - n.
Proof. revert n. by induction l; intros [|i]; simpl; f_equal. Qed.
Lemma drop_all l : drop (length l) l = [].
Robbert Krebbers
committed
Lemma drop_all_alt l n : n = length l → drop n l = [].
Proof. intros. subst. by rewrite drop_all. 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. apply drop_alter. 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 [|?]; simpl; auto using f_equal. Qed.
(** ** Properties of the [replicate] function *)
Robbert Krebbers
committed
Lemma replicate_length n x : length (replicate n x) = n.
Robbert Krebbers
committed
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.
Robbert Krebbers
committed
Proof. revert i. induction n; intros [|?]; naive_solver auto with lia. Qed.
Lemma replicate_S n x : replicate (S n) x = x :: replicate n x.
Proof. done. Qed.
Robbert Krebbers
committed
Lemma replicate_plus n m x :
replicate (n + m) x = replicate n x ++ replicate m x.
Proof. induction n; simpl; f_equal; auto. Qed.
Robbert Krebbers
committed
Lemma take_replicate n m x : take n (replicate m x) = replicate (min n m) x.
Proof. revert m. by induction n; intros [|?]; simpl; f_equal. Qed.
Robbert Krebbers
committed
Lemma take_replicate_plus n m x : take n (replicate (n + m) x) = replicate n x.
Proof. by rewrite take_replicate, min_l by lia. Qed.
Robbert Krebbers
committed
Lemma drop_replicate n m x : drop n (replicate m x) = replicate (m - n) x.
Proof. revert m. by induction n; intros [|?]; simpl; f_equal. Qed.
Robbert Krebbers
committed
Lemma drop_replicate_plus n m x : drop n (replicate (n + m) x) = replicate m x.
Proof. rewrite drop_replicate. f_equal. lia. Qed.
Robbert Krebbers
committed
Lemma reverse_replicate n x : reverse (replicate n x) = replicate n x.
Robbert Krebbers
committed
induction n as [|n IH]; [done|]. simpl. rewrite reverse_cons, IH.
change [x] with (replicate 1 x). by rewrite <-replicate_plus, plus_comm.
(** ** Properties of the [resize] function *)
Robbert Krebbers
committed
Lemma resize_spec l n x : resize n x l = take n l ++ replicate (n - length l) x.
Proof. revert n. induction l; intros [|?]; simpl; f_equal; auto. Qed.
Lemma resize_0 l x : resize 0 x l = [].
Robbert Krebbers
committed
Lemma resize_nil n x : resize n x [] = replicate n x.
Proof. rewrite resize_spec. rewrite take_nil. simpl. f_equal. lia. Qed.
Robbert Krebbers
committed
Lemma resize_ge l n x :
length l ≤ n → resize n x l = l ++ replicate (n - length l) x.
Proof. intros. by rewrite resize_spec, take_ge. Qed.
Robbert Krebbers
committed
Lemma resize_le l n x : n ≤ length l → resize n x l = take n l.
Proof.
intros. rewrite resize_spec, (proj2 (NPeano.Nat.sub_0_le _ _)) by done.
Robbert Krebbers
committed
simpl. by rewrite (right_id_L [] (++)).
Robbert Krebbers
committed
Lemma resize_all l x : resize (length l) x l = l.
Proof. intros. by rewrite resize_le, take_ge. Qed.
Robbert Krebbers
committed
Lemma resize_all_alt l n x : n = length l → resize n x l = l.
Proof. intros. subst. by rewrite resize_all. Qed.
Robbert Krebbers
committed
Lemma resize_plus l n m x :
resize (n + m) x l = resize n x l ++ resize m x (drop n l).
Proof.
Robbert Krebbers
committed
revert n m. induction l; intros [|?] [|?]; simpl; f_equal; auto.
* by rewrite plus_0_r, (right_id_L [] (++)).
Robbert Krebbers
committed
Lemma resize_plus_eq l n m x :
length l = n → resize (n + m) x l = l ++ replicate m x.
Robbert Krebbers
committed
intros. subst. by rewrite resize_plus, resize_all, drop_all, resize_nil.
Robbert Krebbers
committed
Lemma resize_app_le l1 l2 n x :
n ≤ length l1 → resize n x (l1 ++ l2) = resize n x l1.
Robbert Krebbers
committed
intros. by rewrite !resize_le, take_app_le by (rewrite ?app_length; lia).
Robbert Krebbers
committed
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
committed
intros. rewrite !resize_spec, take_app_ge, (associative_L (++)) by done.
do 2 f_equal. rewrite app_length. lia.
Qed.
Robbert Krebbers
committed
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 [|?]; simpl; f_equal; auto. Qed.
Robbert Krebbers
committed
Lemma resize_resize l n m x : n ≤ m → resize n x (resize m x l) = resize n x l.
Proof.
revert n m. induction l; simpl.
* intros. by rewrite !resize_nil, resize_replicate.
* intros [|?] [|?] ?; simpl; f_equal; auto with lia.
Qed.
Robbert Krebbers
committed
Lemma resize_idempotent l n x : resize n x (resize n x l) = resize n x l.
Robbert Krebbers
committed
Lemma resize_take_le l n m x : n ≤ m → resize n x (take m l) = resize n x l.
Robbert Krebbers
committed
revert n m. induction l; intros [|?] [|?] ?; simpl; f_equal; auto with lia.
Robbert Krebbers
committed
Lemma resize_take_eq l n x : resize n x (take n l) = resize n x l.
Robbert Krebbers
committed
Lemma take_resize l n m x : take n (resize m x l) = resize (min n m) x l.
Proof.
revert n m.
induction l; intros [|?] [|?]; simpl; f_equal; auto using take_replicate.
Qed.
Robbert Krebbers
committed
Lemma take_resize_le l n m x : n ≤ m → take n (resize m x l) = resize n x l.
Proof. intros. by rewrite take_resize, Min.min_l. Qed.
Robbert Krebbers
committed
Lemma take_resize_eq l n x : take n (resize n x l) = resize n x l.
Proof. intros. by rewrite take_resize, Min.min_l. Qed.
Robbert Krebbers
committed
Lemma take_length_resize l n x :
length l ≤ n → take (length l) (resize n x l) = l.
Proof. intros. by rewrite take_resize_le, resize_all. Qed.
Robbert Krebbers
committed
Lemma take_length_resize_alt l n m x :
m = length l → m ≤ n → take m (resize n x l) = l.
Proof. intros. subst. by apply take_length_resize. Qed.
Robbert Krebbers
committed
Lemma take_resize_plus l n m x : take n (resize (n + m) x l) = resize n x l.
Proof. by rewrite take_resize, min_l by lia. Qed.
Robbert Krebbers
committed
Lemma drop_resize_le l n m x :
n ≤ m → drop n (resize m x l) = resize (m - n) x (drop n l).
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
committed
Lemma drop_resize_plus l n m x :
drop n (resize (n + m) x l) = resize m x (drop n l).
Proof. rewrite drop_resize_le by lia. f_equal. lia. Qed.
Robbert Krebbers
committed
(** ** Properties of the [Permutation] predicate *)
Lemma Permutation_nil l : l ≡ₚ [] ↔ l = [].
Proof. split. by intro; apply Permutation_nil. by intro; subst. Qed.
Lemma Permutation_singleton l x : l ≡ₚ [x] ↔ l = [x].
Proof. split. by intro; apply Permutation_length_1_inv. by intro; subst. Qed.
Definition Permutation_skip := @perm_skip A.
Definition Permutation_swap := @perm_swap A.
Definition Permutation_singleton_inj := @Permutation_length_1 A.
Global Existing Instance Permutation_app'_Proper.
Global Instance: Proper ((≡ₚ) ==> (=)) (@length A).
Proof. induction 1; simpl; auto with lia. Qed.
Global Instance: Commutative (≡ₚ) (@app A).
Robbert Krebbers
committed
intros l1. induction l1 as [|x l1 IH]; intros l2; simpl.
* by rewrite (right_id_L [] (++)).
* rewrite Permutation_middle, IH. simpl. by rewrite Permutation_middle.
Robbert Krebbers
committed
Global Instance: ∀ x : A, Injective (≡ₚ) (≡ₚ) (x ::).
Proof. red. eauto using Permutation_cons_inv. Qed.
Global Instance: ∀ k : list A, Injective (≡ₚ) (≡ₚ) (k ++).
Robbert Krebbers
committed
red. induction k as [|x k IH]; intros l1 l2; simpl; auto.
intros. by apply IH, (injective (x ::)).
Robbert Krebbers
committed
Global Instance: ∀ k : list A, Injective (≡ₚ) (≡ₚ) (++ k).
Robbert Krebbers
committed
intros k l1 l2. rewrite !(commutative (++) _ k).
by apply (injective (k ++)).
Robbert Krebbers
committed
(** ** Properties of the [prefix_of] and [suffix_of] predicates *)
Global Instance: PreOrder (@prefix_of A).
Proof.
split.
Robbert Krebbers
committed
* intros ?. eexists []. by rewrite (right_id_L [] (++)).
* intros ??? [k1 ?] [k2 ?].
exists (k1 ++ k2). subst. by rewrite (associative_L (++)).
Qed.
Robbert Krebbers
committed
Lemma prefix_of_nil l : [] `prefix_of` l.
Proof. by exists l. Qed.
Lemma prefix_of_nil_not x l : ¬x :: l `prefix_of` [].
Proof. by intros [k E]. Qed.
Lemma prefix_of_cons x l1 l2 : l1 `prefix_of` l2 → x :: l1 `prefix_of` x :: l2.
Proof. intros [k E]. exists k. by subst. Qed.
Lemma prefix_of_cons_alt x y l1 l2 :
x = y → l1 `prefix_of` l2 → x :: l1 `prefix_of` y :: l2.
Proof. intro. subst. apply prefix_of_cons. Qed.
Lemma prefix_of_cons_inv_1 x y l1 l2 : x :: l1 `prefix_of` y :: l2 → x = y.
Proof. intros [k E]. by injection E. Qed.
Lemma prefix_of_cons_inv_2 x y l1 l2 :
x :: l1 `prefix_of` y :: l2 → l1 `prefix_of` l2.
Proof. intros [k E]. exists k. by injection E. Qed.
Lemma prefix_of_app k l1 l2 : l1 `prefix_of` l2 → k ++ l1 `prefix_of` k ++ l2.
Proof. intros [k' ?]. subst. exists k'. by rewrite (associative_L (++)). Qed.
Lemma prefix_of_app_alt k1 k2 l1 l2 :
k1 = k2 → l1 `prefix_of` l2 → k1 ++ l1 `prefix_of` k2 ++ l2.
Proof. intro. subst. apply prefix_of_app. Qed.
Lemma prefix_of_app_l l1 l2 l3 : l1 ++ l3 `prefix_of` l2 → l1 `prefix_of` l2.
Robbert Krebbers
committed
intros [k ?]. red. exists (l3 ++ k). subst. by rewrite <-(associative_L (++)).
Robbert Krebbers
committed
Lemma prefix_of_app_r l1 l2 l3 : l1 `prefix_of` l2 → l1 `prefix_of` l2 ++ l3.
Robbert Krebbers
committed
intros [k ?]. exists (k ++ l3). subst. by rewrite (associative_L (++)).
Qed.
Robbert Krebbers
committed
Lemma prefix_of_length l1 l2 : l1 `prefix_of` l2 → length l1 ≤ length l2.
Proof. intros [??]. subst. rewrite app_length. lia. Qed.
Lemma prefix_of_snoc_not l x : ¬l ++ [x] `prefix_of` l.
Proof. intros [??]. discriminate_list_equality. Qed.
Global Instance: PreOrder (@suffix_of A).
Proof.
split.
Robbert Krebbers
committed
* intros ?. by eexists [].
* intros ??? [k1 ?] [k2 ?].
exists (k2 ++ k1). subst. by rewrite (associative_L (++)).
Qed.
Robbert Krebbers
committed
Global Instance prefix_of_dec `{∀ x y, Decision (x = y)} : ∀ l1 l2,
Decision (l1 `prefix_of` l2) := fix go l1 l2 :=
match l1, l2 return { l1 `prefix_of` l2 } + { ¬l1 `prefix_of` l2 } with