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