# HG changeset patch # User nipkow # Date 861779696 -7200 # Node ID a3b73ba44a11a99b134bfeff1cb631eb1bf148ac # Parent 4be22c30096663edb35f2966fb92ca2e7c17e72f Tidied up. diff -r 4be22c300966 -r a3b73ba44a11 src/HOL/List.ML --- a/src/HOL/List.ML Tue Apr 22 18:05:42 1997 +0200 +++ b/src/HOL/List.ML Wed Apr 23 09:14:56 1997 +0200 @@ -6,20 +6,13 @@ List lemmas *) -open List; - -AddIffs list.distinct; -AddIffs list.inject; - -bind_thm("Cons_inject", (hd list.inject) RS iffD1 RS conjE); - -goal List.thy "!x. xs ~= x#xs"; +goal thy "!x. xs ~= x#xs"; by (list.induct_tac "xs" 1); by (ALLGOALS Asm_simp_tac); qed_spec_mp "not_Cons_self"; Addsimps [not_Cons_self]; -goal List.thy "(xs ~= []) = (? y ys. xs = y#ys)"; +goal thy "(xs ~= []) = (? y ys. xs = y#ys)"; by (list.induct_tac "xs" 1); by (Simp_tac 1); by (Asm_simp_tac 1); @@ -28,7 +21,7 @@ (** list_case **) -goal List.thy +goal thy "P(list_case a f xs) = ((xs=[] --> P(a)) & \ \ (!y ys. xs=y#ys --> P(f y ys)))"; by (list.induct_tac "xs" 1); @@ -36,12 +29,12 @@ by (Blast_tac 1); qed "expand_list_case"; -val prems = goal List.thy "[| P([]); !!x xs. P(x#xs) |] ==> P(xs)"; +val prems = goal thy "[| P([]); !!x xs. P(x#xs) |] ==> P(xs)"; by(list.induct_tac "xs" 1); by(REPEAT(resolve_tac prems 1)); qed "list_cases"; -goal List.thy "(xs=[] --> P([])) & (!y ys. xs=y#ys --> P(y#ys)) --> P(xs)"; +goal thy "(xs=[] --> P([])) & (!y ys. xs=y#ys --> P(y#ys)) --> P(xs)"; by (list.induct_tac "xs" 1); by (Blast_tac 1); by (Blast_tac 1); @@ -51,38 +44,38 @@ (** @ - append **) -goal List.thy "(xs@ys)@zs = xs@(ys@zs)"; +goal thy "(xs@ys)@zs = xs@(ys@zs)"; by (list.induct_tac "xs" 1); by (ALLGOALS Asm_simp_tac); qed "append_assoc"; Addsimps [append_assoc]; -goal List.thy "xs @ [] = xs"; +goal thy "xs @ [] = xs"; by (list.induct_tac "xs" 1); by (ALLGOALS Asm_simp_tac); qed "append_Nil2"; Addsimps [append_Nil2]; -goal List.thy "(xs@ys = []) = (xs=[] & ys=[])"; +goal thy "(xs@ys = []) = (xs=[] & ys=[])"; by (list.induct_tac "xs" 1); by (ALLGOALS Asm_simp_tac); qed "append_is_Nil_conv"; AddIffs [append_is_Nil_conv]; -goal List.thy "([] = xs@ys) = (xs=[] & ys=[])"; +goal thy "([] = xs@ys) = (xs=[] & ys=[])"; by (list.induct_tac "xs" 1); by (ALLGOALS Asm_simp_tac); by(Blast_tac 1); qed "Nil_is_append_conv"; AddIffs [Nil_is_append_conv]; -goal List.thy "(xs @ ys = xs @ zs) = (ys=zs)"; +goal thy "(xs @ ys = xs @ zs) = (ys=zs)"; by (list.induct_tac "xs" 1); by (ALLGOALS Asm_simp_tac); qed "same_append_eq"; AddIffs [same_append_eq]; -goal List.thy "!ys. (xs @ [x] = ys @ [y]) = (xs = ys & x = y)"; +goal thy "!ys. (xs @ [x] = ys @ [y]) = (xs = ys & x = y)"; by(list.induct_tac "xs" 1); br allI 1; by(list.induct_tac "ys" 1); @@ -93,62 +86,62 @@ qed_spec_mp "append1_eq_conv"; AddIffs [append1_eq_conv]; -goal List.thy "xs ~= [] --> hd xs # tl xs = xs"; +goal thy "xs ~= [] --> hd xs # tl xs = xs"; by(list.induct_tac "xs" 1); by(ALLGOALS Asm_simp_tac); qed_spec_mp "hd_Cons_tl"; Addsimps [hd_Cons_tl]; -goal List.thy "hd(xs@ys) = (if xs=[] then hd ys else hd xs)"; +goal thy "hd(xs@ys) = (if xs=[] then hd ys else hd xs)"; by (list.induct_tac "xs" 1); by (ALLGOALS Asm_simp_tac); qed "hd_append"; -goal List.thy "tl(xs@ys) = (case xs of [] => tl(ys) | z#zs => zs@ys)"; +goal thy "tl(xs@ys) = (case xs of [] => tl(ys) | z#zs => zs@ys)"; by(simp_tac (!simpset setloop(split_tac[expand_list_case])) 1); qed "tl_append"; (** map **) -goal List.thy +goal thy "(!x. x : set_of_list xs --> f x = g x) --> map f xs = map g xs"; by(list.induct_tac "xs" 1); by(ALLGOALS Asm_simp_tac); bind_thm("map_ext", impI RS (allI RS (result() RS mp))); -goal List.thy "map (%x.x) = (%xs.xs)"; +goal thy "map (%x.x) = (%xs.xs)"; by (rtac ext 1); by (list.induct_tac "xs" 1); by (ALLGOALS Asm_simp_tac); qed "map_ident"; Addsimps[map_ident]; -goal List.thy "map f (xs@ys) = map f xs @ map f ys"; +goal thy "map f (xs@ys) = map f xs @ map f ys"; by (list.induct_tac "xs" 1); by (ALLGOALS Asm_simp_tac); qed "map_append"; Addsimps[map_append]; -goalw List.thy [o_def] "map (f o g) xs = map f (map g xs)"; +goalw thy [o_def] "map (f o g) xs = map f (map g xs)"; by (list.induct_tac "xs" 1); by (ALLGOALS Asm_simp_tac); qed "map_compose"; Addsimps[map_compose]; -goal List.thy "rev(map f xs) = map f (rev xs)"; +goal thy "rev(map f xs) = map f (rev xs)"; by (list.induct_tac "xs" 1); by (ALLGOALS Asm_simp_tac); qed "rev_map"; (** rev **) -goal List.thy "rev(xs@ys) = rev(ys) @ rev(xs)"; +goal thy "rev(xs@ys) = rev(ys) @ rev(xs)"; by (list.induct_tac "xs" 1); by (ALLGOALS Asm_simp_tac); qed "rev_append"; Addsimps[rev_append]; -goal List.thy "rev(rev l) = l"; +goal thy "rev(rev l) = l"; by (list.induct_tac "l" 1); by (ALLGOALS Asm_simp_tac); qed "rev_rev_ident"; @@ -157,13 +150,13 @@ (** mem **) -goal List.thy "x mem (xs@ys) = (x mem xs | x mem ys)"; +goal thy "x mem (xs@ys) = (x mem xs | x mem ys)"; by (list.induct_tac "xs" 1); by (ALLGOALS (asm_simp_tac (!simpset setloop (split_tac [expand_if])))); qed "mem_append"; Addsimps[mem_append]; -goal List.thy "x mem [x:xs.P(x)] = (x mem xs & P(x))"; +goal thy "x mem [x:xs.P(x)] = (x mem xs & P(x))"; by (list.induct_tac "xs" 1); by (ALLGOALS (asm_simp_tac (!simpset setloop (split_tac [expand_if])))); qed "mem_filter"; @@ -184,25 +177,25 @@ by (Blast_tac 1); qed "set_of_list_mem_eq"; -goal List.thy "set_of_list l <= set_of_list (x#l)"; +goal thy "set_of_list l <= set_of_list (x#l)"; by (Simp_tac 1); by (Blast_tac 1); qed "set_of_list_subset_Cons"; -goal List.thy "(set_of_list xs = {}) = (xs = [])"; +goal thy "(set_of_list xs = {}) = (xs = [])"; by(list.induct_tac "xs" 1); by(ALLGOALS Asm_simp_tac); qed "set_of_list_empty"; Addsimps [set_of_list_empty]; -goal List.thy "set_of_list(rev xs) = set_of_list(xs)"; +goal thy "set_of_list(rev xs) = set_of_list(xs)"; by(list.induct_tac "xs" 1); by(ALLGOALS Asm_simp_tac); by(Blast_tac 1); qed "set_of_list_rev"; Addsimps [set_of_list_rev]; -goal List.thy "set_of_list(map f xs) = f``(set_of_list xs)"; +goal thy "set_of_list(map f xs) = f``(set_of_list xs)"; by(list.induct_tac "xs" 1); by(ALLGOALS Asm_simp_tac); qed "set_of_list_map"; @@ -211,19 +204,19 @@ (** list_all **) -goal List.thy "list_all (%x.True) xs = True"; +goal thy "list_all (%x.True) xs = True"; by (list.induct_tac "xs" 1); by (ALLGOALS Asm_simp_tac); qed "list_all_True"; Addsimps [list_all_True]; -goal List.thy "list_all p (xs@ys) = (list_all p xs & list_all p ys)"; +goal 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); qed "list_all_append"; Addsimps [list_all_append]; -goal List.thy "list_all P xs = (!x. x mem xs --> P(x))"; +goal thy "list_all P xs = (!x. x mem xs --> P(x))"; by (list.induct_tac "xs" 1); by (ALLGOALS (asm_simp_tac (!simpset setloop (split_tac [expand_if])))); by (Blast_tac 1); @@ -232,7 +225,7 @@ (** filter **) -goal List.thy "[x:xs@ys . P] = [x:xs . P] @ [y:ys . P]"; +goal thy "[x:xs@ys . P] = [x:xs . P] @ [y:ys . P]"; by(list.induct_tac "xs" 1); by(ALLGOALS (asm_simp_tac (!simpset setloop (split_tac [expand_if])))); qed "filter_append"; @@ -241,44 +234,44 @@ (** concat **) -goal List.thy "concat(xs@ys) = concat(xs)@concat(ys)"; +goal thy "concat(xs@ys) = concat(xs)@concat(ys)"; by (list.induct_tac "xs" 1); by (ALLGOALS Asm_simp_tac); qed"concat_append"; Addsimps [concat_append]; -goal List.thy "rev(concat ls) = concat (map rev (rev ls))"; +goal thy "rev(concat ls) = concat (map rev (rev ls))"; by (list.induct_tac "ls" 1); by (ALLGOALS Asm_simp_tac); qed "rev_concat"; (** length **) -goal List.thy "length(xs@ys) = length(xs)+length(ys)"; +goal thy "length(xs@ys) = length(xs)+length(ys)"; by (list.induct_tac "xs" 1); by (ALLGOALS Asm_simp_tac); qed"length_append"; Addsimps [length_append]; -goal List.thy "length (map f l) = length l"; +goal thy "length (map f l) = length l"; by (list.induct_tac "l" 1); by (ALLGOALS Simp_tac); qed "length_map"; Addsimps [length_map]; -goal List.thy "length(rev xs) = length(xs)"; +goal thy "length(rev xs) = length(xs)"; by (list.induct_tac "xs" 1); by (ALLGOALS Asm_simp_tac); qed "length_rev"; Addsimps [length_rev]; -goal List.thy "(length xs = 0) = (xs = [])"; +goal thy "(length xs = 0) = (xs = [])"; by(list.induct_tac "xs" 1); by(ALLGOALS Asm_simp_tac); qed "length_0_conv"; AddIffs [length_0_conv]; -goal List.thy "(0 < length xs) = (xs ~= [])"; +goal thy "(0 < length xs) = (xs ~= [])"; by(list.induct_tac "xs" 1); by(ALLGOALS Asm_simp_tac); qed "length_greater_0_conv"; @@ -287,7 +280,7 @@ (** nth **) -goal List.thy +goal thy "!xs. nth n (xs@ys) = \ \ (if n < length xs then nth n xs else nth (n - length xs) ys)"; by(nat_ind_tac "n" 1); @@ -300,7 +293,7 @@ by(ALLGOALS Asm_simp_tac); qed_spec_mp "nth_append"; -goal List.thy "!n. n < length xs --> nth n (map f xs) = f (nth n xs)"; +goal thy "!n. n < length xs --> nth n (map f xs) = f (nth n xs)"; by (list.induct_tac "xs" 1); (* case [] *) by (Asm_full_simp_tac 1); @@ -311,7 +304,7 @@ qed_spec_mp "nth_map"; Addsimps [nth_map]; -goal List.thy "!n. n < length xs --> list_all P xs --> P(nth n xs)"; +goal thy "!n. n < length xs --> list_all P xs --> P(nth n xs)"; by (list.induct_tac "xs" 1); (* case [] *) by (Simp_tac 1); @@ -321,7 +314,7 @@ by (ALLGOALS Asm_full_simp_tac); qed_spec_mp "list_all_nth"; -goal List.thy "!n. n < length xs --> (nth n xs) mem xs"; +goal thy "!n. n < length xs --> (nth n xs) mem xs"; by (list.induct_tac "xs" 1); (* case [] *) by (Simp_tac 1); @@ -360,7 +353,7 @@ Delsimps [take_Cons,drop_Cons]; Addsimps [take_0,take_Suc_Cons,drop_0,drop_Suc_Cons]; -goal List.thy "!xs. length(take n xs) = min (length xs) n"; +goal thy "!xs. length(take n xs) = min (length xs) n"; by(nat_ind_tac "n" 1); by(ALLGOALS Asm_simp_tac); br allI 1; @@ -369,7 +362,7 @@ qed_spec_mp "length_take"; Addsimps [length_take]; -goal List.thy "!xs. length(drop n xs) = (length xs - n)"; +goal thy "!xs. length(drop n xs) = (length xs - n)"; by(nat_ind_tac "n" 1); by(ALLGOALS Asm_simp_tac); br allI 1; @@ -378,7 +371,7 @@ qed_spec_mp "length_drop"; Addsimps [length_drop]; -goal List.thy "!xs. length xs <= n --> take n xs = xs"; +goal thy "!xs. length xs <= n --> take n xs = xs"; by(nat_ind_tac "n" 1); by(ALLGOALS Asm_simp_tac); br allI 1; @@ -386,7 +379,7 @@ by(ALLGOALS Asm_simp_tac); qed_spec_mp "take_all"; -goal List.thy "!xs. length xs <= n --> drop n xs = []"; +goal thy "!xs. length xs <= n --> drop n xs = []"; by(nat_ind_tac "n" 1); by(ALLGOALS Asm_simp_tac); br allI 1; @@ -394,7 +387,7 @@ by(ALLGOALS Asm_simp_tac); qed_spec_mp "drop_all"; -goal List.thy +goal thy "!xs. take n (xs @ ys) = (take n xs @ take (n - length xs) ys)"; by(nat_ind_tac "n" 1); by(ALLGOALS Asm_simp_tac); @@ -404,7 +397,7 @@ qed_spec_mp "take_append"; Addsimps [take_append]; -goal List.thy "!xs. drop n (xs@ys) = drop n xs @ drop (n - length xs) ys"; +goal thy "!xs. drop n (xs@ys) = drop n xs @ drop (n - length xs) ys"; by(nat_ind_tac "n" 1); by(ALLGOALS Asm_simp_tac); br allI 1; @@ -413,7 +406,7 @@ qed_spec_mp "drop_append"; Addsimps [drop_append]; -goal List.thy "!xs n. take n (take m xs) = take (min n m) xs"; +goal thy "!xs n. take n (take m xs) = take (min n m) xs"; by(nat_ind_tac "m" 1); by(ALLGOALS Asm_simp_tac); br allI 1; @@ -424,7 +417,7 @@ by(ALLGOALS Asm_simp_tac); qed_spec_mp "take_take"; -goal List.thy "!xs. drop n (drop m xs) = drop (n + m) xs"; +goal thy "!xs. drop n (drop m xs) = drop (n + m) xs"; by(nat_ind_tac "m" 1); by(ALLGOALS Asm_simp_tac); br allI 1; @@ -432,7 +425,7 @@ by(ALLGOALS Asm_simp_tac); qed_spec_mp "drop_drop"; -goal List.thy "!xs n. take n (drop m xs) = drop m (take (n + m) xs)"; +goal thy "!xs n. take n (drop m xs) = drop m (take (n + m) xs)"; by(nat_ind_tac "m" 1); by(ALLGOALS Asm_simp_tac); br allI 1; @@ -440,7 +433,7 @@ by(ALLGOALS Asm_simp_tac); qed_spec_mp "take_drop"; -goal List.thy "!xs. take n (map f xs) = map f (take n xs)"; +goal thy "!xs. take n (map f xs) = map f (take n xs)"; by(nat_ind_tac "n" 1); by(ALLGOALS Asm_simp_tac); br allI 1; @@ -448,7 +441,7 @@ by(ALLGOALS Asm_simp_tac); qed_spec_mp "take_map"; -goal List.thy "!xs. drop n (map f xs) = map f (drop n xs)"; +goal thy "!xs. drop n (map f xs) = map f (drop n xs)"; by(nat_ind_tac "n" 1); by(ALLGOALS Asm_simp_tac); br allI 1; @@ -456,7 +449,7 @@ by(ALLGOALS Asm_simp_tac); qed_spec_mp "drop_map"; -goal List.thy +goal thy "!n i. i < n --> nth i (take n xs) = nth i xs"; by(list.induct_tac "xs" 1); by(ALLGOALS Asm_simp_tac); @@ -468,7 +461,7 @@ qed_spec_mp "nth_take"; Addsimps [nth_take]; -goal List.thy +goal thy "!xs i. n + i < length xs --> nth i (drop n xs) = nth (n + i) xs"; by(nat_ind_tac "n" 1); by(ALLGOALS Asm_simp_tac); @@ -480,7 +473,7 @@ (** takeWhile & dropWhile **) -goal List.thy +goal thy "x:set_of_list xs & ~P(x) --> takeWhile P (xs @ ys) = takeWhile P xs"; by(list.induct_tac "xs" 1); by(Simp_tac 1); @@ -489,7 +482,7 @@ bind_thm("takeWhile_append1", conjI RS (result() RS mp)); Addsimps [takeWhile_append1]; -goal List.thy +goal thy "(!x:set_of_list xs.P(x)) --> takeWhile P (xs @ ys) = xs @ takeWhile P ys"; by(list.induct_tac "xs" 1); by(Simp_tac 1); @@ -497,7 +490,7 @@ bind_thm("takeWhile_append2", ballI RS (result() RS mp)); Addsimps [takeWhile_append2]; -goal List.thy +goal thy "x:set_of_list xs & ~P(x) --> dropWhile P (xs @ ys) = (dropWhile P xs)@ys"; by(list.induct_tac "xs" 1); by(Simp_tac 1); @@ -506,7 +499,7 @@ bind_thm("dropWhile_append1", conjI RS (result() RS mp)); Addsimps [dropWhile_append1]; -goal List.thy +goal thy "(!x:set_of_list xs.P(x)) --> dropWhile P (xs @ ys) = dropWhile P ys"; by(list.induct_tac "xs" 1); by(Simp_tac 1); @@ -514,7 +507,7 @@ bind_thm("dropWhile_append2", ballI RS (result() RS mp)); Addsimps [dropWhile_append2]; -goal List.thy "x:set_of_list(takeWhile P xs) --> x:set_of_list xs & P x"; +goal thy "x:set_of_list(takeWhile P xs) --> x:set_of_list xs & P x"; by(list.induct_tac "xs" 1); by(Simp_tac 1); by(asm_full_simp_tac (!simpset setloop (split_tac[expand_if])) 1);