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