Newer
Older
Robbert Krebbers
committed
Proof.
split.
- rewrite submseteq_sublist_r. intros (l'&E&Hl').
Robbert Krebbers
committed
rewrite sublist_app_r in Hl'. destruct Hl' as (l1&l2&?&?&?); subst.
exists l1, l2. eauto using sublist_submseteq.
- intros (?&?&E&?&?). rewrite E. eauto using submseteq_app.
Robbert Krebbers
committed
Qed.
Lemma submseteq_app_l l1 l2 k :
l1 ++ l2 ⊆+ k ↔ ∃ k1 k2, k ≡ₚ k1 ++ k2 ∧ l1 ⊆+ k1 ∧ l2 ⊆+ k2.
Robbert Krebbers
committed
Proof.
split.
- rewrite submseteq_sublist_l. intros (l'&Hl'&E).
Robbert Krebbers
committed
rewrite sublist_app_l in Hl'. destruct Hl' as (k1&k2&?&?&?); subst.
exists k1, k2. split. done. eauto using sublist_submseteq.
- intros (?&?&E&?&?). rewrite E. eauto using submseteq_app.
Robbert Krebbers
committed
Qed.
Lemma submseteq_app_inv_l l1 l2 k : k ++ l1 ⊆+ k ++ l2 → l1 ⊆+ l2.
Robbert Krebbers
committed
Proof.
induction k as [|y k IH]; simpl; [done |]. rewrite submseteq_cons_l.
intros (?&E%(inj _)&?). apply IH. by rewrite E.
Robbert Krebbers
committed
Qed.
Lemma submseteq_app_inv_r l1 l2 k : l1 ++ k ⊆+ l2 ++ k → l1 ⊆+ l2.
Robbert Krebbers
committed
Proof.
revert l1 l2. induction k as [|y k IH]; intros l1 l2.
{ by rewrite !(right_id_L [] (++)). }
intros. feed pose proof (IH (l1 ++ [y]) (l2 ++ [y])) as Hl12.
{ by rewrite <-!(assoc_L (++)). }
rewrite submseteq_app_l in Hl12. destruct Hl12 as (k1&k2&E1&?&Hk2).
rewrite submseteq_cons_l in Hk2. destruct Hk2 as (k2'&E2&?).
rewrite E2, (Permutation_cons_append k2'), (assoc_L (++)) in E1.
apply Permutation_app_inv_r in E1. rewrite E1. eauto using submseteq_inserts_r.
Robbert Krebbers
committed
Qed.
Lemma submseteq_cons_middle x l k1 k2 : l ⊆+ k1 ++ k2 → x :: l ⊆+ k1 ++ x :: k2.
Proof. rewrite <-Permutation_middle. by apply submseteq_skip. Qed.
Lemma submseteq_app_middle l1 l2 k1 k2 :
l2 ⊆+ k1 ++ k2 → l1 ++ l2 ⊆+ k1 ++ l1 ++ k2.
Robbert Krebbers
committed
Proof.
rewrite !(assoc (++)), (comm (++) k1 l1), <-(assoc_L (++)).
Robbert Krebbers
committed
Qed.
Lemma submseteq_middle l k1 k2 : l ⊆+ k1 ++ l ++ k2.
Proof. by apply submseteq_inserts_l, submseteq_inserts_r. Qed.
Robbert Krebbers
committed
Lemma Permutation_alt l1 l2 : l1 ≡ₚ l2 ↔ length l1 = length l2 ∧ l1 ⊆+ l2.
Robbert Krebbers
committed
Proof.
split.
- intros [??]; auto using submseteq_Permutation_length_eq.
Robbert Krebbers
committed
Qed.
Lemma NoDup_submseteq l k : NoDup l → (∀ x, x ∈ l → x ∈ k) → l ⊆+ k.
Proof.
intros Hl. revert k. induction Hl as [|x l Hx ? IH].
{ intros k Hk. by apply submseteq_nil_l. }
intros k Hlk. destruct (elem_of_list_split k x) as (l1&l2&?); subst.
{ apply Hlk. by constructor. }
rewrite <-Permutation_middle. apply submseteq_skip, IH.
intros y Hy. rewrite elem_of_app.
specialize (Hlk y). rewrite elem_of_app, !elem_of_cons in Hlk.
by destruct Hlk as [?|[?|?]]; subst; eauto.
Qed.
Lemma NoDup_Permutation l k : NoDup l → NoDup k → (∀ x, x ∈ l ↔ x ∈ k) → l ≡ₚ k.
Proof.
intros. apply (anti_symm submseteq); apply NoDup_submseteq; naive_solver.
Context `{!EqDecision A}.
Robbert Krebbers
committed
Lemma list_remove_Permutation l1 l2 k1 x :
l1 ≡ₚ l2 → list_remove x l1 = Some k1 →
∃ k2, list_remove x l2 = Some k2 ∧ k1 ≡ₚ k2.
Proof.
intros Hl. revert k1. induction Hl
as [|y l1 l2 ? IH|y1 y2 l|l1 l2 l3 ? IH1 ? IH2]; simpl; intros k1 Hk1.
- case_decide; simplify_eq; eauto.
destruct (list_remove x l1) as [l|] eqn:?; simplify_eq.
destruct (IH l) as (?&?&?); simplify_option_eq; eauto.
- simplify_option_eq; eauto using Permutation_swap.
- destruct (IH1 k1) as (k2&?&?); trivial.
Robbert Krebbers
committed
destruct (IH2 k2) as (k3&?&?); trivial.
exists k3. split; eauto. by trans k2.
Robbert Krebbers
committed
Qed.
Lemma list_remove_Some l k x : list_remove x l = Some k → l ≡ₚ x :: k.
Proof.
revert k. induction l as [|y l IH]; simpl; intros k ?; [done |].
simplify_option_eq; auto. by rewrite Permutation_swap, <-IH.
Robbert Krebbers
committed
Qed.
Lemma list_remove_Some_inv l k x :
l ≡ₚ x :: k → ∃ k', list_remove x l = Some k' ∧ k ≡ₚ k'.
Proof.
intros. destruct (list_remove_Permutation (x :: k) l k x) as (k'&?&?).
- done.
- simpl; by case_decide.
- by exists k'.
Robbert Krebbers
committed
Qed.
Lemma list_remove_list_submseteq l1 l2 :
l1 ⊆+ l2 ↔ is_Some (list_remove_list l1 l2).
Robbert Krebbers
committed
Proof.
- revert l2. induction l1 as [|x l1 IH]; simpl.
Robbert Krebbers
committed
{ intros l2 _. by exists l2. }
intros l2. rewrite submseteq_cons_l. intros (k&Hk&?).
Robbert Krebbers
committed
destruct (list_remove_Some_inv l2 k x) as (k2&?&Hk2); trivial.
simplify_option_eq. apply IH. by rewrite <-Hk2.
- intros [k Hk]. revert l2 k Hk.
Robbert Krebbers
committed
induction l1 as [|x l1 IH]; simpl; intros l2 k.
destruct (list_remove x l2) as [k'|] eqn:?; intros; simplify_eq.
rewrite submseteq_cons_l. eauto using list_remove_Some.
Robbert Krebbers
committed
Qed.
Robbert Krebbers
committed
Global Instance submseteq_dec : RelDecision (submseteq : relation (list A)).
Robbert Krebbers
committed
Proof.
Robbert Krebbers
committed
refine (λ l1 l2, cast_if (decide (is_Some (list_remove_list l1 l2))));
abstract (rewrite list_remove_list_submseteq; tauto).
Robbert Krebbers
committed
Defined.
Robbert Krebbers
committed
Global Instance Permutation_dec : RelDecision (Permutation : relation (list A)).
Robbert Krebbers
committed
Proof.
Robbert Krebbers
committed
refine (λ l1 l2, cast_if_and
(decide (length l1 = length l2)) (decide (l1 ⊆+ l2)));
Robbert Krebbers
committed
abstract (rewrite Permutation_alt; tauto).
Defined.
(** ** Properties of [included] *)
Global Instance list_subseteq_po : PreOrder (@subseteq (list A) _).
Proof. split; firstorder. Qed.
Lemma list_subseteq_nil l : [] ⊆ l.
Proof. intros x. by rewrite elem_of_nil. Qed.
Robbert Krebbers
committed
(** ** Properties of the [Forall] and [Exists] predicate *)
Lemma Forall_Exists_dec (P Q : A → Prop) (dec : ∀ x, {P x} + {Q x}) :
∀ l, {Forall P l} + {Exists Q l}.
Proof.
refine (
fix go l :=
match l return {Forall P l} + {Exists Q l} with
| [] => left _
| x :: l => cast_if_and (dec x) (go l)
end); clear go; intuition.
Defined.
Robbert Krebbers
committed
Definition Forall_nil_2 := @Forall_nil A.
Definition Forall_cons_2 := @Forall_cons A.
Global Instance Forall_proper:
Proper (pointwise_relation _ (↔) ==> (=) ==> (↔)) (@Forall A).
Proof. split; subst; induction 1; constructor; by firstorder auto. Qed.
Global Instance Exists_proper:
Proper (pointwise_relation _ (↔) ==> (=) ==> (↔)) (@Exists A).
Proof. split; subst; induction 1; constructor; by firstorder auto. Qed.
Robbert Krebbers
committed
Lemma Forall_forall l : Forall P l ↔ ∀ x, x ∈ l → P x.
split; [induction 1; inversion 1; subst; auto|].
intros Hin; induction l as [|x l IH]; constructor; [apply Hin; constructor|].
apply IH. intros ??. apply Hin. by constructor.
Lemma Forall_nil : Forall P [] ↔ True.
Proof. done. Qed.
Lemma Forall_cons_1 x l : Forall P (x :: l) → P x ∧ Forall P l.
Proof. by inversion 1. Qed.
Lemma Forall_cons x l : Forall P (x :: l) ↔ P x ∧ Forall P l.
Proof. split. by inversion 1. intros [??]. by constructor. Qed.
Lemma Forall_singleton x : Forall P [x] ↔ P x.
Proof. rewrite Forall_cons, Forall_nil; tauto. Qed.
Lemma Forall_app_2 l1 l2 : Forall P l1 → Forall P l2 → Forall P (l1 ++ l2).
Proof. induction 1; simpl; auto. Qed.
Lemma Forall_app l1 l2 : Forall P (l1 ++ l2) ↔ Forall P l1 ∧ Forall P l2.
Proof.
split; [induction l1; inversion 1; intuition|].
intros [??]; auto using Forall_app_2.
Qed.
Lemma Forall_true l : (∀ x, P x) → Forall P l.
Proof. intros ?. induction l; auto. Defined.
Lemma Forall_impl (Q : A → Prop) l :
Forall P l → (∀ x, P x → Q x) → Forall Q l.
Proof. intros H ?. induction H; auto. Defined.
Lemma Forall_iff l (Q : A → Prop) :
Robbert Krebbers
committed
(∀ x, P x ↔ Q x) → Forall P l ↔ Forall Q l.
Proof. intros H. apply Forall_proper. red; apply H. done. Qed.
Lemma Forall_not l : length l ≠ 0 → Forall (not ∘ P) l → ¬Forall P l.
Proof. by destruct 2; inversion 1. Qed.
Lemma Forall_and {Q} l : Forall (λ x, P x ∧ Q x) l ↔ Forall P l ∧ Forall Q l.
Proof.
split; [induction 1; constructor; naive_solver|].
intros [Hl Hl']; revert Hl'; induction Hl; inversion_clear 1; auto.
Qed.
Lemma Forall_and_l {Q} l : Forall (λ x, P x ∧ Q x) l → Forall P l.
Proof. rewrite Forall_and; tauto. Qed.
Lemma Forall_and_r {Q} l : Forall (λ x, P x ∧ Q x) l → Forall Q l.
Proof. rewrite Forall_and; tauto. Qed.
Lemma Forall_delete l i : Forall P l → Forall P (delete i l).
Robbert Krebbers
committed
Proof. intros H. revert i. by induction H; intros [|i]; try constructor. Qed.
Lemma Forall_lookup l : Forall P l ↔ ∀ i x, l !! i = Some x → P x.
Robbert Krebbers
committed
rewrite Forall_forall. setoid_rewrite elem_of_list_lookup. naive_solver.
Robbert Krebbers
committed
Lemma Forall_lookup_1 l i x : Forall P l → l !! i = Some x → P x.
Robbert Krebbers
committed
Lemma Forall_lookup_2 l : (∀ i x, l !! i = Some x → P x) → Forall P l.
Proof. by rewrite Forall_lookup. Qed.
Lemma Forall_tail l : Forall P l → Forall P (tail l).
Proof. destruct 1; simpl; auto. Qed.
Forall P l → (∀ x, l!!i = Some x → P x → P (f x)) → Forall P (alter f i l).
Robbert Krebbers
committed
intros Hl. revert i. induction Hl; simpl; intros [|i]; constructor; auto.
Lemma Forall_alter_inv f l i :
Forall P (alter f i l) → (∀ x, l!!i = Some x → P (f x) → P x) → Forall P l.
Proof.
revert i. induction l; intros [|?]; simpl;
inversion_clear 1; constructor; eauto.
Qed.
Lemma Forall_insert l i x : Forall P l → P x → Forall P (<[i:=x]>l).
Proof. rewrite list_insert_alter; auto using Forall_alter. Qed.
Lemma Forall_inserts l i k :
Forall P l → Forall P k → Forall P (list_inserts i k l).
Proof.
intros Hl Hk; revert i.
induction Hk; simpl; auto using Forall_insert.
Qed.
Robbert Krebbers
committed
Lemma Forall_replicate n x : P x → Forall P (replicate n x).
Proof. induction n; simpl; constructor; auto. Qed.
Lemma Forall_replicate_eq n (x : A) : Forall (x =) (replicate n x).

Ralf Jung
committed
Proof using -(P). induction n; simpl; constructor; auto. Qed.
Robbert Krebbers
committed
Lemma Forall_take n l : Forall P l → Forall P (take n l).
Proof. intros Hl. revert n. induction Hl; intros [|?]; simpl; auto. Qed.
Lemma Forall_drop n l : Forall P l → Forall P (drop n l).
Proof. intros Hl. revert n. induction Hl; intros [|?]; simpl; auto. Qed.
Lemma Forall_resize n x l : P x → Forall P l → Forall P (resize n x l).
Proof.
intros ? Hl. revert n.
induction Hl; intros [|?]; simpl; auto using Forall_replicate.
Qed.
Lemma Forall_resize_inv n x l :
length l ≤ n → Forall P (resize n x l) → Forall P l.
Proof. intros ?. rewrite resize_ge, Forall_app by done. by intros []. Qed.
Lemma Forall_sublist_lookup l i n k :
sublist_lookup i n l = Some k → Forall P l → Forall P k.
Proof.
unfold sublist_lookup. intros; simplify_option_eq.
auto using Forall_take, Forall_drop.
Qed.
Lemma Forall_sublist_alter f l i n k :
Forall P l → sublist_lookup i n l = Some k → Forall P (f k) →
Forall P (sublist_alter f i n l).
unfold sublist_alter, sublist_lookup. intros; simplify_option_eq.
Lemma Forall_sublist_alter_inv f l i n k :
sublist_lookup i n l = Some k →
Forall P (sublist_alter f i n l) → Forall P (f k).
Proof.
unfold sublist_alter, sublist_lookup. intros ?; simplify_option_eq.
Lemma Forall_reshape l szs : Forall P l → Forall (Forall P) (reshape szs l).
Proof.
revert l. induction szs; simpl; auto using Forall_take, Forall_drop.
Lemma Forall_rev_ind (Q : list A → Prop) :
Q [] → (∀ x l, P x → Forall P l → Q l → Q (l ++ [x])) →
∀ l, Forall P l → Q l.
Proof.
intros ?? l. induction l using rev_ind; auto.
rewrite Forall_app, Forall_singleton; intros [??]; auto.
Qed.
Robbert Krebbers
committed
Lemma Exists_exists l : Exists P l ↔ ∃ x, x ∈ l ∧ P x.
- induction 1 as [x|y ?? [x [??]]]; exists x; by repeat constructor.
- intros [x [Hin ?]]. induction l; [by destruct (not_elem_of_nil x)|].
inversion Hin; subst. by left. right; auto.
Qed.
Lemma Exists_inv x l : Exists P (x :: l) → P x ∨ Exists P l.
Proof. inversion 1; intuition trivial. Qed.
Lemma Exists_app l1 l2 : Exists P (l1 ++ l2) ↔ Exists P l1 ∨ Exists P l2.
Proof.
split.
- induction l1; inversion 1; intuition.
- intros [H|H]; [induction H | induction l1]; simpl; intuition.
Lemma Exists_impl (Q : A → Prop) l :
Exists P l → (∀ x, P x → Q x) → Exists Q l.
Proof. intros H ?. induction H; auto. Defined.
Robbert Krebbers
committed
Lemma Exists_not_Forall l : Exists (not ∘ P) l → ¬Forall P l.
Proof. induction 1; inversion_clear 1; contradiction. Qed.
Lemma Forall_not_Exists l : Forall (not ∘ P) l → ¬Exists P l.
Proof. induction 1; inversion_clear 1; contradiction. Qed.
Lemma Forall_list_difference `{!EqDecision A} l k :
Forall P l → Forall P (list_difference l k).
Proof.
rewrite !Forall_forall.
intros ? x; rewrite elem_of_list_difference; naive_solver.
Qed.
Lemma Forall_list_union `{!EqDecision A} l k :
Forall P l → Forall P k → Forall P (list_union l k).
Proof. intros. apply Forall_app; auto using Forall_list_difference. Qed.
Lemma Forall_list_intersection `{!EqDecision A} l k :
Forall P l → Forall P (list_intersection l k).
Proof.
rewrite !Forall_forall.
intros ? x; rewrite elem_of_list_intersection; naive_solver.
Qed.
Context {dec : ∀ x, Decision (P x)}.
Lemma not_Forall_Exists l : ¬Forall P l → Exists (not ∘ P) l.
Proof. intro. by destruct (Forall_Exists_dec P (not ∘ P) dec l). Qed.
Lemma not_Exists_Forall l : ¬Exists P l → Forall (not ∘ P) l.
by destruct (Forall_Exists_dec (not ∘ P) P
(λ x : A, swap_if (decide (P x))) l).
Qed.
Global Instance Forall_dec l : Decision (Forall P l) :=
match Forall_Exists_dec P (not ∘ P) dec l with
| left H => left H
| right H => right (Exists_not_Forall _ H)
end.
Global Instance Exists_dec l : Decision (Exists P l) :=
match Forall_Exists_dec (not ∘ P) P (λ x, swap_if (decide (P x))) l with
| left H => right (Forall_not_Exists _ H)
| right H => left H
End Forall_Exists.
Lemma replicate_as_Forall (x : A) n l :
replicate n x = l ↔ length l = n ∧ Forall (x =) l.
Proof. rewrite replicate_as_elem_of, Forall_forall. naive_solver. Qed.
Lemma replicate_as_Forall_2 (x : A) n l :
length l = n → Forall (x =) l → replicate n x = l.
Proof. by rewrite replicate_as_Forall. Qed.
Robbert Krebbers
committed
Lemma Forall_swap {A B} (Q : A → B → Prop) l1 l2 :
Forall (λ y, Forall (Q y) l1) l2 ↔ Forall (λ x, Forall (flip Q x) l2) l1.
Proof. repeat setoid_rewrite Forall_forall. simpl. split; eauto. Qed.
Lemma Forall_seq (P : nat → Prop) i n :
Forall P (seq i n) ↔ ∀ j, i ≤ j < i + n → P j.
Proof.
rewrite Forall_lookup. split.
- intros H j [??]. apply (H (j - i)).
rewrite lookup_seq; auto with f_equal lia.
- intros H j x Hj. apply lookup_seq_inv in Hj.
destruct Hj; subst. auto with lia.
Qed.
Robbert Krebbers
committed
Robbert Krebbers
committed
Lemma Forall2_same_length {A B} (l : list A) (k : list B) :
Forall2 (λ _ _, True) l k ↔ length l = length k.
Proof.
split; [by induction 1; f_equal/=|].
revert k. induction l; intros [|??] ?; simplify_eq/=; auto.
Qed.
(** ** Properties of the [Forall2] predicate *)
Robbert Krebbers
committed
Lemma Forall_Forall2 {A} (Q : A → A → Prop) l :
Forall (λ x, Q x x) l → Forall2 Q l l.
Proof. induction 1; constructor; auto. Qed.
Robbert Krebbers
committed
Lemma Forall2_forall `{Inhabited A} B C (Q : A → B → C → Prop) l k :
Forall2 (λ x y, ∀ z, Q z x y) l k ↔ ∀ z, Forall2 (Q z) l k.
Proof.
split; [induction 1; constructor; auto|].
intros Hlk. induction (Hlk inhabitant) as [|x y l k _ _ IH]; constructor.
- intros z. by feed inversion (Hlk z).
- apply IH. intros z. by feed inversion (Hlk z).
Qed.
Lemma Forall2_Forall {A} P (l1 l2 : list A) :
Forall2 P l1 l2 → Forall (curry P) (zip l1 l2).
Proof. induction 1; constructor; auto. Qed.
Implicit Types x : A.
Implicit Types y : B.
Implicit Types l : list A.
Implicit Types k : list B.
Lemma Forall2_length l k : Forall2 P l k → length l = length k.
Proof. by induction 1; f_equal/=. Qed.
Lemma Forall2_length_l l k n : Forall2 P l k → length l = n → length k = n.
Proof. intros ? <-; symmetry. by apply Forall2_length. Qed.
Lemma Forall2_length_r l k n : Forall2 P l k → length k = n → length l = n.
Proof. intros ? <-. by apply Forall2_length. Qed.
Lemma Forall2_true l k : (∀ x y, P x y) → length l = length k → Forall2 P l k.
Proof. rewrite <-Forall2_same_length. induction 2; naive_solver. Qed.
Lemma Forall2_flip l k : Forall2 (flip P) k l ↔ Forall2 P l k.
Proof. split; induction 1; constructor; auto. Qed.
Lemma Forall2_transitive {C} (Q : B → C → Prop) (R : A → C → Prop) l k lC :
(∀ x y z, P x y → Q y z → R x z) →
Forall2 P l k → Forall2 Q k lC → Forall2 R l lC.
Proof. intros ? Hl. revert lC. induction Hl; inversion_clear 1; eauto. Qed.
Lemma Forall2_impl (Q : A → B → Prop) l k :
Forall2 P l k → (∀ x y, P x y → Q x y) → Forall2 Q l k.
Proof. intros H ?. induction H; auto. Defined.
Lemma Forall2_unique l k1 k2 :
Forall2 P l k1 → Forall2 P l k2 →
(∀ x y1 y2, P x y1 → P x y2 → y1 = y2) → k1 = k2.
Proof.
intros H. revert k2. induction H; inversion_clear 1; intros; f_equal; eauto.
Qed.
Lemma Forall_Forall2_l l k :
length l = length k → Forall (λ x, ∀ y, P x y) l → Forall2 P l k.
Proof. rewrite <-Forall2_same_length. induction 1; inversion 1; auto. Qed.
Lemma Forall_Forall2_r l k :
length l = length k → Forall (λ y, ∀ x, P x y) k → Forall2 P l k.
Proof. rewrite <-Forall2_same_length. induction 1; inversion 1; auto. Qed.
Lemma Forall2_Forall_l (Q : A → Prop) l k :
Forall2 P l k → Forall (λ y, ∀ x, P x y → Q x) k → Forall Q l.
Proof. induction 1; inversion_clear 1; eauto. Qed.
Lemma Forall2_Forall_r (Q : B → Prop) l k :
Forall2 P l k → Forall (λ x, ∀ y, P x y → Q y) l → Forall Q k.
Proof. induction 1; inversion_clear 1; eauto. Qed.
Robbert Krebbers
committed
Lemma Forall2_nil_inv_l k : Forall2 P [] k → k = [].
Lemma Forall2_nil_inv_r l : Forall2 P l [] → l = [].
Lemma Forall2_cons_inv x l y k :
Forall2 P (x :: l) (y :: k) → P x y ∧ Forall2 P l k.
Lemma Forall2_cons_inv_l x l k :
Forall2 P (x :: l) k → ∃ y k', P x y ∧ Forall2 P l k' ∧ k = y :: k'.
Lemma Forall2_cons_inv_r l k y :
Forall2 P l (y :: k) → ∃ x l', P x y ∧ Forall2 P l' k ∧ l = x :: l'.
Lemma Forall2_cons_nil_inv x l : Forall2 P (x :: l) [] → False.
Lemma Forall2_nil_cons_inv y k : Forall2 P [] (y :: k) → False.
Lemma Forall2_app_l l1 l2 k :
Forall2 P l1 (take (length l1) k) → Forall2 P l2 (drop (length l1) k) →
Forall2 P (l1 ++ l2) k.
Proof. intros. rewrite <-(take_drop (length l1) k). by apply Forall2_app. Qed.
Lemma Forall2_app_r l k1 k2 :
Forall2 P (take (length k1) l) k1 → Forall2 P (drop (length k1) l) k2 →
Forall2 P l (k1 ++ k2).
Proof. intros. rewrite <-(take_drop (length k1) l). by apply Forall2_app. Qed.
Lemma Forall2_app_inv l1 l2 k1 k2 :
length l1 = length k1 →
Robbert Krebbers
committed
Forall2 P (l1 ++ l2) (k1 ++ k2) → Forall2 P l1 k1 ∧ Forall2 P l2 k2.
Proof.
rewrite <-Forall2_same_length. induction 1; inversion 1; naive_solver.
Qed.
Lemma Forall2_app_inv_l l1 l2 k :
Forall2 P (l1 ++ l2) k ↔
∃ k1 k2, Forall2 P l1 k1 ∧ Forall2 P l2 k2 ∧ k = k1 ++ k2.
Proof.
split; [|intros (?&?&?&?&->); by apply Forall2_app].
revert k. induction l1; inversion 1; naive_solver.
Qed.
Lemma Forall2_app_inv_r l k1 k2 :
Forall2 P l (k1 ++ k2) ↔
∃ l1 l2, Forall2 P l1 k1 ∧ Forall2 P l2 k2 ∧ l = l1 ++ l2.
Proof.
split; [|intros (?&?&?&?&->); by apply Forall2_app].
revert l. induction k1; inversion 1; naive_solver.
Qed.
Lemma Forall2_tail l k : Forall2 P l k → Forall2 P (tail l) (tail k).
Proof. destruct 1; simpl; auto. Qed.
Lemma Forall2_take l k n : Forall2 P l k → Forall2 P (take n l) (take n k).
Proof. intros Hl. revert n. induction Hl; intros [|?]; simpl; auto. Qed.
Lemma Forall2_drop l k n : Forall2 P l k → Forall2 P (drop n l) (drop n k).
Proof. intros Hl. revert n. induction Hl; intros [|?]; simpl; auto. Qed.
Lemma Forall2_lookup l k :
Forall2 P l k ↔ ∀ i, option_Forall2 P (l !! i) (k !! i).
split; [induction 1; intros [|?]; simpl; try constructor; eauto|].
revert k. induction l as [|x l IH]; intros [| y k] H.
- done.
- feed inversion (H 0).
- feed inversion (H 0).
- constructor; [by feed inversion (H 0)|]. apply (IH _ $ λ i, H (S i)).
Lemma Forall2_lookup_lr l k i x y :
Forall2 P l k → l !! i = Some x → k !! i = Some y → P x y.
Proof. rewrite Forall2_lookup; intros H; destruct (H i); naive_solver. Qed.
Lemma Forall2_lookup_l l k i x :
Forall2 P l k → l !! i = Some x → ∃ y, k !! i = Some y ∧ P x y.
Proof. rewrite Forall2_lookup; intros H; destruct (H i); naive_solver. Qed.
Lemma Forall2_lookup_r l k i y :
Forall2 P l k → k !! i = Some y → ∃ x, l !! i = Some x ∧ P x y.
Proof. rewrite Forall2_lookup; intros H; destruct (H i); naive_solver. Qed.
Lemma Forall2_same_length_lookup_2 l k :
length l = length k →
(∀ i x y, l !! i = Some x → k !! i = Some y → P x y) → Forall2 P l k.
Proof.
rewrite <-Forall2_same_length. intros Hl Hlookup.
induction Hl as [|?????? IH]; constructor; [by apply (Hlookup 0)|].
apply IH. apply (λ i, Hlookup (S i)).
Qed.
Lemma Forall2_same_length_lookup l k :
Forall2 P l k ↔
length l = length k ∧
(∀ i x y, l !! i = Some x → k !! i = Some y → P x y).
Proof.
naive_solver eauto using Forall2_length,
Forall2_lookup_lr, Forall2_same_length_lookup_2.
Lemma Forall2_alter_l f l k i :
Forall2 P l k →
(∀ x y, l !! i = Some x → k !! i = Some y → P x y → P (f x) y) →
Forall2 P (alter f i l) k.
Proof. intros Hl. revert i. induction Hl; intros [|]; constructor; auto. Qed.
Lemma Forall2_alter_r f l k i :
Forall2 P l k →
(∀ x y, l !! i = Some x → k !! i = Some y → P x y → P x (f y)) →
Forall2 P l (alter f i k).
Proof. intros Hl. revert i. induction Hl; intros [|]; constructor; auto. Qed.
Lemma Forall2_alter f g l k i :
Forall2 P l k →
(∀ x y, l !! i = Some x → k !! i = Some y → P x y → P (f x) (g y)) →
Forall2 P (alter f i l) (alter g i k).
Proof. intros Hl. revert i. induction Hl; intros [|]; constructor; auto. Qed.
Lemma Forall2_insert l k x y i :
Forall2 P l k → P x y → Forall2 P (<[i:=x]> l) (<[i:=y]> k).
Proof. intros Hl. revert i. induction Hl; intros [|]; constructor; auto. Qed.
Lemma Forall2_inserts l k l' k' i :
Forall2 P l k → Forall2 P l' k' →
Forall2 P (list_inserts i l' l) (list_inserts i k' k).
Proof. intros ? H. revert i. induction H; eauto using Forall2_insert. Qed.
Lemma Forall2_delete l k i :
Forall2 P l k → Forall2 P (delete i l) (delete i k).
Proof. intros Hl. revert i. induction Hl; intros [|]; simpl; intuition. Qed.
Lemma Forall2_option_list mx my :
option_Forall2 P mx my → Forall2 P (option_list mx) (option_list my).
Proof. destruct 1; by constructor. Qed.
Lemma Forall2_filter Q1 Q2 `{∀ x, Decision (Q1 x), ∀ y, Decision (Q2 y)} l k:
(∀ x y, P x y → Q1 x ↔ Q2 y) →
Forall2 P l k → Forall2 P (filter Q1 l) (filter Q2 k).
Proof.
intros HP; induction 1 as [|x y l k]; unfold filter; simpl; auto.
simplify_option_eq by (by rewrite <-(HP x y)); repeat constructor; auto.
Qed.
Lemma Forall2_replicate_l k n x :
length k = n → Forall (P x) k → Forall2 P (replicate n x) k.
Proof. intros <-. induction 1; simpl; auto. Qed.
Lemma Forall2_replicate_r l n y :
length l = n → Forall (flip P y) l → Forall2 P l (replicate n y).
Proof. intros <-. induction 1; simpl; auto. Qed.
Lemma Forall2_replicate n x y :
P x y → Forall2 P (replicate n x) (replicate n y).
Proof. induction n; simpl; constructor; auto. Qed.
Lemma Forall2_reverse l k : Forall2 P l k → Forall2 P (reverse l) (reverse k).
Proof.
induction 1; rewrite ?reverse_nil, ?reverse_cons; eauto using Forall2_app.
Qed.
Lemma Forall2_last l k : Forall2 P l k → option_Forall2 P (last l) (last k).
Proof. induction 1 as [|????? []]; simpl; repeat constructor; auto. Qed.
Lemma Forall2_resize l k x y n :
P x y → Forall2 P l k → Forall2 P (resize n x l) (resize n y k).
Proof.
intros. rewrite !resize_spec, (Forall2_length l k) by done.
auto using Forall2_app, Forall2_take, Forall2_replicate.
Qed.
Lemma Forall2_resize_l l k x y n m :
P x y → Forall (flip P y) l →
Forall2 P (resize n x l) k → Forall2 P (resize m x l) (resize m y k).
intros. destruct (decide (m ≤ n)).
{ rewrite <-(resize_resize l m n) by done. by apply Forall2_resize. }
intros. assert (n = length k); subst.
{ by rewrite <-(Forall2_length (resize n x l) k), resize_length. }
rewrite (le_plus_minus (length k) m), !resize_plus,
resize_all, drop_all, resize_nil by lia.
auto using Forall2_app, Forall2_replicate_r,
Forall_resize, Forall_drop, resize_length.
Lemma Forall2_resize_r l k x y n m :
P x y → Forall (P x) k →
Forall2 P l (resize n y k) → Forall2 P (resize m x l) (resize m y k).
intros. destruct (decide (m ≤ n)).
{ rewrite <-(resize_resize k m n) by done. by apply Forall2_resize. }
assert (n = length l); subst.
{ by rewrite (Forall2_length l (resize n y k)), resize_length. }
rewrite (le_plus_minus (length l) m), !resize_plus,
resize_all, drop_all, resize_nil by lia.
auto using Forall2_app, Forall2_replicate_l,
Forall_resize, Forall_drop, resize_length.
Lemma Forall2_resize_r_flip l k x y n m :
P x y → Forall (P x) k →
length k = m → Forall2 P l (resize n y k) → Forall2 P (resize m x l) k.
Proof.
intros ?? <- ?. rewrite <-(resize_all k y) at 2.
apply Forall2_resize_r with n; auto using Forall_true.
Qed.
Lemma Forall2_sublist_lookup_l l k n i l' :
Forall2 P l k → sublist_lookup n i l = Some l' →
∃ k', sublist_lookup n i k = Some k' ∧ Forall2 P l' k'.
unfold sublist_lookup. intros Hlk Hl.
exists (take i (drop n k)); simplify_option_eq.
- auto using Forall2_take, Forall2_drop.
- apply Forall2_length in Hlk; lia.
Lemma Forall2_sublist_lookup_r l k n i k' :
Forall2 P l k → sublist_lookup n i k = Some k' →
∃ l', sublist_lookup n i l = Some l' ∧ Forall2 P l' k'.
erewrite Forall2_length by eauto; intros; simplify_option_eq.
Lemma Forall2_sublist_alter f g l k i n l' k' :
Forall2 P l k → sublist_lookup i n l = Some l' →
sublist_lookup i n k = Some k' → Forall2 P (f l') (g k') →
Forall2 P (sublist_alter f i n l) (sublist_alter g i n k).
erewrite Forall2_length by eauto; intros; simplify_option_eq.
auto using Forall2_app, Forall2_drop, Forall2_take.
Lemma Forall2_sublist_alter_l f l k i n l' k' :
Forall2 P l k → sublist_lookup i n l = Some l' →
sublist_lookup i n k = Some k' → Forall2 P (f l') k' →
Forall2 P (sublist_alter f i n l) k.
Proof.
erewrite <-Forall2_length by eauto; intros; simplify_option_eq.
apply Forall2_app_l;
rewrite ?take_length_le by lia; auto using Forall2_take.
apply Forall2_app_l; erewrite Forall2_length, take_length,
drop_length, <-Forall2_length, Min.min_l by eauto with lia; [done|].
rewrite drop_drop; auto using Forall2_drop.
Global Instance Forall2_dec `{dec : ∀ x y, Decision (P x y)} :
Robbert Krebbers
committed
RelDecision (Forall2 P).
fix go l k : Decision (Forall2 P l k) :=
match l, k with
| x :: l, y :: k => cast_if_and (decide (P x y)) (go l k)
end); clear dec go; abstract first [by constructor | by inversion 1].
Global Instance: Reflexive R → Reflexive (Forall2 R).
Proof. intros ? l. induction l; by constructor. Qed.
Global Instance: Symmetric R → Symmetric (Forall2 R).
Proof. intros. induction 1; constructor; auto. Qed.
Global Instance: Transitive R → Transitive (Forall2 R).
Proof. intros ????. apply Forall2_transitive. by apply @transitivity. Qed.
Robbert Krebbers
committed
Global Instance: Equivalence R → Equivalence (Forall2 R).
Proof. split; apply _. Qed.
Global Instance: PreOrder R → PreOrder (Forall2 R).
Proof. split; apply _. Qed.
Global Instance: AntiSymm (=) R → AntiSymm (=) (Forall2 R).
Proof. induction 2; inversion_clear 1; f_equal; auto. Qed.
Robbert Krebbers
committed
Global Instance: Proper (R ==> Forall2 R ==> Forall2 R) (::).
Proof. by constructor. Qed.
Global Instance: Proper (Forall2 R ==> Forall2 R ==> Forall2 R) (++).
Proof. repeat intro. by apply Forall2_app. Qed.
Global Instance: Proper (Forall2 R ==> (=)) length.
Proof. repeat intro. eauto using Forall2_length. Qed.
Global Instance: Proper (Forall2 R ==> Forall2 R) tail.
Proof. repeat intro. eauto using Forall2_tail. Qed.
Robbert Krebbers
committed
Global Instance: Proper (Forall2 R ==> Forall2 R) (take n).
Proof. repeat intro. eauto using Forall2_take. Qed.
Global Instance: Proper (Forall2 R ==> Forall2 R) (drop n).
Proof. repeat intro. eauto using Forall2_drop. Qed.
2694
2695
2696
2697
2698
2699
2700
2701
2702
2703
2704
2705
2706
2707
2708
2709
2710
2711
2712
2713
2714
2715
2716
2717
2718
2719
Global Instance: Proper (Forall2 R ==> option_Forall2 R) (lookup i).
Proof. repeat intro. by apply Forall2_lookup. Qed.
Global Instance:
Proper (R ==> R) f → Proper (Forall2 R ==> Forall2 R) (alter f i).
Proof. repeat intro. eauto using Forall2_alter. Qed.
Global Instance: Proper (R ==> Forall2 R ==> Forall2 R) (insert i).
Proof. repeat intro. eauto using Forall2_insert. Qed.
Global Instance:
Proper (Forall2 R ==> Forall2 R ==> Forall2 R) (list_inserts i).
Proof. repeat intro. eauto using Forall2_inserts. Qed.
Global Instance: Proper (Forall2 R ==> Forall2 R) (delete i).
Proof. repeat intro. eauto using Forall2_delete. Qed.
Global Instance: Proper (option_Forall2 R ==> Forall2 R) option_list.
Proof. repeat intro. eauto using Forall2_option_list. Qed.
Global Instance: ∀ P `{∀ x, Decision (P x)},
Proper (R ==> iff) P → Proper (Forall2 R ==> Forall2 R) (filter P).
Proof. repeat intro; eauto using Forall2_filter. Qed.
Global Instance: Proper (R ==> Forall2 R) (replicate n).
Proof. repeat intro. eauto using Forall2_replicate. Qed.
Global Instance: Proper (Forall2 R ==> Forall2 R) reverse.
Proof. repeat intro. eauto using Forall2_reverse. Qed.
Global Instance: Proper (Forall2 R ==> option_Forall2 R) last.
Proof. repeat intro. eauto using Forall2_last. Qed.
Robbert Krebbers
committed
Global Instance: Proper (R ==> Forall2 R ==> Forall2 R) (resize n).
Proof. repeat intro. eauto using Forall2_resize. Qed.
Robbert Krebbers
committed
Section Forall3.
Context {A B C} (P : A → B → C → Prop).
Hint Extern 0 (Forall3 _ _ _ _) => constructor.
Forall3 P l1 k1 k1' → Forall3 P l2 k2 k2' →
Forall3 P (l1 ++ l2) (k1 ++ k2) (k1' ++ k2').
Proof. induction 1; simpl; auto. Qed.
Lemma Forall3_cons_inv_l x l k k' :
Forall3 P (x :: l) k k' → ∃ y k2 z k2',
k = y :: k2 ∧ k' = z :: k2' ∧ P x y z ∧ Forall3 P l k2 k2'.
Proof. inversion_clear 1; naive_solver. Qed.
Lemma Forall3_app_inv_l l1 l2 k k' :
Forall3 P (l1 ++ l2) k k' → ∃ k1 k2 k1' k2',
k = k1 ++ k2 ∧ k' = k1' ++ k2' ∧ Forall3 P l1 k1 k1' ∧ Forall3 P l2 k2 k2'.
Proof.
revert k k'. induction l1 as [|x l1 IH]; simpl; inversion_clear 1.
- by repeat eexists; eauto.
- by repeat eexists; eauto.
- edestruct IH as (?&?&?&?&?&?&?&?); eauto; naive_solver.
Qed.
Lemma Forall3_cons_inv_m l y k k' :
Forall3 P l (y :: k) k' → ∃ x l2 z k2',
l = x :: l2 ∧ k' = z :: k2' ∧ P x y z ∧ Forall3 P l2 k k2'.
Proof. inversion_clear 1; naive_solver. Qed.
Lemma Forall3_app_inv_m l k1 k2 k' :
Forall3 P l (k1 ++ k2) k' → ∃ l1 l2 k1' k2',
l = l1 ++ l2 ∧ k' = k1' ++ k2' ∧ Forall3 P l1 k1 k1' ∧ Forall3 P l2 k2 k2'.
Proof.
revert l k'. induction k1 as [|x k1 IH]; simpl; inversion_clear 1.
- by repeat eexists; eauto.
- by repeat eexists; eauto.
- edestruct IH as (?&?&?&?&?&?&?&?); eauto; naive_solver.
Qed.
Lemma Forall3_cons_inv_r l k z k' :
Forall3 P l k (z :: k') → ∃ x l2 y k2,
l = x :: l2 ∧ k = y :: k2 ∧ P x y z ∧ Forall3 P l2 k2 k'.
Proof. inversion_clear 1; naive_solver. Qed.
Lemma Forall3_app_inv_r l k k1' k2' :
Forall3 P l k (k1' ++ k2') → ∃ l1 l2 k1 k2,
l = l1 ++ l2 ∧ k = k1 ++ k2 ∧ Forall3 P l1 k1 k1' ∧ Forall3 P l2 k2 k2'.
revert l k. induction k1' as [|x k1' IH]; simpl; inversion_clear 1.
- by repeat eexists; eauto.
- by repeat eexists; eauto.
- edestruct IH as (?&?&?&?&?&?&?&?); eauto; naive_solver.
Lemma Forall3_impl (Q : A → B → C → Prop) l k k' :
Forall3 P l k k' → (∀ x y z, P x y z → Q x y z) → Forall3 Q l k k'.
Proof. intros Hl ?; induction Hl; auto. Defined.
Lemma Forall3_length_lm l k k' : Forall3 P l k k' → length l = length k.
Proof. by induction 1; f_equal/=. Qed.
Lemma Forall3_length_lr l k k' : Forall3 P l k k' → length l = length k'.
Proof. by induction 1; f_equal/=. Qed.
Lemma Forall3_lookup_lmr l k k' i x y z :
Forall3 P l k k' →
l !! i = Some x → k !! i = Some y → k' !! i = Some z → P x y z.
Proof.
intros H. revert i. induction H; intros [|?] ???; simplify_eq/=; eauto.
Lemma Forall3_lookup_l l k k' i x :
Forall3 P l k k' → l !! i = Some x →
∃ y z, k !! i = Some y ∧ k' !! i = Some z ∧ P x y z.
Proof.
intros H. revert i. induction H; intros [|?] ?; simplify_eq/=; eauto.
Lemma Forall3_lookup_m l k k' i y :
Forall3 P l k k' → k !! i = Some y →
∃ x z, l !! i = Some x ∧ k' !! i = Some z ∧ P x y z.
Proof.
intros H. revert i. induction H; intros [|?] ?; simplify_eq/=; eauto.
Lemma Forall3_lookup_r l k k' i z :
Forall3 P l k k' → k' !! i = Some z →
∃ x y, l !! i = Some x ∧ k !! i = Some y ∧ P x y z.
Proof.
intros H. revert i. induction H; intros [|?] ?; simplify_eq/=; eauto.
Lemma Forall3_alter_lm f g l k k' i :
Forall3 P l k k' →
(∀ x y z, l !! i = Some x → k !! i = Some y → k' !! i = Some z →
P x y z → P (f x) (g y) z) →
Proof. intros Hl. revert i. induction Hl; intros [|]; auto. Qed.
End Forall3.
(** Setoids *)
Section setoid.

Ralf Jung
committed
Context `{Equiv A}.
Implicit Types l k : list A.
Lemma equiv_Forall2 l k : l ≡ k ↔ Forall2 (≡) l k.
Proof. split; induction 1; constructor; auto. Qed.
Lemma list_equiv_lookup l k : l ≡ k ↔ ∀ i, l !! i ≡ k !! i.
Proof.
rewrite equiv_Forall2, Forall2_lookup.
by setoid_rewrite equiv_option_Forall2.
Qed.
Global Instance list_equivalence :
Equivalence ((≡) : relation A) → Equivalence ((≡) : relation (list A)).
Proof.
split.
- intros l. by apply equiv_Forall2.
- intros l k; rewrite !equiv_Forall2; by intros.
- intros l1 l2 l3; rewrite !equiv_Forall2; intros; by trans l2.
Qed.
Global Instance list_leibniz `{!LeibnizEquiv A} : LeibnizEquiv (list A).
Proof. induction 1; f_equal; fold_leibniz; auto. Qed.
Global Instance cons_proper : Proper ((≡) ==> (≡) ==> (≡)) (@cons A).
Proof. by constructor. Qed.
Global Instance app_proper : Proper ((≡) ==> (≡) ==> (≡)) (@app A).
Proof. induction 1; intros ???; simpl; try constructor; auto. Qed.
Global Instance length_proper : Proper ((≡) ==> (=)) (@length A).
Proof. induction 1; f_equal/=; auto. Qed.
Global Instance tail_proper : Proper ((≡) ==> (≡)) (@tail A).
Proof. destruct 1; try constructor; auto. Qed.
Global Instance take_proper n : Proper ((≡) ==> (≡)) (@take A n).
Proof. induction n; destruct 1; constructor; auto. Qed.
Global Instance drop_proper n : Proper ((≡) ==> (≡)) (@drop A n).
Proof. induction n; destruct 1; simpl; try constructor; auto. Qed.
Global Instance list_lookup_proper i :
Proper ((≡) ==> (≡)) (lookup (M:=list A) i).
Proof. induction i; destruct 1; simpl; try constructor; auto. Qed.
Global Instance list_alter_proper f i :
Proper ((≡) ==> (≡)) f → Proper ((≡) ==> (≡)) (alter (M:=list A) f i).
Proof. intros. induction i; destruct 1; constructor; eauto. Qed.
Global Instance list_insert_proper i :
Proper ((≡) ==> (≡) ==> (≡)) (insert (M:=list A) i).
Proof. intros ???; induction i; destruct 1; constructor; eauto. Qed.
Global Instance list_inserts_proper i :
Proper ((≡) ==> (≡) ==> (≡)) (@list_inserts A i).
intros k1 k2 Hk; revert i.
induction Hk; intros ????; simpl; try f_equiv; naive_solver.
Qed.
Global Instance list_delete_proper i :
Proper ((≡) ==> (≡)) (delete (M:=list A) i).
Proof. induction i; destruct 1; try constructor; eauto. Qed.
Global Instance option_list_proper : Proper ((≡) ==> (≡)) (@option_list A).
Proof. destruct 1; repeat constructor; auto. Qed.
Global Instance list_filter_proper P `{∀ x, Decision (P x)} :
Proper ((≡) ==> iff) P → Proper ((≡) ==> (≡)) (filter (B:=list A) P).
Proof. intros ???. rewrite !equiv_Forall2. by apply Forall2_filter. Qed.
Global Instance replicate_proper n : Proper ((≡) ==> (≡)) (@replicate A n).
Proof. induction n; constructor; auto. Qed.
Global Instance reverse_proper : Proper ((≡) ==> (≡)) (@reverse A).
Proof.
induction 1; rewrite ?reverse_cons; simpl; [constructor|].
apply app_proper; repeat constructor; auto.
Qed.
Global Instance last_proper : Proper ((≡) ==> (≡)) (@last A).
Proof. induction 1 as [|????? []]; simpl; repeat constructor; auto. Qed.
Global Instance resize_proper n : Proper ((≡) ==> (≡) ==> (≡)) (@resize A n).
Proof.
induction n; destruct 2; simpl; repeat (constructor || f_equiv); auto.
Qed.
(** * Properties of the monadic operations *)
Robbert Krebbers
committed
Lemma list_fmap_id {A} (l : list A) : id <$> l = l.
Proof. induction l; f_equal/=; auto. Qed.
Section fmap.
Context {A B : Type} (f : A → B).
Implicit Types l : list A.
Lemma list_fmap_compose {C} (g : B → C) l : g ∘ f <$> l = g <$> (f <$> l).
Proof. induction l; f_equal/=; auto. Qed.
Lemma list_fmap_ext (g : A → B) (l1 l2 : list A) :
(∀ x, f x = g x) → l1 = l2 → fmap f l1 = fmap g l2.
Proof. intros ? <-. induction l1; f_equal/=; auto. Qed.

Ralf Jung
committed
Lemma list_fmap_equiv_ext `{Equiv B} (g : A → B) l :
(∀ x, f x ≡ g x) → f <$> l ≡ g <$> l.
Proof. induction l; constructor; auto. Qed.
Global Instance: Inj (=) (=) f → Inj (=) (=) (fmap f).
Proof.
intros ? l1. induction l1 as [|x l1 IH]; [by intros [|??]|].
intros [|??]; intros; f_equal/=; simplify_eq; auto.
Qed.
Definition fmap_nil : f <$> [] = [] := eq_refl.
Definition fmap_cons x l : f <$> x :: l = f x :: (f <$> l) := eq_refl.
Lemma fmap_app l1 l2 : f <$> l1 ++ l2 = (f <$> l1) ++ (f <$> l2).
Proof. by induction l1; f_equal/=. Qed.
Robbert Krebbers
committed
Lemma fmap_nil_inv k : f <$> k = [] → k = [].
Proof. by destruct k. Qed.
Lemma fmap_cons_inv y l k :
Robbert Krebbers
committed
f <$> l = y :: k → ∃ x l', y = f x ∧ k = f <$> l' ∧ l = x :: l'.
Proof. intros. destruct l; simplify_eq/=; eauto. Qed.
Lemma fmap_app_inv l k1 k2 :
Robbert Krebbers
committed
f <$> l = k1 ++ k2 → ∃ l1 l2, k1 = f <$> l1 ∧ k2 = f <$> l2 ∧ l = l1 ++ l2.
Proof.
revert l. induction k1 as [|y k1 IH]; simpl; [intros l ?; by eexists [],l|].
intros [|x l] ?; simplify_eq/=.
destruct (IH l) as (l1&l2&->&->&->); [done|]. by exists (x :: l1), l2.
Qed.
Lemma fmap_length l : length (f <$> l) = length l.
Proof. by induction l; f_equal/=. Qed.
Lemma fmap_reverse l : f <$> reverse l = reverse (f <$> l).
induction l as [|?? IH]; csimpl; by rewrite ?reverse_cons, ?fmap_app, ?IH.
Lemma fmap_tail l : f <$> tail l = tail (f <$> l).
Proof. by destruct l. Qed.
Lemma fmap_last l : last (f <$> l) = f <$> last l.
Proof. induction l as [|? []]; simpl; auto. Qed.
Robbert Krebbers
committed
Lemma fmap_replicate n x : f <$> replicate n x = replicate n (f x).
Proof. by induction n; f_equal/=. Qed.
Lemma fmap_take n l : f <$> take n l = take n (f <$> l).
Proof. revert n. by induction l; intros [|?]; f_equal/=. Qed.
Lemma fmap_drop n l : f <$> drop n l = drop n (f <$> l).
Proof. revert n. by induction l; intros [|?]; f_equal/=. Qed.
Lemma fmap_resize n x l : f <$> resize n x l = resize n (f x) (f <$> l).
Proof.
revert n. induction l; intros [|?]; f_equal/=; auto using fmap_replicate.
Lemma const_fmap (l : list A) (y : B) :
(∀ x, f x = y) → f <$> l = replicate (length l) y.
Proof. intros; induction l; f_equal/=; auto. Qed.
Lemma list_lookup_fmap l i : (f <$> l) !! i = f <$> (l !! i).
Proof. revert i. induction l; intros [|n]; by try revert n. Qed.
Lemma list_lookup_fmap_inv l i x :
(f <$> l) !! i = Some x → ∃ y, x = f y ∧ l !! i = Some y.
intros Hi. rewrite list_lookup_fmap in Hi.
destruct (l !! i) eqn:?; simplify_eq/=; eauto.
Lemma list_alter_fmap (g : A → A) (h : B → B) l i :
Robbert Krebbers
committed
Forall (λ x, f (g x) = h (f x)) l → f <$> alter g i l = alter h i (f <$> l).
Proof. intros Hl. revert i. by induction Hl; intros [|i]; f_equal/=. Qed.
Lemma elem_of_list_fmap_1 l x : x ∈ l → f x ∈ f <$> l.
Proof. induction 1; csimpl; rewrite elem_of_cons; intuition. Qed.
Lemma elem_of_list_fmap_1_alt l x y : x ∈ l → y = f x → y ∈ f <$> l.
Proof. intros. subst. by apply elem_of_list_fmap_1. Qed.
Lemma elem_of_list_fmap_2 l x : x ∈ f <$> l → ∃ y, x = f y ∧ y ∈ l.
induction l as [|y l IH]; simpl; inversion_clear 1.
- exists y. split; [done | by left].
- destruct IH as [z [??]]. done. exists z. split; [done | by right].
Lemma elem_of_list_fmap l x : x ∈ f <$> l ↔ ∃ y, x = f y ∧ y ∈ l.
Proof.
naive_solver eauto using elem_of_list_fmap_1_alt, elem_of_list_fmap_2.
Qed.
Lemma NoDup_fmap_1 l : NoDup (f <$> l) → NoDup l.
Proof.
induction l; simpl; inversion_clear 1; constructor; auto.
rewrite elem_of_list_fmap in *. naive_solver.
Qed.
Lemma NoDup_fmap_2 `{!Inj (=) (=) f} l : NoDup l → NoDup (f <$> l).
Proof.
Robbert Krebbers
committed
induction 1; simpl; constructor; trivial. rewrite elem_of_list_fmap.
intros [y [Hxy ?]]. apply (inj f) in Hxy. by subst.
Qed.
Lemma NoDup_fmap `{!Inj (=) (=) f} l : NoDup (f <$> l) ↔ NoDup l.
Proof. split; auto using NoDup_fmap_1, NoDup_fmap_2. Qed.
Robbert Krebbers
committed
Global Instance fmap_sublist: Proper (sublist ==> sublist) (fmap f).
Proof. induction 1; simpl; econstructor; eauto. Qed.
Global Instance fmap_submseteq: Proper (submseteq ==> submseteq) (fmap f).
Robbert Krebbers
committed
Proof. induction 1; simpl; econstructor; eauto. Qed.
Global Instance fmap_Permutation: Proper ((≡ₚ) ==> (≡ₚ)) (fmap f).
Proof. induction 1; simpl; econstructor; eauto. Qed.
Lemma Forall_fmap_ext_1 (g : A → B) (l : list A) :
Forall (λ x, f x = g x) l → fmap f l = fmap g l.
Proof. by induction 1; f_equal/=. Qed.
Lemma Forall_fmap_ext (g : A → B) (l : list A) :
Forall (λ x, f x = g x) l ↔ fmap f l = fmap g l.