Newer
Older
1001
1002
1003
1004
1005
1006
1007
1008
1009
1010
1011
1012
1013
1014
1015
1016
1017
1018
1019
1020
1021
1022
1023
1024
1025
1026
1027
1028
revert l k. assert (∀ l i,
n ≠ 0 → (n | length l) → ¬n * i `div` n + n ≤ length l → length l ≤ i).
{ intros l i ? [j ->] Hjn. apply Nat.nlt_ge; contradict Hjn.
rewrite <-Nat.mul_succ_r, (Nat.mul_comm n).
apply Nat.mul_le_mono_r, Nat.le_succ_l, Nat.div_lt_upper_bound; lia. }
intros l k Hl Hk Hlookup. destruct (decide (n = 0)) as [->|].
{ by rewrite (nil_length_inv l),
(nil_length_inv k) by eauto using Nat.divide_0_l. }
apply list_eq; intros i. specialize (Hlookup (i `div` n)).
rewrite (Nat.mul_comm _ n) in Hlookup.
unfold sublist_lookup in *; simplify_option_equality;
[|by rewrite !lookup_ge_None_2 by auto].
apply (f_equal (!! i `mod` n)) in Hlookup.
by rewrite !lookup_take, !lookup_drop, <-!Nat.div_mod in Hlookup
by (auto using Nat.mod_upper_bound with lia).
Qed.
Lemma sublist_eq_same_length l k j n :
length l = j * n → length k = j * n →
(∀ i,i < j → sublist_lookup (i * n) n l = sublist_lookup (i * n) n k) → l = k.
Proof.
intros Hl Hk ?. destruct (decide (n = 0)) as [->|].
{ by rewrite (nil_length_inv l), (nil_length_inv k) by lia. }
apply sublist_eq with n; [by exists j|by exists j|].
intros i. destruct (decide (i < j)); [by auto|].
assert (∀ m, m = j * n → m < i * n + n).
{ intros ? ->. replace (i * n + n) with (S i * n) by lia.
apply Nat.mul_lt_mono_pos_r; lia. }
by rewrite !sublist_lookup_None by auto.
Lemma sublist_lookup_reshape l i n m :
0 < n → length l = m * n →
reshape (replicate m n) l !! i = sublist_lookup (i * n) n l.
intros Hn Hl. unfold sublist_lookup. apply option_eq; intros x; split.
* intros Hx. case_option_guard as Hi.
{ f_equal. clear Hi. revert i l Hl Hx.
induction m as [|m IH]; intros [|i] l ??; simplify_equality'; auto.
rewrite <-drop_drop. apply IH; rewrite ?drop_length; auto with lia. }
destruct Hi. rewrite Hl, <-Nat.mul_succ_l.
apply Nat.mul_le_mono_r, Nat.le_succ_l. apply lookup_lt_Some in Hx.
by rewrite reshape_length, replicate_length in Hx.
* intros Hx. case_option_guard as Hi; simplify_equality'.
revert i l Hl Hi. induction m as [|m IH]; [auto with lia|].
intros [|i] l ??; simpl; [done|]. rewrite <-drop_drop.
rewrite IH; rewrite ?drop_length; auto with lia.
1047
1048
1049
1050
1051
1052
1053
1054
1055
1056
1057
1058
1059
1060
1061
1062
1063
1064
1065
1066
1067
1068
1069
1070
1071
1072
1073
1074
1075
1076
1077
1078
1079
1080
1081
1082
1083
1084
1085
1086
1087
1088
1089
1090
1091
1092
1093
1094
1095
1096
1097
1098
1099
1100
1101
1102
1103
1104
1105
1106
1107
1108
Lemma sublist_lookup_compose l1 l2 l3 i n j m :
sublist_lookup i n l1 = Some l2 → sublist_lookup j m l2 = Some l3 →
sublist_lookup (i + j) m l1 = Some l3.
Proof.
unfold sublist_lookup; intros; simplify_option_equality;
repeat match goal with
| H : _ ≤ length _ |- _ => rewrite take_length, drop_length in H
end; rewrite <-?take_drop_commute, ?drop_drop, ?take_take, ?Min.min_l by lia;
auto with lia.
Qed.
Lemma sublist_alter_length f i n l :
(∀ k, sublist_lookup i n l = Some k → length (f k) = n) →
i + n ≤ length l → length (sublist_alter f i n l) = length l.
Proof.
unfold sublist_alter. intros Hk ?. rewrite !app_length, Hk, !take_length,
!drop_length by auto using sublist_lookup_Some; lia.
Qed.
Lemma sublist_lookup_alter f i n l :
(∀ k, sublist_lookup i n l = Some k → length (f k) = n) →
i + n ≤ length l →
sublist_lookup i n (sublist_alter f i n l) = f <$> sublist_lookup i n l.
Proof.
intros Hk ?. unfold sublist_lookup. rewrite sublist_alter_length by done.
case_option_guard; f_equal'; unfold sublist_alter.
rewrite drop_app_alt by (rewrite take_length; lia).
by rewrite take_app_alt by (rewrite Hk; eauto using sublist_lookup_Some).
Qed.
Lemma sublist_lookup_alter_ne f l i j n :
(∀ k, sublist_lookup j n l = Some k → length (f k) = n) →
j + n ≤ length l → i + n ≤ j ∨ j + n ≤ i →
sublist_lookup i n (sublist_alter f j n l) = sublist_lookup i n l.
Proof.
intros Hk Hij ?. unfold sublist_lookup. rewrite sublist_alter_length by done.
case_option_guard; f_equal'; unfold sublist_alter. apply list_eq; intros ii.
destruct (decide (ii < n)); [|by rewrite !lookup_take_ge by lia].
rewrite !lookup_take, !lookup_drop by done.
destruct (decide (i + ii < j)).
{ by rewrite lookup_app_l, lookup_take by (rewrite ?take_length; lia). }
rewrite lookup_app_minus_r by (rewrite take_length; lia).
rewrite take_length_le by lia. rewrite lookup_app_minus_r
by (rewrite Hk; eauto using sublist_lookup_Some; lia).
rewrite Hk, lookup_drop by eauto using sublist_lookup_Some. f_equal; lia.
Qed.
Lemma sublist_alter_all f l n : length l = n → sublist_alter f 0 n l = f l.
Proof.
intros <-. unfold sublist_alter; simpl.
by rewrite drop_all, (right_id_L [] (++)), take_ge.
Qed.
(** ** Properties of the [mask] function *)
Lemma mask_nil f βs : mask f βs (@nil A) = [].
Proof. by destruct βs. Qed.
Lemma mask_length f βs l : length (mask f βs l) = length l.
Proof. revert βs. induction l; intros [|??]; f_equal'; auto. Qed.
Lemma mask_true f l n : length l ≤ n → mask f (replicate n true) l = f <$> l.
Proof. revert n. induction l; intros [|?] ?; f_equal'; auto with lia. Qed.
Lemma mask_false f l n : mask f (replicate n false) l = l.
Proof. revert l. induction n; intros [|??]; f_equal'; auto. Qed.
Lemma mask_app f βs1 βs2 l :
mask f (βs1 ++ βs2) l
= mask f βs1 (take (length βs1) l) ++ mask f βs2 (drop (length βs1) l).
Proof. revert l. induction βs1;intros [|??]; f_equal'; auto using mask_nil. Qed.
Lemma take_mask f βs l n : take n (mask f βs l) = mask f (take n βs) (take n l).
Proof. revert n βs. induction l; intros [|?] [|[] ?]; f_equal'; auto. Qed.
Lemma drop_mask f βs l n : drop n (mask f βs l) = mask f (drop n βs) (drop n l).
Proof.
revert n βs. induction l; intros [|?] [|[] ?]; f_equal'; auto using mask_nil.
Qed.
Lemma sublist_lookup_mask f βs l i n :
sublist_lookup i n (mask f βs l)
= mask f (take n (drop i βs)) <$> sublist_lookup i n l.
Proof.
unfold sublist_lookup; rewrite mask_length; simplify_option_equality; auto.
by rewrite drop_mask, take_mask.
(** ** Properties of the [seq] function *)
Lemma fmap_seq j n : S <$> seq j n = seq (S j) n.
Proof. revert j. induction n; intros; f_equal'; auto. Qed.
Lemma lookup_seq j n i : i < n → seq j n !! i = Some (j + i).
Proof.
revert j i. induction n as [|n IH]; intros j [|i] ?; simpl; auto with lia.
rewrite IH; auto with lia.
Qed.
Lemma lookup_seq_ge j n i : n ≤ i → seq j n !! i = None.
Proof. revert j i. induction n; intros j [|i] ?; simpl; auto with lia. Qed.
Lemma lookup_seq_inv j n i j' : seq j n !! i = Some j' → j' = j + i ∧ i < n.
Proof.
destruct (le_lt_dec n i); [by rewrite lookup_seq_ge|].
rewrite lookup_seq by done. intuition congruence.
Robbert Krebbers
committed
(** ** Properties of the [Permutation] predicate *)
Lemma Permutation_nil l : l ≡ₚ [] ↔ l = [].
Proof. split. by intro; apply Permutation_nil. by intros ->. Qed.
Robbert Krebbers
committed
Lemma Permutation_singleton l x : l ≡ₚ [x] ↔ l = [x].
Proof. split. by intro; apply Permutation_length_1_inv. by intros ->. Qed.
Robbert Krebbers
committed
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).
intros k l1 l2. rewrite !(commutative (++) _ k). by apply (injective (k ++)).
Lemma replicate_Permutation n x l : replicate n x ≡ₚ l → replicate n x = l.
Proof.
intros Hl. apply replicate_as_elem_of. split.
* by rewrite <-Hl, replicate_length.
* intros y. rewrite <-Hl. by apply elem_of_replicate_inv.
Qed.
Lemma reverse_Permutation l : reverse l ≡ₚ l.
Proof.
induction l as [|x l IH]; [done|].
by rewrite reverse_cons, (commutative (++)), IH.
Qed.
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). by rewrite (associative_L (++)).
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 ?]. Qed.
Robbert Krebbers
committed
Lemma prefix_of_cons x l1 l2 : l1 `prefix_of` l2 → x :: l1 `prefix_of` x :: l2.
Proof. intros [k ->]. by exists k. Qed.
Robbert Krebbers
committed
Lemma prefix_of_cons_alt x y l1 l2 :
x = y → l1 `prefix_of` l2 → x :: l1 `prefix_of` y :: l2.
Proof. intros ->. apply prefix_of_cons. Qed.
Robbert Krebbers
committed
Lemma prefix_of_cons_inv_1 x y l1 l2 : x :: l1 `prefix_of` y :: l2 → x = y.
Proof. by intros [k ?]; simplify_equality'. Qed.
Robbert Krebbers
committed
Lemma prefix_of_cons_inv_2 x y l1 l2 :
x :: l1 `prefix_of` y :: l2 → l1 `prefix_of` l2.
Proof. intros [k ?]; simplify_equality. by exists k. Qed.
Robbert Krebbers
committed
Lemma prefix_of_app k l1 l2 : l1 `prefix_of` l2 → k ++ l1 `prefix_of` k ++ l2.
Proof. intros [k' ->]. exists k'. by rewrite (associative_L (++)). Qed.
Robbert Krebbers
committed
Lemma prefix_of_app_alt k1 k2 l1 l2 :
k1 = k2 → l1 `prefix_of` l2 → k1 ++ l1 `prefix_of` k2 ++ l2.
Proof. intros ->. apply prefix_of_app. Qed.
Robbert Krebbers
committed
Lemma prefix_of_app_l l1 l2 l3 : l1 ++ l3 `prefix_of` l2 → l1 `prefix_of` l2.
Proof. intros [k ->]. exists (l3 ++ k). by rewrite (associative_L (++)). Qed.
Robbert Krebbers
committed
Lemma prefix_of_app_r l1 l2 l3 : l1 `prefix_of` l2 → l1 `prefix_of` l2 ++ l3.
Proof. intros [k ->]. exists (k ++ l3). by rewrite (associative_L (++)). Qed.
Robbert Krebbers
committed
Lemma prefix_of_length l1 l2 : l1 `prefix_of` l2 → length l1 ≤ length l2.
Proof. intros [? ->]. rewrite app_length. lia. Qed.
Robbert Krebbers
committed
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). by rewrite (associative_L (++)).
Robbert Krebbers
committed
Global Instance prefix_of_dec `{∀ x y, Decision (x = y)} : ∀ l1 l2,
Decision (l1 `prefix_of` l2) := fix go l1 l2 :=
Robbert Krebbers
committed
match l1, l2 return { l1 `prefix_of` l2 } + { ¬l1 `prefix_of` l2 } with
| [], _ => left (prefix_of_nil _)
| _, [] => right (prefix_of_nil_not _ _)
| x :: l1, y :: l2 =>
match decide_rel (=) x y with
| left Exy =>
match go l1 l2 with
| left Hl1l2 => left (prefix_of_cons_alt _ _ _ _ Exy Hl1l2)
| right Hl1l2 => right (Hl1l2 ∘ prefix_of_cons_inv_2 _ _ _ _)
end
| right Exy => right (Exy ∘ prefix_of_cons_inv_1 _ _ _ _)
end
end.
Robbert Krebbers
committed
Section prefix_ops.
Context `{∀ x y, Decision (x = y)}.
Lemma max_prefix_of_fst l1 l2 :
l1 = (max_prefix_of l1 l2).2 ++ (max_prefix_of l1 l2).1.1.
Robbert Krebbers
committed
Proof.
revert l2. induction l1; intros [|??]; simpl;
repeat case_decide; f_equal'; auto.
Robbert Krebbers
committed
Qed.
Lemma max_prefix_of_fst_alt l1 l2 k1 k2 k3 :
max_prefix_of l1 l2 = (k1, k2, k3) → l1 = k3 ++ k1.
Proof.
intros. pose proof (max_prefix_of_fst l1 l2).
Robbert Krebbers
committed
by destruct (max_prefix_of l1 l2) as [[]?]; simplify_equality.
Qed.
Lemma max_prefix_of_fst_prefix l1 l2 : (max_prefix_of l1 l2).2 `prefix_of` l1.
Robbert Krebbers
committed
Proof. eexists. apply max_prefix_of_fst. Qed.
Lemma max_prefix_of_fst_prefix_alt l1 l2 k1 k2 k3 :
max_prefix_of l1 l2 = (k1, k2, k3) → k3 `prefix_of` l1.
Proof. eexists. eauto using max_prefix_of_fst_alt. Qed.
Lemma max_prefix_of_snd l1 l2 :
l2 = (max_prefix_of l1 l2).2 ++ (max_prefix_of l1 l2).1.2.
Robbert Krebbers
committed
Proof.
revert l2. induction l1; intros [|??]; simpl;
repeat case_decide; f_equal'; auto.
Robbert Krebbers
committed
Qed.
Lemma max_prefix_of_snd_alt l1 l2 k1 k2 k3 :
max_prefix_of l1 l2 = (k1, k2, k3) → l2 = k3 ++ k2.
Proof.
intro. pose proof (max_prefix_of_snd l1 l2).
by destruct (max_prefix_of l1 l2) as [[]?]; simplify_equality.
Qed.
Lemma max_prefix_of_snd_prefix l1 l2 : (max_prefix_of l1 l2).2 `prefix_of` l2.
Robbert Krebbers
committed
Proof. eexists. apply max_prefix_of_snd. Qed.
Lemma max_prefix_of_snd_prefix_alt l1 l2 k1 k2 k3 :
max_prefix_of l1 l2 = (k1,k2,k3) → k3 `prefix_of` l2.
Proof. eexists. eauto using max_prefix_of_snd_alt. Qed.
Lemma max_prefix_of_max l1 l2 k :
k `prefix_of` l1 → k `prefix_of` l2 → k `prefix_of` (max_prefix_of l1 l2).2.
Robbert Krebbers
committed
Proof.
intros [l1' ->] [l2' ->]. by induction k; simpl; repeat case_decide;
simpl; auto using prefix_of_nil, prefix_of_cons.
Robbert Krebbers
committed
Qed.
Lemma max_prefix_of_max_alt l1 l2 k1 k2 k3 k :
max_prefix_of l1 l2 = (k1,k2,k3) →
k `prefix_of` l1 → k `prefix_of` l2 → k `prefix_of` k3.
Proof.
intro. pose proof (max_prefix_of_max l1 l2 k).
by destruct (max_prefix_of l1 l2) as [[]?]; simplify_equality.
Qed.
Lemma max_prefix_of_max_snoc l1 l2 k1 k2 k3 x1 x2 :
max_prefix_of l1 l2 = (x1 :: k1, x2 :: k2, k3) → x1 ≠ x2.
Proof.
intros Hl ->. destruct (prefix_of_snoc_not k3 x2).
Robbert Krebbers
committed
1289
1290
1291
1292
1293
1294
1295
1296
1297
1298
1299
1300
1301
1302
1303
1304
1305
1306
1307
1308
1309
1310
1311
1312
1313
1314
eapply max_prefix_of_max_alt; eauto.
* rewrite (max_prefix_of_fst_alt _ _ _ _ _ Hl).
apply prefix_of_app, prefix_of_cons, prefix_of_nil.
* rewrite (max_prefix_of_snd_alt _ _ _ _ _ Hl).
apply prefix_of_app, prefix_of_cons, prefix_of_nil.
Qed.
End prefix_ops.
Lemma prefix_suffix_reverse l1 l2 :
l1 `prefix_of` l2 ↔ reverse l1 `suffix_of` reverse l2.
Proof.
split; intros [k E]; exists (reverse k).
* by rewrite E, reverse_app.
* by rewrite <-(reverse_involutive l2), E, reverse_app, reverse_involutive.
Qed.
Lemma suffix_prefix_reverse l1 l2 :
l1 `suffix_of` l2 ↔ reverse l1 `prefix_of` reverse l2.
Proof. by rewrite prefix_suffix_reverse, !reverse_involutive. Qed.
Lemma suffix_of_nil l : [] `suffix_of` l.
Proof. exists l. by rewrite (right_id_L [] (++)). Qed.
Lemma suffix_of_nil_inv l : l `suffix_of` [] → l = [].
Proof. by intros [[|?] ?]; simplify_list_equality. Qed.
Lemma suffix_of_cons_nil_inv x l : ¬x :: l `suffix_of` [].
Proof. by intros [[] ?]. Qed.
Lemma suffix_of_snoc l1 l2 x :
l1 `suffix_of` l2 → l1 ++ [x] `suffix_of` l2 ++ [x].
Proof. intros [k ->]. exists k. by rewrite (associative_L (++)). Qed.
Robbert Krebbers
committed
Lemma suffix_of_snoc_alt x y l1 l2 :
x = y → l1 `suffix_of` l2 → l1 ++ [x] `suffix_of` l2 ++ [y].
Proof. intros ->. apply suffix_of_snoc. Qed.
Robbert Krebbers
committed
Lemma suffix_of_app l1 l2 k : l1 `suffix_of` l2 → l1 ++ k `suffix_of` l2 ++ k.
Proof. intros [k' ->]. exists k'. by rewrite (associative_L (++)). Qed.
Robbert Krebbers
committed
Lemma suffix_of_app_alt l1 l2 k1 k2 :
k1 = k2 → l1 `suffix_of` l2 → l1 ++ k1 `suffix_of` l2 ++ k2.
Proof. intros ->. apply suffix_of_app. Qed.
Robbert Krebbers
committed
Lemma suffix_of_snoc_inv_1 x y l1 l2 :
l1 ++ [x] `suffix_of` l2 ++ [y] → x = y.
Proof.
intros [k' E]. rewrite (associative_L (++)) in E. by simplify_list_equality.
Qed.
Lemma suffix_of_snoc_inv_2 x y l1 l2 :
l1 ++ [x] `suffix_of` l2 ++ [y] → l1 `suffix_of` l2.
Proof.
intros [k' E]. exists k'. rewrite (associative_L (++)) in E.
by simplify_list_equality.
Qed.
Lemma suffix_of_app_inv l1 l2 k :
l1 ++ k `suffix_of` l2 ++ k → l1 `suffix_of` l2.
Proof.
intros [k' E]. exists k'. rewrite (associative_L (++)) in E.
by simplify_list_equality.
Qed.
Lemma suffix_of_cons_l l1 l2 x : x :: l1 `suffix_of` l2 → l1 `suffix_of` l2.
Proof. intros [k ->]. exists (k ++ [x]). by rewrite <-(associative_L (++)). Qed.
Robbert Krebbers
committed
Lemma suffix_of_app_l l1 l2 l3 : l3 ++ l1 `suffix_of` l2 → l1 `suffix_of` l2.
Proof. intros [k ->]. exists (k ++ l3). by rewrite <-(associative_L (++)). Qed.
Robbert Krebbers
committed
Lemma suffix_of_cons_r l1 l2 x : l1 `suffix_of` l2 → l1 `suffix_of` x :: l2.
Proof. intros [k ->]. by exists (x :: k). Qed.
Robbert Krebbers
committed
Lemma suffix_of_app_r l1 l2 l3 : l1 `suffix_of` l2 → l1 `suffix_of` l3 ++ l2.
Proof. intros [k ->]. exists (l3 ++ k). by rewrite (associative_L (++)). Qed.
Robbert Krebbers
committed
Lemma suffix_of_cons_inv l1 l2 x y :
x :: l1 `suffix_of` y :: l2 → x :: l1 = y :: l2 ∨ x :: l1 `suffix_of` l2.
Proof.
intros [[|? k] E]; [by left|].
Robbert Krebbers
committed
right. simplify_equality. by apply suffix_of_app_r.
Qed.
Lemma suffix_of_length l1 l2 : l1 `suffix_of` l2 → length l1 ≤ length l2.
Proof. intros [? ->]. rewrite app_length. lia. Qed.
Robbert Krebbers
committed
Lemma suffix_of_cons_not x l : ¬x :: l `suffix_of` l.
Proof. intros [??]. discriminate_list_equality. Qed.
Global Instance suffix_of_dec `{∀ x y, Decision (x = y)} l1 l2 :
Decision (l1 `suffix_of` l2).
Proof.
refine (cast_if (decide_rel prefix_of (reverse l1) (reverse l2)));
abstract (by rewrite suffix_prefix_reverse).
Defined.
Section max_suffix_of.
Context `{∀ x y, Decision (x = y)}.
Lemma max_suffix_of_fst l1 l2 :
l1 = (max_suffix_of l1 l2).1.1 ++ (max_suffix_of l1 l2).2.
Robbert Krebbers
committed
Proof.
rewrite <-(reverse_involutive l1) at 1.
rewrite (max_prefix_of_fst (reverse l1) (reverse l2)). unfold max_suffix_of.
destruct (max_prefix_of (reverse l1) (reverse l2)) as ((?&?)&?); simpl.
by rewrite reverse_app.
Qed.
Lemma max_suffix_of_fst_alt l1 l2 k1 k2 k3 :
max_suffix_of l1 l2 = (k1, k2, k3) → l1 = k1 ++ k3.
Proof.
intro. pose proof (max_suffix_of_fst l1 l2).
by destruct (max_suffix_of l1 l2) as [[]?]; simplify_equality.
Qed.
Lemma max_suffix_of_fst_suffix l1 l2 : (max_suffix_of l1 l2).2 `suffix_of` l1.
Robbert Krebbers
committed
Proof. eexists. apply max_suffix_of_fst. Qed.
Lemma max_suffix_of_fst_suffix_alt l1 l2 k1 k2 k3 :
max_suffix_of l1 l2 = (k1, k2, k3) → k3 `suffix_of` l1.
Proof. eexists. eauto using max_suffix_of_fst_alt. Qed.
Lemma max_suffix_of_snd l1 l2 :
l2 = (max_suffix_of l1 l2).1.2 ++ (max_suffix_of l1 l2).2.
Robbert Krebbers
committed
Proof.
rewrite <-(reverse_involutive l2) at 1.
rewrite (max_prefix_of_snd (reverse l1) (reverse l2)). unfold max_suffix_of.
Robbert Krebbers
committed
destruct (max_prefix_of (reverse l1) (reverse l2)) as ((?&?)&?); simpl.
by rewrite reverse_app.
Qed.
Lemma max_suffix_of_snd_alt l1 l2 k1 k2 k3 :
max_suffix_of l1 l2 = (k1,k2,k3) → l2 = k2 ++ k3.
Proof.
intro. pose proof (max_suffix_of_snd l1 l2).
by destruct (max_suffix_of l1 l2) as [[]?]; simplify_equality.
Qed.
Lemma max_suffix_of_snd_suffix l1 l2 : (max_suffix_of l1 l2).2 `suffix_of` l2.
Robbert Krebbers
committed
Proof. eexists. apply max_suffix_of_snd. Qed.
Lemma max_suffix_of_snd_suffix_alt l1 l2 k1 k2 k3 :
max_suffix_of l1 l2 = (k1,k2,k3) → k3 `suffix_of` l2.
Proof. eexists. eauto using max_suffix_of_snd_alt. Qed.
Lemma max_suffix_of_max l1 l2 k :
k `suffix_of` l1 → k `suffix_of` l2 → k `suffix_of` (max_suffix_of l1 l2).2.
Robbert Krebbers
committed
Proof.
generalize (max_prefix_of_max (reverse l1) (reverse l2)).
rewrite !suffix_prefix_reverse. unfold max_suffix_of.
destruct (max_prefix_of (reverse l1) (reverse l2)) as ((?&?)&?); simpl.
rewrite reverse_involutive. auto.
Qed.
Lemma max_suffix_of_max_alt l1 l2 k1 k2 k3 k :
max_suffix_of l1 l2 = (k1, k2, k3) →
k `suffix_of` l1 → k `suffix_of` l2 → k `suffix_of` k3.
Proof.
intro. pose proof (max_suffix_of_max l1 l2 k).
by destruct (max_suffix_of l1 l2) as [[]?]; simplify_equality.
Qed.
Lemma max_suffix_of_max_snoc l1 l2 k1 k2 k3 x1 x2 :
max_suffix_of l1 l2 = (k1 ++ [x1], k2 ++ [x2], k3) → x1 ≠ x2.
Proof.
intros Hl ->. destruct (suffix_of_cons_not x2 k3).
Robbert Krebbers
committed
eapply max_suffix_of_max_alt; eauto.
* rewrite (max_suffix_of_fst_alt _ _ _ _ _ Hl).
by apply (suffix_of_app [x2]), suffix_of_app_r.
* rewrite (max_suffix_of_snd_alt _ _ _ _ _ Hl).
by apply (suffix_of_app [x2]), suffix_of_app_r.
Qed.
End max_suffix_of.
(** ** Properties of the [sublist] predicate *)
Lemma sublist_length l1 l2 : l1 `sublist` l2 → length l1 ≤ length l2.
Proof. induction 1; simpl; auto with arith. Qed.
Lemma sublist_nil_l l : [] `sublist` l.
Proof. induction l; try constructor; auto. Qed.
Lemma sublist_nil_r l : l `sublist` [] ↔ l = [].
Proof. split. by inversion 1. intros ->. constructor. Qed.
Robbert Krebbers
committed
Lemma sublist_app l1 l2 k1 k2 :
l1 `sublist` l2 → k1 `sublist` k2 → l1 ++ k1 `sublist` l2 ++ k2.
Proof. induction 1; simpl; try constructor; auto. Qed.
Lemma sublist_inserts_l k l1 l2 : l1 `sublist` l2 → l1 `sublist` k ++ l2.
Proof. induction k; try constructor; auto. Qed.
Lemma sublist_inserts_r k l1 l2 : l1 `sublist` l2 → l1 `sublist` l2 ++ k.
Proof. induction 1; simpl; try constructor; auto using sublist_nil_l. Qed.
Lemma sublist_cons_r x l k :
l `sublist` x :: k ↔ l `sublist` k ∨ ∃ l', l = x :: l' ∧ l' `sublist` k.
Proof. split. inversion 1; eauto. intros [?|(?&->&?)]; constructor; auto. Qed.
Robbert Krebbers
committed
Lemma sublist_cons_l x l k :
x :: l `sublist` k ↔ ∃ k1 k2, k = k1 ++ x :: k2 ∧ l `sublist` k2.
Proof.
split.
* intros Hlk. induction k as [|y k IH]; inversion Hlk.
+ eexists [], k. by repeat constructor.
+ destruct IH as (k1&k2&->&?); auto. by exists (y :: k1) k2.
* intros (k1&k2&->&?). by apply sublist_inserts_l, sublist_skip.
Robbert Krebbers
committed
Qed.
Lemma sublist_app_r l k1 k2 :
l `sublist` k1 ++ k2 ↔
∃ l1 l2, l = l1 ++ l2 ∧ l1 `sublist` k1 ∧ l2 `sublist` k2.
Proof.
split.
* revert l k2. induction k1 as [|y k1 IH]; intros l k2; simpl.
{ eexists [], l. by repeat constructor. }
rewrite sublist_cons_r. intros [?|(l' & ? &?)]; subst.
+ destruct (IH l k2) as (l1&l2&?&?&?); trivial; subst.
Robbert Krebbers
committed
1470
1471
1472
1473
1474
1475
1476
1477
1478
1479
1480
1481
1482
1483
1484
1485
1486
1487
1488
1489
1490
1491
+ destruct (IH l' k2) as (l1&l2&?&?&?); trivial; subst.
exists (y :: l1) l2. auto using sublist_skip.
* intros (?&?&?&?&?); subst. auto using sublist_app.
Qed.
Lemma sublist_app_l l1 l2 k :
l1 ++ l2 `sublist` k ↔
∃ k1 k2, k = k1 ++ k2 ∧ l1 `sublist` k1 ∧ l2 `sublist` k2.
Proof.
split.
* revert l2 k. induction l1 as [|x l1 IH]; intros l2 k; simpl.
{ eexists [], k. by repeat constructor. }
rewrite sublist_cons_l. intros (k1 & k2 &?&?); subst.
destruct (IH l2 k2) as (h1 & h2 &?&?&?); trivial; subst.
exists (k1 ++ x :: h1) h2. rewrite <-(associative_L (++)).
auto using sublist_inserts_l, sublist_skip.
* intros (?&?&?&?&?); subst. auto using sublist_app.
Qed.
Lemma sublist_app_inv_l k l1 l2 : k ++ l1 `sublist` k ++ l2 → l1 `sublist` l2.
Proof.
induction k as [|y k IH]; simpl; [done |].
rewrite sublist_cons_r. intros [Hl12|(?&?&?)]; [|simplify_equality; eauto].
rewrite sublist_cons_l in Hl12. destruct Hl12 as (k1&k2&Hk&?).
apply IH. rewrite Hk. eauto using sublist_inserts_l, sublist_cons.
Robbert Krebbers
committed
1493
1494
1495
1496
1497
1498
1499
1500
1501
1502
1503
1504
1505
1506
1507
1508
1509
1510
1511
1512
1513
Qed.
Lemma sublist_app_inv_r k l1 l2 : l1 ++ k `sublist` l2 ++ k → l1 `sublist` l2.
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 <-!(associative_L (++)). }
rewrite sublist_app_l in Hl12. destruct Hl12 as (k1&k2&E&?&Hk2).
destruct k2 as [|z k2] using rev_ind; [inversion Hk2|].
rewrite (associative_L (++)) in E. simplify_list_equality.
eauto using sublist_inserts_r.
Qed.
Global Instance: PartialOrder (@sublist A).
Proof.
split; [split|].
* intros l. induction l; constructor; auto.
* intros l1 l2 l3 Hl12. revert l3. induction Hl12.
+ auto using sublist_nil_l.
+ intros ?. rewrite sublist_cons_l. intros (?&?&?&?); subst.
eauto using sublist_inserts_l, sublist_skip.
+ intros ?. rewrite sublist_cons_l. intros (?&?&?&?); subst.
eauto using sublist_inserts_l, sublist_cons.
Robbert Krebbers
committed
* intros l1 l2 Hl12 Hl21. apply sublist_length in Hl21.
induction Hl12; f_equal'; auto with arith.
Robbert Krebbers
committed
1517
1518
1519
1520
1521
1522
1523
1524
1525
1526
1527
1528
1529
1530
1531
1532
1533
1534
1535
1536
1537
1538
1539
1540
1541
apply sublist_length in Hl12. lia.
Qed.
Lemma sublist_take l i : take i l `sublist` l.
Proof. rewrite <-(take_drop i l) at 2. by apply sublist_inserts_r. Qed.
Lemma sublist_drop l i : drop i l `sublist` l.
Proof. rewrite <-(take_drop i l) at 2. by apply sublist_inserts_l. Qed.
Lemma sublist_delete l i : delete i l `sublist` l.
Proof. revert i. by induction l; intros [|?]; simpl; constructor. Qed.
Lemma sublist_delete_list l is : delete_list is l `sublist` l.
Proof.
induction is as [|i is IH]; simpl; [done |].
transitivity (delete_list is l); auto using sublist_delete.
Qed.
Lemma sublist_alt l1 l2 : l1 `sublist` l2 ↔ ∃ is, l1 = delete_list is l2.
Proof.
split.
* intros Hl12. cut (∀ k, ∃ is, k ++ l1 = delete_list is (k ++ l2)).
{ intros help. apply (help []). }
induction Hl12 as [|x l1 l2 _ IH|x l1 l2 _ IH]; intros k.
+ by eexists [].
+ destruct (IH (k ++ [x])) as [is His]. exists is.
by rewrite <-!(associative_L (++)) in His.
+ destruct (IH k) as [is His]. exists (is ++ [length k]).
unfold delete_list. rewrite fold_right_app. simpl.
by rewrite delete_middle.
* intros [is ->]. apply sublist_delete_list.
Robbert Krebbers
committed
1543
1544
1545
1546
1547
1548
1549
1550
1551
1552
1553
1554
1555
1556
1557
1558
1559
1560
1561
1562
1563
1564
1565
1566
1567
1568
1569
1570
1571
1572
1573
1574
1575
1576
1577
1578
1579
1580
1581
1582
1583
1584
1585
1586
1587
1588
1589
1590
Qed.
Lemma Permutation_sublist l1 l2 l3 :
l1 ≡ₚ l2 → l2 `sublist` l3 → ∃ l4, l1 `sublist` l4 ∧ l4 ≡ₚ l3.
Proof.
intros Hl1l2. revert l3.
induction Hl1l2 as [|x l1 l2 ? IH|x y l1|l1 l1' l2 ? IH1 ? IH2].
* intros l3. by exists l3.
* intros l3. rewrite sublist_cons_l. intros (l3'&l3''&?&?); subst.
destruct (IH l3'') as (l4&?&Hl4); auto. exists (l3' ++ x :: l4).
split. by apply sublist_inserts_l, sublist_skip. by rewrite Hl4.
* intros l3. rewrite sublist_cons_l. intros (l3'&l3''&?& Hl3); subst.
rewrite sublist_cons_l in Hl3. destruct Hl3 as (l5'&l5''&?& Hl5); subst.
exists (l3' ++ y :: l5' ++ x :: l5''). split.
- by do 2 apply sublist_inserts_l, sublist_skip.
- by rewrite !Permutation_middle, Permutation_swap.
* intros l3 ?. destruct (IH2 l3) as (l3'&?&?); trivial.
destruct (IH1 l3') as (l3'' &?&?); trivial. exists l3''.
split. done. etransitivity; eauto.
Qed.
Lemma sublist_Permutation l1 l2 l3 :
l1 `sublist` l2 → l2 ≡ₚ l3 → ∃ l4, l1 ≡ₚ l4 ∧ l4 `sublist` l3.
Proof.
intros Hl1l2 Hl2l3. revert l1 Hl1l2.
induction Hl2l3 as [|x l2 l3 ? IH|x y l2|l2 l2' l3 ? IH1 ? IH2].
* intros l1. by exists l1.
* intros l1. rewrite sublist_cons_r. intros [?|(l1'&l1''&?)]; subst.
{ destruct (IH l1) as (l4&?&?); trivial.
exists l4. split. done. by constructor. }
destruct (IH l1') as (l4&?&Hl4); auto. exists (x :: l4).
split. by constructor. by constructor.
* intros l1. rewrite sublist_cons_r. intros [Hl1|(l1'&l1''&Hl1)]; subst.
{ exists l1. split; [done|]. rewrite sublist_cons_r in Hl1.
destruct Hl1 as [?|(l1'&?&?)]; subst; by repeat constructor. }
rewrite sublist_cons_r in Hl1. destruct Hl1 as [?|(l1''&?&?)]; subst.
+ exists (y :: l1'). by repeat constructor.
+ exists (x :: y :: l1''). by repeat constructor.
* intros l1 ?. destruct (IH1 l1) as (l3'&?&?); trivial.
destruct (IH2 l3') as (l3'' &?&?); trivial. exists l3''.
split; [|done]. etransitivity; eauto.
Qed.
(** Properties of the [contains] predicate *)
Lemma contains_length l1 l2 : l1 `contains` l2 → length l1 ≤ length l2.
Proof. induction 1; simpl; auto with lia. Qed.
Lemma contains_nil_l l : [] `contains` l.
Proof. induction l; constructor; auto. Qed.
Lemma contains_nil_r l : l `contains` [] ↔ l = [].
Proof.
split; [|intros ->; constructor].
Robbert Krebbers
committed
intros Hl. apply contains_length in Hl. destruct l; simpl in *; auto with lia.
Qed.
Global Instance: PreOrder (@contains A).
Proof.
split.
* intros l. induction l; constructor; auto.
* red. apply contains_trans.
Qed.
Lemma Permutation_contains l1 l2 : l1 ≡ₚ l2 → l1 `contains` l2.
Proof. induction 1; econstructor; eauto. Qed.
Lemma sublist_contains l1 l2 : l1 `sublist` l2 → l1 `contains` l2.
Proof. induction 1; constructor; auto. Qed.
Lemma contains_Permutation l1 l2 : l1 `contains` l2 → ∃ k, l2 ≡ₚ l1 ++ k.
Proof.
induction 1 as
[|x y l ? [k Hk]| |x l1 l2 ? [k Hk]|l1 l2 l3 ? [k Hk] ? [k' Hk']].
* by eexists [].
* exists k. by rewrite Hk.
* eexists []. rewrite (right_id_L [] (++)). by constructor.
* exists (x :: k). by rewrite Hk, Permutation_middle.
* exists (k ++ k'). by rewrite Hk', Hk, (associative_L (++)).
Qed.
Lemma contains_Permutation_length_le l1 l2 :
Robbert Krebbers
committed
length l2 ≤ length l1 → l1 `contains` l2 → l1 ≡ₚ l2.
Proof.
intros Hl21 Hl12. destruct (contains_Permutation l1 l2) as [[|??] Hk]; auto.
* by rewrite Hk, (right_id_L [] (++)).
* rewrite Hk, app_length in Hl21; simpl in Hl21; lia.
Qed.
Lemma contains_Permutation_length_eq l1 l2 :
Robbert Krebbers
committed
length l2 = length l1 → l1 `contains` l2 → l1 ≡ₚ l2.
Proof. intro. apply contains_Permutation_length_le. lia. Qed.
Robbert Krebbers
committed
Global Instance: Proper ((≡ₚ) ==> (≡ₚ) ==> iff) (@contains A).
Proof.
intros l1 l2 ? k1 k2 ?. split; intros.
* transitivity l1. by apply Permutation_contains.
transitivity k1. done. by apply Permutation_contains.
* transitivity l2. by apply Permutation_contains.
transitivity k2. done. by apply Permutation_contains.
Qed.
Global Instance: AntiSymmetric (≡ₚ) (@contains A).
Proof. red. auto using contains_Permutation_length_le, contains_length. Qed.
Robbert Krebbers
committed
1634
1635
1636
1637
1638
1639
1640
1641
1642
1643
1644
1645
1646
1647
1648
1649
1650
1651
1652
1653
1654
1655
1656
1657
1658
1659
1660
1661
1662
1663
1664
1665
1666
1667
1668
1669
1670
1671
1672
1673
1674
1675
1676
1677
1678
1679
1680
1681
1682
1683
1684
1685
1686
1687
1688
1689
1690
1691
1692
1693
1694
1695
1696
1697
1698
1699
1700
1701
1702
1703
1704
1705
1706
1707
1708
1709
1710
1711
1712
1713
1714
1715
1716
1717
1718
Lemma contains_take l i : take i l `contains` l.
Proof. auto using sublist_take, sublist_contains. Qed.
Lemma contains_drop l i : drop i l `contains` l.
Proof. auto using sublist_drop, sublist_contains. Qed.
Lemma contains_delete l i : delete i l `contains` l.
Proof. auto using sublist_delete, sublist_contains. Qed.
Lemma contains_delete_list l is : delete_list is l `sublist` l.
Proof. auto using sublist_delete_list, sublist_contains. Qed.
Lemma contains_sublist_l l1 l3 :
l1 `contains` l3 ↔ ∃ l2, l1 `sublist` l2 ∧ l2 ≡ₚ l3.
Proof.
split.
{ intros Hl13. elim Hl13; clear l1 l3 Hl13.
* by eexists [].
* intros x l1 l3 ? (l2&?&?). exists (x :: l2). by repeat constructor.
* intros x y l. exists (y :: x :: l). by repeat constructor.
* intros x l1 l3 ? (l2&?&?). exists (x :: l2). by repeat constructor.
* intros l1 l3 l5 ? (l2&?&?) ? (l4&?&?).
destruct (Permutation_sublist l2 l3 l4) as (l3'&?&?); trivial.
exists l3'. split; etransitivity; eauto. }
intros (l2&?&?).
transitivity l2; auto using sublist_contains, Permutation_contains.
Qed.
Lemma contains_sublist_r l1 l3 :
l1 `contains` l3 ↔ ∃ l2, l1 ≡ₚ l2 ∧ l2 `sublist` l3.
Proof.
rewrite contains_sublist_l.
split; intros (l2&?&?); eauto using sublist_Permutation, Permutation_sublist.
Qed.
Lemma contains_inserts_l k l1 l2 : l1 `contains` l2 → l1 `contains` k ++ l2.
Proof. induction k; try constructor; auto. Qed.
Lemma contains_inserts_r k l1 l2 : l1 `contains` l2 → l1 `contains` l2 ++ k.
Proof. rewrite (commutative (++)). apply contains_inserts_l. Qed.
Lemma contains_skips_l k l1 l2 : l1 `contains` l2 → k ++ l1 `contains` k ++ l2.
Proof. induction k; try constructor; auto. Qed.
Lemma contains_skips_r k l1 l2 : l1 `contains` l2 → l1 ++ k `contains` l2 ++ k.
Proof. rewrite !(commutative (++) _ k). apply contains_skips_l. Qed.
Lemma contains_app l1 l2 k1 k2 :
l1 `contains` l2 → k1 `contains` k2 → l1 ++ k1 `contains` l2 ++ k2.
Proof.
transitivity (l1 ++ k2); auto using contains_skips_l, contains_skips_r.
Qed.
Lemma contains_cons_r x l k :
l `contains` x :: k ↔ l `contains` k ∨ ∃ l', l ≡ₚ x :: l' ∧ l' `contains` k.
Proof.
split.
* rewrite contains_sublist_r. intros (l'&E&Hl').
rewrite sublist_cons_r in Hl'. destruct Hl' as [?|(?&?&?)]; subst.
+ left. rewrite E. eauto using sublist_contains.
+ right. eauto using sublist_contains.
* intros [?|(?&E&?)]; [|rewrite E]; by constructor.
Qed.
Lemma contains_cons_l x l k :
x :: l `contains` k ↔ ∃ k', k ≡ₚ x :: k' ∧ l `contains` k'.
Proof.
split.
* rewrite contains_sublist_l. intros (l'&Hl'&E).
rewrite sublist_cons_l in Hl'. destruct Hl' as (k1&k2&?&?); subst.
exists (k1 ++ k2). split; eauto using contains_inserts_l, sublist_contains.
by rewrite Permutation_middle.
* intros (?&E&?). rewrite E. by constructor.
Qed.
Lemma contains_app_r l k1 k2 :
l `contains` k1 ++ k2 ↔ ∃ l1 l2,
l ≡ₚ l1 ++ l2 ∧ l1 `contains` k1 ∧ l2 `contains` k2.
Proof.
split.
* rewrite contains_sublist_r. intros (l'&E&Hl').
rewrite sublist_app_r in Hl'. destruct Hl' as (l1&l2&?&?&?); subst.
exists l1 l2. eauto using sublist_contains.
* intros (?&?&E&?&?). rewrite E. eauto using contains_app.
Qed.
Lemma contains_app_l l1 l2 k :
l1 ++ l2 `contains` k ↔ ∃ k1 k2,
k ≡ₚ k1 ++ k2 ∧ l1 `contains` k1 ∧ l2 `contains` k2.
Proof.
split.
* rewrite contains_sublist_l. intros (l'&Hl'&E).
rewrite sublist_app_l in Hl'. destruct Hl' as (k1&k2&?&?&?); subst.
exists k1 k2. split. done. eauto using sublist_contains.
* intros (?&?&E&?&?). rewrite E. eauto using contains_app.
Qed.
Lemma contains_app_inv_l l1 l2 k :
k ++ l1 `contains` k ++ l2 → l1 `contains` l2.
Proof.
induction k as [|y k IH]; simpl; [done |]. rewrite contains_cons_l.
intros (?&E&?). apply Permutation_cons_inv in E. apply IH. by rewrite E.
Robbert Krebbers
committed
1721
1722
1723
1724
1725
1726
1727
1728
1729
1730
1731
1732
1733
1734
1735
1736
1737
1738
1739
1740
1741
1742
1743
1744
1745
1746
1747
1748
Qed.
Lemma contains_app_inv_r l1 l2 k :
l1 ++ k `contains` l2 ++ k → l1 `contains` l2.
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 <-!(associative_L (++)). }
rewrite contains_app_l in Hl12. destruct Hl12 as (k1&k2&E1&?&Hk2).
rewrite contains_cons_l in Hk2. destruct Hk2 as (k2'&E2&?).
rewrite E2, (Permutation_cons_append k2'), (associative_L (++)) in E1.
apply Permutation_app_inv_r in E1. rewrite E1. eauto using contains_inserts_r.
Qed.
Lemma contains_cons_middle x l k1 k2 :
l `contains` k1 ++ k2 → x :: l `contains` k1 ++ x :: k2.
Proof. rewrite <-Permutation_middle. by apply contains_skip. Qed.
Lemma contains_app_middle l1 l2 k1 k2 :
l2 `contains` k1 ++ k2 → l1 ++ l2 `contains` k1 ++ l1 ++ k2.
Proof.
rewrite !(associative (++)), (commutative (++) k1 l1), <-(associative_L (++)).
by apply contains_skips_l.
Qed.
Lemma contains_middle l k1 k2 : l `contains` k1 ++ l ++ k2.
Proof. by apply contains_inserts_l, contains_inserts_r. Qed.
Lemma Permutation_alt l1 l2 :
l1 ≡ₚ l2 ↔ length l1 = length l2 ∧ l1 `contains` l2.
Proof.
split.
* by intros Hl; rewrite Hl.
* intros [??]; auto using contains_Permutation_length_eq.
Robbert Krebbers
committed
Qed.
Lemma NoDup_contains l k : NoDup l → (∀ x, x ∈ l → x ∈ k) → l `contains` k.
Proof.
intros Hl. revert k. induction Hl as [|x l Hx ? IH].
{ intros k Hk. by apply contains_nil_l. }
intros k Hlk. destruct (elem_of_list_split k x) as (l1&l2&?); subst.
{ apply Hlk. by constructor. }
rewrite <-Permutation_middle. apply contains_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_symmetric contains); apply NoDup_contains; naive_solver.
Qed.
Robbert Krebbers
committed
Section contains_dec.
Context `{∀ x y, Decision (x = y)}.
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.
Robbert Krebbers
committed
* done.
* case_decide; simplify_equality; eauto.
destruct (list_remove x l1) as [l|] eqn:?; simplify_equality.
destruct (IH l) as (?&?&?); simplify_option_equality; eauto.
* simplify_option_equality; eauto using Permutation_swap.
Robbert Krebbers
committed
* destruct (IH1 k1) as (k2&?&?); trivial.
destruct (IH2 k2) as (k3&?&?); trivial.
exists k3. split; eauto. by transitivity k2.
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_equality; 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'.
Qed.
Lemma list_remove_list_contains l1 l2 :
l1 `contains` l2 ↔ is_Some (list_remove_list l1 l2).
Proof.
Robbert Krebbers
committed
1805
1806
1807
1808
1809
1810
1811
1812
1813
1814
1815
1816
1817
1818
1819
1820
1821
1822
1823
1824
1825
1826
1827
* revert l2. induction l1 as [|x l1 IH]; simpl.
{ intros l2 _. by exists l2. }
intros l2. rewrite contains_cons_l. intros (k&Hk&?).
destruct (list_remove_Some_inv l2 k x) as (k2&?&Hk2); trivial.
simplify_option_equality. apply IH. by rewrite <-Hk2.
* intros [k Hk]. revert l2 k Hk.
induction l1 as [|x l1 IH]; simpl; intros l2 k.
{ intros. apply contains_nil_l. }
destruct (list_remove x l2) as [k'|] eqn:?; intros; simplify_equality.
rewrite contains_cons_l. eauto using list_remove_Some.
Qed.
Global Instance contains_dec l1 l2 : Decision (l1 `contains` l2).
Proof.
refine (cast_if (decide (is_Some (list_remove_list l1 l2))));
abstract (rewrite list_remove_list_contains; tauto).
Defined.
Global Instance Permutation_dec l1 l2 : Decision (l1 ≡ₚ l2).
Proof.
refine (cast_if_and
(decide (length l1 = length l2)) (decide (l1 `contains` l2)));
abstract (rewrite Permutation_alt; tauto).
Defined.
End contains_dec.
Robbert Krebbers
committed
(** ** Properties of the [Forall] and [Exists] predicate *)
Lemma Forall_Exists_dec {A} {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.
Context {A} (P : A → Prop).
Robbert Krebbers
committed
Definition Forall_nil_2 := @Forall_nil A.
Definition Forall_cons_2 := @Forall_cons A.
Lemma Forall_forall l : Forall P l ↔ ∀ x, x ∈ l → P x.
{ induction 1; inversion 1; subst; auto. }
intros Hin. induction l; constructor.
* apply Hin. constructor.
* apply IHl. 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. induction l; auto. Qed.
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.
Global Instance Forall_proper:
Proper (pointwise_relation _ (↔) ==> (=) ==> (↔)) (@Forall A).
Proof. split; subst; induction 1; constructor (by firstorder auto). Qed.
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_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.
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.
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).
Proof. 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_equality.
auto using Forall_take, Forall_drop.
Qed.
Lemma Forall_sublist_alter f l i n :
i + n ≤ length l → Forall P l →
(∀ k, sublist_lookup i n l = Some k → Forall P k → Forall P (f k)) →
Forall P (sublist_alter f i n l).
unfold sublist_alter.
auto 8 using Forall_app_2, Forall_drop, Forall_take, sublist_lookup_Some.
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.
Qed.
(*
Lemma Forall_mask f βs l: P x → P y → Forall P l → Forall P (mask f βs l).
Proof.
intros ??. revert l. induction βs as [|[]]; intros ? [|????]; simpl; auto.
Qed.
Lemma Forall_resize_mask x y n βs l :
P y → Forall P (resize n x l) → length βs ≤ n → Forall P (mask x y βs l).
Proof.
intros ?. revert n l. induction βs as [|β βs IH]; [constructor|].
intros [|n] [|z l]; simpl; inversion_clear 1; intros; try lia.
* constructor. by destruct β. apply IH with n. by rewrite resize_nil. lia.
* constructor. by destruct β. apply IH with n; auto with lia.
Qed.
*)
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.
Robbert Krebbers
committed
* 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.
Global Instance Exists_proper:
Proper (pointwise_relation _ (↔) ==> (=) ==> (↔)) (@Exists A).
Proof. split; subst; induction 1; constructor (by firstorder auto). Qed.
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.
Context {dec : ∀ x, Decision (P x)}.
Lemma not_Forall_Exists l : ¬Forall P l → Exists (not ∘ P) l.
Proof. intro. destruct (Forall_Exists_dec dec l); intuition. Qed.
Lemma not_Exists_Forall l : ¬Exists P l → Forall (not ∘ P) l.
Proof. by destruct (Forall_Exists_dec (λ x, swap_if (decide (P x))) l). Qed.
Global Instance Forall_dec l : Decision (Forall P l) :=
match Forall_Exists_dec 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 (λ x, swap_if (decide (P x))) l with
| left H => right (Forall_not_Exists _ H)
| right H => left H
End Forall_Exists.