| author | wenzelm | 
| Thu, 08 Nov 2001 23:52:56 +0100 | |
| changeset 12106 | 4a8558dbb6a0 | 
| parent 11145 | 3e47692e3a3e | 
| child 12789 | 459b5de466b2 | 
| 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 | (*** Aspects of the datatype definition ***) | 
| 0 | 10 | |
| 11 | (*An elimination rule, for type-checking*) | |
| 9907 | 12 | bind_thm ("ConsE", list.mk_cases "Cons(a,l) : list(A)");
 | 
| 0 | 13 | |
| 14 | (*Proving freeness results*) | |
| 9907 | 15 | bind_thm ("Cons_iff", list.mk_free "Cons(a,l)=Cons(a',l') <-> a=a' & l=l'");
 | 
| 16 | bind_thm ("Nil_Cons_iff", list.mk_free "~ Nil=Cons(a,l)");
 | |
| 0 | 17 | |
| 5067 | 18 | Goal "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 | 19 | let open list; val rew = rewrite_rule con_defs in | 
| 4091 | 20 | 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 | 21 | end; | 
| 760 | 22 | qed "list_unfold"; | 
| 435 | 23 | |
| 0 | 24 | (** Lemmas to justify using "list" in other recursive type definitions **) | 
| 25 | ||
| 5137 | 26 | Goalw list.defs "A<=B ==> list(A) <= list(B)"; | 
| 0 | 27 | by (rtac lfp_mono 1); | 
| 516 | 28 | by (REPEAT (rtac list.bnd_mono 1)); | 
| 0 | 29 | by (REPEAT (ares_tac (univ_mono::basic_monos) 1)); | 
| 760 | 30 | qed "list_mono"; | 
| 0 | 31 | |
| 32 | (*There is a similar proof by list induction.*) | |
| 5067 | 33 | Goalw (list.defs@list.con_defs) "list(univ(A)) <= univ(A)"; | 
| 0 | 34 | by (rtac lfp_lowerbound 1); | 
| 35 | by (rtac (A_subset_univ RS univ_mono) 2); | |
| 4091 | 36 | by (blast_tac (claset() addSIs [zero_in_univ, Inl_in_univ, Inr_in_univ, | 
| 6112 | 37 | Pair_in_univ]) 1); | 
| 760 | 38 | qed "list_univ"; | 
| 0 | 39 | |
| 908 
1f99a44c10cb
Updated comment about list_subset_univ and list_into_univ.
 lcp parents: 
782diff
changeset | 40 | (*These two theorems justify datatypes involving list(nat), list(A), ...*) | 
| 6112 | 41 | bind_thm ("list_subset_univ", [list_mono, list_univ] MRS subset_trans);
 | 
| 0 | 42 | |
| 5137 | 43 | Goal "[| l: list(A); A <= univ(B) |] ==> l: univ(B)"; | 
| 435 | 44 | by (REPEAT (ares_tac [list_subset_univ RS subsetD] 1)); | 
| 760 | 45 | qed "list_into_univ"; | 
| 435 | 46 | |
| 5321 | 47 | val major::prems = Goal | 
| 0 | 48 | "[| l: list(A); \ | 
| 15 
6c6d2f6e3185
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 49 | \ c: C(Nil); \ | 
| 
6c6d2f6e3185
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 50 | \ !!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 | 51 | \ |] ==> list_case(c,h,l) : C(l)"; | 
| 516 | 52 | by (rtac (major RS list.induct) 1); | 
| 6053 
8a1059aa01f0
new inductive, datatype and primrec packages, etc.
 paulson parents: 
5529diff
changeset | 53 | by (ALLGOALS (asm_simp_tac (simpset() addsimps prems))); | 
| 760 | 54 | qed "list_case_type"; | 
| 0 | 55 | |
| 56 | ||
| 516 | 57 | (*** List functions ***) | 
| 58 | ||
| 5137 | 59 | Goal "l: list(A) ==> tl(l) : list(A)"; | 
| 6065 | 60 | by (exhaust_tac "l" 1); | 
| 4091 | 61 | by (ALLGOALS (asm_simp_tac (simpset() addsimps list.intrs))); | 
| 760 | 62 | qed "tl_type"; | 
| 516 | 63 | |
| 64 | (** drop **) | |
| 65 | ||
| 6070 | 66 | Goal "i:nat ==> drop(i, Nil) = Nil"; | 
| 67 | by (induct_tac "i" 1); | |
| 2469 | 68 | by (ALLGOALS Asm_simp_tac); | 
| 760 | 69 | qed "drop_Nil"; | 
| 516 | 70 | |
| 6070 | 71 | Goal "i:nat ==> drop(succ(i), Cons(a,l)) = drop(i,l)"; | 
| 2493 | 72 | by (rtac sym 1); | 
| 6070 | 73 | by (induct_tac "i" 1); | 
| 2469 | 74 | by (Simp_tac 1); | 
| 75 | by (Asm_simp_tac 1); | |
| 760 | 76 | qed "drop_succ_Cons"; | 
| 516 | 77 | |
| 6070 | 78 | Addsimps [drop_Nil, drop_succ_Cons]; | 
| 2469 | 79 | |
| 6070 | 80 | Goal "[| i:nat; l: list(A) |] ==> drop(i,l) : list(A)"; | 
| 81 | by (induct_tac "i" 1); | |
| 4091 | 82 | by (ALLGOALS (asm_simp_tac (simpset() addsimps [tl_type]))); | 
| 760 | 83 | qed "drop_type"; | 
| 516 | 84 | |
| 6070 | 85 | Delsimps [drop_SUCC]; | 
| 86 | ||
| 516 | 87 | |
| 6053 
8a1059aa01f0
new inductive, datatype and primrec packages, etc.
 paulson parents: 
5529diff
changeset | 88 | (** Type checking -- proved by induction, as usual **) | 
| 516 | 89 | |
| 5321 | 90 | val prems = Goal | 
| 516 | 91 | "[| l: list(A); \ | 
| 92 | \ c: C(Nil); \ | |
| 93 | \ !!x y r. [| x:A; y: list(A); r: C(y) |] ==> h(x,y,r): C(Cons(x,y)) \ | |
| 6053 
8a1059aa01f0
new inductive, datatype and primrec packages, etc.
 paulson parents: 
5529diff
changeset | 94 | \ |] ==> list_rec(c,h,l) : C(l)"; | 
| 6065 | 95 | by (cut_facts_tac prems 1); | 
| 96 | by (induct_tac "l" 1); | |
| 4091 | 97 | by (ALLGOALS (asm_simp_tac (simpset() addsimps prems))); | 
| 760 | 98 | qed "list_rec_type"; | 
| 516 | 99 | |
| 100 | (** map **) | |
| 101 | ||
| 11145 | 102 | val prems = Goalw [thm "map_list_def"] | 
| 516 | 103 | "[| l: list(A); !!x. x: A ==> h(x): B |] ==> map(h,l) : list(B)"; | 
| 104 | by (REPEAT (ares_tac (prems @ list.intrs @ [list_rec_type]) 1)); | |
| 760 | 105 | qed "map_type"; | 
| 516 | 106 | |
| 5321 | 107 | Goal "l: list(A) ==> map(h,l) : list({h(u). u:A})";
 | 
| 108 | by (etac map_type 1); | |
| 516 | 109 | by (etac RepFunI 1); | 
| 760 | 110 | qed "map_type2"; | 
| 516 | 111 | |
| 112 | (** length **) | |
| 113 | ||
| 11145 | 114 | Goalw [thm "length_list_def"] | 
| 5147 
825877190618
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5137diff
changeset | 115 | "l: list(A) ==> length(l) : nat"; | 
| 516 | 116 | by (REPEAT (ares_tac [list_rec_type, nat_0I, nat_succI] 1)); | 
| 760 | 117 | qed "length_type"; | 
| 516 | 118 | |
| 119 | (** app **) | |
| 120 | ||
| 11145 | 121 | Goalw [thm "op @_list_def"] | 
| 5147 
825877190618
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5137diff
changeset | 122 | "[| xs: list(A); ys: list(A) |] ==> xs@ys : list(A)"; | 
| 516 | 123 | by (REPEAT (ares_tac [list_rec_type, list.Cons_I] 1)); | 
| 760 | 124 | qed "app_type"; | 
| 516 | 125 | |
| 126 | (** rev **) | |
| 127 | ||
| 11145 | 128 | Goalw [thm "rev_list_def"] | 
| 5147 
825877190618
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5137diff
changeset | 129 | "xs: list(A) ==> rev(xs) : list(A)"; | 
| 516 | 130 | by (REPEAT (ares_tac (list.intrs @ [list_rec_type, app_type]) 1)); | 
| 760 | 131 | qed "rev_type"; | 
| 516 | 132 | |
| 133 | ||
| 134 | (** flat **) | |
| 135 | ||
| 11145 | 136 | Goalw [thm "flat_list_def"] | 
| 5147 
825877190618
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5137diff
changeset | 137 | "ls: list(list(A)) ==> flat(ls) : list(A)"; | 
| 516 | 138 | by (REPEAT (ares_tac (list.intrs @ [list_rec_type, app_type]) 1)); | 
| 760 | 139 | qed "flat_type"; | 
| 516 | 140 | |
| 141 | ||
| 1926 | 142 | (** set_of_list **) | 
| 143 | ||
| 11145 | 144 | Goalw [thm "set_of_list_list_def"] | 
| 5147 
825877190618
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5137diff
changeset | 145 | "l: list(A) ==> set_of_list(l) : Pow(A)"; | 
| 2033 | 146 | by (etac list_rec_type 1); | 
| 3016 | 147 | by (ALLGOALS (Blast_tac)); | 
| 1926 | 148 | qed "set_of_list_type"; | 
| 149 | ||
| 5147 
825877190618
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5137diff
changeset | 150 | Goal "xs: list(A) ==> \ | 
| 1926 | 151 | \ set_of_list (xs@ys) = set_of_list(xs) Un set_of_list(ys)"; | 
| 152 | by (etac list.induct 1); | |
| 4091 | 153 | by (ALLGOALS (asm_simp_tac (simpset() addsimps [Un_cons]))); | 
| 1926 | 154 | qed "set_of_list_append"; | 
| 155 | ||
| 156 | ||
| 516 | 157 | (** list_add **) | 
| 158 | ||
| 11145 | 159 | Goalw [thm "list_add_list_def"] | 
| 5147 
825877190618
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5137diff
changeset | 160 | "xs: list(nat) ==> list_add(xs) : nat"; | 
| 516 | 161 | by (REPEAT (ares_tac [list_rec_type, nat_0I, add_type] 1)); | 
| 760 | 162 | qed "list_add_type"; | 
| 516 | 163 | |
| 9907 | 164 | bind_thms ("list_typechecks",
 | 
| 516 | 165 | list.intrs @ | 
| 166 | [list_rec_type, map_type, map_type2, app_type, length_type, | |
| 9907 | 167 | rev_type, flat_type, list_add_type]); | 
| 516 | 168 | |
| 6153 | 169 | AddTCs list_typechecks; | 
| 516 | 170 | |
| 171 | ||
| 172 | (*** theorems about map ***) | |
| 173 | ||
| 5321 | 174 | Goal "l: list(A) ==> map(%u. u, l) = l"; | 
| 6065 | 175 | by (induct_tac "l" 1); | 
| 3016 | 176 | by (ALLGOALS Asm_simp_tac); | 
| 760 | 177 | qed "map_ident"; | 
| 6112 | 178 | Addsimps [map_ident]; | 
| 516 | 179 | |
| 5321 | 180 | Goal "l: list(A) ==> map(h, map(j,l)) = map(%u. h(j(u)), l)"; | 
| 6065 | 181 | by (induct_tac "l" 1); | 
| 3016 | 182 | by (ALLGOALS Asm_simp_tac); | 
| 760 | 183 | qed "map_compose"; | 
| 516 | 184 | |
| 5321 | 185 | Goal "xs: list(A) ==> map(h, xs@ys) = map(h,xs) @ map(h,ys)"; | 
| 6065 | 186 | by (induct_tac "xs" 1); | 
| 3016 | 187 | by (ALLGOALS Asm_simp_tac); | 
| 760 | 188 | qed "map_app_distrib"; | 
| 516 | 189 | |
| 5321 | 190 | Goal "ls: list(list(A)) ==> map(h, flat(ls)) = flat(map(map(h),ls))"; | 
| 6065 | 191 | by (induct_tac "ls" 1); | 
| 4091 | 192 | by (ALLGOALS (asm_simp_tac (simpset() addsimps [map_app_distrib]))); | 
| 760 | 193 | qed "map_flat"; | 
| 516 | 194 | |
| 5321 | 195 | Goal "l: list(A) ==> \ | 
| 6053 
8a1059aa01f0
new inductive, datatype and primrec packages, etc.
 paulson parents: 
5529diff
changeset | 196 | \ list_rec(c, d, map(h,l)) = \ | 
| 
8a1059aa01f0
new inductive, datatype and primrec packages, etc.
 paulson parents: 
5529diff
changeset | 197 | \ list_rec(c, %x xs r. d(h(x), map(h,xs), r), l)"; | 
| 6065 | 198 | by (induct_tac "l" 1); | 
| 3016 | 199 | by (ALLGOALS Asm_simp_tac); | 
| 760 | 200 | qed "list_rec_map"; | 
| 516 | 201 | |
| 202 | (** theorems about list(Collect(A,P)) -- used in ex/term.ML **) | |
| 203 | ||
| 204 | (* c : list(Collect(B,P)) ==> c : list(B) *) | |
| 6112 | 205 | bind_thm ("list_CollectD", Collect_subset RS list_mono RS subsetD);
 | 
| 516 | 206 | |
| 5321 | 207 | Goal "l: list({x:A. h(x)=j(x)}) ==> map(h,l) = map(j,l)";
 | 
| 6065 | 208 | by (induct_tac "l" 1); | 
| 3016 | 209 | by (ALLGOALS Asm_simp_tac); | 
| 760 | 210 | qed "map_list_Collect"; | 
| 516 | 211 | |
| 212 | (*** theorems about length ***) | |
| 213 | ||
| 5321 | 214 | Goal "xs: list(A) ==> length(map(h,xs)) = length(xs)"; | 
| 6065 | 215 | by (induct_tac "xs" 1); | 
| 3016 | 216 | by (ALLGOALS Asm_simp_tac); | 
| 760 | 217 | qed "length_map"; | 
| 516 | 218 | |
| 9491 
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
 paulson parents: 
6153diff
changeset | 219 | Goal "[| xs: list(A); ys: list(A) |] \ | 
| 
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
 paulson parents: 
6153diff
changeset | 220 | \ ==> length(xs@ys) = length(xs) #+ length(ys)"; | 
| 6065 | 221 | by (induct_tac "xs" 1); | 
| 3016 | 222 | by (ALLGOALS Asm_simp_tac); | 
| 760 | 223 | qed "length_app"; | 
| 516 | 224 | |
| 5321 | 225 | Goal "xs: list(A) ==> length(rev(xs)) = length(xs)"; | 
| 6065 | 226 | by (induct_tac "xs" 1); | 
| 6112 | 227 | by (ALLGOALS (asm_simp_tac (simpset() addsimps [length_app]))); | 
| 760 | 228 | qed "length_rev"; | 
| 516 | 229 | |
| 5321 | 230 | Goal "ls: list(list(A)) ==> length(flat(ls)) = list_add(map(length,ls))"; | 
| 6065 | 231 | by (induct_tac "ls" 1); | 
| 4091 | 232 | by (ALLGOALS (asm_simp_tac (simpset() addsimps [length_app]))); | 
| 760 | 233 | qed "length_flat"; | 
| 516 | 234 | |
| 235 | (** Length and drop **) | |
| 236 | ||
| 237 | (*Lemma for the inductive step of drop_length*) | |
| 5147 
825877190618
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5137diff
changeset | 238 | Goal "xs: list(A) ==> \ | 
| 516 | 239 | \ ALL x. EX z zs. drop(length(xs), Cons(x,xs)) = Cons(z,zs)"; | 
| 240 | by (etac list.induct 1); | |
| 2469 | 241 | by (ALLGOALS Asm_simp_tac); | 
| 3016 | 242 | by (Blast_tac 1); | 
| 6112 | 243 | qed_spec_mp "drop_length_Cons"; | 
| 516 | 244 | |
| 6112 | 245 | Goal "l: list(A) ==> ALL i:length(l). (EX z zs. drop(i,l) = Cons(z,zs))"; | 
| 516 | 246 | by (etac list.induct 1); | 
| 2469 | 247 | by (ALLGOALS Asm_simp_tac); | 
| 6112 | 248 | by Safe_tac; | 
| 516 | 249 | by (etac drop_length_Cons 1); | 
| 250 | by (rtac natE 1); | |
| 251 | by (etac ([asm_rl, length_type, Ord_nat] MRS Ord_trans) 1); | |
| 252 | by (assume_tac 1); | |
| 3016 | 253 | by (ALLGOALS Asm_simp_tac); | 
| 4091 | 254 | by (ALLGOALS (blast_tac (claset() addIs [succ_in_naturalD, length_type]))); | 
| 6112 | 255 | qed_spec_mp "drop_length"; | 
| 516 | 256 | |
| 257 | ||
| 258 | (*** theorems about app ***) | |
| 259 | ||
| 5321 | 260 | Goal "xs: list(A) ==> xs@Nil=xs"; | 
| 261 | by (etac list.induct 1); | |
| 3016 | 262 | by (ALLGOALS Asm_simp_tac); | 
| 760 | 263 | qed "app_right_Nil"; | 
| 6112 | 264 | Addsimps [app_right_Nil]; | 
| 516 | 265 | |
| 5321 | 266 | Goal "xs: list(A) ==> (xs@ys)@zs = xs@(ys@zs)"; | 
| 6065 | 267 | by (induct_tac "xs" 1); | 
| 3016 | 268 | by (ALLGOALS Asm_simp_tac); | 
| 760 | 269 | qed "app_assoc"; | 
| 516 | 270 | |
| 5321 | 271 | Goal "ls: list(list(A)) ==> flat(ls@ms) = flat(ls)@flat(ms)"; | 
| 6065 | 272 | by (induct_tac "ls" 1); | 
| 4091 | 273 | by (ALLGOALS (asm_simp_tac (simpset() addsimps [app_assoc]))); | 
| 760 | 274 | qed "flat_app_distrib"; | 
| 516 | 275 | |
| 276 | (*** theorems about rev ***) | |
| 277 | ||
| 5321 | 278 | Goal "l: list(A) ==> rev(map(h,l)) = map(h,rev(l))"; | 
| 6065 | 279 | by (induct_tac "l" 1); | 
| 4091 | 280 | by (ALLGOALS (asm_simp_tac (simpset() addsimps [map_app_distrib]))); | 
| 760 | 281 | qed "rev_map_distrib"; | 
| 516 | 282 | |
| 283 | (*Simplifier needs the premises as assumptions because rewriting will not | |
| 284 | instantiate the variable ?A in the rules' typing conditions; note that | |
| 285 | rev_type does not instantiate ?A. Only the premises do. | |
| 286 | *) | |
| 5147 
825877190618
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5137diff
changeset | 287 | Goal "[| xs: list(A); ys: list(A) |] ==> rev(xs@ys) = rev(ys)@rev(xs)"; | 
| 516 | 288 | by (etac list.induct 1); | 
| 6112 | 289 | by (ALLGOALS (asm_simp_tac (simpset() addsimps [app_assoc]))); | 
| 760 | 290 | qed "rev_app_distrib"; | 
| 516 | 291 | |
| 5321 | 292 | Goal "l: list(A) ==> rev(rev(l))=l"; | 
| 6065 | 293 | by (induct_tac "l" 1); | 
| 4091 | 294 | by (ALLGOALS (asm_simp_tac (simpset() addsimps [rev_app_distrib]))); | 
| 760 | 295 | qed "rev_rev_ident"; | 
| 6112 | 296 | Addsimps [rev_rev_ident]; | 
| 516 | 297 | |
| 5321 | 298 | Goal "ls: list(list(A)) ==> rev(flat(ls)) = flat(map(rev,rev(ls)))"; | 
| 6065 | 299 | by (induct_tac "ls" 1); | 
| 6112 | 300 | by (ALLGOALS | 
| 301 | (asm_simp_tac (simpset() addsimps | |
| 302 | [map_app_distrib, flat_app_distrib, rev_app_distrib]))); | |
| 760 | 303 | qed "rev_flat"; | 
| 516 | 304 | |
| 305 | ||
| 306 | (*** theorems about list_add ***) | |
| 307 | ||
| 5321 | 308 | Goal "[| xs: list(nat); ys: list(nat) |] ==> \ | 
| 516 | 309 | \ list_add(xs@ys) = list_add(ys) #+ list_add(xs)"; | 
| 6065 | 310 | by (induct_tac "xs" 1); | 
| 9548 | 311 | by (ALLGOALS Asm_simp_tac); | 
| 760 | 312 | qed "list_add_app"; | 
| 516 | 313 | |
| 5321 | 314 | Goal "l: list(nat) ==> list_add(rev(l)) = list_add(l)"; | 
| 6065 | 315 | by (induct_tac "l" 1); | 
| 9548 | 316 | by (ALLGOALS (asm_simp_tac (simpset() addsimps [list_add_app]))); | 
| 760 | 317 | qed "list_add_rev"; | 
| 516 | 318 | |
| 5321 | 319 | Goal "ls: list(list(nat)) ==> list_add(flat(ls)) = list_add(map(list_add,ls))"; | 
| 6065 | 320 | by (induct_tac "ls" 1); | 
| 4091 | 321 | by (ALLGOALS (asm_simp_tac (simpset() addsimps [list_add_app]))); | 
| 516 | 322 | by (REPEAT (ares_tac [refl, list_add_type, map_type, add_commute] 1)); | 
| 760 | 323 | qed "list_add_flat"; | 
| 516 | 324 | |
| 325 | (** New induction rule **) | |
| 326 | ||
| 5321 | 327 | val major::prems = Goal | 
| 516 | 328 | "[| l: list(A); \ | 
| 329 | \ P(Nil); \ | |
| 330 | \ !!x y. [| x: A; y: list(A); P(y) |] ==> P(y @ [x]) \ | |
| 331 | \ |] ==> P(l)"; | |
| 332 | by (rtac (major RS rev_rev_ident RS subst) 1); | |
| 333 | by (rtac (major RS rev_type RS list.induct) 1); | |
| 4091 | 334 | by (ALLGOALS (asm_simp_tac (simpset() addsimps prems))); | 
| 760 | 335 | qed "list_append_induct"; | 
| 516 | 336 |