| author | paulson | 
| Tue, 21 Oct 1997 10:36:23 +0200 | |
| changeset 3960 | 7a38fae985f9 | 
| parent 3840 | e0baea4d485a | 
| child 4091 | 771b1f6422a8 | 
| permissions | -rw-r--r-- | 
| 1461 | 1 | (* Title: ZF/List.ML | 
| 0 | 2 | ID: $Id$ | 
| 1461 | 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 0 | 4 | Copyright 1993 University of Cambridge | 
| 5 | ||
| 6 | Datatype definition of Lists | |
| 7 | *) | |
| 8 | ||
| 516 | 9 | open List; | 
| 0 | 10 | |
| 516 | 11 | (*** Aspects of the datatype definition ***) | 
| 0 | 12 | |
| 2469 | 13 | Addsimps list.case_eqns; | 
| 14 | ||
| 15 | ||
| 0 | 16 | (*An elimination rule, for type-checking*) | 
| 516 | 17 | val ConsE = list.mk_cases list.con_defs "Cons(a,l) : list(A)"; | 
| 0 | 18 | |
| 19 | (*Proving freeness results*) | |
| 516 | 20 | val Cons_iff = list.mk_free "Cons(a,l)=Cons(a',l') <-> a=a' & l=l'"; | 
| 21 | val Nil_Cons_iff = list.mk_free "~ Nil=Cons(a,l)"; | |
| 0 | 22 | |
| 23 | (*Perform induction on l, then prove the major premise using prems. *) | |
| 24 | fun list_ind_tac a prems i = | |
| 516 | 25 |     EVERY [res_inst_tac [("x",a)] list.induct i,
 | 
| 1461 | 26 | rename_last_tac a ["1"] (i+2), | 
| 27 | ares_tac prems i]; | |
| 0 | 28 | |
| 435 | 29 | goal List.thy "list(A) = {0} + (A * list(A))";
 | 
| 525 
e62519a8497d
ZF/List, ex/Brouwer,Data,LList,Ntree,TF,Term: much simplified proof of _unfold
 lcp parents: 
516diff
changeset | 30 | let open list; val rew = rewrite_rule con_defs in | 
| 3016 | 31 | by (blast_tac (!claset addSIs (map rew intrs) addEs [rew elim]) 1) | 
| 525 
e62519a8497d
ZF/List, ex/Brouwer,Data,LList,Ntree,TF,Term: much simplified proof of _unfold
 lcp parents: 
516diff
changeset | 32 | end; | 
| 760 | 33 | qed "list_unfold"; | 
| 435 | 34 | |
| 0 | 35 | (** Lemmas to justify using "list" in other recursive type definitions **) | 
| 36 | ||
| 516 | 37 | goalw List.thy list.defs "!!A B. A<=B ==> list(A) <= list(B)"; | 
| 0 | 38 | by (rtac lfp_mono 1); | 
| 516 | 39 | by (REPEAT (rtac list.bnd_mono 1)); | 
| 0 | 40 | by (REPEAT (ares_tac (univ_mono::basic_monos) 1)); | 
| 760 | 41 | qed "list_mono"; | 
| 0 | 42 | |
| 43 | (*There is a similar proof by list induction.*) | |
| 516 | 44 | goalw List.thy (list.defs@list.con_defs) "list(univ(A)) <= univ(A)"; | 
| 0 | 45 | by (rtac lfp_lowerbound 1); | 
| 46 | by (rtac (A_subset_univ RS univ_mono) 2); | |
| 3016 | 47 | by (blast_tac (!claset addSIs [zero_in_univ, Inl_in_univ, Inr_in_univ, | 
| 1461 | 48 | Pair_in_univ]) 1); | 
| 760 | 49 | qed "list_univ"; | 
| 0 | 50 | |
| 908 
1f99a44c10cb
Updated comment about list_subset_univ and list_into_univ.
 lcp parents: 
782diff
changeset | 51 | (*These two theorems justify datatypes involving list(nat), list(A), ...*) | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 52 | bind_thm ("list_subset_univ", ([list_mono, list_univ] MRS subset_trans));
 | 
| 0 | 53 | |
| 435 | 54 | goal List.thy "!!l A B. [| l: list(A); A <= univ(B) |] ==> l: univ(B)"; | 
| 55 | by (REPEAT (ares_tac [list_subset_univ RS subsetD] 1)); | |
| 760 | 56 | qed "list_into_univ"; | 
| 435 | 57 | |
| 0 | 58 | val major::prems = goal List.thy | 
| 59 | "[| l: list(A); \ | |
| 15 
6c6d2f6e3185
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 60 | \ c: C(Nil); \ | 
| 
6c6d2f6e3185
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 61 | \ !!x y. [| x: A; y: list(A) |] ==> h(x,y): C(Cons(x,y)) \ | 
| 
6c6d2f6e3185
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 62 | \ |] ==> list_case(c,h,l) : C(l)"; | 
| 516 | 63 | by (rtac (major RS list.induct) 1); | 
| 2469 | 64 | by (ALLGOALS (asm_simp_tac (!simpset addsimps (list.case_eqns @ prems)))); | 
| 760 | 65 | qed "list_case_type"; | 
| 0 | 66 | |
| 67 | ||
| 68 | (** For recursion **) | |
| 69 | ||
| 516 | 70 | goalw List.thy list.con_defs "rank(a) < rank(Cons(a,l))"; | 
| 6 
8ce8c4d13d4d
Installation of new simplifier for ZF.  Deleted all congruence rules not
 lcp parents: 
0diff
changeset | 71 | by (simp_tac rank_ss 1); | 
| 760 | 72 | qed "rank_Cons1"; | 
| 0 | 73 | |
| 516 | 74 | goalw List.thy list.con_defs "rank(l) < rank(Cons(a,l))"; | 
| 6 
8ce8c4d13d4d
Installation of new simplifier for ZF.  Deleted all congruence rules not
 lcp parents: 
0diff
changeset | 75 | by (simp_tac rank_ss 1); | 
| 760 | 76 | qed "rank_Cons2"; | 
| 0 | 77 | |
| 516 | 78 | |
| 79 | (*** List functions ***) | |
| 80 | ||
| 81 | (** hd and tl **) | |
| 82 | ||
| 83 | goalw List.thy [hd_def] "hd(Cons(a,l)) = a"; | |
| 84 | by (resolve_tac list.case_eqns 1); | |
| 760 | 85 | qed "hd_Cons"; | 
| 516 | 86 | |
| 87 | goalw List.thy [tl_def] "tl(Nil) = Nil"; | |
| 88 | by (resolve_tac list.case_eqns 1); | |
| 760 | 89 | qed "tl_Nil"; | 
| 516 | 90 | |
| 91 | goalw List.thy [tl_def] "tl(Cons(a,l)) = l"; | |
| 92 | by (resolve_tac list.case_eqns 1); | |
| 760 | 93 | qed "tl_Cons"; | 
| 516 | 94 | |
| 2469 | 95 | Addsimps [hd_Cons, tl_Nil, tl_Cons]; | 
| 96 | ||
| 516 | 97 | goal List.thy "!!l. l: list(A) ==> tl(l) : list(A)"; | 
| 98 | by (etac list.elim 1); | |
| 2469 | 99 | by (ALLGOALS (asm_simp_tac (!simpset addsimps list.intrs))); | 
| 760 | 100 | qed "tl_type"; | 
| 516 | 101 | |
| 102 | (** drop **) | |
| 103 | ||
| 104 | goalw List.thy [drop_def] "drop(0, l) = l"; | |
| 105 | by (rtac rec_0 1); | |
| 760 | 106 | qed "drop_0"; | 
| 516 | 107 | |
| 108 | goalw List.thy [drop_def] "!!i. i:nat ==> drop(i, Nil) = Nil"; | |
| 109 | by (etac nat_induct 1); | |
| 2469 | 110 | by (ALLGOALS Asm_simp_tac); | 
| 760 | 111 | qed "drop_Nil"; | 
| 516 | 112 | |
| 113 | goalw List.thy [drop_def] | |
| 114 | "!!i. i:nat ==> drop(succ(i), Cons(a,l)) = drop(i,l)"; | |
| 2493 | 115 | by (rtac sym 1); | 
| 516 | 116 | by (etac nat_induct 1); | 
| 2469 | 117 | by (Simp_tac 1); | 
| 118 | by (Asm_simp_tac 1); | |
| 760 | 119 | qed "drop_succ_Cons"; | 
| 516 | 120 | |
| 2469 | 121 | Addsimps [drop_0, drop_Nil, drop_succ_Cons]; | 
| 122 | ||
| 516 | 123 | goalw List.thy [drop_def] | 
| 124 | "!!i l. [| i:nat; l: list(A) |] ==> drop(i,l) : list(A)"; | |
| 125 | by (etac nat_induct 1); | |
| 2469 | 126 | by (ALLGOALS (asm_simp_tac (!simpset addsimps [tl_type]))); | 
| 760 | 127 | qed "drop_type"; | 
| 516 | 128 | |
| 129 | (** list_rec -- by Vset recursion **) | |
| 130 | ||
| 131 | goal List.thy "list_rec(Nil,c,h) = c"; | |
| 132 | by (rtac (list_rec_def RS def_Vrec RS trans) 1); | |
| 2469 | 133 | by (simp_tac (!simpset addsimps list.case_eqns) 1); | 
| 760 | 134 | qed "list_rec_Nil"; | 
| 516 | 135 | |
| 136 | goal List.thy "list_rec(Cons(a,l), c, h) = h(a, l, list_rec(l,c,h))"; | |
| 137 | by (rtac (list_rec_def RS def_Vrec RS trans) 1); | |
| 138 | by (simp_tac (rank_ss addsimps (rank_Cons2::list.case_eqns)) 1); | |
| 760 | 139 | qed "list_rec_Cons"; | 
| 516 | 140 | |
| 2469 | 141 | Addsimps [list_rec_Nil, list_rec_Cons]; | 
| 142 | ||
| 143 | ||
| 516 | 144 | (*Type checking -- proved by induction, as usual*) | 
| 145 | val prems = goal List.thy | |
| 146 | "[| l: list(A); \ | |
| 147 | \ c: C(Nil); \ | |
| 148 | \ !!x y r. [| x:A; y: list(A); r: C(y) |] ==> h(x,y,r): C(Cons(x,y)) \ | |
| 149 | \ |] ==> list_rec(l,c,h) : C(l)"; | |
| 150 | by (list_ind_tac "l" prems 1); | |
| 2469 | 151 | by (ALLGOALS (asm_simp_tac (!simpset addsimps prems))); | 
| 760 | 152 | qed "list_rec_type"; | 
| 516 | 153 | |
| 154 | (** Versions for use with definitions **) | |
| 155 | ||
| 156 | val [rew] = goal List.thy | |
| 157 | "[| !!l. j(l)==list_rec(l, c, h) |] ==> j(Nil) = c"; | |
| 158 | by (rewtac rew); | |
| 159 | by (rtac list_rec_Nil 1); | |
| 760 | 160 | qed "def_list_rec_Nil"; | 
| 516 | 161 | |
| 162 | val [rew] = goal List.thy | |
| 163 | "[| !!l. j(l)==list_rec(l, c, h) |] ==> j(Cons(a,l)) = h(a,l,j(l))"; | |
| 164 | by (rewtac rew); | |
| 165 | by (rtac list_rec_Cons 1); | |
| 760 | 166 | qed "def_list_rec_Cons"; | 
| 516 | 167 | |
| 168 | fun list_recs def = map standard | |
| 1461 | 169 | ([def] RL [def_list_rec_Nil, def_list_rec_Cons]); | 
| 516 | 170 | |
| 171 | (** map **) | |
| 172 | ||
| 173 | val [map_Nil,map_Cons] = list_recs map_def; | |
| 2469 | 174 | Addsimps [map_Nil, map_Cons]; | 
| 516 | 175 | |
| 176 | val prems = goalw List.thy [map_def] | |
| 177 | "[| l: list(A); !!x. x: A ==> h(x): B |] ==> map(h,l) : list(B)"; | |
| 178 | by (REPEAT (ares_tac (prems @ list.intrs @ [list_rec_type]) 1)); | |
| 760 | 179 | qed "map_type"; | 
| 516 | 180 | |
| 181 | val [major] = goal List.thy "l: list(A) ==> map(h,l) : list({h(u). u:A})";
 | |
| 182 | by (rtac (major RS map_type) 1); | |
| 183 | by (etac RepFunI 1); | |
| 760 | 184 | qed "map_type2"; | 
| 516 | 185 | |
| 186 | (** length **) | |
| 187 | ||
| 188 | val [length_Nil,length_Cons] = list_recs length_def; | |
| 2469 | 189 | Addsimps [length_Nil,length_Cons]; | 
| 516 | 190 | |
| 191 | goalw List.thy [length_def] | |
| 192 | "!!l. l: list(A) ==> length(l) : nat"; | |
| 193 | by (REPEAT (ares_tac [list_rec_type, nat_0I, nat_succI] 1)); | |
| 760 | 194 | qed "length_type"; | 
| 516 | 195 | |
| 196 | (** app **) | |
| 197 | ||
| 198 | val [app_Nil,app_Cons] = list_recs app_def; | |
| 2469 | 199 | Addsimps [app_Nil, app_Cons]; | 
| 516 | 200 | |
| 201 | goalw List.thy [app_def] | |
| 202 | "!!xs ys. [| xs: list(A); ys: list(A) |] ==> xs@ys : list(A)"; | |
| 203 | by (REPEAT (ares_tac [list_rec_type, list.Cons_I] 1)); | |
| 760 | 204 | qed "app_type"; | 
| 516 | 205 | |
| 206 | (** rev **) | |
| 207 | ||
| 208 | val [rev_Nil,rev_Cons] = list_recs rev_def; | |
| 2469 | 209 | Addsimps [rev_Nil,rev_Cons] ; | 
| 516 | 210 | |
| 211 | goalw List.thy [rev_def] | |
| 212 | "!!xs. xs: list(A) ==> rev(xs) : list(A)"; | |
| 213 | by (REPEAT (ares_tac (list.intrs @ [list_rec_type, app_type]) 1)); | |
| 760 | 214 | qed "rev_type"; | 
| 516 | 215 | |
| 216 | ||
| 217 | (** flat **) | |
| 218 | ||
| 219 | val [flat_Nil,flat_Cons] = list_recs flat_def; | |
| 2469 | 220 | Addsimps [flat_Nil,flat_Cons]; | 
| 516 | 221 | |
| 222 | goalw List.thy [flat_def] | |
| 223 | "!!ls. ls: list(list(A)) ==> flat(ls) : list(A)"; | |
| 224 | by (REPEAT (ares_tac (list.intrs @ [list_rec_type, app_type]) 1)); | |
| 760 | 225 | qed "flat_type"; | 
| 516 | 226 | |
| 227 | ||
| 1926 | 228 | (** set_of_list **) | 
| 229 | ||
| 230 | val [set_of_list_Nil,set_of_list_Cons] = list_recs set_of_list_def; | |
| 2469 | 231 | Addsimps [set_of_list_Nil,set_of_list_Cons]; | 
| 1926 | 232 | |
| 233 | goalw List.thy [set_of_list_def] | |
| 234 | "!!l. l: list(A) ==> set_of_list(l) : Pow(A)"; | |
| 2033 | 235 | by (etac list_rec_type 1); | 
| 3016 | 236 | by (ALLGOALS (Blast_tac)); | 
| 1926 | 237 | qed "set_of_list_type"; | 
| 238 | ||
| 239 | goal List.thy | |
| 240 | "!!l. xs: list(A) ==> \ | |
| 241 | \ set_of_list (xs@ys) = set_of_list(xs) Un set_of_list(ys)"; | |
| 242 | by (etac list.induct 1); | |
| 2469 | 243 | by (ALLGOALS (asm_simp_tac (!simpset addsimps [Un_cons]))); | 
| 1926 | 244 | qed "set_of_list_append"; | 
| 245 | ||
| 246 | ||
| 516 | 247 | (** list_add **) | 
| 248 | ||
| 249 | val [list_add_Nil,list_add_Cons] = list_recs list_add_def; | |
| 2469 | 250 | Addsimps [list_add_Nil,list_add_Cons]; | 
| 516 | 251 | |
| 252 | goalw List.thy [list_add_def] | |
| 253 | "!!xs. xs: list(nat) ==> list_add(xs) : nat"; | |
| 254 | by (REPEAT (ares_tac [list_rec_type, nat_0I, add_type] 1)); | |
| 760 | 255 | qed "list_add_type"; | 
| 516 | 256 | |
| 257 | val list_typechecks = | |
| 258 | list.intrs @ | |
| 259 | [list_rec_type, map_type, map_type2, app_type, length_type, | |
| 260 | rev_type, flat_type, list_add_type]; | |
| 261 | ||
| 2637 
e9b203f854ae
reflecting my recent changes of the simplifier and classical reasoner
 oheimb parents: 
2496diff
changeset | 262 | simpset := !simpset setSolver (type_auto_tac list_typechecks); | 
| 516 | 263 | |
| 264 | ||
| 265 | (*** theorems about map ***) | |
| 266 | ||
| 267 | val prems = goal List.thy | |
| 3840 | 268 | "l: list(A) ==> map(%u. u, l) = l"; | 
| 516 | 269 | by (list_ind_tac "l" prems 1); | 
| 3016 | 270 | by (ALLGOALS Asm_simp_tac); | 
| 760 | 271 | qed "map_ident"; | 
| 516 | 272 | |
| 273 | val prems = goal List.thy | |
| 3840 | 274 | "l: list(A) ==> map(h, map(j,l)) = map(%u. h(j(u)), l)"; | 
| 516 | 275 | by (list_ind_tac "l" prems 1); | 
| 3016 | 276 | by (ALLGOALS Asm_simp_tac); | 
| 760 | 277 | qed "map_compose"; | 
| 516 | 278 | |
| 279 | val prems = goal List.thy | |
| 280 | "xs: list(A) ==> map(h, xs@ys) = map(h,xs) @ map(h,ys)"; | |
| 281 | by (list_ind_tac "xs" prems 1); | |
| 3016 | 282 | by (ALLGOALS Asm_simp_tac); | 
| 760 | 283 | qed "map_app_distrib"; | 
| 516 | 284 | |
| 285 | val prems = goal List.thy | |
| 286 | "ls: list(list(A)) ==> map(h, flat(ls)) = flat(map(map(h),ls))"; | |
| 287 | by (list_ind_tac "ls" prems 1); | |
| 2469 | 288 | by (ALLGOALS (asm_simp_tac (!simpset addsimps [map_app_distrib]))); | 
| 760 | 289 | qed "map_flat"; | 
| 516 | 290 | |
| 291 | val prems = goal List.thy | |
| 292 | "l: list(A) ==> \ | |
| 293 | \ list_rec(map(h,l), c, d) = \ | |
| 294 | \ list_rec(l, c, %x xs r. d(h(x), map(h,xs), r))"; | |
| 295 | by (list_ind_tac "l" prems 1); | |
| 3016 | 296 | by (ALLGOALS Asm_simp_tac); | 
| 760 | 297 | qed "list_rec_map"; | 
| 516 | 298 | |
| 299 | (** theorems about list(Collect(A,P)) -- used in ex/term.ML **) | |
| 300 | ||
| 301 | (* c : list(Collect(B,P)) ==> c : list(B) *) | |
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 302 | bind_thm ("list_CollectD", (Collect_subset RS list_mono RS subsetD));
 | 
| 516 | 303 | |
| 304 | val prems = goal List.thy | |
| 305 |     "l: list({x:A. h(x)=j(x)}) ==> map(h,l) = map(j,l)";
 | |
| 306 | by (list_ind_tac "l" prems 1); | |
| 3016 | 307 | by (ALLGOALS Asm_simp_tac); | 
| 760 | 308 | qed "map_list_Collect"; | 
| 516 | 309 | |
| 310 | (*** theorems about length ***) | |
| 311 | ||
| 312 | val prems = goal List.thy | |
| 313 | "xs: list(A) ==> length(map(h,xs)) = length(xs)"; | |
| 314 | by (list_ind_tac "xs" prems 1); | |
| 3016 | 315 | by (ALLGOALS Asm_simp_tac); | 
| 760 | 316 | qed "length_map"; | 
| 516 | 317 | |
| 318 | val prems = goal List.thy | |
| 319 | "xs: list(A) ==> length(xs@ys) = length(xs) #+ length(ys)"; | |
| 320 | by (list_ind_tac "xs" prems 1); | |
| 3016 | 321 | by (ALLGOALS Asm_simp_tac); | 
| 760 | 322 | qed "length_app"; | 
| 516 | 323 | |
| 324 | (* [| m: nat; n: nat |] ==> m #+ succ(n) = succ(n) #+ m | |
| 325 | Used for rewriting below*) | |
| 326 | val add_commute_succ = nat_succI RSN (2,add_commute); | |
| 327 | ||
| 328 | val prems = goal List.thy | |
| 329 | "xs: list(A) ==> length(rev(xs)) = length(xs)"; | |
| 330 | by (list_ind_tac "xs" prems 1); | |
| 2469 | 331 | by (ALLGOALS (asm_simp_tac (!simpset addsimps [length_app, add_commute_succ]))); | 
| 760 | 332 | qed "length_rev"; | 
| 516 | 333 | |
| 334 | val prems = goal List.thy | |
| 335 | "ls: list(list(A)) ==> length(flat(ls)) = list_add(map(length,ls))"; | |
| 336 | by (list_ind_tac "ls" prems 1); | |
| 2469 | 337 | by (ALLGOALS (asm_simp_tac (!simpset addsimps [length_app]))); | 
| 760 | 338 | qed "length_flat"; | 
| 516 | 339 | |
| 340 | (** Length and drop **) | |
| 341 | ||
| 342 | (*Lemma for the inductive step of drop_length*) | |
| 343 | goal List.thy | |
| 344 | "!!xs. xs: list(A) ==> \ | |
| 345 | \ ALL x. EX z zs. drop(length(xs), Cons(x,xs)) = Cons(z,zs)"; | |
| 346 | by (etac list.induct 1); | |
| 2469 | 347 | by (ALLGOALS Asm_simp_tac); | 
| 3016 | 348 | by (Blast_tac 1); | 
| 760 | 349 | qed "drop_length_Cons_lemma"; | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 350 | bind_thm ("drop_length_Cons", (drop_length_Cons_lemma RS spec));
 | 
| 516 | 351 | |
| 352 | goal List.thy | |
| 353 | "!!l. l: list(A) ==> ALL i: length(l). EX z zs. drop(i,l) = Cons(z,zs)"; | |
| 354 | by (etac list.induct 1); | |
| 2469 | 355 | by (ALLGOALS Asm_simp_tac); | 
| 516 | 356 | by (rtac conjI 1); | 
| 357 | by (etac drop_length_Cons 1); | |
| 358 | by (rtac ballI 1); | |
| 359 | by (rtac natE 1); | |
| 360 | by (etac ([asm_rl, length_type, Ord_nat] MRS Ord_trans) 1); | |
| 361 | by (assume_tac 1); | |
| 3016 | 362 | by (ALLGOALS Asm_simp_tac); | 
| 363 | by (ALLGOALS (blast_tac (!claset addIs [succ_in_naturalD, length_type]))); | |
| 760 | 364 | qed "drop_length_lemma"; | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 365 | bind_thm ("drop_length", (drop_length_lemma RS bspec));
 | 
| 516 | 366 | |
| 367 | ||
| 368 | (*** theorems about app ***) | |
| 369 | ||
| 370 | val [major] = goal List.thy "xs: list(A) ==> xs@Nil=xs"; | |
| 371 | by (rtac (major RS list.induct) 1); | |
| 3016 | 372 | by (ALLGOALS Asm_simp_tac); | 
| 760 | 373 | qed "app_right_Nil"; | 
| 516 | 374 | |
| 375 | val prems = goal List.thy "xs: list(A) ==> (xs@ys)@zs = xs@(ys@zs)"; | |
| 376 | by (list_ind_tac "xs" prems 1); | |
| 3016 | 377 | by (ALLGOALS Asm_simp_tac); | 
| 760 | 378 | qed "app_assoc"; | 
| 516 | 379 | |
| 380 | val prems = goal List.thy | |
| 381 | "ls: list(list(A)) ==> flat(ls@ms) = flat(ls)@flat(ms)"; | |
| 382 | by (list_ind_tac "ls" prems 1); | |
| 2469 | 383 | by (ALLGOALS (asm_simp_tac (!simpset addsimps [app_assoc]))); | 
| 760 | 384 | qed "flat_app_distrib"; | 
| 516 | 385 | |
| 386 | (*** theorems about rev ***) | |
| 387 | ||
| 388 | val prems = goal List.thy "l: list(A) ==> rev(map(h,l)) = map(h,rev(l))"; | |
| 389 | by (list_ind_tac "l" prems 1); | |
| 2469 | 390 | by (ALLGOALS (asm_simp_tac (!simpset addsimps [map_app_distrib]))); | 
| 760 | 391 | qed "rev_map_distrib"; | 
| 516 | 392 | |
| 393 | (*Simplifier needs the premises as assumptions because rewriting will not | |
| 394 | instantiate the variable ?A in the rules' typing conditions; note that | |
| 395 | rev_type does not instantiate ?A. Only the premises do. | |
| 396 | *) | |
| 397 | goal List.thy | |
| 398 | "!!xs. [| xs: list(A); ys: list(A) |] ==> rev(xs@ys) = rev(ys)@rev(xs)"; | |
| 399 | by (etac list.induct 1); | |
| 2469 | 400 | by (ALLGOALS (asm_simp_tac (!simpset addsimps [app_right_Nil,app_assoc]))); | 
| 760 | 401 | qed "rev_app_distrib"; | 
| 516 | 402 | |
| 403 | val prems = goal List.thy "l: list(A) ==> rev(rev(l))=l"; | |
| 404 | by (list_ind_tac "l" prems 1); | |
| 2469 | 405 | by (ALLGOALS (asm_simp_tac (!simpset addsimps [rev_app_distrib]))); | 
| 760 | 406 | qed "rev_rev_ident"; | 
| 516 | 407 | |
| 408 | val prems = goal List.thy | |
| 409 | "ls: list(list(A)) ==> rev(flat(ls)) = flat(map(rev,rev(ls)))"; | |
| 410 | by (list_ind_tac "ls" prems 1); | |
| 2469 | 411 | by (ALLGOALS (asm_simp_tac (!simpset addsimps | 
| 516 | 412 | [map_app_distrib, flat_app_distrib, rev_app_distrib, app_right_Nil]))); | 
| 760 | 413 | qed "rev_flat"; | 
| 516 | 414 | |
| 415 | ||
| 416 | (*** theorems about list_add ***) | |
| 417 | ||
| 418 | val prems = goal List.thy | |
| 419 | "[| xs: list(nat); ys: list(nat) |] ==> \ | |
| 420 | \ list_add(xs@ys) = list_add(ys) #+ list_add(xs)"; | |
| 421 | by (cut_facts_tac prems 1); | |
| 422 | by (list_ind_tac "xs" prems 1); | |
| 423 | by (ALLGOALS | |
| 2469 | 424 | (asm_simp_tac (!simpset addsimps [add_0_right, add_assoc RS sym]))); | 
| 516 | 425 | by (rtac (add_commute RS subst_context) 1); | 
| 426 | by (REPEAT (ares_tac [refl, list_add_type] 1)); | |
| 760 | 427 | qed "list_add_app"; | 
| 516 | 428 | |
| 429 | val prems = goal List.thy | |
| 430 | "l: list(nat) ==> list_add(rev(l)) = list_add(l)"; | |
| 431 | by (list_ind_tac "l" prems 1); | |
| 432 | by (ALLGOALS | |
| 2469 | 433 | (asm_simp_tac (!simpset addsimps [list_add_app, add_0_right]))); | 
| 760 | 434 | qed "list_add_rev"; | 
| 516 | 435 | |
| 436 | val prems = goal List.thy | |
| 437 | "ls: list(list(nat)) ==> list_add(flat(ls)) = list_add(map(list_add,ls))"; | |
| 438 | by (list_ind_tac "ls" prems 1); | |
| 2469 | 439 | by (ALLGOALS (asm_simp_tac (!simpset addsimps [list_add_app]))); | 
| 516 | 440 | by (REPEAT (ares_tac [refl, list_add_type, map_type, add_commute] 1)); | 
| 760 | 441 | qed "list_add_flat"; | 
| 516 | 442 | |
| 443 | (** New induction rule **) | |
| 444 | ||
| 445 | val major::prems = goal List.thy | |
| 446 | "[| l: list(A); \ | |
| 447 | \ P(Nil); \ | |
| 448 | \ !!x y. [| x: A; y: list(A); P(y) |] ==> P(y @ [x]) \ | |
| 449 | \ |] ==> P(l)"; | |
| 450 | by (rtac (major RS rev_rev_ident RS subst) 1); | |
| 451 | by (rtac (major RS rev_type RS list.induct) 1); | |
| 2469 | 452 | by (ALLGOALS (asm_simp_tac (!simpset addsimps prems))); | 
| 760 | 453 | qed "list_append_induct"; | 
| 516 | 454 |