aboutsummaryrefslogtreecommitdiff
path: root/src/Util/ListUtil.v
diff options
context:
space:
mode:
Diffstat (limited to 'src/Util/ListUtil.v')
-rw-r--r--src/Util/ListUtil.v49
1 files changed, 49 insertions, 0 deletions
diff --git a/src/Util/ListUtil.v b/src/Util/ListUtil.v
index 1f9a62457..36d8a3ad3 100644
--- a/src/Util/ListUtil.v
+++ b/src/Util/ListUtil.v
@@ -433,6 +433,22 @@ Proof.
auto.
Qed.
+Lemma nth_default_cons_S : forall {A} us (u0 : A) n d,
+ nth_default d (u0 :: us) (S n) = nth_default d us n.
+Proof.
+ boring.
+Qed.
+
+Lemma nth_error_Some_nth_default : forall {T} i x (l : list T), (i < length l)%nat ->
+ nth_error l i = Some (nth_default x l i).
+Proof.
+ intros ? ? ? ? i_lt_length.
+ destruct (nth_error_length_exists_value _ _ i_lt_length) as [k nth_err_k].
+ unfold nth_default.
+ rewrite nth_err_k.
+ reflexivity.
+Qed.
+
Lemma set_nth_cons : forall {T} (x u0 : T) us, set_nth 0 x (u0 :: us) = x :: us.
Proof.
auto.
@@ -533,3 +549,36 @@ Lemma cons_eq_tail : forall {T} (x y:T) xs ys, x::xs = y::ys -> xs=ys.
Proof.
intros; solve_by_inversion.
Qed.
+
+Lemma map_nth_default_always {A B} (f : A -> B) (n : nat) (x : A) (l : list A)
+ : nth_default (f x) (map f l) n = f (nth_default x l n).
+Proof.
+ revert n; induction l; simpl; intro n; destruct n; [ try reflexivity.. ].
+ nth_tac.
+Qed.
+
+Lemma fold_right_and_True_forall_In_iff : forall {T} (l : list T) (P : T -> Prop),
+ (forall x, In x l -> P x) <-> fold_right and True (map P l).
+Proof.
+ induction l; intros; simpl; try tauto.
+ rewrite <- IHl.
+ intuition (subst; auto).
+Qed.
+
+Lemma fold_right_invariant : forall {A} P (f: A -> A -> A) l x,
+ P x -> (forall y, In y l -> forall z, P z -> P (f y z)) ->
+ P (fold_right f x l).
+Proof.
+ induction l; intros ? ? step; auto.
+ simpl.
+ apply step; try apply in_eq.
+ apply IHl; auto.
+ intros y in_y_l.
+ apply (in_cons a) in in_y_l.
+ auto.
+Qed.
+
+Lemma In_firstn : forall {T} n l (x : T), In x (firstn n l) -> In x l.
+Proof.
+ induction n; destruct l; boring.
+Qed.