author | clasohm |
Wed, 02 Nov 1994 11:50:09 +0100 | |
changeset 156 | fd1be45b64bf |
parent 128 | 89669c58e506 |
child 171 | 16c4ea954511 |
permissions | -rw-r--r-- |
113 | 1 |
(* Title: HOL/List |
0 | 2 |
ID: $Id$ |
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
4 |
Copyright 1993 University of Cambridge |
|
5 |
||
128 | 6 |
Definition of type 'a list by a least fixed point |
0 | 7 |
*) |
8 |
||
9 |
open List; |
|
10 |
||
128 | 11 |
val list_con_defs = [NIL_def, CONS_def]; |
0 | 12 |
|
128 | 13 |
goal List.thy "list(A) = {Numb(0)} <+> (A <*> list(A))"; |
14 |
let val rew = rewrite_rule list_con_defs in |
|
15 |
by (fast_tac (univ_cs addSIs (equalityI :: map rew list.intrs) |
|
16 |
addEs [rew list.elim]) 1) |
|
17 |
end; |
|
18 |
val list_unfold = result(); |
|
0 | 19 |
|
128 | 20 |
(*This justifies using list in other recursive type definitions*) |
21 |
goalw List.thy list.defs "!!A B. A<=B ==> list(A) <= list(B)"; |
|
22 |
by (rtac lfp_mono 1); |
|
23 |
by (REPEAT (ares_tac basic_monos 1)); |
|
24 |
val list_mono = result(); |
|
0 | 25 |
|
128 | 26 |
(*Type checking -- list creates well-founded sets*) |
27 |
goalw List.thy (list_con_defs @ list.defs) "list(sexp) <= sexp"; |
|
28 |
by (rtac lfp_lowerbound 1); |
|
29 |
by (fast_tac (univ_cs addIs sexp.intrs@[sexp_In0I,sexp_In1I]) 1); |
|
30 |
val list_sexp = result(); |
|
31 |
||
32 |
(* A <= sexp ==> list(A) <= sexp *) |
|
33 |
val list_subset_sexp = standard ([list_mono, list_sexp] MRS subset_trans); |
|
0 | 34 |
|
35 |
(*Induction for the type 'a list *) |
|
36 |
val prems = goalw List.thy [Nil_def,Cons_def] |
|
37 |
"[| P(Nil); \ |
|
48
21291189b51e
changed "." to "$" and Cons to infix "#" to eliminate ambiguity
clasohm
parents:
44
diff
changeset
|
38 |
\ !!x xs. P(xs) ==> P(x # xs) |] ==> P(l)"; |
128 | 39 |
by (rtac (Rep_list_inverse RS subst) 1); (*types force good instantiation*) |
40 |
by (rtac (Rep_list RS list.induct) 1); |
|
0 | 41 |
by (REPEAT (ares_tac prems 1 |
128 | 42 |
ORELSE eresolve_tac [rangeE, ssubst, Abs_list_inverse RS subst] 1)); |
0 | 43 |
val list_induct = result(); |
44 |
||
45 |
(*Perform induction on xs. *) |
|
46 |
fun list_ind_tac a M = |
|
47 |
EVERY [res_inst_tac [("l",a)] list_induct M, |
|
48 |
rename_last_tac a ["1"] (M+1)]; |
|
49 |
||
50 |
(*** Isomorphisms ***) |
|
51 |
||
128 | 52 |
goal List.thy "inj(Rep_list)"; |
0 | 53 |
by (rtac inj_inverseI 1); |
128 | 54 |
by (rtac Rep_list_inverse 1); |
55 |
val inj_Rep_list = result(); |
|
0 | 56 |
|
128 | 57 |
goal List.thy "inj_onto(Abs_list,list(range(Leaf)))"; |
0 | 58 |
by (rtac inj_onto_inverseI 1); |
128 | 59 |
by (etac Abs_list_inverse 1); |
60 |
val inj_onto_Abs_list = result(); |
|
0 | 61 |
|
62 |
(** Distinctness of constructors **) |
|
63 |
||
128 | 64 |
goalw List.thy list_con_defs "CONS(M,N) ~= NIL"; |
0 | 65 |
by (rtac In1_not_In0 1); |
66 |
val CONS_not_NIL = result(); |
|
67 |
val NIL_not_CONS = standard (CONS_not_NIL RS not_sym); |
|
68 |
||
69 |
val CONS_neq_NIL = standard (CONS_not_NIL RS notE); |
|
70 |
val NIL_neq_CONS = sym RS CONS_neq_NIL; |
|
71 |
||
48
21291189b51e
changed "." to "$" and Cons to infix "#" to eliminate ambiguity
clasohm
parents:
44
diff
changeset
|
72 |
goalw List.thy [Nil_def,Cons_def] "x # xs ~= Nil"; |
128 | 73 |
by (rtac (CONS_not_NIL RS (inj_onto_Abs_list RS inj_onto_contraD)) 1); |
74 |
by (REPEAT (resolve_tac (list.intrs @ [rangeI, Rep_list]) 1)); |
|
0 | 75 |
val Cons_not_Nil = result(); |
76 |
||
77 |
val Nil_not_Cons = standard (Cons_not_Nil RS not_sym); |
|
78 |
||
79 |
val Cons_neq_Nil = standard (Cons_not_Nil RS notE); |
|
80 |
val Nil_neq_Cons = sym RS Cons_neq_Nil; |
|
81 |
||
82 |
(** Injectiveness of CONS and Cons **) |
|
83 |
||
84 |
goalw List.thy [CONS_def] "(CONS(K,M)=CONS(L,N)) = (K=L & M=N)"; |
|
85 |
by (fast_tac (HOL_cs addSEs [Scons_inject, make_elim In1_inject]) 1); |
|
86 |
val CONS_CONS_eq = result(); |
|
87 |
||
88 |
val CONS_inject = standard (CONS_CONS_eq RS iffD1 RS conjE); |
|
89 |
||
90 |
(*For reasoning about abstract list constructors*) |
|
128 | 91 |
val list_cs = set_cs addIs [Rep_list] @ list.intrs |
0 | 92 |
addSEs [CONS_neq_NIL,NIL_neq_CONS,CONS_inject] |
128 | 93 |
addSDs [inj_onto_Abs_list RS inj_ontoD, |
94 |
inj_Rep_list RS injD, Leaf_inject]; |
|
0 | 95 |
|
48
21291189b51e
changed "." to "$" and Cons to infix "#" to eliminate ambiguity
clasohm
parents:
44
diff
changeset
|
96 |
goalw List.thy [Cons_def] "(x#xs=y#ys) = (x=y & xs=ys)"; |
128 | 97 |
by (fast_tac list_cs 1); |
0 | 98 |
val Cons_Cons_eq = result(); |
99 |
val Cons_inject = standard (Cons_Cons_eq RS iffD1 RS conjE); |
|
100 |
||
128 | 101 |
val [major] = goal List.thy "CONS(M,N): list(A) ==> M: A & N: list(A)"; |
0 | 102 |
by (rtac (major RS setup_induction) 1); |
128 | 103 |
by (etac list.induct 1); |
104 |
by (ALLGOALS (fast_tac list_cs)); |
|
0 | 105 |
val CONS_D = result(); |
106 |
||
107 |
val prems = goalw List.thy [CONS_def,In1_def] |
|
128 | 108 |
"CONS(M,N): sexp ==> M: sexp & N: sexp"; |
0 | 109 |
by (cut_facts_tac prems 1); |
110 |
by (fast_tac (set_cs addSDs [Scons_D]) 1); |
|
128 | 111 |
val sexp_CONS_D = result(); |
0 | 112 |
|
113 |
||
114 |
(*Basic ss with constructors and their freeness*) |
|
115 |
val list_free_simps = [Cons_not_Nil, Nil_not_Cons, Cons_Cons_eq, |
|
128 | 116 |
CONS_not_NIL, NIL_not_CONS, CONS_CONS_eq] |
117 |
@ list.intrs; |
|
0 | 118 |
val list_free_ss = HOL_ss addsimps list_free_simps; |
119 |
||
128 | 120 |
goal List.thy "!!N. N: list(A) ==> !M. N ~= CONS(M,N)"; |
121 |
by (etac list.induct 1); |
|
0 | 122 |
by (ALLGOALS (asm_simp_tac list_free_ss)); |
123 |
val not_CONS_self = result(); |
|
124 |
||
48
21291189b51e
changed "." to "$" and Cons to infix "#" to eliminate ambiguity
clasohm
parents:
44
diff
changeset
|
125 |
goal List.thy "!x. l ~= x#l"; |
0 | 126 |
by (list_ind_tac "l" 1); |
127 |
by (ALLGOALS (asm_simp_tac list_free_ss)); |
|
128 |
val not_Cons_self = result(); |
|
129 |
||
130 |
||
48
21291189b51e
changed "." to "$" and Cons to infix "#" to eliminate ambiguity
clasohm
parents:
44
diff
changeset
|
131 |
goal List.thy "(xs ~= []) = (? y ys. xs = y#ys)"; |
40 | 132 |
by(list_ind_tac "xs" 1); |
133 |
by(simp_tac list_free_ss 1); |
|
134 |
by(asm_simp_tac list_free_ss 1); |
|
135 |
by(REPEAT(resolve_tac [exI,refl,conjI] 1)); |
|
136 |
val neq_Nil_conv = result(); |
|
137 |
||
0 | 138 |
(** Conversion rules for List_case: case analysis operator **) |
139 |
||
113 | 140 |
goalw List.thy [List_case_def,NIL_def] "List_case(c, h, NIL) = c"; |
0 | 141 |
by (rtac Case_In0 1); |
142 |
val List_case_NIL = result(); |
|
143 |
||
113 | 144 |
goalw List.thy [List_case_def,CONS_def] "List_case(c, h, CONS(M,N)) = h(M,N)"; |
0 | 145 |
by (simp_tac (HOL_ss addsimps [Split,Case_In1]) 1); |
146 |
val List_case_CONS = result(); |
|
147 |
||
128 | 148 |
(*** List_rec -- by wf recursion on pred_sexp ***) |
149 |
||
150 |
(* The trancl(pred_sexp) is essential because pred_sexp_CONS_I1,2 would not |
|
151 |
hold if pred_sexp^+ were changed to pred_sexp. *) |
|
0 | 152 |
|
128 | 153 |
val List_rec_unfold = [List_rec_def, wf_pred_sexp RS wf_trancl] MRS def_wfrec |
154 |
|> standard; |
|
0 | 155 |
|
128 | 156 |
(** pred_sexp lemmas **) |
0 | 157 |
|
158 |
goalw List.thy [CONS_def,In1_def] |
|
128 | 159 |
"!!M. [| M: sexp; N: sexp |] ==> <M, CONS(M,N)> : pred_sexp^+"; |
160 |
by (asm_simp_tac pred_sexp_ss 1); |
|
161 |
val pred_sexp_CONS_I1 = result(); |
|
0 | 162 |
|
163 |
goalw List.thy [CONS_def,In1_def] |
|
128 | 164 |
"!!M. [| M: sexp; N: sexp |] ==> <N, CONS(M,N)> : pred_sexp^+"; |
165 |
by (asm_simp_tac pred_sexp_ss 1); |
|
166 |
val pred_sexp_CONS_I2 = result(); |
|
0 | 167 |
|
168 |
val [prem] = goal List.thy |
|
128 | 169 |
"<CONS(M1,M2), N> : pred_sexp^+ ==> \ |
170 |
\ <M1,N> : pred_sexp^+ & <M2,N> : pred_sexp^+"; |
|
171 |
by (rtac (prem RS (pred_sexp_subset_Sigma RS trancl_subset_Sigma RS |
|
0 | 172 |
subsetD RS SigmaE2)) 1); |
128 | 173 |
by (etac (sexp_CONS_D RS conjE) 1); |
174 |
by (REPEAT (ares_tac [conjI, pred_sexp_CONS_I1, pred_sexp_CONS_I2, |
|
0 | 175 |
prem RSN (2, trans_trancl RS transD)] 1)); |
128 | 176 |
val pred_sexp_CONS_D = result(); |
0 | 177 |
|
178 |
(** Conversion rules for List_rec **) |
|
179 |
||
180 |
goal List.thy "List_rec(NIL,c,h) = c"; |
|
181 |
by (rtac (List_rec_unfold RS trans) 1); |
|
113 | 182 |
by (simp_tac (HOL_ss addsimps [List_case_NIL]) 1); |
0 | 183 |
val List_rec_NIL = result(); |
184 |
||
128 | 185 |
goal List.thy "!!M. [| M: sexp; N: sexp |] ==> \ |
0 | 186 |
\ List_rec(CONS(M,N), c, h) = h(M, N, List_rec(N,c,h))"; |
187 |
by (rtac (List_rec_unfold RS trans) 1); |
|
113 | 188 |
by (asm_simp_tac |
128 | 189 |
(HOL_ss addsimps [List_case_CONS, list.CONS_I, pred_sexp_CONS_I2, |
190 |
cut_apply])1); |
|
0 | 191 |
val List_rec_CONS = result(); |
192 |
||
193 |
(*** list_rec -- by List_rec ***) |
|
194 |
||
128 | 195 |
val Rep_list_in_sexp = |
196 |
[range_Leaf_subset_sexp RS list_subset_sexp, Rep_list] MRS subsetD; |
|
0 | 197 |
|
198 |
local |
|
199 |
val list_rec_simps = list_free_simps @ |
|
200 |
[List_rec_NIL, List_rec_CONS, |
|
128 | 201 |
Abs_list_inverse, Rep_list_inverse, |
202 |
Rep_list, rangeI, inj_Leaf, Inv_f_f, |
|
203 |
sexp.LeafI, Rep_list_in_sexp] |
|
0 | 204 |
in |
205 |
val list_rec_Nil = prove_goalw List.thy [list_rec_def, Nil_def] |
|
206 |
"list_rec(Nil,c,h) = c" |
|
207 |
(fn _=> [simp_tac (HOL_ss addsimps list_rec_simps) 1]); |
|
208 |
||
209 |
val list_rec_Cons = prove_goalw List.thy [list_rec_def, Cons_def] |
|
48
21291189b51e
changed "." to "$" and Cons to infix "#" to eliminate ambiguity
clasohm
parents:
44
diff
changeset
|
210 |
"list_rec(a#l, c, h) = h(a, l, list_rec(l,c,h))" |
0 | 211 |
(fn _=> [simp_tac (HOL_ss addsimps list_rec_simps) 1]); |
212 |
end; |
|
213 |
||
214 |
val list_simps = [List_rec_NIL, List_rec_CONS, |
|
215 |
list_rec_Nil, list_rec_Cons]; |
|
216 |
val list_ss = list_free_ss addsimps list_simps; |
|
217 |
||
218 |
||
219 |
(*Type checking. Useful?*) |
|
128 | 220 |
val major::A_subset_sexp::prems = goal List.thy |
221 |
"[| M: list(A); \ |
|
222 |
\ A<=sexp; \ |
|
0 | 223 |
\ c: C(NIL); \ |
128 | 224 |
\ !!x y r. [| x: A; y: list(A); r: C(y) |] ==> h(x,y,r): C(CONS(x,y)) \ |
225 |
\ |] ==> List_rec(M,c,h) : C(M :: 'a item)"; |
|
226 |
val sexp_ListA_I = A_subset_sexp RS list_subset_sexp RS subsetD; |
|
227 |
val sexp_A_I = A_subset_sexp RS subsetD; |
|
228 |
by (rtac (major RS list.induct) 1); |
|
229 |
by (ALLGOALS(asm_simp_tac (list_ss addsimps ([sexp_A_I,sexp_ListA_I]@prems)))); |
|
0 | 230 |
val List_rec_type = result(); |
231 |
||
232 |
(** Generalized map functionals **) |
|
233 |
||
234 |
goalw List.thy [Rep_map_def] "Rep_map(f,Nil) = NIL"; |
|
235 |
by (rtac list_rec_Nil 1); |
|
236 |
val Rep_map_Nil = result(); |
|
237 |
||
238 |
goalw List.thy [Rep_map_def] |
|
48
21291189b51e
changed "." to "$" and Cons to infix "#" to eliminate ambiguity
clasohm
parents:
44
diff
changeset
|
239 |
"Rep_map(f, x#xs) = CONS(f(x), Rep_map(f,xs))"; |
0 | 240 |
by (rtac list_rec_Cons 1); |
241 |
val Rep_map_Cons = result(); |
|
242 |
||
128 | 243 |
goalw List.thy [Rep_map_def] "!!f. (!!x. f(x): A) ==> Rep_map(f,xs): list(A)"; |
0 | 244 |
by (rtac list_induct 1); |
245 |
by(ALLGOALS(asm_simp_tac list_ss)); |
|
246 |
val Rep_map_type = result(); |
|
247 |
||
248 |
goalw List.thy [Abs_map_def] "Abs_map(g,NIL) = Nil"; |
|
249 |
by (rtac List_rec_NIL 1); |
|
250 |
val Abs_map_NIL = result(); |
|
251 |
||
252 |
val prems = goalw List.thy [Abs_map_def] |
|
128 | 253 |
"[| M: sexp; N: sexp |] ==> \ |
48
21291189b51e
changed "." to "$" and Cons to infix "#" to eliminate ambiguity
clasohm
parents:
44
diff
changeset
|
254 |
\ Abs_map(g, CONS(M,N)) = g(M) # Abs_map(g,N)"; |
0 | 255 |
by (REPEAT (resolve_tac (List_rec_CONS::prems) 1)); |
256 |
val Abs_map_CONS = result(); |
|
257 |
||
34 | 258 |
(*These 2 rules ease the use of primitive recursion. NOTE USE OF == *) |
259 |
val [rew] = goal List.thy |
|
260 |
"[| !!xs. f(xs) == list_rec(xs,c,h) |] ==> f([]) = c"; |
|
261 |
by (rewtac rew); |
|
0 | 262 |
by (rtac list_rec_Nil 1); |
34 | 263 |
val def_list_rec_Nil = result(); |
0 | 264 |
|
34 | 265 |
val [rew] = goal List.thy |
48
21291189b51e
changed "." to "$" and Cons to infix "#" to eliminate ambiguity
clasohm
parents:
44
diff
changeset
|
266 |
"[| !!xs. f(xs) == list_rec(xs,c,h) |] ==> f(x#xs) = h(x,xs,f(xs))"; |
34 | 267 |
by (rewtac rew); |
0 | 268 |
by (rtac list_rec_Cons 1); |
34 | 269 |
val def_list_rec_Cons = result(); |
0 | 270 |
|
34 | 271 |
fun list_recs def = |
272 |
[standard (def RS def_list_rec_Nil), |
|
273 |
standard (def RS def_list_rec_Cons)]; |
|
0 | 274 |
|
34 | 275 |
(*** Unfolding the basic combinators ***) |
0 | 276 |
|
34 | 277 |
val [null_Nil,null_Cons] = list_recs null_def; |
278 |
val [_,hd_Cons] = list_recs hd_def; |
|
279 |
val [_,tl_Cons] = list_recs tl_def; |
|
40 | 280 |
val [ttl_Nil,ttl_Cons] = list_recs ttl_def; |
34 | 281 |
val [append_Nil,append_Cons] = list_recs append_def; |
282 |
val [mem_Nil, mem_Cons] = list_recs mem_def; |
|
283 |
val [map_Nil,map_Cons] = list_recs map_def; |
|
284 |
val [list_case_Nil,list_case_Cons] = list_recs list_case_def; |
|
285 |
val [filter_Nil,filter_Cons] = list_recs filter_def; |
|
286 |
val [list_all_Nil,list_all_Cons] = list_recs list_all_def; |
|
0 | 287 |
|
34 | 288 |
val list_ss = arith_ss addsimps |
289 |
[Cons_not_Nil, Nil_not_Cons, Cons_Cons_eq, |
|
290 |
list_rec_Nil, list_rec_Cons, |
|
40 | 291 |
null_Nil, null_Cons, hd_Cons, tl_Cons, ttl_Nil, ttl_Cons, |
34 | 292 |
mem_Nil, mem_Cons, |
293 |
list_case_Nil, list_case_Cons, |
|
294 |
append_Nil, append_Cons, |
|
295 |
map_Nil, map_Cons, |
|
296 |
list_all_Nil, list_all_Cons, |
|
297 |
filter_Nil, filter_Cons]; |
|
298 |
||
0 | 299 |
|
13 | 300 |
(** @ - append **) |
301 |
||
302 |
goal List.thy "(xs@ys)@zs = xs@(ys@zs)"; |
|
34 | 303 |
by(list_ind_tac "xs" 1); |
304 |
by(ALLGOALS(asm_simp_tac list_ss)); |
|
13 | 305 |
val append_assoc = result(); |
0 | 306 |
|
40 | 307 |
goal List.thy "xs @ [] = xs"; |
308 |
by(list_ind_tac "xs" 1); |
|
309 |
by(ALLGOALS(asm_simp_tac list_ss)); |
|
310 |
val append_Nil2 = result(); |
|
311 |
||
34 | 312 |
(** mem **) |
313 |
||
314 |
goal List.thy "x mem (xs@ys) = (x mem xs | x mem ys)"; |
|
315 |
by(list_ind_tac "xs" 1); |
|
316 |
by(ALLGOALS(asm_simp_tac (list_ss setloop (split_tac [expand_if])))); |
|
317 |
val mem_append = result(); |
|
318 |
||
319 |
goal List.thy "x mem [x:xs.P(x)] = (x mem xs & P(x))"; |
|
320 |
by(list_ind_tac "xs" 1); |
|
321 |
by(ALLGOALS(asm_simp_tac (list_ss setloop (split_tac [expand_if])))); |
|
322 |
val mem_filter = result(); |
|
323 |
||
324 |
(** list_all **) |
|
0 | 325 |
|
34 | 326 |
goal List.thy "(Alls x:xs.True) = True"; |
327 |
by(list_ind_tac "xs" 1); |
|
328 |
by(ALLGOALS(asm_simp_tac list_ss)); |
|
329 |
val list_all_True = result(); |
|
0 | 330 |
|
34 | 331 |
goal List.thy "list_all(p,xs@ys) = (list_all(p,xs) & list_all(p,ys))"; |
332 |
by(list_ind_tac "xs" 1); |
|
333 |
by(ALLGOALS(asm_simp_tac list_ss)); |
|
334 |
val list_all_conj = result(); |
|
335 |
||
336 |
goal List.thy "(Alls x:xs.P(x)) = (!x. x mem xs --> P(x))"; |
|
337 |
by(list_ind_tac "xs" 1); |
|
338 |
by(ALLGOALS(asm_simp_tac (list_ss setloop (split_tac [expand_if])))); |
|
339 |
by(fast_tac HOL_cs 1); |
|
340 |
val list_all_mem_conv = result(); |
|
341 |
||
342 |
||
343 |
(** The functional "map" **) |
|
0 | 344 |
|
345 |
val map_simps = [Abs_map_NIL, Abs_map_CONS, |
|
346 |
Rep_map_Nil, Rep_map_Cons, |
|
347 |
map_Nil, map_Cons]; |
|
348 |
val map_ss = list_free_ss addsimps map_simps; |
|
349 |
||
128 | 350 |
val [major,A_subset_sexp,minor] = goal List.thy |
351 |
"[| M: list(A); A<=sexp; !!z. z: A ==> f(g(z)) = z |] \ |
|
0 | 352 |
\ ==> Rep_map(f, Abs_map(g,M)) = M"; |
128 | 353 |
by (rtac (major RS list.induct) 1); |
354 |
by (ALLGOALS (asm_simp_tac(map_ss addsimps [sexp_A_I,sexp_ListA_I,minor]))); |
|
0 | 355 |
val Abs_map_inverse = result(); |
356 |
||
357 |
(*Rep_map_inverse is obtained via Abs_Rep_map and map_ident*) |
|
358 |
||
40 | 359 |
(** list_case **) |
360 |
||
361 |
goal List.thy |
|
113 | 362 |
"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
|
363 |
\ (!y ys. xs=y#ys --> P(f(y,ys))))"; |
40 | 364 |
by(list_ind_tac "xs" 1); |
44 | 365 |
by(ALLGOALS(asm_simp_tac list_ss)); |
40 | 366 |
by(fast_tac HOL_cs 1); |
367 |
val expand_list_case = result(); |
|
368 |
||
0 | 369 |
|
370 |
(** Additional mapping lemmas **) |
|
371 |
||
372 |
goal List.thy "map(%x.x, xs) = xs"; |
|
373 |
by (list_ind_tac "xs" 1); |
|
374 |
by (ALLGOALS (asm_simp_tac map_ss)); |
|
375 |
val map_ident = result(); |
|
376 |
||
83 | 377 |
goal List.thy "map(f, xs@ys) = map(f,xs) @ map(f,ys)"; |
378 |
by (list_ind_tac "xs" 1); |
|
379 |
by (ALLGOALS (asm_simp_tac (map_ss addsimps [append_Nil,append_Cons]))); |
|
380 |
val map_append = result(); |
|
381 |
||
382 |
goalw List.thy [o_def] "map(f o g, xs) = map(f, map(g, xs))"; |
|
383 |
by (list_ind_tac "xs" 1); |
|
384 |
by (ALLGOALS (asm_simp_tac map_ss)); |
|
385 |
val map_compose = result(); |
|
386 |
||
128 | 387 |
goal List.thy "!!f. (!!x. f(x): sexp) ==> \ |
0 | 388 |
\ Abs_map(g, Rep_map(f,xs)) = map(%t. g(f(t)), xs)"; |
389 |
by (list_ind_tac "xs" 1); |
|
390 |
by(ALLGOALS(asm_simp_tac(map_ss addsimps |
|
128 | 391 |
[Rep_map_type,list_sexp RS subsetD]))); |
0 | 392 |
val Abs_Rep_map = result(); |
20 | 393 |
|
34 | 394 |
val list_ss = list_ss addsimps |
40 | 395 |
[mem_append, mem_filter, append_assoc, append_Nil2, map_ident, |
34 | 396 |
list_all_True, list_all_conj]; |
40 | 397 |