src/HOL/List.thy
changeset 37455 059ee3176686
parent 37451 3058c918e7a3
parent 37454 9132a5955127
child 37456 0a1cc2675958
equal deleted inserted replaced
37453:44a307746163 37455:059ee3176686
  3188 by (fast dest!: not0_implies_Suc intro!: set_replicate_Suc)
  3188 by (fast dest!: not0_implies_Suc intro!: set_replicate_Suc)
  3189 
  3189 
  3190 lemma set_replicate_conv_if: "set (replicate n x) = (if n = 0 then {} else {x})"
  3190 lemma set_replicate_conv_if: "set (replicate n x) = (if n = 0 then {} else {x})"
  3191 by auto
  3191 by auto
  3192 
  3192 
       
  3193 lemma Ball_set_replicate[simp]:
       
  3194   "(ALL x : set(replicate n a). P x) = (P a | n=0)"
       
  3195 by(simp add: set_replicate_conv_if)
       
  3196 
       
  3197 lemma Bex_set_replicate[simp]:
       
  3198   "(EX x : set(replicate n a). P x) = (P a & n\<noteq>0)"
       
  3199 by(simp add: set_replicate_conv_if)
       
  3200 
  3193 lemma in_set_replicateD: "x : set (replicate n y) ==> x = y"
  3201 lemma in_set_replicateD: "x : set (replicate n y) ==> x = y"
  3194 by (simp add: set_replicate_conv_if split: split_if_asm)
  3202 by (simp add: set_replicate_conv_if split: split_if_asm)
  3195 
  3203 
  3196 lemma replicate_append_same:
  3204 lemma replicate_append_same:
  3197   "replicate i x @ [x] = x # replicate i x"
  3205   "replicate i x @ [x] = x # replicate i x"
  4704   "list_all P (rev xs) \<longleftrightarrow> list_all P xs"
  4712   "list_all P (rev xs) \<longleftrightarrow> list_all P xs"
  4705 by (simp add: list_all_iff)
  4713 by (simp add: list_all_iff)
  4706 
  4714 
  4707 lemma list_all_length:
  4715 lemma list_all_length:
  4708   "list_all P xs \<longleftrightarrow> (\<forall>n < length xs. P (xs ! n))"
  4716   "list_all P xs \<longleftrightarrow> (\<forall>n < length xs. P (xs ! n))"
  4709   unfolding list_all_iff by (auto intro: all_nth_imp_all_set)
  4717 unfolding list_all_iff by (auto intro: all_nth_imp_all_set)
       
  4718 
       
  4719 lemma list_all_cong[fundef_cong]:
       
  4720   "xs = ys \<Longrightarrow> (\<And>x. x \<in> set ys \<Longrightarrow> f x = g x) \<Longrightarrow> list_all f xs = list_all g ys"
       
  4721 by (simp add: list_all_iff)
  4710 
  4722 
  4711 lemma list_ex_iff [code_post]:
  4723 lemma list_ex_iff [code_post]:
  4712   "list_ex P xs \<longleftrightarrow> (\<exists>x \<in> set xs. P x)"
  4724   "list_ex P xs \<longleftrightarrow> (\<exists>x \<in> set xs. P x)"
  4713 by (induct xs) simp_all
  4725 by (induct xs) simp_all
  4714 
  4726 
  4715 lemmas list_bex_code [code_unfold] =
  4727 lemmas list_bex_code [code_unfold] =
  4716   list_ex_iff [symmetric]
  4728   list_ex_iff [symmetric]
  4717 
  4729 
  4718 lemma list_ex_length:
  4730 lemma list_ex_length:
  4719   "list_ex P xs \<longleftrightarrow> (\<exists>n < length xs. P (xs ! n))"
  4731   "list_ex P xs \<longleftrightarrow> (\<exists>n < length xs. P (xs ! n))"
  4720   unfolding list_ex_iff set_conv_nth by auto
  4732 unfolding list_ex_iff set_conv_nth by auto
       
  4733 
       
  4734 lemma list_ex_cong[fundef_cong]:
       
  4735   "xs = ys \<Longrightarrow> (\<And>x. x \<in> set ys \<Longrightarrow> f x = g x) \<Longrightarrow> list_ex f xs = list_ex g ys"
       
  4736 by (simp add: list_ex_iff)
  4721 
  4737 
  4722 lemma filtermap_conv:
  4738 lemma filtermap_conv:
  4723    "filtermap f xs = map (\<lambda>x. the (f x)) (filter (\<lambda>x. f x \<noteq> None) xs)"
  4739    "filtermap f xs = map (\<lambda>x. the (f x)) (filter (\<lambda>x. f x \<noteq> None) xs)"
  4724 by (induct xs) (simp_all split: option.split) 
  4740 by (induct xs) (simp_all split: option.split) 
  4725 
  4741