aboutsummaryrefslogtreecommitdiff
diff options
context:
space:
mode:
-rw-r--r--doc/changelog/10-standard-library/12237-list-more-filter-incl.rst2
-rw-r--r--theories/Lists/List.v4
2 files changed, 5 insertions, 1 deletions
diff --git a/doc/changelog/10-standard-library/12237-list-more-filter-incl.rst b/doc/changelog/10-standard-library/12237-list-more-filter-incl.rst
index 7b4b0e86d3..37aaf697b5 100644
--- a/doc/changelog/10-standard-library/12237-list-more-filter-incl.rst
+++ b/doc/changelog/10-standard-library/12237-list-more-filter-incl.rst
@@ -1,4 +1,4 @@
- **Added:**
- new lemmas in ``List``: ``incl_map``, ``incl_filter``, ``NoDup_filter``
+ new lemmas in ``List``: ``incl_map``, ``incl_filter``, ``NoDup_filter``, ``incl_Forall_in_iff``
(`#12237 <https://github.com/coq/coq/pull/12237>`_,
by Olivier Laurent).
diff --git a/theories/Lists/List.v b/theories/Lists/List.v
index e4b87507f4..bdaaa91828 100644
--- a/theories/Lists/List.v
+++ b/theories/Lists/List.v
@@ -2968,6 +2968,10 @@ Section Exists_Forall.
now apply neg_Forall_Exists_neg.
Defined.
+ Lemma incl_Forall_in_iff l l' :
+ incl l l' <-> Forall (fun x => In x l') l.
+ Proof. now rewrite Forall_forall; split. Qed.
+
End Exists_Forall.
Hint Constructors Exists : core.