# HG changeset patch # User nipkow # Date 1505317254 -7200 # Node ID 8f4d252ce2fe45d92819d277e1390c6302c95798 # Parent e9be3d6995f9795e316fa5fe4aabf036f5e42948 added lemma; zip_with -> map2 diff -r e9be3d6995f9 -r 8f4d252ce2fe src/HOL/Library/Stirling.thy --- a/src/HOL/Library/Stirling.thy Tue Sep 12 20:40:46 2017 +0200 +++ b/src/HOL/Library/Stirling.thy Wed Sep 13 17:40:54 2017 +0200 @@ -246,7 +246,7 @@ \ definition zip_with_prev :: "('a \ 'a \ 'b) \ 'a \ 'a list \ 'b list" - where "zip_with_prev f x xs = zip_with f (x # xs) xs" + where "zip_with_prev f x xs = map2 f (x # xs) xs" lemma zip_with_prev_altdef: "zip_with_prev f x xs = diff -r e9be3d6995f9 -r 8f4d252ce2fe src/HOL/List.thy --- a/src/HOL/List.thy Tue Sep 12 20:40:46 2017 +0200 +++ b/src/HOL/List.thy Wed Sep 13 17:40:54 2017 +0200 @@ -151,8 +151,8 @@ \ \Warning: simpset does not contain this definition, but separate theorems for \xs = []\ and \xs = z # zs\\ -abbreviation zip_with :: "('a \ 'b \ 'c) \ 'a list \ 'b list \ 'c list" where -"zip_with f xs ys \ map (\(x,y). f x y) (zip xs ys)" +abbreviation map2 :: "('a \ 'b \ 'c) \ 'a list \ 'b list \ 'c list" where +"map2 f xs ys \ map (\(x,y). f x y) (zip xs ys)" primrec product :: "'a list \ 'b list \ ('a \ 'b) list" where "product [] _ = []" | @@ -4396,12 +4396,12 @@ done lemma nths_shift_lemma: - "map fst [p<-zip xs [i.. distinct (nths xs I)" - by (induct xs arbitrary: I) (auto simp: nths_Cons) - +by (induct xs arbitrary: I) (auto simp: nths_Cons) lemma nths_upt_eq_take [simp]: "nths l {.. P(xs!i)}" +by(induction xs) (auto simp: nths_Cons) lemma filter_in_nths: - "distinct xs \ filter (%x. x \ set (nths xs s)) xs = nths xs s" + "distinct xs \ filter (%x. x \ set (nths xs s)) xs = nths xs s" proof (induct xs arbitrary: s) case Nil thus ?case by simp next