author | nipkow |
Sun, 02 Apr 1995 10:43:59 +0200 | |
changeset 995 | 95c148a7b9c4 |
parent 974 | 68d58134fad6 |
child 1169 | 5873833cf37f |
permissions | -rw-r--r-- |
923 | 1 |
(* Title: HOL/List |
2 |
ID: $Id$ |
|
3 |
Author: Tobias Nipkow |
|
4 |
Copyright 1994 TU Muenchen |
|
5 |
||
6 |
List lemmas |
|
7 |
*) |
|
8 |
||
9 |
open List; |
|
10 |
||
11 |
val [Nil_not_Cons,Cons_not_Nil] = list.distinct; |
|
12 |
||
13 |
bind_thm("Cons_neq_Nil", Cons_not_Nil RS notE); |
|
14 |
bind_thm("Nil_neq_Cons", sym RS Cons_neq_Nil); |
|
15 |
||
16 |
bind_thm("Cons_inject", (hd list.inject) RS iffD1 RS conjE); |
|
17 |
||
18 |
val list_ss = HOL_ss addsimps list.simps; |
|
19 |
||
20 |
goal List.thy "!x. xs ~= x#xs"; |
|
21 |
by (list.induct_tac "xs" 1); |
|
22 |
by (ALLGOALS (asm_simp_tac list_ss)); |
|
23 |
qed "not_Cons_self"; |
|
24 |
||
25 |
goal List.thy "(xs ~= []) = (? y ys. xs = y#ys)"; |
|
26 |
by (list.induct_tac "xs" 1); |
|
27 |
by(simp_tac list_ss 1); |
|
28 |
by(asm_simp_tac list_ss 1); |
|
29 |
by(REPEAT(resolve_tac [exI,refl,conjI] 1)); |
|
30 |
qed "neq_Nil_conv"; |
|
31 |
||
32 |
val list_ss = arith_ss addsimps list.simps @ |
|
33 |
[null_Nil, null_Cons, hd_Cons, tl_Cons, ttl_Nil, ttl_Cons, |
|
34 |
mem_Nil, mem_Cons, |
|
35 |
append_Nil, append_Cons, |
|
36 |
map_Nil, map_Cons, |
|
37 |
flat_Nil, flat_Cons, |
|
38 |
list_all_Nil, list_all_Cons, |
|
962 | 39 |
filter_Nil, filter_Cons, |
974 | 40 |
foldl_Nil, foldl_Cons, |
962 | 41 |
length_Nil, length_Cons]; |
923 | 42 |
|
43 |
||
44 |
(** @ - append **) |
|
45 |
||
46 |
goal List.thy "(xs@ys)@zs = xs@(ys@zs)"; |
|
47 |
by (list.induct_tac "xs" 1); |
|
48 |
by(ALLGOALS(asm_simp_tac list_ss)); |
|
49 |
qed "append_assoc"; |
|
50 |
||
51 |
goal List.thy "xs @ [] = xs"; |
|
52 |
by (list.induct_tac "xs" 1); |
|
53 |
by(ALLGOALS(asm_simp_tac list_ss)); |
|
54 |
qed "append_Nil2"; |
|
55 |
||
56 |
goal List.thy "(xs@ys = []) = (xs=[] & ys=[])"; |
|
57 |
by (list.induct_tac "xs" 1); |
|
58 |
by(ALLGOALS(asm_simp_tac list_ss)); |
|
59 |
qed "append_is_Nil"; |
|
60 |
||
61 |
goal List.thy "(xs @ ys = xs @ zs) = (ys=zs)"; |
|
62 |
by (list.induct_tac "xs" 1); |
|
63 |
by(ALLGOALS(asm_simp_tac list_ss)); |
|
64 |
qed "same_append_eq"; |
|
65 |
||
66 |
||
67 |
(** mem **) |
|
68 |
||
69 |
goal List.thy "x mem (xs@ys) = (x mem xs | x mem ys)"; |
|
70 |
by (list.induct_tac "xs" 1); |
|
71 |
by(ALLGOALS(asm_simp_tac (list_ss setloop (split_tac [expand_if])))); |
|
72 |
qed "mem_append"; |
|
73 |
||
74 |
goal List.thy "x mem [x:xs.P(x)] = (x mem xs & P(x))"; |
|
75 |
by (list.induct_tac "xs" 1); |
|
76 |
by(ALLGOALS(asm_simp_tac (list_ss setloop (split_tac [expand_if])))); |
|
77 |
qed "mem_filter"; |
|
78 |
||
79 |
(** list_all **) |
|
80 |
||
81 |
goal List.thy "(Alls x:xs.True) = True"; |
|
82 |
by (list.induct_tac "xs" 1); |
|
83 |
by(ALLGOALS(asm_simp_tac list_ss)); |
|
84 |
qed "list_all_True"; |
|
85 |
||
86 |
goal List.thy "list_all p (xs@ys) = (list_all p xs & list_all p ys)"; |
|
87 |
by (list.induct_tac "xs" 1); |
|
88 |
by(ALLGOALS(asm_simp_tac list_ss)); |
|
89 |
qed "list_all_conj"; |
|
90 |
||
91 |
goal List.thy "(Alls x:xs.P(x)) = (!x. x mem xs --> P(x))"; |
|
92 |
by (list.induct_tac "xs" 1); |
|
93 |
by(ALLGOALS(asm_simp_tac (list_ss setloop (split_tac [expand_if])))); |
|
94 |
by(fast_tac HOL_cs 1); |
|
95 |
qed "list_all_mem_conv"; |
|
96 |
||
97 |
||
98 |
(** list_case **) |
|
99 |
||
100 |
goal List.thy |
|
101 |
"P(list_case a f xs) = ((xs=[] --> P(a)) & \ |
|
102 |
\ (!y ys. xs=y#ys --> P(f y ys)))"; |
|
103 |
by (list.induct_tac "xs" 1); |
|
104 |
by(ALLGOALS(asm_simp_tac list_ss)); |
|
105 |
by(fast_tac HOL_cs 1); |
|
106 |
qed "expand_list_case"; |
|
107 |
||
108 |
goal List.thy "(xs=[] --> P([])) & (!y ys. xs=y#ys --> P(y#ys)) --> P(xs)"; |
|
109 |
by(list.induct_tac "xs" 1); |
|
110 |
by(fast_tac HOL_cs 1); |
|
111 |
by(fast_tac HOL_cs 1); |
|
112 |
bind_thm("list_eq_cases", |
|
113 |
impI RSN (2,allI RSN (2,allI RSN (2,impI RS (conjI RS (result() RS mp)))))); |
|
114 |
||
115 |
(** flat **) |
|
116 |
||
117 |
goal List.thy "flat(xs@ys) = flat(xs)@flat(ys)"; |
|
118 |
by (list.induct_tac "xs" 1); |
|
119 |
by(ALLGOALS(asm_simp_tac (list_ss addsimps [append_assoc]))); |
|
120 |
qed"flat_append"; |
|
121 |
||
962 | 122 |
(** length **) |
123 |
||
124 |
goal List.thy "length(xs@ys) = length(xs)+length(ys)"; |
|
125 |
by (list.induct_tac "xs" 1); |
|
126 |
by(ALLGOALS(asm_simp_tac list_ss)); |
|
127 |
qed"length_append"; |
|
128 |
||
923 | 129 |
(** nth **) |
130 |
||
131 |
val [nth_0,nth_Suc] = nat_recs nth_def; |
|
132 |
store_thm("nth_0",nth_0); |
|
133 |
store_thm("nth_Suc",nth_Suc); |
|
134 |
||
135 |
(** Additional mapping lemmas **) |
|
136 |
||
995
95c148a7b9c4
generalized map (%x.x) xs = xs to map (%x.x) = (%xs.xs)
nipkow
parents:
974
diff
changeset
|
137 |
goal List.thy "map (%x.x) = (%xs.xs)"; |
95c148a7b9c4
generalized map (%x.x) xs = xs to map (%x.x) = (%xs.xs)
nipkow
parents:
974
diff
changeset
|
138 |
by (rtac ext 1); |
923 | 139 |
by (list.induct_tac "xs" 1); |
140 |
by (ALLGOALS (asm_simp_tac list_ss)); |
|
141 |
qed "map_ident"; |
|
142 |
||
143 |
goal List.thy "map f (xs@ys) = map f xs @ map f ys"; |
|
144 |
by (list.induct_tac "xs" 1); |
|
145 |
by (ALLGOALS (asm_simp_tac list_ss)); |
|
146 |
qed "map_append"; |
|
147 |
||
148 |
goalw List.thy [o_def] "map (f o g) xs = map f (map g xs)"; |
|
149 |
by (list.induct_tac "xs" 1); |
|
150 |
by (ALLGOALS (asm_simp_tac list_ss)); |
|
151 |
qed "map_compose"; |
|
152 |
||
153 |
val list_ss = list_ss addsimps |
|
154 |
[not_Cons_self, append_assoc, append_Nil2, append_is_Nil, same_append_eq, |
|
155 |
mem_append, mem_filter, |
|
156 |
map_ident, map_append, map_compose, |
|
962 | 157 |
flat_append, length_append, list_all_True, list_all_conj, nth_0, nth_Suc]; |
923 | 158 |