| author | wenzelm | 
| Thu, 06 Dec 2001 00:37:59 +0100 | |
| changeset 12395 | d6913de7655f | 
| parent 11868 | 56db9f3a6b3e | 
| child 12486 | 0ed8bdd883e0 | 
| permissions | -rw-r--r-- | 
| 1465 | 1 | (* Title: HOL/List | 
| 923 | 2 | ID: $Id$ | 
| 1465 | 3 | Author: Tobias Nipkow | 
| 923 | 4 | Copyright 1994 TU Muenchen | 
| 5 | ||
| 6 | List lemmas | |
| 7 | *) | |
| 8 | ||
| 4935 | 9 | Goal "!x. xs ~= x#xs"; | 
| 3040 
7d48671753da
Introduced a generic "induct_tac" which picks up the right induction scheme
 nipkow parents: 
3011diff
changeset | 10 | by (induct_tac "xs" 1); | 
| 5316 | 11 | by Auto_tac; | 
| 2608 | 12 | qed_spec_mp "not_Cons_self"; | 
| 3574 | 13 | bind_thm("not_Cons_self2",not_Cons_self RS not_sym);
 | 
| 14 | Addsimps [not_Cons_self,not_Cons_self2]; | |
| 923 | 15 | |
| 4935 | 16 | Goal "(xs ~= []) = (? y ys. xs = y#ys)"; | 
| 3040 
7d48671753da
Introduced a generic "induct_tac" which picks up the right induction scheme
 nipkow parents: 
3011diff
changeset | 17 | by (induct_tac "xs" 1); | 
| 5316 | 18 | by Auto_tac; | 
| 923 | 19 | qed "neq_Nil_conv"; | 
| 20 | ||
| 4830 | 21 | (* Induction over the length of a list: *) | 
| 4935 | 22 | val [prem] = Goal | 
| 4911 | 23 | "(!!xs. (!ys. length ys < length xs --> P ys) ==> P xs) ==> P(xs)"; | 
| 5132 | 24 | by (rtac measure_induct 1 THEN etac prem 1); | 
| 4911 | 25 | qed "length_induct"; | 
| 26 | ||
| 923 | 27 | |
| 3468 | 28 | (** "lists": the list-forming operator over sets **) | 
| 3342 
ec3b55fcb165
New operator "lists" for formalizing sets of lists
 paulson parents: 
3292diff
changeset | 29 | |
| 5043 | 30 | Goalw lists.defs "A<=B ==> lists A <= lists B"; | 
| 3342 
ec3b55fcb165
New operator "lists" for formalizing sets of lists
 paulson parents: 
3292diff
changeset | 31 | by (rtac lfp_mono 1); | 
| 
ec3b55fcb165
New operator "lists" for formalizing sets of lists
 paulson parents: 
3292diff
changeset | 32 | by (REPEAT (ares_tac basic_monos 1)); | 
| 
ec3b55fcb165
New operator "lists" for formalizing sets of lists
 paulson parents: 
3292diff
changeset | 33 | qed "lists_mono"; | 
| 3196 | 34 | |
| 9108 | 35 | bind_thm ("listsE", lists.mk_cases "x#l : lists A");
 | 
| 3468 | 36 | AddSEs [listsE]; | 
| 37 | AddSIs lists.intrs; | |
| 38 | ||
| 5043 | 39 | Goal "l: lists A ==> l: lists B --> l: lists (A Int B)"; | 
| 3468 | 40 | by (etac lists.induct 1); | 
| 41 | by (ALLGOALS Blast_tac); | |
| 42 | qed_spec_mp "lists_IntI"; | |
| 43 | ||
| 4935 | 44 | Goal "lists (A Int B) = lists A Int lists B"; | 
| 4423 | 45 | by (rtac (mono_Int RS equalityI) 1); | 
| 4089 | 46 | by (simp_tac (simpset() addsimps [mono_def, lists_mono]) 1); | 
| 47 | by (blast_tac (claset() addSIs [lists_IntI]) 1); | |
| 3468 | 48 | qed "lists_Int_eq"; | 
| 49 | Addsimps [lists_Int_eq]; | |
| 50 | ||
| 9268 | 51 | Goal "(xs@ys : lists A) = (xs : lists A & ys : lists A)"; | 
| 52 | by(induct_tac "xs" 1); | |
| 53 | by(Auto_tac); | |
| 54 | qed "append_in_lists_conv"; | |
| 55 | AddIffs [append_in_lists_conv]; | |
| 2608 | 56 | |
| 3860 | 57 | (** length **) | 
| 58 | (* needs to come before "@" because of thm append_eq_append_conv *) | |
| 59 | ||
| 60 | section "length"; | |
| 61 | ||
| 4935 | 62 | Goal "length(xs@ys) = length(xs)+length(ys)"; | 
| 3860 | 63 | by (induct_tac "xs" 1); | 
| 5316 | 64 | by Auto_tac; | 
| 3860 | 65 | qed"length_append"; | 
| 66 | Addsimps [length_append]; | |
| 67 | ||
| 5129 | 68 | Goal "length (map f xs) = length xs"; | 
| 69 | by (induct_tac "xs" 1); | |
| 5316 | 70 | by Auto_tac; | 
| 3860 | 71 | qed "length_map"; | 
| 72 | Addsimps [length_map]; | |
| 73 | ||
| 4935 | 74 | Goal "length(rev xs) = length(xs)"; | 
| 3860 | 75 | by (induct_tac "xs" 1); | 
| 5316 | 76 | by Auto_tac; | 
| 3860 | 77 | qed "length_rev"; | 
| 78 | Addsimps [length_rev]; | |
| 79 | ||
| 7028 | 80 | Goal "length(tl xs) = (length xs) - 1"; | 
| 8442 
96023903c2df
case_tac now subsumes both boolean and datatype cases;
 wenzelm parents: 
8423diff
changeset | 81 | by (case_tac "xs" 1); | 
| 5316 | 82 | by Auto_tac; | 
| 3896 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 nipkow parents: 
3860diff
changeset | 83 | qed "length_tl"; | 
| 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 nipkow parents: 
3860diff
changeset | 84 | Addsimps [length_tl]; | 
| 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 nipkow parents: 
3860diff
changeset | 85 | |
| 4935 | 86 | Goal "(length xs = 0) = (xs = [])"; | 
| 3860 | 87 | by (induct_tac "xs" 1); | 
| 5316 | 88 | by Auto_tac; | 
| 3860 | 89 | qed "length_0_conv"; | 
| 90 | AddIffs [length_0_conv]; | |
| 91 | ||
| 4935 | 92 | Goal "(0 < length xs) = (xs ~= [])"; | 
| 3860 | 93 | by (induct_tac "xs" 1); | 
| 5316 | 94 | by Auto_tac; | 
| 3860 | 95 | qed "length_greater_0_conv"; | 
| 96 | AddIffs [length_greater_0_conv]; | |
| 97 | ||
| 5296 | 98 | Goal "(length xs = Suc n) = (? y ys. xs = y#ys & length ys = n)"; | 
| 99 | by (induct_tac "xs" 1); | |
| 6813 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 paulson parents: 
6794diff
changeset | 100 | by Auto_tac; | 
| 5296 | 101 | qed "length_Suc_conv"; | 
| 102 | ||
| 923 | 103 | (** @ - append **) | 
| 104 | ||
| 3467 | 105 | section "@ - append"; | 
| 106 | ||
| 4935 | 107 | Goal "(xs@ys)@zs = xs@(ys@zs)"; | 
| 3040 
7d48671753da
Introduced a generic "induct_tac" which picks up the right induction scheme
 nipkow parents: 
3011diff
changeset | 108 | by (induct_tac "xs" 1); | 
| 5316 | 109 | by Auto_tac; | 
| 923 | 110 | qed "append_assoc"; | 
| 2512 | 111 | Addsimps [append_assoc]; | 
| 923 | 112 | |
| 4935 | 113 | Goal "xs @ [] = xs"; | 
| 3040 
7d48671753da
Introduced a generic "induct_tac" which picks up the right induction scheme
 nipkow parents: 
3011diff
changeset | 114 | by (induct_tac "xs" 1); | 
| 5316 | 115 | by Auto_tac; | 
| 923 | 116 | qed "append_Nil2"; | 
| 2512 | 117 | Addsimps [append_Nil2]; | 
| 923 | 118 | |
| 4935 | 119 | Goal "(xs@ys = []) = (xs=[] & ys=[])"; | 
| 3040 
7d48671753da
Introduced a generic "induct_tac" which picks up the right induction scheme
 nipkow parents: 
3011diff
changeset | 120 | by (induct_tac "xs" 1); | 
| 5316 | 121 | by Auto_tac; | 
| 2608 | 122 | qed "append_is_Nil_conv"; | 
| 123 | AddIffs [append_is_Nil_conv]; | |
| 124 | ||
| 4935 | 125 | Goal "([] = xs@ys) = (xs=[] & ys=[])"; | 
| 3040 
7d48671753da
Introduced a generic "induct_tac" which picks up the right induction scheme
 nipkow parents: 
3011diff
changeset | 126 | by (induct_tac "xs" 1); | 
| 5316 | 127 | by Auto_tac; | 
| 2608 | 128 | qed "Nil_is_append_conv"; | 
| 129 | AddIffs [Nil_is_append_conv]; | |
| 923 | 130 | |
| 4935 | 131 | Goal "(xs @ ys = xs) = (ys=[])"; | 
| 3574 | 132 | by (induct_tac "xs" 1); | 
| 5316 | 133 | by Auto_tac; | 
| 3574 | 134 | qed "append_self_conv"; | 
| 135 | ||
| 4935 | 136 | Goal "(xs = xs @ ys) = (ys=[])"; | 
| 3574 | 137 | by (induct_tac "xs" 1); | 
| 5316 | 138 | by Auto_tac; | 
| 3574 | 139 | qed "self_append_conv"; | 
| 140 | AddIffs [append_self_conv,self_append_conv]; | |
| 141 | ||
| 4935 | 142 | Goal "!ys. length xs = length ys | length us = length vs \ | 
| 3860 | 143 | \ --> (xs@us = ys@vs) = (xs=ys & us=vs)"; | 
| 4423 | 144 | by (induct_tac "xs" 1); | 
| 145 | by (rtac allI 1); | |
| 8442 
96023903c2df
case_tac now subsumes both boolean and datatype cases;
 wenzelm parents: 
8423diff
changeset | 146 | by (case_tac "ys" 1); | 
| 4423 | 147 | by (Asm_simp_tac 1); | 
| 5641 | 148 | by (Force_tac 1); | 
| 4423 | 149 | by (rtac allI 1); | 
| 8442 
96023903c2df
case_tac now subsumes both boolean and datatype cases;
 wenzelm parents: 
8423diff
changeset | 150 | by (case_tac "ys" 1); | 
| 5641 | 151 | by (Force_tac 1); | 
| 4423 | 152 | by (Asm_simp_tac 1); | 
| 3860 | 153 | qed_spec_mp "append_eq_append_conv"; | 
| 154 | Addsimps [append_eq_append_conv]; | |
| 155 | ||
| 4935 | 156 | Goal "(xs @ ys = xs @ zs) = (ys=zs)"; | 
| 3896 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 nipkow parents: 
3860diff
changeset | 157 | by (Simp_tac 1); | 
| 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 nipkow parents: 
3860diff
changeset | 158 | qed "same_append_eq"; | 
| 3860 | 159 | |
| 4935 | 160 | Goal "(xs @ [x] = ys @ [y]) = (xs = ys & x = y)"; | 
| 3896 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 nipkow parents: 
3860diff
changeset | 161 | by (Simp_tac 1); | 
| 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 nipkow parents: 
3860diff
changeset | 162 | qed "append1_eq_conv"; | 
| 2608 | 163 | |
| 4935 | 164 | Goal "(ys @ xs = zs @ xs) = (ys=zs)"; | 
| 3896 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 nipkow parents: 
3860diff
changeset | 165 | by (Simp_tac 1); | 
| 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 nipkow parents: 
3860diff
changeset | 166 | qed "append_same_eq"; | 
| 2608 | 167 | |
| 9003 | 168 | AddIffs [same_append_eq, append1_eq_conv, append_same_eq]; | 
| 3571 | 169 | |
| 4935 | 170 | Goal "(xs @ ys = ys) = (xs=[])"; | 
| 5132 | 171 | by (cut_inst_tac [("zs","[]")] append_same_eq 1);
 | 
| 5316 | 172 | by Auto_tac; | 
| 4647 | 173 | qed "append_self_conv2"; | 
| 174 | ||
| 4935 | 175 | Goal "(ys = xs @ ys) = (xs=[])"; | 
| 5132 | 176 | by (simp_tac (simpset() addsimps | 
| 4647 | 177 |      [simplify (simpset()) (read_instantiate[("ys","[]")]append_same_eq)]) 1);
 | 
| 5132 | 178 | by (Blast_tac 1); | 
| 4647 | 179 | qed "self_append_conv2"; | 
| 180 | AddIffs [append_self_conv2,self_append_conv2]; | |
| 181 | ||
| 4935 | 182 | Goal "xs ~= [] --> hd xs # tl xs = xs"; | 
| 3457 | 183 | by (induct_tac "xs" 1); | 
| 5316 | 184 | by Auto_tac; | 
| 2608 | 185 | qed_spec_mp "hd_Cons_tl"; | 
| 186 | Addsimps [hd_Cons_tl]; | |
| 923 | 187 | |
| 4935 | 188 | Goal "hd(xs@ys) = (if xs=[] then hd ys else hd xs)"; | 
| 3040 
7d48671753da
Introduced a generic "induct_tac" which picks up the right induction scheme
 nipkow parents: 
3011diff
changeset | 189 | by (induct_tac "xs" 1); | 
| 5316 | 190 | by Auto_tac; | 
| 1327 
6c29cfab679c
added new arithmetic lemmas and the functions take and drop.
 nipkow parents: 
1301diff
changeset | 191 | qed "hd_append"; | 
| 923 | 192 | |
| 5043 | 193 | Goal "xs ~= [] ==> hd(xs @ ys) = hd xs"; | 
| 4089 | 194 | by (asm_simp_tac (simpset() addsimps [hd_append] | 
| 5183 | 195 | addsplits [list.split]) 1); | 
| 3571 | 196 | qed "hd_append2"; | 
| 197 | Addsimps [hd_append2]; | |
| 198 | ||
| 4935 | 199 | Goal "tl(xs@ys) = (case xs of [] => tl(ys) | z#zs => zs@ys)"; | 
| 5183 | 200 | by (simp_tac (simpset() addsplits [list.split]) 1); | 
| 2608 | 201 | qed "tl_append"; | 
| 202 | ||
| 5043 | 203 | Goal "xs ~= [] ==> tl(xs @ ys) = (tl xs) @ ys"; | 
| 4089 | 204 | by (asm_simp_tac (simpset() addsimps [tl_append] | 
| 5183 | 205 | addsplits [list.split]) 1); | 
| 3571 | 206 | qed "tl_append2"; | 
| 207 | Addsimps [tl_append2]; | |
| 208 | ||
| 5272 | 209 | (* trivial rules for solving @-equations automatically *) | 
| 210 | ||
| 211 | Goal "xs = ys ==> xs = [] @ ys"; | |
| 5318 | 212 | by (Asm_simp_tac 1); | 
| 5272 | 213 | qed "eq_Nil_appendI"; | 
| 214 | ||
| 215 | Goal "[| x#xs1 = ys; xs = xs1 @ zs |] ==> x#xs = ys@zs"; | |
| 5318 | 216 | by (dtac sym 1); | 
| 217 | by (Asm_simp_tac 1); | |
| 5272 | 218 | qed "Cons_eq_appendI"; | 
| 219 | ||
| 220 | Goal "[| xs@xs1 = zs; ys = xs1 @ us |] ==> xs@ys = zs@us"; | |
| 5318 | 221 | by (dtac sym 1); | 
| 222 | by (Asm_simp_tac 1); | |
| 5272 | 223 | qed "append_eq_appendI"; | 
| 224 | ||
| 4830 | 225 | |
| 5427 | 226 | (*** | 
| 227 | Simplification procedure for all list equalities. | |
| 228 | Currently only tries to rearranges @ to see if | |
| 229 | - both lists end in a singleton list, | |
| 230 | - or both lists end in the same list. | |
| 231 | ***) | |
| 232 | local | |
| 233 | ||
| 234 | val list_eq_pattern = | |
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11770diff
changeset | 235 |   Thm.read_cterm (Theory.sign_of (the_context ())) ("(xs::'a list) = ys",HOLogic.boolT)
 | 
| 5427 | 236 | |
| 7224 | 237 | fun last (cons as Const("List.list.Cons",_) $ _ $ xs) =
 | 
| 238 |       (case xs of Const("List.list.Nil",_) => cons | _ => last xs)
 | |
| 5427 | 239 |   | last (Const("List.op @",_) $ _ $ ys) = last ys
 | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11770diff
changeset | 240 | | last t = t | 
| 5427 | 241 | |
| 7224 | 242 | fun list1 (Const("List.list.Cons",_) $ _ $ Const("List.list.Nil",_)) = true
 | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11770diff
changeset | 243 | | list1 _ = false | 
| 5427 | 244 | |
| 7224 | 245 | fun butlast ((cons as Const("List.list.Cons",_) $ x) $ xs) =
 | 
| 246 |       (case xs of Const("List.list.Nil",_) => xs | _ => cons $ butlast xs)
 | |
| 5427 | 247 |   | butlast ((app as Const("List.op @",_) $ xs) $ ys) = app $ butlast ys
 | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11770diff
changeset | 248 |   | butlast xs = Const("List.list.Nil",fastype_of xs)
 | 
| 5427 | 249 | |
| 250 | val rearr_tac = | |
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11770diff
changeset | 251 | simp_tac (HOL_basic_ss addsimps [append_assoc,append_Nil,append_Cons]) | 
| 5427 | 252 | |
| 253 | fun list_eq sg _ (F as (eq as Const(_,eqT)) $ lhs $ rhs) = | |
| 254 | let | |
| 255 | val lastl = last lhs and lastr = last rhs | |
| 256 | fun rearr conv = | |
| 257 | let val lhs1 = butlast lhs and rhs1 = butlast rhs | |
| 258 | val Type(_,listT::_) = eqT | |
| 259 | val appT = [listT,listT] ---> listT | |
| 260 |           val app = Const("List.op @",appT)
 | |
| 261 | val F2 = eq $ (app$lhs1$lastl) $ (app$rhs1$lastr) | |
| 262 | val ct = cterm_of sg (HOLogic.mk_Trueprop(HOLogic.mk_eq(F,F2))) | |
| 263 | val thm = prove_goalw_cterm [] ct (K [rearr_tac 1]) | |
| 264 | handle ERROR => | |
| 265 |             error("The error(s) above occurred while trying to prove " ^
 | |
| 266 | string_of_cterm ct) | |
| 267 | in Some((conv RS (thm RS trans)) RS eq_reflection) end | |
| 268 | ||
| 269 | in if list1 lastl andalso list1 lastr | |
| 270 | then rearr append1_eq_conv | |
| 271 | else | |
| 272 | if lastl aconv lastr | |
| 273 | then rearr append_same_eq | |
| 274 | else None | |
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11770diff
changeset | 275 | end | 
| 5427 | 276 | in | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11770diff
changeset | 277 | val list_eq_simproc = mk_simproc "list_eq" [list_eq_pattern] list_eq | 
| 5427 | 278 | end; | 
| 279 | ||
| 280 | Addsimprocs [list_eq_simproc]; | |
| 281 | ||
| 282 | ||
| 2608 | 283 | (** map **) | 
| 284 | ||
| 3467 | 285 | section "map"; | 
| 286 | ||
| 5278 | 287 | Goal "(!x. x : set xs --> f x = g x) --> map f xs = map g xs"; | 
| 3457 | 288 | by (induct_tac "xs" 1); | 
| 5316 | 289 | by Auto_tac; | 
| 2608 | 290 | bind_thm("map_ext", impI RS (allI RS (result() RS mp)));
 | 
| 291 | ||
| 4935 | 292 | Goal "map (%x. x) = (%xs. xs)"; | 
| 2608 | 293 | by (rtac ext 1); | 
| 3040 
7d48671753da
Introduced a generic "induct_tac" which picks up the right induction scheme
 nipkow parents: 
3011diff
changeset | 294 | by (induct_tac "xs" 1); | 
| 5316 | 295 | by Auto_tac; | 
| 2608 | 296 | qed "map_ident"; | 
| 297 | Addsimps[map_ident]; | |
| 298 | ||
| 4935 | 299 | Goal "map f (xs@ys) = map f xs @ map f ys"; | 
| 3040 
7d48671753da
Introduced a generic "induct_tac" which picks up the right induction scheme
 nipkow parents: 
3011diff
changeset | 300 | by (induct_tac "xs" 1); | 
| 5316 | 301 | by Auto_tac; | 
| 2608 | 302 | qed "map_append"; | 
| 303 | Addsimps[map_append]; | |
| 304 | ||
| 4935 | 305 | Goalw [o_def] "map (f o g) xs = map f (map g xs)"; | 
| 3040 
7d48671753da
Introduced a generic "induct_tac" which picks up the right induction scheme
 nipkow parents: 
3011diff
changeset | 306 | by (induct_tac "xs" 1); | 
| 5316 | 307 | by Auto_tac; | 
| 2608 | 308 | qed "map_compose"; | 
| 9700 | 309 | (*Addsimps[map_compose];*) | 
| 2608 | 310 | |
| 4935 | 311 | Goal "rev(map f xs) = map f (rev xs)"; | 
| 3040 
7d48671753da
Introduced a generic "induct_tac" which picks up the right induction scheme
 nipkow parents: 
3011diff
changeset | 312 | by (induct_tac "xs" 1); | 
| 5316 | 313 | by Auto_tac; | 
| 2608 | 314 | qed "rev_map"; | 
| 315 | ||
| 3589 
244daa75f890
Added function `replicate' and lemmas map_cong and set_replicate.
 nipkow parents: 
3586diff
changeset | 316 | (* a congruence rule for map: *) | 
| 6451 | 317 | Goal "xs=ys ==> (!x. x : set ys --> f x = g x) --> map f xs = map g ys"; | 
| 4423 | 318 | by (hyp_subst_tac 1); | 
| 319 | by (induct_tac "ys" 1); | |
| 5316 | 320 | by Auto_tac; | 
| 6451 | 321 | bind_thm("map_cong", impI RSN (2,allI RSN (2, result() RS mp)));
 | 
| 3589 
244daa75f890
Added function `replicate' and lemmas map_cong and set_replicate.
 nipkow parents: 
3586diff
changeset | 322 | |
| 4935 | 323 | Goal "(map f xs = []) = (xs = [])"; | 
| 8442 
96023903c2df
case_tac now subsumes both boolean and datatype cases;
 wenzelm parents: 
8423diff
changeset | 324 | by (case_tac "xs" 1); | 
| 5316 | 325 | by Auto_tac; | 
| 3860 | 326 | qed "map_is_Nil_conv"; | 
| 327 | AddIffs [map_is_Nil_conv]; | |
| 328 | ||
| 4935 | 329 | Goal "([] = map f xs) = (xs = [])"; | 
| 8442 
96023903c2df
case_tac now subsumes both boolean and datatype cases;
 wenzelm parents: 
8423diff
changeset | 330 | by (case_tac "xs" 1); | 
| 5316 | 331 | by Auto_tac; | 
| 3860 | 332 | qed "Nil_is_map_conv"; | 
| 333 | AddIffs [Nil_is_map_conv]; | |
| 334 | ||
| 8009 | 335 | Goal "(map f xs = y#ys) = (? x xs'. xs = x#xs' & f x = y & map f xs' = ys)"; | 
| 8442 
96023903c2df
case_tac now subsumes both boolean and datatype cases;
 wenzelm parents: 
8423diff
changeset | 336 | by (case_tac "xs" 1); | 
| 8009 | 337 | by (ALLGOALS Asm_simp_tac); | 
| 338 | qed "map_eq_Cons"; | |
| 339 | ||
| 340 | Goal "!xs. map f xs = map f ys --> (!x y. f x = f y --> x=y) --> xs=ys"; | |
| 341 | by (induct_tac "ys" 1); | |
| 342 | by (Asm_simp_tac 1); | |
| 343 | by (fast_tac (claset() addss (simpset() addsimps [map_eq_Cons])) 1); | |
| 344 | qed_spec_mp "map_injective"; | |
| 345 | ||
| 346 | Goal "inj f ==> inj (map f)"; | |
| 8064 | 347 | by (blast_tac (claset() addDs [map_injective,injD] addIs [injI]) 1); | 
| 8009 | 348 | qed "inj_mapI"; | 
| 349 | ||
| 350 | Goalw [inj_on_def] "inj (map f) ==> inj f"; | |
| 8064 | 351 | by (Clarify_tac 1); | 
| 352 | by (eres_inst_tac [("x","[x]")] ballE 1);
 | |
| 353 |  by (eres_inst_tac [("x","[y]")] ballE 1);
 | |
| 354 | by (Asm_full_simp_tac 1); | |
| 355 | by (Blast_tac 1); | |
| 356 | by (Blast_tac 1); | |
| 8009 | 357 | qed "inj_mapD"; | 
| 358 | ||
| 359 | Goal "inj (map f) = inj f"; | |
| 8064 | 360 | by (blast_tac (claset() addDs [inj_mapD] addIs [inj_mapI]) 1); | 
| 8009 | 361 | qed "inj_map"; | 
| 3860 | 362 | |
| 1169 | 363 | (** rev **) | 
| 364 | ||
| 3467 | 365 | section "rev"; | 
| 366 | ||
| 4935 | 367 | Goal "rev(xs@ys) = rev(ys) @ rev(xs)"; | 
| 3040 
7d48671753da
Introduced a generic "induct_tac" which picks up the right induction scheme
 nipkow parents: 
3011diff
changeset | 368 | by (induct_tac "xs" 1); | 
| 5316 | 369 | by Auto_tac; | 
| 1169 | 370 | qed "rev_append"; | 
| 2512 | 371 | Addsimps[rev_append]; | 
| 1169 | 372 | |
| 4935 | 373 | Goal "rev(rev l) = l"; | 
| 3040 
7d48671753da
Introduced a generic "induct_tac" which picks up the right induction scheme
 nipkow parents: 
3011diff
changeset | 374 | by (induct_tac "l" 1); | 
| 5316 | 375 | by Auto_tac; | 
| 1169 | 376 | qed "rev_rev_ident"; | 
| 2512 | 377 | Addsimps[rev_rev_ident]; | 
| 1169 | 378 | |
| 4935 | 379 | Goal "(rev xs = []) = (xs = [])"; | 
| 4423 | 380 | by (induct_tac "xs" 1); | 
| 5316 | 381 | by Auto_tac; | 
| 3860 | 382 | qed "rev_is_Nil_conv"; | 
| 383 | AddIffs [rev_is_Nil_conv]; | |
| 384 | ||
| 4935 | 385 | Goal "([] = rev xs) = (xs = [])"; | 
| 4423 | 386 | by (induct_tac "xs" 1); | 
| 5316 | 387 | by Auto_tac; | 
| 3860 | 388 | qed "Nil_is_rev_conv"; | 
| 389 | AddIffs [Nil_is_rev_conv]; | |
| 390 | ||
| 6820 | 391 | Goal "!ys. (rev xs = rev ys) = (xs = ys)"; | 
| 6831 | 392 | by (induct_tac "xs" 1); | 
| 6820 | 393 | by (Force_tac 1); | 
| 6831 | 394 | by (rtac allI 1); | 
| 8442 
96023903c2df
case_tac now subsumes both boolean and datatype cases;
 wenzelm parents: 
8423diff
changeset | 395 | by (case_tac "ys" 1); | 
| 6820 | 396 | by (Asm_simp_tac 1); | 
| 397 | by (Force_tac 1); | |
| 398 | qed_spec_mp "rev_is_rev_conv"; | |
| 399 | AddIffs [rev_is_rev_conv]; | |
| 400 | ||
| 4935 | 401 | val prems = Goal "[| P []; !!x xs. P xs ==> P(xs@[x]) |] ==> P xs"; | 
| 5132 | 402 | by (stac (rev_rev_ident RS sym) 1); | 
| 6162 | 403 | by (res_inst_tac [("list", "rev xs")] list.induct 1);
 | 
| 5132 | 404 | by (ALLGOALS Simp_tac); | 
| 405 | by (resolve_tac prems 1); | |
| 406 | by (eresolve_tac prems 1); | |
| 4935 | 407 | qed "rev_induct"; | 
| 408 | ||
| 9747 | 409 | val rev_induct_tac = induct_thm_tac rev_induct; | 
| 5272 | 410 | |
| 4935 | 411 | Goal "(xs = [] --> P) --> (!ys y. xs = ys@[y] --> P) --> P"; | 
| 9747 | 412 | by (rev_induct_tac "xs" 1); | 
| 5316 | 413 | by Auto_tac; | 
| 10385 | 414 | qed "rev_exhaust_aux"; | 
| 415 | ||
| 11770 | 416 | bind_thm ("rev_exhaust", ObjectLogic.rulify rev_exhaust_aux);
 | 
| 4935 | 417 | |
| 2608 | 418 | |
| 3465 | 419 | (** set **) | 
| 1812 | 420 | |
| 3467 | 421 | section "set"; | 
| 422 | ||
| 7032 | 423 | Goal "finite (set xs)"; | 
| 424 | by (induct_tac "xs" 1); | |
| 425 | by Auto_tac; | |
| 426 | qed "finite_set"; | |
| 427 | AddIffs [finite_set]; | |
| 5296 | 428 | |
| 4935 | 429 | Goal "set (xs@ys) = (set xs Un set ys)"; | 
| 3040 
7d48671753da
Introduced a generic "induct_tac" which picks up the right induction scheme
 nipkow parents: 
3011diff
changeset | 430 | by (induct_tac "xs" 1); | 
| 5316 | 431 | by Auto_tac; | 
| 3647 
a64c8fbcd98f
Renamed theorems of the form set_of_list_XXX to set_XXX
 paulson parents: 
3589diff
changeset | 432 | qed "set_append"; | 
| 
a64c8fbcd98f
Renamed theorems of the form set_of_list_XXX to set_XXX
 paulson parents: 
3589diff
changeset | 433 | Addsimps[set_append]; | 
| 1812 | 434 | |
| 4935 | 435 | Goal "set l <= set (x#l)"; | 
| 5316 | 436 | by Auto_tac; | 
| 3647 
a64c8fbcd98f
Renamed theorems of the form set_of_list_XXX to set_XXX
 paulson parents: 
3589diff
changeset | 437 | qed "set_subset_Cons"; | 
| 1936 | 438 | |
| 4935 | 439 | Goal "(set xs = {}) = (xs = [])";
 | 
| 3457 | 440 | by (induct_tac "xs" 1); | 
| 5316 | 441 | by Auto_tac; | 
| 3647 
a64c8fbcd98f
Renamed theorems of the form set_of_list_XXX to set_XXX
 paulson parents: 
3589diff
changeset | 442 | qed "set_empty"; | 
| 
a64c8fbcd98f
Renamed theorems of the form set_of_list_XXX to set_XXX
 paulson parents: 
3589diff
changeset | 443 | Addsimps [set_empty]; | 
| 2608 | 444 | |
| 4935 | 445 | Goal "set(rev xs) = set(xs)"; | 
| 3457 | 446 | by (induct_tac "xs" 1); | 
| 5316 | 447 | by Auto_tac; | 
| 3647 
a64c8fbcd98f
Renamed theorems of the form set_of_list_XXX to set_XXX
 paulson parents: 
3589diff
changeset | 448 | qed "set_rev"; | 
| 
a64c8fbcd98f
Renamed theorems of the form set_of_list_XXX to set_XXX
 paulson parents: 
3589diff
changeset | 449 | Addsimps [set_rev]; | 
| 2608 | 450 | |
| 10832 | 451 | Goal "set(map f xs) = f`(set xs)"; | 
| 3457 | 452 | by (induct_tac "xs" 1); | 
| 5316 | 453 | by Auto_tac; | 
| 3647 
a64c8fbcd98f
Renamed theorems of the form set_of_list_XXX to set_XXX
 paulson parents: 
3589diff
changeset | 454 | qed "set_map"; | 
| 
a64c8fbcd98f
Renamed theorems of the form set_of_list_XXX to set_XXX
 paulson parents: 
3589diff
changeset | 455 | Addsimps [set_map]; | 
| 2608 | 456 | |
| 6433 | 457 | Goal "set(filter P xs) = {x. x : set xs & P x}";
 | 
| 6813 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 paulson parents: 
6794diff
changeset | 458 | by (induct_tac "xs" 1); | 
| 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 paulson parents: 
6794diff
changeset | 459 | by Auto_tac; | 
| 6433 | 460 | qed "set_filter"; | 
| 461 | Addsimps [set_filter]; | |
| 8009 | 462 | |
| 6433 | 463 | Goal "set[i..j(] = {k. i <= k & k < j}";
 | 
| 6813 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 paulson parents: 
6794diff
changeset | 464 | by (induct_tac "j" 1); | 
| 9187 
68ecc04785f1
fixed proof to cope with the default of equalityCE instead of equalityE
 paulson parents: 
9108diff
changeset | 465 | by (ALLGOALS Asm_simp_tac); | 
| 
68ecc04785f1
fixed proof to cope with the default of equalityCE instead of equalityE
 paulson parents: 
9108diff
changeset | 466 | by (etac ssubst 1); | 
| 6813 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 paulson parents: 
6794diff
changeset | 467 | by Auto_tac; | 
| 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 paulson parents: 
6794diff
changeset | 468 | by (arith_tac 1); | 
| 6433 | 469 | qed "set_upt"; | 
| 470 | Addsimps [set_upt]; | |
| 471 | ||
| 5272 | 472 | Goal "(x : set xs) = (? ys zs. xs = ys@x#zs)"; | 
| 5318 | 473 | by (induct_tac "xs" 1); | 
| 474 | by (Simp_tac 1); | |
| 475 | by (Asm_simp_tac 1); | |
| 476 | by (rtac iffI 1); | |
| 477 | by (blast_tac (claset() addIs [eq_Nil_appendI,Cons_eq_appendI]) 1); | |
| 478 | by (REPEAT(etac exE 1)); | |
| 8442 
96023903c2df
case_tac now subsumes both boolean and datatype cases;
 wenzelm parents: 
8423diff
changeset | 479 | by (case_tac "ys" 1); | 
| 5316 | 480 | by Auto_tac; | 
| 5272 | 481 | qed "in_set_conv_decomp"; | 
| 482 | ||
| 8009 | 483 | |
| 5272 | 484 | (* eliminate `lists' in favour of `set' *) | 
| 485 | ||
| 486 | Goal "(xs : lists A) = (!x : set xs. x : A)"; | |
| 5318 | 487 | by (induct_tac "xs" 1); | 
| 5316 | 488 | by Auto_tac; | 
| 5272 | 489 | qed "in_lists_conv_set"; | 
| 490 | ||
| 491 | bind_thm("in_listsD",in_lists_conv_set RS iffD1);
 | |
| 492 | AddSDs [in_listsD]; | |
| 493 | bind_thm("in_listsI",in_lists_conv_set RS iffD2);
 | |
| 494 | AddSIs [in_listsI]; | |
| 1812 | 495 | |
| 5518 | 496 | (** mem **) | 
| 497 | ||
| 498 | section "mem"; | |
| 499 | ||
| 500 | Goal "(x mem xs) = (x: set xs)"; | |
| 501 | by (induct_tac "xs" 1); | |
| 502 | by Auto_tac; | |
| 503 | qed "set_mem_eq"; | |
| 504 | ||
| 505 | ||
| 923 | 506 | (** list_all **) | 
| 507 | ||
| 3467 | 508 | section "list_all"; | 
| 509 | ||
| 5518 | 510 | Goal "list_all P xs = (!x:set xs. P x)"; | 
| 511 | by (induct_tac "xs" 1); | |
| 512 | by Auto_tac; | |
| 513 | qed "list_all_conv"; | |
| 514 | ||
| 5443 
e2459d18ff47
changed constants mem and list_all to mere translations
 oheimb parents: 
5427diff
changeset | 515 | Goal "list_all P (xs@ys) = (list_all P xs & list_all P ys)"; | 
| 3040 
7d48671753da
Introduced a generic "induct_tac" which picks up the right induction scheme
 nipkow parents: 
3011diff
changeset | 516 | by (induct_tac "xs" 1); | 
| 5316 | 517 | by Auto_tac; | 
| 2512 | 518 | qed "list_all_append"; | 
| 519 | Addsimps [list_all_append]; | |
| 923 | 520 | |
| 521 | ||
| 2608 | 522 | (** filter **) | 
| 923 | 523 | |
| 3467 | 524 | section "filter"; | 
| 525 | ||
| 4935 | 526 | Goal "filter P (xs@ys) = filter P xs @ filter P ys"; | 
| 3457 | 527 | by (induct_tac "xs" 1); | 
| 5316 | 528 | by Auto_tac; | 
| 2608 | 529 | qed "filter_append"; | 
| 530 | Addsimps [filter_append]; | |
| 531 | ||
| 4935 | 532 | Goal "filter (%x. True) xs = xs"; | 
| 4605 | 533 | by (induct_tac "xs" 1); | 
| 5316 | 534 | by Auto_tac; | 
| 4605 | 535 | qed "filter_True"; | 
| 536 | Addsimps [filter_True]; | |
| 537 | ||
| 4935 | 538 | Goal "filter (%x. False) xs = []"; | 
| 4605 | 539 | by (induct_tac "xs" 1); | 
| 5316 | 540 | by Auto_tac; | 
| 4605 | 541 | qed "filter_False"; | 
| 542 | Addsimps [filter_False]; | |
| 543 | ||
| 4935 | 544 | Goal "length (filter P xs) <= length xs"; | 
| 3457 | 545 | by (induct_tac "xs" 1); | 
| 5316 | 546 | by Auto_tac; | 
| 8741 
61bc5ed22b62
removal of less_SucI, le_SucI from default simpset
 paulson parents: 
8442diff
changeset | 547 | by (asm_simp_tac (simpset() addsimps [le_SucI]) 1); | 
| 4605 | 548 | qed "length_filter"; | 
| 5443 
e2459d18ff47
changed constants mem and list_all to mere translations
 oheimb parents: 
5427diff
changeset | 549 | Addsimps[length_filter]; | 
| 2608 | 550 | |
| 5443 
e2459d18ff47
changed constants mem and list_all to mere translations
 oheimb parents: 
5427diff
changeset | 551 | Goal "set (filter P xs) <= set xs"; | 
| 
e2459d18ff47
changed constants mem and list_all to mere translations
 oheimb parents: 
5427diff
changeset | 552 | by Auto_tac; | 
| 
e2459d18ff47
changed constants mem and list_all to mere translations
 oheimb parents: 
5427diff
changeset | 553 | qed "filter_is_subset"; | 
| 
e2459d18ff47
changed constants mem and list_all to mere translations
 oheimb parents: 
5427diff
changeset | 554 | Addsimps [filter_is_subset]; | 
| 
e2459d18ff47
changed constants mem and list_all to mere translations
 oheimb parents: 
5427diff
changeset | 555 | |
| 2608 | 556 | |
| 3467 | 557 | section "concat"; | 
| 558 | ||
| 4935 | 559 | Goal "concat(xs@ys) = concat(xs)@concat(ys)"; | 
| 3040 
7d48671753da
Introduced a generic "induct_tac" which picks up the right induction scheme
 nipkow parents: 
3011diff
changeset | 560 | by (induct_tac "xs" 1); | 
| 5316 | 561 | by Auto_tac; | 
| 2608 | 562 | qed"concat_append"; | 
| 563 | Addsimps [concat_append]; | |
| 2512 | 564 | |
| 4935 | 565 | Goal "(concat xss = []) = (!xs:set xss. xs=[])"; | 
| 4423 | 566 | by (induct_tac "xss" 1); | 
| 5316 | 567 | by Auto_tac; | 
| 3896 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 nipkow parents: 
3860diff
changeset | 568 | qed "concat_eq_Nil_conv"; | 
| 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 nipkow parents: 
3860diff
changeset | 569 | AddIffs [concat_eq_Nil_conv]; | 
| 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 nipkow parents: 
3860diff
changeset | 570 | |
| 4935 | 571 | Goal "([] = concat xss) = (!xs:set xss. xs=[])"; | 
| 4423 | 572 | by (induct_tac "xss" 1); | 
| 5316 | 573 | by Auto_tac; | 
| 3896 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 nipkow parents: 
3860diff
changeset | 574 | qed "Nil_eq_concat_conv"; | 
| 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 nipkow parents: 
3860diff
changeset | 575 | AddIffs [Nil_eq_concat_conv]; | 
| 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 nipkow parents: 
3860diff
changeset | 576 | |
| 10832 | 577 | Goal "set(concat xs) = Union(set ` set xs)"; | 
| 3467 | 578 | by (induct_tac "xs" 1); | 
| 5316 | 579 | by Auto_tac; | 
| 3647 
a64c8fbcd98f
Renamed theorems of the form set_of_list_XXX to set_XXX
 paulson parents: 
3589diff
changeset | 580 | qed"set_concat"; | 
| 
a64c8fbcd98f
Renamed theorems of the form set_of_list_XXX to set_XXX
 paulson parents: 
3589diff
changeset | 581 | Addsimps [set_concat]; | 
| 3467 | 582 | |
| 4935 | 583 | Goal "map f (concat xs) = concat (map (map f) xs)"; | 
| 3467 | 584 | by (induct_tac "xs" 1); | 
| 5316 | 585 | by Auto_tac; | 
| 3467 | 586 | qed "map_concat"; | 
| 587 | ||
| 4935 | 588 | Goal "filter p (concat xs) = concat (map (filter p) xs)"; | 
| 3467 | 589 | by (induct_tac "xs" 1); | 
| 5316 | 590 | by Auto_tac; | 
| 3467 | 591 | qed"filter_concat"; | 
| 592 | ||
| 4935 | 593 | Goal "rev(concat xs) = concat (map rev (rev xs))"; | 
| 3467 | 594 | by (induct_tac "xs" 1); | 
| 5316 | 595 | by Auto_tac; | 
| 2608 | 596 | qed "rev_concat"; | 
| 923 | 597 | |
| 598 | (** nth **) | |
| 599 | ||
| 3467 | 600 | section "nth"; | 
| 601 | ||
| 6408 | 602 | Goal "(x#xs)!0 = x"; | 
| 603 | by Auto_tac; | |
| 604 | qed "nth_Cons_0"; | |
| 605 | Addsimps [nth_Cons_0]; | |
| 5644 | 606 | |
| 6408 | 607 | Goal "(x#xs)!(Suc n) = xs!n"; | 
| 608 | by Auto_tac; | |
| 609 | qed "nth_Cons_Suc"; | |
| 610 | Addsimps [nth_Cons_Suc]; | |
| 611 | ||
| 612 | Delsimps (thms "nth.simps"); | |
| 613 | ||
| 614 | Goal "!n. (xs@ys)!n = (if n < length xs then xs!n else ys!(n - length xs))"; | |
| 615 | by (induct_tac "xs" 1); | |
| 3457 | 616 | by (Asm_simp_tac 1); | 
| 617 | by (rtac allI 1); | |
| 8442 
96023903c2df
case_tac now subsumes both boolean and datatype cases;
 wenzelm parents: 
8423diff
changeset | 618 | by (case_tac "n" 1); | 
| 5316 | 619 | by Auto_tac; | 
| 2608 | 620 | qed_spec_mp "nth_append"; | 
| 621 | ||
| 4935 | 622 | Goal "!n. n < length xs --> (map f xs)!n = f(xs!n)"; | 
| 3040 
7d48671753da
Introduced a generic "induct_tac" which picks up the right induction scheme
 nipkow parents: 
3011diff
changeset | 623 | by (induct_tac "xs" 1); | 
| 8118 | 624 | by (Asm_full_simp_tac 1); | 
| 1301 | 625 | by (rtac allI 1); | 
| 5183 | 626 | by (induct_tac "n" 1); | 
| 5316 | 627 | by Auto_tac; | 
| 1485 
240cc98b94a7
Added qed_spec_mp to avoid renaming of bound vars in 'th RS spec'
 nipkow parents: 
1465diff
changeset | 628 | qed_spec_mp "nth_map"; | 
| 1301 | 629 | Addsimps [nth_map]; | 
| 630 | ||
| 8118 | 631 | Goal "set xs = {xs!i |i. i < length xs}";
 | 
| 3040 
7d48671753da
Introduced a generic "induct_tac" which picks up the right induction scheme
 nipkow parents: 
3011diff
changeset | 632 | by (induct_tac "xs" 1); | 
| 8118 | 633 | by (Simp_tac 1); | 
| 8254 | 634 | by (Asm_simp_tac 1); | 
| 635 | by Safe_tac; | |
| 636 |   by (res_inst_tac [("x","0")] exI 1);
 | |
| 8118 | 637 | by (Simp_tac 1); | 
| 8254 | 638 |  by (res_inst_tac [("x","Suc i")] exI 1);
 | 
| 639 | by (Asm_simp_tac 1); | |
| 8442 
96023903c2df
case_tac now subsumes both boolean and datatype cases;
 wenzelm parents: 
8423diff
changeset | 640 | by (case_tac "i" 1); | 
| 8254 | 641 | by (Asm_full_simp_tac 1); | 
| 642 | by (rename_tac "j" 1); | |
| 643 |  by (res_inst_tac [("x","j")] exI 1);
 | |
| 644 | by (Asm_simp_tac 1); | |
| 8118 | 645 | qed "set_conv_nth"; | 
| 646 | ||
| 647 | Goal "n < length xs ==> Ball (set xs) P --> P(xs!n)"; | |
| 648 | by (simp_tac (simpset() addsimps [set_conv_nth]) 1); | |
| 8254 | 649 | by (Blast_tac 1); | 
| 5518 | 650 | qed_spec_mp "list_ball_nth"; | 
| 1301 | 651 | |
| 8118 | 652 | Goal "n < length xs ==> xs!n : set xs"; | 
| 653 | by (simp_tac (simpset() addsimps [set_conv_nth]) 1); | |
| 8254 | 654 | by (Blast_tac 1); | 
| 1485 
240cc98b94a7
Added qed_spec_mp to avoid renaming of bound vars in 'th RS spec'
 nipkow parents: 
1465diff
changeset | 655 | qed_spec_mp "nth_mem"; | 
| 1301 | 656 | Addsimps [nth_mem]; | 
| 657 | ||
| 8009 | 658 | Goal "(!i. i < length xs --> P(xs!i)) --> (!x : set xs. P x)"; | 
| 8118 | 659 | by (simp_tac (simpset() addsimps [set_conv_nth]) 1); | 
| 8254 | 660 | by (Blast_tac 1); | 
| 8009 | 661 | qed_spec_mp "all_nth_imp_all_set"; | 
| 662 | ||
| 663 | Goal "(!x : set xs. P x) = (!i. i<length xs --> P (xs ! i))"; | |
| 8118 | 664 | by (simp_tac (simpset() addsimps [set_conv_nth]) 1); | 
| 8254 | 665 | by (Blast_tac 1); | 
| 8009 | 666 | qed_spec_mp "all_set_conv_all_nth"; | 
| 667 | ||
| 668 | ||
| 5077 
71043526295f
* HOL/List: new function list_update written xs[i:=v] that updates the i-th
 nipkow parents: 
5043diff
changeset | 669 | (** list update **) | 
| 
71043526295f
* HOL/List: new function list_update written xs[i:=v] that updates the i-th
 nipkow parents: 
5043diff
changeset | 670 | |
| 
71043526295f
* HOL/List: new function list_update written xs[i:=v] that updates the i-th
 nipkow parents: 
5043diff
changeset | 671 | section "list update"; | 
| 
71043526295f
* HOL/List: new function list_update written xs[i:=v] that updates the i-th
 nipkow parents: 
5043diff
changeset | 672 | |
| 
71043526295f
* HOL/List: new function list_update written xs[i:=v] that updates the i-th
 nipkow parents: 
5043diff
changeset | 673 | Goal "!i. length(xs[i:=x]) = length xs"; | 
| 
71043526295f
* HOL/List: new function list_update written xs[i:=v] that updates the i-th
 nipkow parents: 
5043diff
changeset | 674 | by (induct_tac "xs" 1); | 
| 
71043526295f
* HOL/List: new function list_update written xs[i:=v] that updates the i-th
 nipkow parents: 
5043diff
changeset | 675 | by (Simp_tac 1); | 
| 5183 | 676 | by (asm_full_simp_tac (simpset() addsplits [nat.split]) 1); | 
| 5077 
71043526295f
* HOL/List: new function list_update written xs[i:=v] that updates the i-th
 nipkow parents: 
5043diff
changeset | 677 | qed_spec_mp "length_list_update"; | 
| 
71043526295f
* HOL/List: new function list_update written xs[i:=v] that updates the i-th
 nipkow parents: 
5043diff
changeset | 678 | Addsimps [length_list_update]; | 
| 
71043526295f
* HOL/List: new function list_update written xs[i:=v] that updates the i-th
 nipkow parents: 
5043diff
changeset | 679 | |
| 5644 | 680 | Goal "!i j. i < length xs --> (xs[i:=x])!j = (if i=j then x else xs!j)"; | 
| 6162 | 681 | by (induct_tac "xs" 1); | 
| 682 | by (Simp_tac 1); | |
| 683 | by (auto_tac (claset(), simpset() addsimps [nth_Cons] addsplits [nat.split])); | |
| 5644 | 684 | qed_spec_mp "nth_list_update"; | 
| 685 | ||
| 8144 | 686 | Goal "i < length xs ==> (xs[i:=x])!i = x"; | 
| 687 | by (asm_simp_tac (simpset() addsimps [nth_list_update]) 1); | |
| 688 | qed "nth_list_update_eq"; | |
| 689 | Addsimps [nth_list_update_eq]; | |
| 690 | ||
| 691 | Goal "!i j. i ~= j --> xs[i:=x]!j = xs!j"; | |
| 692 | by (induct_tac "xs" 1); | |
| 693 | by (Simp_tac 1); | |
| 694 | by (auto_tac (claset(), simpset() addsimps [nth_Cons] addsplits [nat.split])); | |
| 695 | qed_spec_mp "nth_list_update_neq"; | |
| 696 | Addsimps [nth_list_update_neq]; | |
| 697 | ||
| 6433 | 698 | Goal "!i. i < size xs --> xs[i:=x, i:=y] = xs[i:=y]"; | 
| 6813 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 paulson parents: 
6794diff
changeset | 699 | by (induct_tac "xs" 1); | 
| 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 paulson parents: 
6794diff
changeset | 700 | by (Simp_tac 1); | 
| 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 paulson parents: 
6794diff
changeset | 701 | by (asm_simp_tac (simpset() addsplits [nat.split]) 1); | 
| 6433 | 702 | qed_spec_mp "list_update_overwrite"; | 
| 703 | Addsimps [list_update_overwrite]; | |
| 704 | ||
| 705 | Goal "!i < length xs. (xs[i := x] = xs) = (xs!i = x)"; | |
| 6813 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 paulson parents: 
6794diff
changeset | 706 | by (induct_tac "xs" 1); | 
| 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 paulson parents: 
6794diff
changeset | 707 | by (Simp_tac 1); | 
| 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 paulson parents: 
6794diff
changeset | 708 | by (simp_tac (simpset() addsplits [nat.split]) 1); | 
| 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 paulson parents: 
6794diff
changeset | 709 | by (Blast_tac 1); | 
| 6433 | 710 | qed_spec_mp "list_update_same_conv"; | 
| 711 | ||
| 8009 | 712 | Goal "!i xy xs. length xs = length ys --> \ | 
| 713 | \ (zip xs ys)[i:=xy] = zip (xs[i:=fst xy]) (ys[i:=snd xy])"; | |
| 714 | by (induct_tac "ys" 1); | |
| 715 | by Auto_tac; | |
| 8442 
96023903c2df
case_tac now subsumes both boolean and datatype cases;
 wenzelm parents: 
8423diff
changeset | 716 | by (case_tac "xs" 1); | 
| 8009 | 717 | by (auto_tac (claset(), simpset() addsplits [nat.split])); | 
| 718 | qed_spec_mp "update_zip"; | |
| 719 | ||
| 720 | Goal "!i. set(xs[i:=x]) <= insert x (set xs)"; | |
| 721 | by (induct_tac "xs" 1); | |
| 722 | by (asm_full_simp_tac (simpset() addsimps []) 1); | |
| 723 | by (asm_full_simp_tac (simpset() addsplits [nat.split]) 1); | |
| 724 | by (Fast_tac 1); | |
| 8287 | 725 | qed_spec_mp "set_update_subset_insert"; | 
| 8009 | 726 | |
| 8287 | 727 | Goal "[| set xs <= A; x:A |] ==> set(xs[i := x]) <= A"; | 
| 728 | by(fast_tac (claset() addSDs [set_update_subset_insert RS subsetD]) 1); | |
| 729 | qed "set_update_subsetI"; | |
| 5077 
71043526295f
* HOL/List: new function list_update written xs[i:=v] that updates the i-th
 nipkow parents: 
5043diff
changeset | 730 | |
| 3896 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 nipkow parents: 
3860diff
changeset | 731 | (** last & butlast **) | 
| 1327 
6c29cfab679c
added new arithmetic lemmas and the functions take and drop.
 nipkow parents: 
1301diff
changeset | 732 | |
| 5644 | 733 | section "last / butlast"; | 
| 734 | ||
| 4935 | 735 | Goal "last(xs@[x]) = x"; | 
| 4423 | 736 | by (induct_tac "xs" 1); | 
| 5316 | 737 | by Auto_tac; | 
| 3896 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 nipkow parents: 
3860diff
changeset | 738 | qed "last_snoc"; | 
| 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 nipkow parents: 
3860diff
changeset | 739 | Addsimps [last_snoc]; | 
| 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 nipkow parents: 
3860diff
changeset | 740 | |
| 4935 | 741 | Goal "butlast(xs@[x]) = xs"; | 
| 4423 | 742 | by (induct_tac "xs" 1); | 
| 5316 | 743 | by Auto_tac; | 
| 3896 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 nipkow parents: 
3860diff
changeset | 744 | qed "butlast_snoc"; | 
| 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 nipkow parents: 
3860diff
changeset | 745 | Addsimps [butlast_snoc]; | 
| 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 nipkow parents: 
3860diff
changeset | 746 | |
| 4935 | 747 | Goal "length(butlast xs) = length xs - 1"; | 
| 9747 | 748 | by (rev_induct_tac "xs" 1); | 
| 5316 | 749 | by Auto_tac; | 
| 4643 | 750 | qed "length_butlast"; | 
| 751 | Addsimps [length_butlast]; | |
| 752 | ||
| 5278 | 753 | Goal "!ys. butlast (xs@ys) = (if ys=[] then butlast xs else xs@butlast ys)"; | 
| 4423 | 754 | by (induct_tac "xs" 1); | 
| 5316 | 755 | by Auto_tac; | 
| 3896 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 nipkow parents: 
3860diff
changeset | 756 | qed_spec_mp "butlast_append"; | 
| 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 nipkow parents: 
3860diff
changeset | 757 | |
| 8118 | 758 | Goal "xs ~= [] --> butlast xs @ [last xs] = xs"; | 
| 8254 | 759 | by (induct_tac "xs" 1); | 
| 760 | by (ALLGOALS Asm_simp_tac); | |
| 8118 | 761 | qed_spec_mp "append_butlast_last_id"; | 
| 762 | Addsimps [append_butlast_last_id]; | |
| 763 | ||
| 4935 | 764 | Goal "x:set(butlast xs) --> x:set xs"; | 
| 4423 | 765 | by (induct_tac "xs" 1); | 
| 5316 | 766 | by Auto_tac; | 
| 3896 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 nipkow parents: 
3860diff
changeset | 767 | qed_spec_mp "in_set_butlastD"; | 
| 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 nipkow parents: 
3860diff
changeset | 768 | |
| 5448 
40a09282ba14
in_set_butlast_appendI supersedes in_set_butlast_appendI1,2
 paulson parents: 
5443diff
changeset | 769 | Goal "x:set(butlast xs) | x:set(butlast ys) ==> x:set(butlast(xs@ys))"; | 
| 
40a09282ba14
in_set_butlast_appendI supersedes in_set_butlast_appendI1,2
 paulson parents: 
5443diff
changeset | 770 | by (auto_tac (claset() addDs [in_set_butlastD], | 
| 
40a09282ba14
in_set_butlast_appendI supersedes in_set_butlast_appendI1,2
 paulson parents: 
5443diff
changeset | 771 | simpset() addsimps [butlast_append])); | 
| 
40a09282ba14
in_set_butlast_appendI supersedes in_set_butlast_appendI1,2
 paulson parents: 
5443diff
changeset | 772 | qed "in_set_butlast_appendI"; | 
| 3902 | 773 | |
| 2608 | 774 | (** take & drop **) | 
| 775 | section "take & drop"; | |
| 1327 
6c29cfab679c
added new arithmetic lemmas and the functions take and drop.
 nipkow parents: 
1301diff
changeset | 776 | |
| 4935 | 777 | Goal "take 0 xs = []"; | 
| 3040 
7d48671753da
Introduced a generic "induct_tac" which picks up the right induction scheme
 nipkow parents: 
3011diff
changeset | 778 | by (induct_tac "xs" 1); | 
| 5316 | 779 | by Auto_tac; | 
| 1327 
6c29cfab679c
added new arithmetic lemmas and the functions take and drop.
 nipkow parents: 
1301diff
changeset | 780 | qed "take_0"; | 
| 
6c29cfab679c
added new arithmetic lemmas and the functions take and drop.
 nipkow parents: 
1301diff
changeset | 781 | |
| 4935 | 782 | Goal "drop 0 xs = xs"; | 
| 3040 
7d48671753da
Introduced a generic "induct_tac" which picks up the right induction scheme
 nipkow parents: 
3011diff
changeset | 783 | by (induct_tac "xs" 1); | 
| 5316 | 784 | by Auto_tac; | 
| 2608 | 785 | qed "drop_0"; | 
| 786 | ||
| 4935 | 787 | Goal "take (Suc n) (x#xs) = x # take n xs"; | 
| 1552 | 788 | by (Simp_tac 1); | 
| 1419 
a6a034a47a71
defined take/drop by induction over list rather than nat.
 nipkow parents: 
1327diff
changeset | 789 | qed "take_Suc_Cons"; | 
| 1327 
6c29cfab679c
added new arithmetic lemmas and the functions take and drop.
 nipkow parents: 
1301diff
changeset | 790 | |
| 4935 | 791 | Goal "drop (Suc n) (x#xs) = drop n xs"; | 
| 2608 | 792 | by (Simp_tac 1); | 
| 793 | qed "drop_Suc_Cons"; | |
| 794 | ||
| 795 | Delsimps [take_Cons,drop_Cons]; | |
| 796 | Addsimps [take_0,take_Suc_Cons,drop_0,drop_Suc_Cons]; | |
| 797 | ||
| 4935 | 798 | Goal "!xs. length(take n xs) = min (length xs) n"; | 
| 5183 | 799 | by (induct_tac "n" 1); | 
| 5316 | 800 | by Auto_tac; | 
| 8442 
96023903c2df
case_tac now subsumes both boolean and datatype cases;
 wenzelm parents: 
8423diff
changeset | 801 | by (case_tac "xs" 1); | 
| 5316 | 802 | by Auto_tac; | 
| 2608 | 803 | qed_spec_mp "length_take"; | 
| 804 | Addsimps [length_take]; | |
| 923 | 805 | |
| 4935 | 806 | Goal "!xs. length(drop n xs) = (length xs - n)"; | 
| 5183 | 807 | by (induct_tac "n" 1); | 
| 5316 | 808 | by Auto_tac; | 
| 8442 
96023903c2df
case_tac now subsumes both boolean and datatype cases;
 wenzelm parents: 
8423diff
changeset | 809 | by (case_tac "xs" 1); | 
| 5316 | 810 | by Auto_tac; | 
| 2608 | 811 | qed_spec_mp "length_drop"; | 
| 812 | Addsimps [length_drop]; | |
| 813 | ||
| 4935 | 814 | Goal "!xs. length xs <= n --> take n xs = xs"; | 
| 5183 | 815 | by (induct_tac "n" 1); | 
| 5316 | 816 | by Auto_tac; | 
| 8442 
96023903c2df
case_tac now subsumes both boolean and datatype cases;
 wenzelm parents: 
8423diff
changeset | 817 | by (case_tac "xs" 1); | 
| 5316 | 818 | by Auto_tac; | 
| 2608 | 819 | qed_spec_mp "take_all"; | 
| 7246 | 820 | Addsimps [take_all]; | 
| 923 | 821 | |
| 4935 | 822 | Goal "!xs. length xs <= n --> drop n xs = []"; | 
| 5183 | 823 | by (induct_tac "n" 1); | 
| 5316 | 824 | by Auto_tac; | 
| 8442 
96023903c2df
case_tac now subsumes both boolean and datatype cases;
 wenzelm parents: 
8423diff
changeset | 825 | by (case_tac "xs" 1); | 
| 5316 | 826 | by Auto_tac; | 
| 2608 | 827 | qed_spec_mp "drop_all"; | 
| 7246 | 828 | Addsimps [drop_all]; | 
| 2608 | 829 | |
| 5278 | 830 | Goal "!xs. take n (xs @ ys) = (take n xs @ take (n - length xs) ys)"; | 
| 5183 | 831 | by (induct_tac "n" 1); | 
| 5316 | 832 | by Auto_tac; | 
| 8442 
96023903c2df
case_tac now subsumes both boolean and datatype cases;
 wenzelm parents: 
8423diff
changeset | 833 | by (case_tac "xs" 1); | 
| 5316 | 834 | by Auto_tac; | 
| 2608 | 835 | qed_spec_mp "take_append"; | 
| 836 | Addsimps [take_append]; | |
| 837 | ||
| 4935 | 838 | Goal "!xs. drop n (xs@ys) = drop n xs @ drop (n - length xs) ys"; | 
| 5183 | 839 | by (induct_tac "n" 1); | 
| 5316 | 840 | by Auto_tac; | 
| 8442 
96023903c2df
case_tac now subsumes both boolean and datatype cases;
 wenzelm parents: 
8423diff
changeset | 841 | by (case_tac "xs" 1); | 
| 5316 | 842 | by Auto_tac; | 
| 2608 | 843 | qed_spec_mp "drop_append"; | 
| 844 | Addsimps [drop_append]; | |
| 845 | ||
| 4935 | 846 | Goal "!xs n. take n (take m xs) = take (min n m) xs"; | 
| 5183 | 847 | by (induct_tac "m" 1); | 
| 5316 | 848 | by Auto_tac; | 
| 8442 
96023903c2df
case_tac now subsumes both boolean and datatype cases;
 wenzelm parents: 
8423diff
changeset | 849 | by (case_tac "xs" 1); | 
| 5316 | 850 | by Auto_tac; | 
| 8442 
96023903c2df
case_tac now subsumes both boolean and datatype cases;
 wenzelm parents: 
8423diff
changeset | 851 | by (case_tac "na" 1); | 
| 5316 | 852 | by Auto_tac; | 
| 2608 | 853 | qed_spec_mp "take_take"; | 
| 7570 | 854 | Addsimps [take_take]; | 
| 2608 | 855 | |
| 4935 | 856 | Goal "!xs. drop n (drop m xs) = drop (n + m) xs"; | 
| 5183 | 857 | by (induct_tac "m" 1); | 
| 5316 | 858 | by Auto_tac; | 
| 8442 
96023903c2df
case_tac now subsumes both boolean and datatype cases;
 wenzelm parents: 
8423diff
changeset | 859 | by (case_tac "xs" 1); | 
| 5316 | 860 | by Auto_tac; | 
| 2608 | 861 | qed_spec_mp "drop_drop"; | 
| 7570 | 862 | Addsimps [drop_drop]; | 
| 923 | 863 | |
| 4935 | 864 | Goal "!xs n. take n (drop m xs) = drop m (take (n + m) xs)"; | 
| 5183 | 865 | by (induct_tac "m" 1); | 
| 5316 | 866 | by Auto_tac; | 
| 8442 
96023903c2df
case_tac now subsumes both boolean and datatype cases;
 wenzelm parents: 
8423diff
changeset | 867 | by (case_tac "xs" 1); | 
| 5316 | 868 | by Auto_tac; | 
| 2608 | 869 | qed_spec_mp "take_drop"; | 
| 870 | ||
| 6813 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 paulson parents: 
6794diff
changeset | 871 | Goal "!xs. take n xs @ drop n xs = xs"; | 
| 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 paulson parents: 
6794diff
changeset | 872 | by (induct_tac "n" 1); | 
| 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 paulson parents: 
6794diff
changeset | 873 | by Auto_tac; | 
| 8442 
96023903c2df
case_tac now subsumes both boolean and datatype cases;
 wenzelm parents: 
8423diff
changeset | 874 | by (case_tac "xs" 1); | 
| 6813 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 paulson parents: 
6794diff
changeset | 875 | by Auto_tac; | 
| 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 paulson parents: 
6794diff
changeset | 876 | qed_spec_mp "append_take_drop_id"; | 
| 8118 | 877 | Addsimps [append_take_drop_id]; | 
| 6813 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 paulson parents: 
6794diff
changeset | 878 | |
| 4935 | 879 | Goal "!xs. take n (map f xs) = map f (take n xs)"; | 
| 5183 | 880 | by (induct_tac "n" 1); | 
| 5316 | 881 | by Auto_tac; | 
| 8442 
96023903c2df
case_tac now subsumes both boolean and datatype cases;
 wenzelm parents: 
8423diff
changeset | 882 | by (case_tac "xs" 1); | 
| 5316 | 883 | by Auto_tac; | 
| 2608 | 884 | qed_spec_mp "take_map"; | 
| 885 | ||
| 4935 | 886 | Goal "!xs. drop n (map f xs) = map f (drop n xs)"; | 
| 5183 | 887 | by (induct_tac "n" 1); | 
| 5316 | 888 | by Auto_tac; | 
| 8442 
96023903c2df
case_tac now subsumes both boolean and datatype cases;
 wenzelm parents: 
8423diff
changeset | 889 | by (case_tac "xs" 1); | 
| 5316 | 890 | by Auto_tac; | 
| 2608 | 891 | qed_spec_mp "drop_map"; | 
| 892 | ||
| 4935 | 893 | Goal "!n i. i < n --> (take n xs)!i = xs!i"; | 
| 3457 | 894 | by (induct_tac "xs" 1); | 
| 5316 | 895 | by Auto_tac; | 
| 8442 
96023903c2df
case_tac now subsumes both boolean and datatype cases;
 wenzelm parents: 
8423diff
changeset | 896 | by (case_tac "n" 1); | 
| 3457 | 897 | by (Blast_tac 1); | 
| 8442 
96023903c2df
case_tac now subsumes both boolean and datatype cases;
 wenzelm parents: 
8423diff
changeset | 898 | by (case_tac "i" 1); | 
| 5316 | 899 | by Auto_tac; | 
| 2608 | 900 | qed_spec_mp "nth_take"; | 
| 901 | Addsimps [nth_take]; | |
| 923 | 902 | |
| 4935 | 903 | Goal "!xs i. n + i <= length xs --> (drop n xs)!i = xs!(n+i)"; | 
| 5183 | 904 | by (induct_tac "n" 1); | 
| 5316 | 905 | by Auto_tac; | 
| 8442 
96023903c2df
case_tac now subsumes both boolean and datatype cases;
 wenzelm parents: 
8423diff
changeset | 906 | by (case_tac "xs" 1); | 
| 5316 | 907 | by Auto_tac; | 
| 2608 | 908 | qed_spec_mp "nth_drop"; | 
| 909 | Addsimps [nth_drop]; | |
| 910 | ||
| 8118 | 911 | |
| 912 | Goal | |
| 913 | "!zs. (xs@ys = zs) = (xs = take (length xs) zs & ys = drop (length xs) zs)"; | |
| 8254 | 914 | by (induct_tac "xs" 1); | 
| 915 | by (Simp_tac 1); | |
| 916 | by (Asm_full_simp_tac 1); | |
| 917 | by (Clarify_tac 1); | |
| 8442 
96023903c2df
case_tac now subsumes both boolean and datatype cases;
 wenzelm parents: 
8423diff
changeset | 918 | by (case_tac "zs" 1); | 
| 8254 | 919 | by (Auto_tac); | 
| 8118 | 920 | qed_spec_mp "append_eq_conv_conj"; | 
| 921 | ||
| 2608 | 922 | (** takeWhile & dropWhile **) | 
| 923 | ||
| 3467 | 924 | section "takeWhile & dropWhile"; | 
| 925 | ||
| 4935 | 926 | Goal "takeWhile P xs @ dropWhile P xs = xs"; | 
| 3586 | 927 | by (induct_tac "xs" 1); | 
| 5316 | 928 | by Auto_tac; | 
| 3586 | 929 | qed "takeWhile_dropWhile_id"; | 
| 930 | Addsimps [takeWhile_dropWhile_id]; | |
| 931 | ||
| 4935 | 932 | Goal "x:set xs & ~P(x) --> takeWhile P (xs @ ys) = takeWhile P xs"; | 
| 3457 | 933 | by (induct_tac "xs" 1); | 
| 5316 | 934 | by Auto_tac; | 
| 2608 | 935 | bind_thm("takeWhile_append1", conjI RS (result() RS mp));
 | 
| 936 | Addsimps [takeWhile_append1]; | |
| 923 | 937 | |
| 4935 | 938 | Goal "(!x:set xs. P(x)) --> takeWhile P (xs @ ys) = xs @ takeWhile P ys"; | 
| 3457 | 939 | by (induct_tac "xs" 1); | 
| 5316 | 940 | by Auto_tac; | 
| 2608 | 941 | bind_thm("takeWhile_append2", ballI RS (result() RS mp));
 | 
| 942 | Addsimps [takeWhile_append2]; | |
| 1169 | 943 | |
| 11289 | 944 | Goal "~P(x) ==> takeWhile P (xs @ (x#l)) = takeWhile P xs"; | 
| 945 | by (induct_tac "xs" 1); | |
| 946 | by Auto_tac; | |
| 947 | qed "takeWhile_tail"; | |
| 948 | ||
| 4935 | 949 | Goal "x:set xs & ~P(x) --> dropWhile P (xs @ ys) = (dropWhile P xs)@ys"; | 
| 3457 | 950 | by (induct_tac "xs" 1); | 
| 5316 | 951 | by Auto_tac; | 
| 2608 | 952 | bind_thm("dropWhile_append1", conjI RS (result() RS mp));
 | 
| 953 | Addsimps [dropWhile_append1]; | |
| 954 | ||
| 4935 | 955 | Goal "(!x:set xs. P(x)) --> dropWhile P (xs @ ys) = dropWhile P ys"; | 
| 3457 | 956 | by (induct_tac "xs" 1); | 
| 5316 | 957 | by Auto_tac; | 
| 2608 | 958 | bind_thm("dropWhile_append2", ballI RS (result() RS mp));
 | 
| 959 | Addsimps [dropWhile_append2]; | |
| 960 | ||
| 4935 | 961 | Goal "x:set(takeWhile P xs) --> x:set xs & P x"; | 
| 3457 | 962 | by (induct_tac "xs" 1); | 
| 5316 | 963 | by Auto_tac; | 
| 3647 
a64c8fbcd98f
Renamed theorems of the form set_of_list_XXX to set_XXX
 paulson parents: 
3589diff
changeset | 964 | qed_spec_mp"set_take_whileD"; | 
| 2608 | 965 | |
| 6306 | 966 | (** zip **) | 
| 967 | section "zip"; | |
| 968 | ||
| 969 | Goal "zip [] ys = []"; | |
| 6813 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 paulson parents: 
6794diff
changeset | 970 | by (induct_tac "ys" 1); | 
| 6306 | 971 | by Auto_tac; | 
| 972 | qed "zip_Nil"; | |
| 973 | Addsimps [zip_Nil]; | |
| 974 | ||
| 975 | Goal "zip (x#xs) (y#ys) = (x,y)#zip xs ys"; | |
| 6813 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 paulson parents: 
6794diff
changeset | 976 | by (Simp_tac 1); | 
| 6306 | 977 | qed "zip_Cons_Cons"; | 
| 978 | Addsimps [zip_Cons_Cons]; | |
| 979 | ||
| 980 | Delsimps(tl (thms"zip.simps")); | |
| 4605 | 981 | |
| 8118 | 982 | Goal "!xs. length (zip xs ys) = min (length xs) (length ys)"; | 
| 8009 | 983 | by (induct_tac "ys" 1); | 
| 984 | by (Simp_tac 1); | |
| 985 | by (Clarify_tac 1); | |
| 8442 
96023903c2df
case_tac now subsumes both boolean and datatype cases;
 wenzelm parents: 
8423diff
changeset | 986 | by (case_tac "xs" 1); | 
| 8064 | 987 | by (Auto_tac); | 
| 8009 | 988 | qed_spec_mp "length_zip"; | 
| 989 | Addsimps [length_zip]; | |
| 990 | ||
| 991 | Goal | |
| 8118 | 992 | "!xs. zip (xs@ys) zs = \ | 
| 993 | \ zip xs (take (length xs) zs) @ zip ys (drop (length xs) zs)"; | |
| 8254 | 994 | by (induct_tac "zs" 1); | 
| 995 | by (Simp_tac 1); | |
| 8064 | 996 | by (Clarify_tac 1); | 
| 8442 
96023903c2df
case_tac now subsumes both boolean and datatype cases;
 wenzelm parents: 
8423diff
changeset | 997 | by (case_tac "xs" 1); | 
| 8254 | 998 | by (Asm_simp_tac 1); | 
| 999 | by (Asm_simp_tac 1); | |
| 8118 | 1000 | qed_spec_mp "zip_append1"; | 
| 1001 | ||
| 1002 | Goal | |
| 1003 | "!ys. zip xs (ys@zs) = \ | |
| 1004 | \ zip (take (length ys) xs) ys @ zip (drop (length ys) xs) zs"; | |
| 8254 | 1005 | by (induct_tac "xs" 1); | 
| 1006 | by (Simp_tac 1); | |
| 8118 | 1007 | by (Clarify_tac 1); | 
| 8442 
96023903c2df
case_tac now subsumes both boolean and datatype cases;
 wenzelm parents: 
8423diff
changeset | 1008 | by (case_tac "ys" 1); | 
| 8254 | 1009 | by (Asm_simp_tac 1); | 
| 1010 | by (Asm_simp_tac 1); | |
| 8118 | 1011 | qed_spec_mp "zip_append2"; | 
| 1012 | ||
| 1013 | Goal | |
| 1014 | "[| length xs = length us; length ys = length vs |] ==> \ | |
| 1015 | \ zip (xs@ys) (us@vs) = zip xs us @ zip ys vs"; | |
| 8254 | 1016 | by (asm_simp_tac (simpset() addsimps [zip_append1]) 1); | 
| 8009 | 1017 | qed_spec_mp "zip_append"; | 
| 8118 | 1018 | Addsimps [zip_append]; | 
| 8009 | 1019 | |
| 1020 | Goal "!xs. length xs = length ys --> zip (rev xs) (rev ys) = rev (zip xs ys)"; | |
| 8064 | 1021 | by (induct_tac "ys" 1); | 
| 1022 | by (Asm_full_simp_tac 1); | |
| 1023 | by (Asm_full_simp_tac 1); | |
| 1024 | by (Clarify_tac 1); | |
| 8442 
96023903c2df
case_tac now subsumes both boolean and datatype cases;
 wenzelm parents: 
8423diff
changeset | 1025 | by (case_tac "xs" 1); | 
| 8064 | 1026 | by (Auto_tac); | 
| 8009 | 1027 | qed_spec_mp "zip_rev"; | 
| 1028 | ||
| 8115 | 1029 | |
| 1030 | Goal | |
| 8009 | 1031 | "!i xs. i < length xs --> i < length ys --> (zip xs ys)!i = (xs!i, ys!i)"; | 
| 1032 | by (induct_tac "ys" 1); | |
| 1033 | by (Simp_tac 1); | |
| 1034 | by (Clarify_tac 1); | |
| 8442 
96023903c2df
case_tac now subsumes both boolean and datatype cases;
 wenzelm parents: 
8423diff
changeset | 1035 | by (case_tac "xs" 1); | 
| 8064 | 1036 | by (Auto_tac); | 
| 8009 | 1037 | by (asm_full_simp_tac (simpset() addsimps (thms"nth.simps") addsplits [nat.split]) 1); | 
| 1038 | qed_spec_mp "nth_zip"; | |
| 1039 | Addsimps [nth_zip]; | |
| 1040 | ||
| 8118 | 1041 | Goal "set(zip xs ys) = {(xs!i,ys!i) |i. i < min (length xs) (length ys)}";
 | 
| 1042 | by (simp_tac (simpset() addsimps [set_conv_nth]addcongs [rev_conj_cong]) 1); | |
| 1043 | qed_spec_mp "set_zip"; | |
| 1044 | ||
| 8009 | 1045 | Goal | 
| 1046 | "length xs = length ys ==> zip (xs[i:=x]) (ys[i:=y]) = (zip xs ys)[i:=(x,y)]"; | |
| 8064 | 1047 | by (rtac sym 1); | 
| 1048 | by (asm_simp_tac (simpset() addsimps [update_zip]) 1); | |
| 8009 | 1049 | qed_spec_mp "zip_update"; | 
| 1050 | ||
| 1051 | Goal "!j. zip (replicate i x) (replicate j y) = replicate (min i j) (x,y)"; | |
| 1052 | by (induct_tac "i" 1); | |
| 8064 | 1053 | by (Auto_tac); | 
| 8442 
96023903c2df
case_tac now subsumes both boolean and datatype cases;
 wenzelm parents: 
8423diff
changeset | 1054 | by (case_tac "j" 1); | 
| 8064 | 1055 | by (Auto_tac); | 
| 8009 | 1056 | qed "zip_replicate"; | 
| 1057 | Addsimps [zip_replicate]; | |
| 1058 | ||
| 8115 | 1059 | (** list_all2 **) | 
| 1060 | section "list_all2"; | |
| 1061 | ||
| 1062 | Goalw [list_all2_def] "list_all2 P xs ys ==> length xs = length ys"; | |
| 8254 | 1063 | by (Asm_simp_tac 1); | 
| 8115 | 1064 | qed "list_all2_lengthD"; | 
| 1065 | ||
| 1066 | Goalw [list_all2_def] "list_all2 P [] ys = (ys=[])"; | |
| 1067 | by (Simp_tac 1); | |
| 1068 | qed "list_all2_Nil"; | |
| 1069 | AddIffs [list_all2_Nil]; | |
| 1070 | ||
| 1071 | Goalw [list_all2_def] "list_all2 P xs [] = (xs=[])"; | |
| 1072 | by (Simp_tac 1); | |
| 1073 | qed "list_all2_Nil2"; | |
| 1074 | AddIffs [list_all2_Nil2]; | |
| 1075 | ||
| 1076 | Goalw [list_all2_def] | |
| 1077 | "list_all2 P (x#xs) (y#ys) = (P x y & list_all2 P xs ys)"; | |
| 1078 | by (Auto_tac); | |
| 1079 | qed "list_all2_Cons"; | |
| 1080 | AddIffs[list_all2_Cons]; | |
| 1081 | ||
| 1082 | Goalw [list_all2_def] | |
| 8118 | 1083 | "list_all2 P (x#xs) ys = (? z zs. ys = z#zs & P x z & list_all2 P xs zs)"; | 
| 8442 
96023903c2df
case_tac now subsumes both boolean and datatype cases;
 wenzelm parents: 
8423diff
changeset | 1084 | by (case_tac "ys" 1); | 
| 8254 | 1085 | by (Auto_tac); | 
| 8118 | 1086 | qed "list_all2_Cons1"; | 
| 1087 | ||
| 1088 | Goalw [list_all2_def] | |
| 1089 | "list_all2 P xs (y#ys) = (? z zs. xs = z#zs & P z y & list_all2 P zs ys)"; | |
| 8442 
96023903c2df
case_tac now subsumes both boolean and datatype cases;
 wenzelm parents: 
8423diff
changeset | 1090 | by (case_tac "xs" 1); | 
| 8254 | 1091 | by (Auto_tac); | 
| 8118 | 1092 | qed "list_all2_Cons2"; | 
| 1093 | ||
| 1094 | Goalw [list_all2_def] | |
| 1095 | "list_all2 P (xs@ys) zs = \ | |
| 1096 | \ (EX us vs. zs = us@vs & length us = length xs & length vs = length ys & \ | |
| 1097 | \ list_all2 P xs us & list_all2 P ys vs)"; | |
| 8254 | 1098 | by (simp_tac (simpset() addsimps [zip_append1]) 1); | 
| 1099 | by (rtac iffI 1); | |
| 1100 |  by (res_inst_tac [("x","take (length xs) zs")] exI 1);
 | |
| 1101 |  by (res_inst_tac [("x","drop (length xs) zs")] exI 1);
 | |
| 10709 | 1102 | by (force_tac (claset(), | 
| 1103 | simpset() addsplits [nat_diff_split] addsimps [min_def]) 1); | |
| 8118 | 1104 | by (Clarify_tac 1); | 
| 8254 | 1105 | by (asm_full_simp_tac (simpset() addsimps [ball_Un]) 1); | 
| 8118 | 1106 | qed "list_all2_append1"; | 
| 1107 | ||
| 1108 | Goalw [list_all2_def] | |
| 1109 | "list_all2 P xs (ys@zs) = \ | |
| 1110 | \ (EX us vs. xs = us@vs & length us = length ys & length vs = length zs & \ | |
| 1111 | \ list_all2 P us ys & list_all2 P vs zs)"; | |
| 8254 | 1112 | by (simp_tac (simpset() addsimps [zip_append2]) 1); | 
| 1113 | by (rtac iffI 1); | |
| 1114 |  by (res_inst_tac [("x","take (length ys) xs")] exI 1);
 | |
| 1115 |  by (res_inst_tac [("x","drop (length ys) xs")] exI 1);
 | |
| 10709 | 1116 | by (force_tac (claset(), | 
| 1117 | simpset() addsplits [nat_diff_split] addsimps [min_def]) 1); | |
| 8118 | 1118 | by (Clarify_tac 1); | 
| 8254 | 1119 | by (asm_full_simp_tac (simpset() addsimps [ball_Un]) 1); | 
| 8118 | 1120 | qed "list_all2_append2"; | 
| 1121 | ||
| 1122 | Goalw [list_all2_def] | |
| 8115 | 1123 | "list_all2 P xs ys = \ | 
| 1124 | \ (length xs = length ys & (!i<length xs. P (xs!i) (ys!i)))"; | |
| 8254 | 1125 | by (force_tac (claset(), simpset() addsimps [set_zip]) 1); | 
| 8115 | 1126 | qed "list_all2_conv_all_nth"; | 
| 5272 | 1127 | |
| 11336 | 1128 | Goal "ALL a b c. P1 a b --> P2 b c --> P3 a c ==> \ | 
| 1129 | \ ALL bs cs. list_all2 P1 as bs --> list_all2 P2 bs cs --> list_all2 P3 as cs"; | |
| 1130 | by (induct_tac "as" 1); | |
| 1131 | by (Simp_tac 1); | |
| 1132 | by (rtac allI 1); | |
| 1133 | by (induct_tac "bs" 1); | |
| 1134 | by (Simp_tac 1); | |
| 1135 | by (rtac allI 1); | |
| 1136 | by (induct_tac "cs" 1); | |
| 1137 | by Auto_tac; | |
| 1138 | qed_spec_mp "list_all2_trans"; | |
| 1139 | ||
| 1140 | ||
| 5272 | 1141 | section "foldl"; | 
| 1142 | ||
| 1143 | Goal "!a. foldl f a (xs @ ys) = foldl f (foldl f a xs) ys"; | |
| 5318 | 1144 | by (induct_tac "xs" 1); | 
| 5316 | 1145 | by Auto_tac; | 
| 5272 | 1146 | qed_spec_mp "foldl_append"; | 
| 1147 | Addsimps [foldl_append]; | |
| 1148 | ||
| 1149 | (* Note: `n <= foldl op+ n ns' looks simpler, but is more difficult to use | |
| 1150 | because it requires an additional transitivity step | |
| 1151 | *) | |
| 1152 | Goal "!n::nat. m <= n --> m <= foldl op+ n ns"; | |
| 5318 | 1153 | by (induct_tac "ns" 1); | 
| 6058 | 1154 | by Auto_tac; | 
| 5272 | 1155 | qed_spec_mp "start_le_sum"; | 
| 1156 | ||
| 8935 
548901d05a0e
added type constraint ::nat because 0 is now overloaded
 paulson parents: 
8741diff
changeset | 1157 | Goal "!!n::nat. n : set ns ==> n <= foldl op+ 0 ns"; | 
| 5758 
27a2b36efd95
corrected auto_tac (applications of unsafe wrappers)
 oheimb parents: 
5644diff
changeset | 1158 | by (force_tac (claset() addIs [start_le_sum], | 
| 
27a2b36efd95
corrected auto_tac (applications of unsafe wrappers)
 oheimb parents: 
5644diff
changeset | 1159 | simpset() addsimps [in_set_conv_decomp]) 1); | 
| 5272 | 1160 | qed "elem_le_sum"; | 
| 1161 | ||
| 8935 
548901d05a0e
added type constraint ::nat because 0 is now overloaded
 paulson parents: 
8741diff
changeset | 1162 | Goal "!m::nat. (foldl op+ m ns = 0) = (m=0 & (!n : set ns. n=0))"; | 
| 5318 | 1163 | by (induct_tac "ns" 1); | 
| 5316 | 1164 | by Auto_tac; | 
| 5272 | 1165 | qed_spec_mp "sum_eq_0_conv"; | 
| 1166 | AddIffs [sum_eq_0_conv]; | |
| 1167 | ||
| 5425 | 1168 | (** upto **) | 
| 1169 | ||
| 5427 | 1170 | (* Does not terminate! *) | 
| 1171 | Goal "[i..j(] = (if i<j then i#[Suc i..j(] else [])"; | |
| 6162 | 1172 | by (induct_tac "j" 1); | 
| 5427 | 1173 | by Auto_tac; | 
| 1174 | qed "upt_rec"; | |
| 5425 | 1175 | |
| 5427 | 1176 | Goal "j<=i ==> [i..j(] = []"; | 
| 6162 | 1177 | by (stac upt_rec 1); | 
| 1178 | by (Asm_simp_tac 1); | |
| 5427 | 1179 | qed "upt_conv_Nil"; | 
| 1180 | Addsimps [upt_conv_Nil]; | |
| 1181 | ||
| 8982 
4cb682fc083d
renamed upt_Suc, since that name is needed for its primrec rule
 paulson parents: 
8935diff
changeset | 1182 | (*Only needed if upt_Suc is deleted from the simpset*) | 
| 5427 | 1183 | Goal "i<=j ==> [i..(Suc j)(] = [i..j(]@[j]"; | 
| 1184 | by (Asm_simp_tac 1); | |
| 8982 
4cb682fc083d
renamed upt_Suc, since that name is needed for its primrec rule
 paulson parents: 
8935diff
changeset | 1185 | qed "upt_Suc_append"; | 
| 5427 | 1186 | |
| 1187 | Goal "i<j ==> [i..j(] = i#[Suc i..j(]"; | |
| 6162 | 1188 | by (rtac trans 1); | 
| 1189 | by (stac upt_rec 1); | |
| 1190 | by (rtac refl 2); | |
| 5427 | 1191 | by (Asm_simp_tac 1); | 
| 1192 | qed "upt_conv_Cons"; | |
| 1193 | ||
| 9003 | 1194 | (*LOOPS as a simprule, since j<=j*) | 
| 1195 | Goal "i<=j ==> [i..j+k(] = [i..j(]@[j..j+k(]"; | |
| 1196 | by (induct_tac "k" 1); | |
| 1197 | by Auto_tac; | |
| 1198 | qed "upt_add_eq_append"; | |
| 1199 | ||
| 5427 | 1200 | Goal "length [i..j(] = j-i"; | 
| 6162 | 1201 | by (induct_tac "j" 1); | 
| 5427 | 1202 | by (Simp_tac 1); | 
| 6162 | 1203 | by (asm_simp_tac (simpset() addsimps [Suc_diff_le]) 1); | 
| 5427 | 1204 | qed "length_upt"; | 
| 1205 | Addsimps [length_upt]; | |
| 5425 | 1206 | |
| 5427 | 1207 | Goal "i+k < j --> [i..j(] ! k = i+k"; | 
| 6162 | 1208 | by (induct_tac "j" 1); | 
| 9014 | 1209 | by (asm_simp_tac (simpset() addsimps [less_Suc_eq, nth_append] | 
| 1210 | addsplits [nat_diff_split]) 2); | |
| 1211 | by (Simp_tac 1); | |
| 5427 | 1212 | qed_spec_mp "nth_upt"; | 
| 1213 | Addsimps [nth_upt]; | |
| 5425 | 1214 | |
| 6433 | 1215 | Goal "!i. i+m <= n --> take m [i..n(] = [i..i+m(]"; | 
| 6813 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 paulson parents: 
6794diff
changeset | 1216 | by (induct_tac "m" 1); | 
| 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 paulson parents: 
6794diff
changeset | 1217 | by (Simp_tac 1); | 
| 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 paulson parents: 
6794diff
changeset | 1218 | by (Clarify_tac 1); | 
| 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 paulson parents: 
6794diff
changeset | 1219 | by (stac upt_rec 1); | 
| 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 paulson parents: 
6794diff
changeset | 1220 | by (rtac sym 1); | 
| 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 paulson parents: 
6794diff
changeset | 1221 | by (stac upt_rec 1); | 
| 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 paulson parents: 
6794diff
changeset | 1222 | by (asm_simp_tac (simpset() delsimps (thms"upt.simps")) 1); | 
| 6433 | 1223 | qed_spec_mp "take_upt"; | 
| 1224 | Addsimps [take_upt]; | |
| 1225 | ||
| 9003 | 1226 | Goal "map Suc [m..n(] = [Suc m..n]"; | 
| 6813 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 paulson parents: 
6794diff
changeset | 1227 | by (induct_tac "n" 1); | 
| 9003 | 1228 | by Auto_tac; | 
| 1229 | qed "map_Suc_upt"; | |
| 1230 | ||
| 1231 | Goal "ALL i. i < n-m --> (map f [m..n(]) ! i = f(m+i)"; | |
| 9747 | 1232 | by (induct_thm_tac diff_induct "n m" 1); | 
| 9003 | 1233 | by (stac (map_Suc_upt RS sym) 3); | 
| 1234 | by (auto_tac (claset(), simpset() addsimps [less_diff_conv, nth_upt])); | |
| 6433 | 1235 | qed_spec_mp "nth_map_upt"; | 
| 1236 | ||
| 6813 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 paulson parents: 
6794diff
changeset | 1237 | Goal "ALL xs ys. k <= length xs --> k <= length ys --> \ | 
| 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 paulson parents: 
6794diff
changeset | 1238 | \ (ALL i. i < k --> xs!i = ys!i) \ | 
| 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 paulson parents: 
6794diff
changeset | 1239 | \ --> take k xs = take k ys"; | 
| 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 paulson parents: 
6794diff
changeset | 1240 | by (induct_tac "k" 1); | 
| 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 paulson parents: 
6794diff
changeset | 1241 | by (ALLGOALS (asm_simp_tac (simpset() addsimps [less_Suc_eq_0_disj, | 
| 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 paulson parents: 
6794diff
changeset | 1242 | all_conj_distrib]))); | 
| 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 paulson parents: 
6794diff
changeset | 1243 | by (Clarify_tac 1); | 
| 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 paulson parents: 
6794diff
changeset | 1244 | (*Both lists must be non-empty*) | 
| 8442 
96023903c2df
case_tac now subsumes both boolean and datatype cases;
 wenzelm parents: 
8423diff
changeset | 1245 | by (case_tac "xs" 1); | 
| 
96023903c2df
case_tac now subsumes both boolean and datatype cases;
 wenzelm parents: 
8423diff
changeset | 1246 | by (case_tac "ys" 2); | 
| 6813 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 paulson parents: 
6794diff
changeset | 1247 | by (ALLGOALS Clarify_tac); | 
| 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 paulson parents: 
6794diff
changeset | 1248 | (*prenexing's needed, not miniscoping*) | 
| 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 paulson parents: 
6794diff
changeset | 1249 | by (ALLGOALS (full_simp_tac (simpset() addsimps (all_simps RL [sym]) | 
| 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 paulson parents: 
6794diff
changeset | 1250 | delsimps (all_simps)))); | 
| 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 paulson parents: 
6794diff
changeset | 1251 | by (Blast_tac 1); | 
| 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 paulson parents: 
6794diff
changeset | 1252 | qed_spec_mp "nth_take_lemma"; | 
| 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 paulson parents: 
6794diff
changeset | 1253 | |
| 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 paulson parents: 
6794diff
changeset | 1254 | Goal "[| length xs = length ys; \ | 
| 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 paulson parents: 
6794diff
changeset | 1255 | \ ALL i. i < length xs --> xs!i = ys!i |] \ | 
| 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 paulson parents: 
6794diff
changeset | 1256 | \ ==> xs = ys"; | 
| 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 paulson parents: 
6794diff
changeset | 1257 | by (forward_tac [[le_refl, eq_imp_le] MRS nth_take_lemma] 1); | 
| 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 paulson parents: 
6794diff
changeset | 1258 | by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [take_all]))); | 
| 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 paulson parents: 
6794diff
changeset | 1259 | qed_spec_mp "nth_equalityI"; | 
| 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 paulson parents: 
6794diff
changeset | 1260 | |
| 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 paulson parents: 
6794diff
changeset | 1261 | (*The famous take-lemma*) | 
| 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 paulson parents: 
6794diff
changeset | 1262 | Goal "(ALL i. take i xs = take i ys) ==> xs = ys"; | 
| 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 paulson parents: 
6794diff
changeset | 1263 | by (dres_inst_tac [("x", "max (length xs) (length ys)")] spec 1);
 | 
| 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 paulson parents: 
6794diff
changeset | 1264 | by (full_simp_tac (simpset() addsimps [le_max_iff_disj, take_all]) 1); | 
| 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 paulson parents: 
6794diff
changeset | 1265 | qed_spec_mp "take_equalityI"; | 
| 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 paulson parents: 
6794diff
changeset | 1266 | |
| 5272 | 1267 | |
| 4605 | 1268 | (** nodups & remdups **) | 
| 1269 | section "nodups & remdups"; | |
| 1270 | ||
| 4935 | 1271 | Goal "set(remdups xs) = set xs"; | 
| 4605 | 1272 | by (induct_tac "xs" 1); | 
| 1273 | by (Simp_tac 1); | |
| 4686 | 1274 | by (asm_full_simp_tac (simpset() addsimps [insert_absorb]) 1); | 
| 4605 | 1275 | qed "set_remdups"; | 
| 1276 | Addsimps [set_remdups]; | |
| 1277 | ||
| 4935 | 1278 | Goal "nodups(remdups xs)"; | 
| 4605 | 1279 | by (induct_tac "xs" 1); | 
| 5316 | 1280 | by Auto_tac; | 
| 4605 | 1281 | qed "nodups_remdups"; | 
| 1282 | ||
| 4935 | 1283 | Goal "nodups xs --> nodups (filter P xs)"; | 
| 4605 | 1284 | by (induct_tac "xs" 1); | 
| 5316 | 1285 | by Auto_tac; | 
| 4605 | 1286 | qed_spec_mp "nodups_filter"; | 
| 1287 | ||
| 3589 
244daa75f890
Added function `replicate' and lemmas map_cong and set_replicate.
 nipkow parents: 
3586diff
changeset | 1288 | (** replicate **) | 
| 
244daa75f890
Added function `replicate' and lemmas map_cong and set_replicate.
 nipkow parents: 
3586diff
changeset | 1289 | section "replicate"; | 
| 
244daa75f890
Added function `replicate' and lemmas map_cong and set_replicate.
 nipkow parents: 
3586diff
changeset | 1290 | |
| 6794 | 1291 | Goal "length(replicate n x) = n"; | 
| 6813 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 paulson parents: 
6794diff
changeset | 1292 | by (induct_tac "n" 1); | 
| 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 paulson parents: 
6794diff
changeset | 1293 | by Auto_tac; | 
| 6794 | 1294 | qed "length_replicate"; | 
| 1295 | Addsimps [length_replicate]; | |
| 1296 | ||
| 1297 | Goal "map f (replicate n x) = replicate n (f x)"; | |
| 1298 | by (induct_tac "n" 1); | |
| 6813 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 paulson parents: 
6794diff
changeset | 1299 | by Auto_tac; | 
| 6794 | 1300 | qed "map_replicate"; | 
| 1301 | Addsimps [map_replicate]; | |
| 1302 | ||
| 1303 | Goal "(replicate n x) @ (x#xs) = x # replicate n x @ xs"; | |
| 1304 | by (induct_tac "n" 1); | |
| 6813 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 paulson parents: 
6794diff
changeset | 1305 | by Auto_tac; | 
| 6794 | 1306 | qed "replicate_app_Cons_same"; | 
| 1307 | ||
| 1308 | Goal "rev(replicate n x) = replicate n x"; | |
| 1309 | by (induct_tac "n" 1); | |
| 6813 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 paulson parents: 
6794diff
changeset | 1310 | by (Simp_tac 1); | 
| 6794 | 1311 | by (asm_simp_tac (simpset() addsimps [replicate_app_Cons_same]) 1); | 
| 1312 | qed "rev_replicate"; | |
| 1313 | Addsimps [rev_replicate]; | |
| 1314 | ||
| 8009 | 1315 | Goal "replicate (n+m) x = replicate n x @ replicate m x"; | 
| 1316 | by (induct_tac "n" 1); | |
| 1317 | by Auto_tac; | |
| 1318 | qed "replicate_add"; | |
| 1319 | ||
| 6794 | 1320 | Goal"n ~= 0 --> hd(replicate n x) = x"; | 
| 1321 | by (induct_tac "n" 1); | |
| 6813 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 paulson parents: 
6794diff
changeset | 1322 | by Auto_tac; | 
| 6794 | 1323 | qed_spec_mp "hd_replicate"; | 
| 1324 | Addsimps [hd_replicate]; | |
| 1325 | ||
| 11701 
3d51fbf81c17
sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
 wenzelm parents: 
11336diff
changeset | 1326 | Goal "n ~= 0 --> tl(replicate n x) = replicate (n - 1) x"; | 
| 6794 | 1327 | by (induct_tac "n" 1); | 
| 6813 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 paulson parents: 
6794diff
changeset | 1328 | by Auto_tac; | 
| 6794 | 1329 | qed_spec_mp "tl_replicate"; | 
| 1330 | Addsimps [tl_replicate]; | |
| 1331 | ||
| 1332 | Goal "n ~= 0 --> last(replicate n x) = x"; | |
| 1333 | by (induct_tac "n" 1); | |
| 6813 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 paulson parents: 
6794diff
changeset | 1334 | by Auto_tac; | 
| 6794 | 1335 | qed_spec_mp "last_replicate"; | 
| 1336 | Addsimps [last_replicate]; | |
| 1337 | ||
| 1338 | Goal "!i. i<n --> (replicate n x)!i = x"; | |
| 6813 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 paulson parents: 
6794diff
changeset | 1339 | by (induct_tac "n" 1); | 
| 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 paulson parents: 
6794diff
changeset | 1340 | by (Simp_tac 1); | 
| 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 paulson parents: 
6794diff
changeset | 1341 | by (asm_simp_tac (simpset() addsimps [nth_Cons] addsplits [nat.split]) 1); | 
| 6794 | 1342 | qed_spec_mp "nth_replicate"; | 
| 1343 | Addsimps [nth_replicate]; | |
| 1344 | ||
| 4935 | 1345 | Goal "set(replicate (Suc n) x) = {x}";
 | 
| 4423 | 1346 | by (induct_tac "n" 1); | 
| 5316 | 1347 | by Auto_tac; | 
| 3589 
244daa75f890
Added function `replicate' and lemmas map_cong and set_replicate.
 nipkow parents: 
3586diff
changeset | 1348 | val lemma = result(); | 
| 
244daa75f890
Added function `replicate' and lemmas map_cong and set_replicate.
 nipkow parents: 
3586diff
changeset | 1349 | |
| 5043 | 1350 | Goal "n ~= 0 ==> set(replicate n x) = {x}";
 | 
| 4423 | 1351 | by (fast_tac (claset() addSDs [not0_implies_Suc] addSIs [lemma]) 1); | 
| 3589 
244daa75f890
Added function `replicate' and lemmas map_cong and set_replicate.
 nipkow parents: 
3586diff
changeset | 1352 | qed "set_replicate"; | 
| 
244daa75f890
Added function `replicate' and lemmas map_cong and set_replicate.
 nipkow parents: 
3586diff
changeset | 1353 | Addsimps [set_replicate]; | 
| 5162 | 1354 | |
| 8009 | 1355 | Goal "set(replicate n x) = (if n=0 then {} else {x})";
 | 
| 8064 | 1356 | by (Auto_tac); | 
| 8009 | 1357 | qed "set_replicate_conv_if"; | 
| 1358 | ||
| 1359 | Goal "x : set(replicate n y) --> x=y"; | |
| 8064 | 1360 | by (asm_simp_tac (simpset() addsimps [set_replicate_conv_if]) 1); | 
| 8009 | 1361 | qed_spec_mp "in_set_replicateD"; | 
| 1362 | ||
| 5162 | 1363 | |
| 5281 | 1364 | (*** Lexcicographic orderings on lists ***) | 
| 1365 | section"Lexcicographic orderings on lists"; | |
| 1366 | ||
| 1367 | Goal "wf r ==> wf(lexn r n)"; | |
| 5318 | 1368 | by (induct_tac "n" 1); | 
| 1369 | by (Simp_tac 1); | |
| 1370 | by (Simp_tac 1); | |
| 1371 | by (rtac wf_subset 1); | |
| 1372 | by (rtac Int_lower1 2); | |
| 1373 | by (rtac wf_prod_fun_image 1); | |
| 1374 | by (rtac injI 2); | |
| 6813 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 paulson parents: 
6794diff
changeset | 1375 | by Auto_tac; | 
| 5281 | 1376 | qed "wf_lexn"; | 
| 1377 | ||
| 1378 | Goal "!xs ys. (xs,ys) : lexn r n --> length xs = n & length ys = n"; | |
| 5318 | 1379 | by (induct_tac "n" 1); | 
| 6813 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 paulson parents: 
6794diff
changeset | 1380 | by Auto_tac; | 
| 5281 | 1381 | qed_spec_mp "lexn_length"; | 
| 1382 | ||
| 1383 | Goalw [lex_def] "wf r ==> wf(lex r)"; | |
| 5318 | 1384 | by (rtac wf_UN 1); | 
| 1385 | by (blast_tac (claset() addIs [wf_lexn]) 1); | |
| 1386 | by (Clarify_tac 1); | |
| 1387 | by (rename_tac "m n" 1); | |
| 1388 | by (subgoal_tac "m ~= n" 1); | |
| 1389 | by (Blast_tac 2); | |
| 1390 | by (blast_tac (claset() addDs [lexn_length,not_sym]) 1); | |
| 5281 | 1391 | qed "wf_lex"; | 
| 1392 | AddSIs [wf_lex]; | |
| 1393 | ||
| 11265 | 1394 | |
| 5281 | 1395 | Goal | 
| 1396 | "lexn r n = \ | |
| 1397 | \ {(xs,ys). length xs = n & length ys = n & \
 | |
| 1398 | \ (? xys x y xs' ys'. xs= xys @ x#xs' & ys= xys @ y#ys' & (x,y):r)}"; | |
| 5318 | 1399 | by (induct_tac "n" 1); | 
| 1400 | by (Simp_tac 1); | |
| 1401 | by (Blast_tac 1); | |
| 11265 | 1402 | by (asm_full_simp_tac (simpset() addsimps [image_Collect, lex_prod_def]) 1); | 
| 1403 | by Auto_tac; | |
| 5318 | 1404 | by (Blast_tac 1); | 
| 1405 | by (rename_tac "a xys x xs' y ys'" 1); | |
| 1406 |  by (res_inst_tac [("x","a#xys")] exI 1);
 | |
| 1407 | by (Simp_tac 1); | |
| 8442 
96023903c2df
case_tac now subsumes both boolean and datatype cases;
 wenzelm parents: 
8423diff
changeset | 1408 | by (case_tac "xys" 1); | 
| 11265 | 1409 | by (ALLGOALS Asm_full_simp_tac); | 
| 5318 | 1410 | by (Blast_tac 1); | 
| 5281 | 1411 | qed "lexn_conv"; | 
| 1412 | ||
| 1413 | Goalw [lex_def] | |
| 1414 | "lex r = \ | |
| 1415 | \ {(xs,ys). length xs = length ys & \
 | |
| 1416 | \ (? xys x y xs' ys'. xs= xys @ x#xs' & ys= xys @ y#ys' & (x,y):r)}"; | |
| 5641 | 1417 | by (force_tac (claset(), simpset() addsimps [lexn_conv]) 1); | 
| 5281 | 1418 | qed "lex_conv"; | 
| 1419 | ||
| 1420 | Goalw [lexico_def] "wf r ==> wf(lexico r)"; | |
| 5318 | 1421 | by (Blast_tac 1); | 
| 5281 | 1422 | qed "wf_lexico"; | 
| 1423 | AddSIs [wf_lexico]; | |
| 1424 | ||
| 10709 | 1425 | Goalw [lexico_def,diag_def,lex_prod_def,measure_def,inv_image_def] | 
| 5281 | 1426 | "lexico r = {(xs,ys). length xs < length ys | \
 | 
| 1427 | \ length xs = length ys & (xs,ys) : lex r}"; | |
| 5318 | 1428 | by (Simp_tac 1); | 
| 5281 | 1429 | qed "lexico_conv"; | 
| 1430 | ||
| 5283 | 1431 | Goal "([],ys) ~: lex r"; | 
| 5318 | 1432 | by (simp_tac (simpset() addsimps [lex_conv]) 1); | 
| 5283 | 1433 | qed "Nil_notin_lex"; | 
| 1434 | ||
| 1435 | Goal "(xs,[]) ~: lex r"; | |
| 5318 | 1436 | by (simp_tac (simpset() addsimps [lex_conv]) 1); | 
| 5283 | 1437 | qed "Nil2_notin_lex"; | 
| 1438 | ||
| 1439 | AddIffs [Nil_notin_lex,Nil2_notin_lex]; | |
| 1440 | ||
| 1441 | Goal "((x#xs,y#ys) : lex r) = \ | |
| 1442 | \ ((x,y) : r & length xs = length ys | x=y & (xs,ys) : lex r)"; | |
| 5318 | 1443 | by (simp_tac (simpset() addsimps [lex_conv]) 1); | 
| 1444 | by (rtac iffI 1); | |
| 1445 | by (blast_tac (claset() addIs [Cons_eq_appendI]) 2); | |
| 10709 | 1446 | by (Clarify_tac 1); | 
| 8442 
96023903c2df
case_tac now subsumes both boolean and datatype cases;
 wenzelm parents: 
8423diff
changeset | 1447 | by (case_tac "xys" 1); | 
| 5318 | 1448 | by (Asm_full_simp_tac 1); | 
| 1449 | by (Asm_full_simp_tac 1); | |
| 1450 | by (Blast_tac 1); | |
| 5283 | 1451 | qed "Cons_in_lex"; | 
| 1452 | AddIffs [Cons_in_lex]; | |
| 7032 | 1453 | |
| 1454 | ||
| 9336 | 1455 | (*** sublist (a generalization of nth to sets) ***) | 
| 1456 | ||
| 1457 | Goalw [sublist_def] "sublist l {} = []";
 | |
| 1458 | by Auto_tac; | |
| 1459 | qed "sublist_empty"; | |
| 1460 | ||
| 1461 | Goalw [sublist_def] "sublist [] A = []"; | |
| 1462 | by Auto_tac; | |
| 1463 | qed "sublist_nil"; | |
| 1464 | ||
| 1465 | Goal "map fst [p:zip xs [i..i + length xs(] . snd p : A] = \ | |
| 1466 | \ map fst [p:zip xs [0..length xs(] . snd p + i : A]"; | |
| 9747 | 1467 | by (rev_induct_tac "xs" 1); | 
| 9336 | 1468 | by (asm_simp_tac (simpset() addsimps [add_commute]) 2); | 
| 1469 | by (Simp_tac 1); | |
| 1470 | qed "sublist_shift_lemma"; | |
| 1471 | ||
| 1472 | Goalw [sublist_def] | |
| 1473 |      "sublist (l@l') A = sublist l A @ sublist l' {j. j + length l : A}";
 | |
| 9747 | 1474 | by (rev_induct_tac "l'" 1); | 
| 9336 | 1475 | by (Simp_tac 1); | 
| 1476 | by (asm_simp_tac (simpset() addsimps [inst "i" "0" upt_add_eq_append, | |
| 1477 | zip_append, sublist_shift_lemma]) 1); | |
| 1478 | by (asm_simp_tac (simpset() addsimps [add_commute]) 1); | |
| 1479 | qed "sublist_append"; | |
| 1480 | ||
| 1481 | Addsimps [sublist_empty, sublist_nil]; | |
| 1482 | ||
| 1483 | Goal "sublist (x#l) A = (if 0:A then [x] else []) @ sublist l {j. Suc j : A}";
 | |
| 9747 | 1484 | by (rev_induct_tac "l" 1); | 
| 9336 | 1485 | by (asm_simp_tac (simpset() delsimps [append_Cons] | 
| 1486 | addsimps [append_Cons RS sym, sublist_append]) 2); | |
| 1487 | by (simp_tac (simpset() addsimps [sublist_def]) 1); | |
| 1488 | qed "sublist_Cons"; | |
| 1489 | ||
| 1490 | Goal "sublist [x] A = (if 0 : A then [x] else [])"; | |
| 1491 | by (simp_tac (simpset() addsimps [sublist_Cons]) 1); | |
| 1492 | qed "sublist_singleton"; | |
| 1493 | Addsimps [sublist_singleton]; | |
| 1494 | ||
| 1495 | Goal "sublist l {..n(} = take n l";
 | |
| 9747 | 1496 | by (rev_induct_tac "l" 1); | 
| 9336 | 1497 | by (asm_simp_tac (simpset() addsplits [nat_diff_split] | 
| 1498 | addsimps [sublist_append]) 2); | |
| 1499 | by (Simp_tac 1); | |
| 1500 | qed "sublist_upt_eq_take"; | |
| 1501 | Addsimps [sublist_upt_eq_take]; | |
| 1502 | ||
| 1503 | ||
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11770diff
changeset | 1504 | Goal "take n (x#xs) = (if n=0 then [] else x # take (n - 1) xs)"; | 
| 8442 
96023903c2df
case_tac now subsumes both boolean and datatype cases;
 wenzelm parents: 
8423diff
changeset | 1505 | by (case_tac "n" 1); | 
| 7032 | 1506 | by (ALLGOALS | 
| 1507 | (asm_simp_tac (simpset() addsimps [numeral_0_eq_0, numeral_1_eq_1]))); | |
| 1508 | qed "take_Cons'"; | |
| 1509 | ||
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11770diff
changeset | 1510 | Goal "drop n (x#xs) = (if n=0 then x#xs else drop (n - 1) xs)"; | 
| 8442 
96023903c2df
case_tac now subsumes both boolean and datatype cases;
 wenzelm parents: 
8423diff
changeset | 1511 | by (case_tac "n" 1); | 
| 7032 | 1512 | by (ALLGOALS | 
| 1513 | (asm_simp_tac (simpset() addsimps [numeral_0_eq_0, numeral_1_eq_1]))); | |
| 1514 | qed "drop_Cons'"; | |
| 1515 | ||
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11770diff
changeset | 1516 | Goal "(x#xs)!n = (if n=0 then x else xs!(n - 1))"; | 
| 8442 
96023903c2df
case_tac now subsumes both boolean and datatype cases;
 wenzelm parents: 
8423diff
changeset | 1517 | by (case_tac "n" 1); | 
| 7032 | 1518 | by (ALLGOALS | 
| 1519 | (asm_simp_tac (simpset() addsimps [numeral_0_eq_0, numeral_1_eq_1]))); | |
| 1520 | qed "nth_Cons'"; | |
| 1521 | ||
| 1522 | Addsimps (map (inst "n" "number_of ?v") [take_Cons', drop_Cons', nth_Cons']); | |
| 1523 |