src/HOL/List.thy
changeset 19770 be5c23ebe1eb
parent 19623 12e6cc4382ae
child 19787 b949911ecff5
equal deleted inserted replaced
19769:c40ce2de2020 19770:be5c23ebe1eb
     4 *)
     4 *)
     5 
     5 
     6 header {* The datatype of finite lists *}
     6 header {* The datatype of finite lists *}
     7 
     7 
     8 theory List
     8 theory List
     9 imports PreList
     9 imports PreList FunDef
    10 begin
    10 begin
    11 
    11 
    12 datatype 'a list =
    12 datatype 'a list =
    13     Nil    ("[]")
    13     Nil    ("[]")
    14   | Cons 'a  "'a list"    (infixr "#" 65)
    14   | Cons 'a  "'a list"    (infixr "#" 65)
   496 by (induct xs) auto
   496 by (induct xs) auto
   497 
   497 
   498 lemma map_eq_conv[simp]: "(map f xs = map g xs) = (!x : set xs. f x = g x)"
   498 lemma map_eq_conv[simp]: "(map f xs = map g xs) = (!x : set xs. f x = g x)"
   499 by (induct xs) auto
   499 by (induct xs) auto
   500 
   500 
   501 lemma map_cong [recdef_cong]:
   501 lemma map_cong [fundef_cong, recdef_cong]:
   502 "xs = ys ==> (!!x. x : set ys ==> f x = g x) ==> map f xs = map g ys"
   502 "xs = ys ==> (!!x. x : set ys ==> f x = g x) ==> map f xs = map g ys"
   503 -- {* a congruence rule for @{text map} *}
   503 -- {* a congruence rule for @{text map} *}
   504 by simp
   504 by simp
   505 
   505 
   506 lemma map_is_Nil_conv [iff]: "(map f xs = []) = (xs = [])"
   506 lemma map_is_Nil_conv [iff]: "(map f xs = []) = (xs = [])"
   861 lemma Cons_eq_filter_iff:
   861 lemma Cons_eq_filter_iff:
   862  "(x#xs = filter P ys) =
   862  "(x#xs = filter P ys) =
   863   (\<exists>us vs. ys = us @ x # vs \<and> (\<forall>u\<in>set us. \<not> P u) \<and> P x \<and> xs = filter P vs)"
   863   (\<exists>us vs. ys = us @ x # vs \<and> (\<forall>u\<in>set us. \<not> P u) \<and> P x \<and> xs = filter P vs)"
   864 by(auto dest:Cons_eq_filterD)
   864 by(auto dest:Cons_eq_filterD)
   865 
   865 
   866 lemma filter_cong[recdef_cong]:
   866 lemma filter_cong[fundef_cong, recdef_cong]:
   867  "xs = ys \<Longrightarrow> (\<And>x. x \<in> set ys \<Longrightarrow> P x = Q x) \<Longrightarrow> filter P xs = filter Q ys"
   867  "xs = ys \<Longrightarrow> (\<And>x. x \<in> set ys \<Longrightarrow> P x = Q x) \<Longrightarrow> filter P xs = filter Q ys"
   868 apply simp
   868 apply simp
   869 apply(erule thin_rl)
   869 apply(erule thin_rl)
   870 by (induct ys) simp_all
   870 by (induct ys) simp_all
   871 
   871 
  1361  apply simp
  1361  apply simp
  1362 apply(case_tac xs)
  1362 apply(case_tac xs)
  1363 apply(auto)
  1363 apply(auto)
  1364 done
  1364 done
  1365 
  1365 
  1366 lemma takeWhile_cong [recdef_cong]:
  1366 lemma takeWhile_cong [fundef_cong, recdef_cong]:
  1367   "[| l = k; !!x. x : set l ==> P x = Q x |] 
  1367   "[| l = k; !!x. x : set l ==> P x = Q x |] 
  1368   ==> takeWhile P l = takeWhile Q k"
  1368   ==> takeWhile P l = takeWhile Q k"
  1369   by (induct k fixing: l, simp_all)
  1369   by (induct k fixing: l, simp_all)
  1370 
  1370 
  1371 lemma dropWhile_cong [recdef_cong]:
  1371 lemma dropWhile_cong [fundef_cong, recdef_cong]:
  1372   "[| l = k; !!x. x : set l ==> P x = Q x |] 
  1372   "[| l = k; !!x. x : set l ==> P x = Q x |] 
  1373   ==> dropWhile P l = dropWhile Q k"
  1373   ==> dropWhile P l = dropWhile Q k"
  1374   by (induct k fixing: l, simp_all)
  1374   by (induct k fixing: l, simp_all)
  1375 
  1375 
  1376 
  1376 
  1611 by (induct xs) auto
  1611 by (induct xs) auto
  1612 
  1612 
  1613 lemma foldr_append[simp]: "foldr f (xs @ ys) a = foldr f xs (foldr f ys a)"
  1613 lemma foldr_append[simp]: "foldr f (xs @ ys) a = foldr f xs (foldr f ys a)"
  1614 by (induct xs) auto
  1614 by (induct xs) auto
  1615 
  1615 
  1616 lemma foldl_cong [recdef_cong]:
  1616 lemma foldl_cong [fundef_cong, recdef_cong]:
  1617   "[| a = b; l = k; !!a x. x : set l ==> f a x = g a x |] 
  1617   "[| a = b; l = k; !!a x. x : set l ==> f a x = g a x |] 
  1618   ==> foldl f a l = foldl g b k"
  1618   ==> foldl f a l = foldl g b k"
  1619   by (induct k fixing: a b l, simp_all)
  1619   by (induct k fixing: a b l, simp_all)
  1620 
  1620 
  1621 lemma foldr_cong [recdef_cong]:
  1621 lemma foldr_cong [fundef_cong, recdef_cong]:
  1622   "[| a = b; l = k; !!a x. x : set l ==> f x a = g x a |] 
  1622   "[| a = b; l = k; !!a x. x : set l ==> f x a = g x a |] 
  1623   ==> foldr f l a = foldr g k b"
  1623   ==> foldr f l a = foldr g k b"
  1624   by (induct k fixing: a b l, simp_all)
  1624   by (induct k fixing: a b l, simp_all)
  1625 
  1625 
  1626 lemma foldr_foldl: "foldr f xs a = foldl (%x y. f y x) a (rev xs)"
  1626 lemma foldr_foldl: "foldr f xs a = foldl (%x y. f y x) a (rev xs)"