--- a/src/HOL/Library/Stirling.thy Tue Sep 12 12:14:38 2017 +0200
+++ b/src/HOL/Library/Stirling.thy Tue Sep 12 20:40:46 2017 +0200
@@ -246,7 +246,7 @@
\<close>
definition zip_with_prev :: "('a \<Rightarrow> 'a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'a list \<Rightarrow> 'b list"
- where "zip_with_prev f x xs = map (\<lambda>(x,y). f x y) (zip (x # xs) xs)"
+ where "zip_with_prev f x xs = zip_with f (x # xs) xs"
lemma zip_with_prev_altdef:
"zip_with_prev f x xs =
--- a/src/HOL/List.thy Tue Sep 12 12:14:38 2017 +0200
+++ b/src/HOL/List.thy Tue Sep 12 20:40:46 2017 +0200
@@ -151,6 +151,9 @@
\<comment> \<open>Warning: simpset does not contain this definition, but separate
theorems for \<open>xs = []\<close> and \<open>xs = z # zs\<close>\<close>
+abbreviation zip_with :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'a list \<Rightarrow> 'b list \<Rightarrow> 'c list" where
+"zip_with f xs ys \<equiv> map (\<lambda>(x,y). f x y) (zip xs ys)"
+
primrec product :: "'a list \<Rightarrow> 'b list \<Rightarrow> ('a \<times> 'b) list" where
"product [] _ = []" |
"product (x#xs) ys = map (Pair x) ys @ product xs ys"