author | nipkow |
Sun, 29 Jan 1995 14:02:17 +0100 | |
changeset 204 | 21c405b4039f |
parent 203 | d465d3be2744 |
child 210 | 1a3d3b5b5d15 |
permissions | -rw-r--r-- |
113 | 1 |
(* Title: HOL/List |
0 | 2 |
ID: $Id$ |
196
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
171
diff
changeset
|
3 |
Author: Tobias Nipkow |
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
171
diff
changeset
|
4 |
Copyright 1994 TU Muenchen |
0 | 5 |
|
196
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
171
diff
changeset
|
6 |
List lemmas |
0 | 7 |
*) |
8 |
||
9 |
open List; |
|
10 |
||
196
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
171
diff
changeset
|
11 |
val [Nil_not_Cons,Cons_not_Nil] = list.distinct; |
0 | 12 |
|
202 | 13 |
bind_thm("Cons_neq_Nil", Cons_not_Nil RS notE); |
199 | 14 |
bind_thm("Nil_neq_Cons", sym RS Cons_neq_Nil); |
0 | 15 |
|
202 | 16 |
bind_thm("Cons_inject", (hd list.inject) RS iffD1 RS conjE); |
0 | 17 |
|
196
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
171
diff
changeset
|
18 |
val list_ss = HOL_ss addsimps list.simps; |
0 | 19 |
|
196
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
171
diff
changeset
|
20 |
goal List.thy "!x. xs ~= x#xs"; |
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
171
diff
changeset
|
21 |
by (list.induct_tac "xs" 1); |
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
171
diff
changeset
|
22 |
by (ALLGOALS (asm_simp_tac list_ss)); |
171 | 23 |
qed "not_Cons_self"; |
0 | 24 |
|
48
21291189b51e
changed "." to "$" and Cons to infix "#" to eliminate ambiguity
clasohm
parents:
44
diff
changeset
|
25 |
goal List.thy "(xs ~= []) = (? y ys. xs = y#ys)"; |
196
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
171
diff
changeset
|
26 |
by (list.induct_tac "xs" 1); |
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
171
diff
changeset
|
27 |
by(simp_tac list_ss 1); |
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
171
diff
changeset
|
28 |
by(asm_simp_tac list_ss 1); |
40 | 29 |
by(REPEAT(resolve_tac [exI,refl,conjI] 1)); |
171 | 30 |
qed "neq_Nil_conv"; |
40 | 31 |
|
196
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
171
diff
changeset
|
32 |
val list_ss = arith_ss addsimps list.simps @ |
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
171
diff
changeset
|
33 |
[null_Nil, null_Cons, hd_Cons, tl_Cons, ttl_Nil, ttl_Cons, |
34 | 34 |
mem_Nil, mem_Cons, |
35 |
append_Nil, append_Cons, |
|
36 |
map_Nil, map_Cons, |
|
37 |
list_all_Nil, list_all_Cons, |
|
38 |
filter_Nil, filter_Cons]; |
|
39 |
||
0 | 40 |
|
13 | 41 |
(** @ - append **) |
42 |
||
43 |
goal List.thy "(xs@ys)@zs = xs@(ys@zs)"; |
|
196
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
171
diff
changeset
|
44 |
by (list.induct_tac "xs" 1); |
34 | 45 |
by(ALLGOALS(asm_simp_tac list_ss)); |
171 | 46 |
qed "append_assoc"; |
0 | 47 |
|
40 | 48 |
goal List.thy "xs @ [] = xs"; |
196
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
171
diff
changeset
|
49 |
by (list.induct_tac "xs" 1); |
40 | 50 |
by(ALLGOALS(asm_simp_tac list_ss)); |
171 | 51 |
qed "append_Nil2"; |
40 | 52 |
|
203 | 53 |
goal List.thy "(xs@ys = []) = (xs=[] & ys=[])"; |
54 |
by (list.induct_tac "xs" 1); |
|
55 |
by(ALLGOALS(asm_simp_tac list_ss)); |
|
56 |
qed "append_is_Nil"; |
|
57 |
||
34 | 58 |
(** mem **) |
59 |
||
60 |
goal List.thy "x mem (xs@ys) = (x mem xs | x mem ys)"; |
|
196
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
171
diff
changeset
|
61 |
by (list.induct_tac "xs" 1); |
34 | 62 |
by(ALLGOALS(asm_simp_tac (list_ss setloop (split_tac [expand_if])))); |
171 | 63 |
qed "mem_append"; |
34 | 64 |
|
65 |
goal List.thy "x mem [x:xs.P(x)] = (x mem xs & P(x))"; |
|
196
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
171
diff
changeset
|
66 |
by (list.induct_tac "xs" 1); |
34 | 67 |
by(ALLGOALS(asm_simp_tac (list_ss setloop (split_tac [expand_if])))); |
171 | 68 |
qed "mem_filter"; |
34 | 69 |
|
70 |
(** list_all **) |
|
0 | 71 |
|
34 | 72 |
goal List.thy "(Alls x:xs.True) = True"; |
196
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
171
diff
changeset
|
73 |
by (list.induct_tac "xs" 1); |
34 | 74 |
by(ALLGOALS(asm_simp_tac list_ss)); |
171 | 75 |
qed "list_all_True"; |
0 | 76 |
|
34 | 77 |
goal List.thy "list_all(p,xs@ys) = (list_all(p,xs) & list_all(p,ys))"; |
196
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
171
diff
changeset
|
78 |
by (list.induct_tac "xs" 1); |
34 | 79 |
by(ALLGOALS(asm_simp_tac list_ss)); |
171 | 80 |
qed "list_all_conj"; |
34 | 81 |
|
82 |
goal List.thy "(Alls x:xs.P(x)) = (!x. x mem xs --> P(x))"; |
|
196
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
171
diff
changeset
|
83 |
by (list.induct_tac "xs" 1); |
34 | 84 |
by(ALLGOALS(asm_simp_tac (list_ss setloop (split_tac [expand_if])))); |
85 |
by(fast_tac HOL_cs 1); |
|
171 | 86 |
qed "list_all_mem_conv"; |
34 | 87 |
|
88 |
||
40 | 89 |
(** list_case **) |
90 |
||
91 |
goal List.thy |
|
113 | 92 |
"P(list_case(a,f,xs)) = ((xs=[] --> P(a)) & \ |
48
21291189b51e
changed "." to "$" and Cons to infix "#" to eliminate ambiguity
clasohm
parents:
44
diff
changeset
|
93 |
\ (!y ys. xs=y#ys --> P(f(y,ys))))"; |
196
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
171
diff
changeset
|
94 |
by (list.induct_tac "xs" 1); |
44 | 95 |
by(ALLGOALS(asm_simp_tac list_ss)); |
40 | 96 |
by(fast_tac HOL_cs 1); |
171 | 97 |
qed "expand_list_case"; |
40 | 98 |
|
203 | 99 |
(** nth **) |
100 |
||
101 |
val [nth_0,nth_Suc] = nat_recs nth_def; |
|
102 |
store_thm("nth_0",nth_0); |
|
103 |
store_thm("nth_Suc",nth_Suc); |
|
0 | 104 |
|
105 |
(** Additional mapping lemmas **) |
|
106 |
||
107 |
goal List.thy "map(%x.x, xs) = xs"; |
|
196
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
171
diff
changeset
|
108 |
by (list.induct_tac "xs" 1); |
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
171
diff
changeset
|
109 |
by (ALLGOALS (asm_simp_tac list_ss)); |
171 | 110 |
qed "map_ident"; |
0 | 111 |
|
83 | 112 |
goal List.thy "map(f, xs@ys) = map(f,xs) @ map(f,ys)"; |
196
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
171
diff
changeset
|
113 |
by (list.induct_tac "xs" 1); |
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
171
diff
changeset
|
114 |
by (ALLGOALS (asm_simp_tac list_ss)); |
171 | 115 |
qed "map_append"; |
83 | 116 |
|
117 |
goalw List.thy [o_def] "map(f o g, xs) = map(f, map(g, xs))"; |
|
196
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
171
diff
changeset
|
118 |
by (list.induct_tac "xs" 1); |
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
171
diff
changeset
|
119 |
by (ALLGOALS (asm_simp_tac list_ss)); |
171 | 120 |
qed "map_compose"; |
83 | 121 |
|
34 | 122 |
val list_ss = list_ss addsimps |
203 | 123 |
[not_Cons_self, append_assoc, append_Nil2, append_is_Nil, |
124 |
mem_append, mem_filter, |
|
196
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
171
diff
changeset
|
125 |
map_ident, map_append, map_compose, |
203 | 126 |
list_all_True, list_all_conj, nth_0, nth_Suc]; |
40 | 127 |