merged
authorwenzelm
Tue, 21 Jun 2022 23:36:16 +0200
changeset 75581 29654a8e9374
parent 75564 d32201f08e98 (diff)
parent 75580 5c1c4f537ae8 (current diff)
child 75585 a789c5732f7a
child 75586 b2b097624e4c
merged
NEWS
src/Pure/General/bytes.ML
--- a/NEWS	Tue Jun 21 23:30:19 2022 +0200
+++ b/NEWS	Tue Jun 21 23:36:16 2022 +0200
@@ -94,6 +94,9 @@
       image_mset_filter_mset_swap
       multp_image_mset_image_msetD
 
+* Theory "HOL-Library.Sublist":
+  - Added lemma map_mono_strict_suffix.
+
 * Sledgehammer:
   - Redesigned multithreading to provide more fine grained prover schedules.
     The binary option 'slice' has been replaced by a numeric value 'slices'
--- a/src/HOL/Library/Sublist.thy	Tue Jun 21 23:30:19 2022 +0200
+++ b/src/HOL/Library/Sublist.thy	Tue Jun 21 23:36:16 2022 +0200
@@ -609,6 +609,9 @@
 lemma map_mono_suffix: "suffix xs ys \<Longrightarrow> suffix (map f xs) (map f ys)"
 by (auto elim!: suffixE intro: suffixI)
 
+lemma map_mono_strict_suffix: "strict_suffix xs ys \<Longrightarrow> strict_suffix (map f xs) (map f ys)"
+  by (auto simp: strict_suffix_def suffix_def)
+
 lemma filter_mono_suffix: "suffix xs ys \<Longrightarrow> suffix (filter P xs) (filter P ys)"
 by (auto simp: suffix_def)