# HG changeset patch # User lcp # Date 804436397 -7200 # Node ID 5873833cf37f8f8b3fa02d5146f7cb2df01d19b8 # Parent 74be52691d626a198205930b424389cf7fabb5c6 Added function rev and its properties length_rev, etc. diff -r 74be52691d62 -r 5873833cf37f src/HOL/List.ML --- a/src/HOL/List.ML Thu Jun 29 16:28:40 1995 +0200 +++ b/src/HOL/List.ML Thu Jun 29 16:33:17 1995 +0200 @@ -24,15 +24,16 @@ goal List.thy "(xs ~= []) = (? y ys. xs = y#ys)"; by (list.induct_tac "xs" 1); -by(simp_tac list_ss 1); -by(asm_simp_tac list_ss 1); -by(REPEAT(resolve_tac [exI,refl,conjI] 1)); +by (simp_tac list_ss 1); +by (asm_simp_tac list_ss 1); +by (REPEAT(resolve_tac [exI,refl,conjI] 1)); qed "neq_Nil_conv"; val list_ss = arith_ss addsimps list.simps @ [null_Nil, null_Cons, hd_Cons, tl_Cons, ttl_Nil, ttl_Cons, mem_Nil, mem_Cons, append_Nil, append_Cons, + rev_Nil, rev_Cons, map_Nil, map_Cons, flat_Nil, flat_Cons, list_all_Nil, list_all_Cons, @@ -45,53 +46,66 @@ goal List.thy "(xs@ys)@zs = xs@(ys@zs)"; by (list.induct_tac "xs" 1); -by(ALLGOALS(asm_simp_tac list_ss)); +by (ALLGOALS (asm_simp_tac list_ss)); qed "append_assoc"; goal List.thy "xs @ [] = xs"; by (list.induct_tac "xs" 1); -by(ALLGOALS(asm_simp_tac list_ss)); +by (ALLGOALS (asm_simp_tac list_ss)); qed "append_Nil2"; goal List.thy "(xs@ys = []) = (xs=[] & ys=[])"; by (list.induct_tac "xs" 1); -by(ALLGOALS(asm_simp_tac list_ss)); +by (ALLGOALS (asm_simp_tac list_ss)); qed "append_is_Nil"; goal List.thy "(xs @ ys = xs @ zs) = (ys=zs)"; by (list.induct_tac "xs" 1); -by(ALLGOALS(asm_simp_tac list_ss)); +by (ALLGOALS (asm_simp_tac list_ss)); qed "same_append_eq"; +(** rev **) + +goal List.thy "rev(xs@ys) = rev(ys) @ rev(xs)"; +by (list.induct_tac "xs" 1); +by (ALLGOALS (asm_simp_tac (list_ss addsimps [append_Nil2,append_assoc]))); +qed "rev_append"; + +goal List.thy "rev(rev l) = l"; +by (list.induct_tac "l" 1); +by (ALLGOALS (asm_simp_tac (list_ss addsimps [rev_append]))); +qed "rev_rev_ident"; + + (** mem **) goal List.thy "x mem (xs@ys) = (x mem xs | x mem ys)"; by (list.induct_tac "xs" 1); -by(ALLGOALS(asm_simp_tac (list_ss setloop (split_tac [expand_if])))); +by (ALLGOALS (asm_simp_tac (list_ss setloop (split_tac [expand_if])))); qed "mem_append"; goal List.thy "x mem [x:xs.P(x)] = (x mem xs & P(x))"; by (list.induct_tac "xs" 1); -by(ALLGOALS(asm_simp_tac (list_ss setloop (split_tac [expand_if])))); +by (ALLGOALS (asm_simp_tac (list_ss setloop (split_tac [expand_if])))); qed "mem_filter"; (** list_all **) goal List.thy "(Alls x:xs.True) = True"; by (list.induct_tac "xs" 1); -by(ALLGOALS(asm_simp_tac list_ss)); +by (ALLGOALS (asm_simp_tac list_ss)); qed "list_all_True"; goal List.thy "list_all p (xs@ys) = (list_all p xs & list_all p ys)"; by (list.induct_tac "xs" 1); -by(ALLGOALS(asm_simp_tac list_ss)); +by (ALLGOALS (asm_simp_tac list_ss)); qed "list_all_conj"; goal List.thy "(Alls x:xs.P(x)) = (!x. x mem xs --> P(x))"; by (list.induct_tac "xs" 1); -by(ALLGOALS(asm_simp_tac (list_ss setloop (split_tac [expand_if])))); -by(fast_tac HOL_cs 1); +by (ALLGOALS (asm_simp_tac (list_ss setloop (split_tac [expand_if])))); +by (fast_tac HOL_cs 1); qed "list_all_mem_conv"; @@ -101,14 +115,14 @@ "P(list_case a f xs) = ((xs=[] --> P(a)) & \ \ (!y ys. xs=y#ys --> P(f y ys)))"; by (list.induct_tac "xs" 1); -by(ALLGOALS(asm_simp_tac list_ss)); -by(fast_tac HOL_cs 1); +by (ALLGOALS (asm_simp_tac list_ss)); +by (fast_tac HOL_cs 1); qed "expand_list_case"; goal List.thy "(xs=[] --> P([])) & (!y ys. xs=y#ys --> P(y#ys)) --> P(xs)"; -by(list.induct_tac "xs" 1); -by(fast_tac HOL_cs 1); -by(fast_tac HOL_cs 1); +by (list.induct_tac "xs" 1); +by (fast_tac HOL_cs 1); +by (fast_tac HOL_cs 1); bind_thm("list_eq_cases", impI RSN (2,allI RSN (2,allI RSN (2,impI RS (conjI RS (result() RS mp)))))); @@ -116,16 +130,21 @@ goal List.thy "flat(xs@ys) = flat(xs)@flat(ys)"; by (list.induct_tac "xs" 1); -by(ALLGOALS(asm_simp_tac (list_ss addsimps [append_assoc]))); +by (ALLGOALS (asm_simp_tac (list_ss addsimps [append_assoc]))); qed"flat_append"; (** length **) goal List.thy "length(xs@ys) = length(xs)+length(ys)"; by (list.induct_tac "xs" 1); -by(ALLGOALS(asm_simp_tac list_ss)); +by (ALLGOALS (asm_simp_tac list_ss)); qed"length_append"; +goal List.thy "length(rev xs) = length(xs)"; +by (list.induct_tac "xs" 1); +by (ALLGOALS (asm_simp_tac (list_ss addsimps [length_append]))); +qed "length_rev"; + (** nth **) val [nth_0,nth_Suc] = nat_recs nth_def; @@ -150,6 +169,17 @@ by (ALLGOALS (asm_simp_tac list_ss)); qed "map_compose"; +goal List.thy "rev(map f l) = map f (rev l)"; +by (list.induct_tac "l" 1); +by (ALLGOALS (asm_simp_tac (list_ss addsimps [map_append]))); +qed "rev_map_distrib"; + +goal List.thy "rev(flat ls) = flat (map rev (rev ls))"; +by (list.induct_tac "ls" 1); +by (ALLGOALS (asm_simp_tac (list_ss addsimps + [map_append, flat_append, rev_append, append_Nil2]))); +qed "rev_flat"; + val list_ss = list_ss addsimps [not_Cons_self, append_assoc, append_Nil2, append_is_Nil, same_append_eq, mem_append, mem_filter, diff -r 74be52691d62 -r 5873833cf37f src/HOL/List.thy --- a/src/HOL/List.thy Thu Jun 29 16:28:40 1995 +0200 +++ b/src/HOL/List.thy Thu Jun 29 16:33:17 1995 +0200 @@ -20,6 +20,7 @@ list_all :: "('a => bool) => ('a list => bool)" map :: "('a=>'b) => ('a list => 'b list)" "@" :: "['a list, 'a list] => 'a list" (infixr 65) + rev :: "'a list => 'a list" filter :: "['a => bool, 'a list] => 'a list" foldl :: "[['b,'a] => 'b, 'b, 'a list] => 'b" length :: "'a list => nat" @@ -66,6 +67,9 @@ primrec "op @" list append_Nil "[] @ ys = ys" append_Cons "(x#xs)@ys = x#(xs@ys)" +primrec rev list + rev_Nil "rev([]) = []" + rev_Cons "rev(x#xs) = rev(xs) @ [x]" primrec filter list filter_Nil "filter P [] = []" filter_Cons "filter P (x#xs) = (if P x then x#filter P xs else filter P xs)"