| author | paulson | 
| Fri, 25 Sep 1998 14:05:13 +0200 | |
| changeset 5565 | 301a3a4d3dc7 | 
| parent 5537 | c2bd39a2c0ee | 
| child 5641 | 5266f09db46c | 
| 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 | |
| 3468 | 35 | val listsE = lists.mk_cases list.simps "x#l : lists A"; | 
| 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 | ||
| 3196 | 51 | |
| 4643 | 52 | (** Case analysis **) | 
| 53 | section "Case analysis"; | |
| 2608 | 54 | |
| 4935 | 55 | val prems = Goal "[| P([]); !!x xs. P(x#xs) |] ==> P(xs)"; | 
| 3457 | 56 | by (induct_tac "xs" 1); | 
| 57 | by (REPEAT(resolve_tac prems 1)); | |
| 2608 | 58 | qed "list_cases"; | 
| 59 | ||
| 4935 | 60 | Goal "(xs=[] --> P([])) & (!y ys. xs=y#ys --> P(y#ys)) --> P(xs)"; | 
| 3040 
7d48671753da
Introduced a generic "induct_tac" which picks up the right induction scheme
 nipkow parents: 
3011diff
changeset | 61 | by (induct_tac "xs" 1); | 
| 2891 | 62 | by (Blast_tac 1); | 
| 63 | by (Blast_tac 1); | |
| 2608 | 64 | bind_thm("list_eq_cases",
 | 
| 65 | impI RSN (2,allI RSN (2,allI RSN (2,impI RS (conjI RS (result() RS mp)))))); | |
| 66 | ||
| 3860 | 67 | (** length **) | 
| 68 | (* needs to come before "@" because of thm append_eq_append_conv *) | |
| 69 | ||
| 70 | section "length"; | |
| 71 | ||
| 4935 | 72 | Goal "length(xs@ys) = length(xs)+length(ys)"; | 
| 3860 | 73 | by (induct_tac "xs" 1); | 
| 5316 | 74 | by Auto_tac; | 
| 3860 | 75 | qed"length_append"; | 
| 76 | Addsimps [length_append]; | |
| 77 | ||
| 5129 | 78 | Goal "length (map f xs) = length xs"; | 
| 79 | by (induct_tac "xs" 1); | |
| 5316 | 80 | by Auto_tac; | 
| 3860 | 81 | qed "length_map"; | 
| 82 | Addsimps [length_map]; | |
| 83 | ||
| 4935 | 84 | Goal "length(rev xs) = length(xs)"; | 
| 3860 | 85 | by (induct_tac "xs" 1); | 
| 5316 | 86 | by Auto_tac; | 
| 3860 | 87 | qed "length_rev"; | 
| 88 | Addsimps [length_rev]; | |
| 89 | ||
| 5043 | 90 | Goal "xs ~= [] ==> length(tl xs) = (length xs) - 1"; | 
| 4423 | 91 | by (exhaust_tac "xs" 1); | 
| 5316 | 92 | by Auto_tac; | 
| 3896 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 nipkow parents: 
3860diff
changeset | 93 | qed "length_tl"; | 
| 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 nipkow parents: 
3860diff
changeset | 94 | Addsimps [length_tl]; | 
| 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 nipkow parents: 
3860diff
changeset | 95 | |
| 4935 | 96 | Goal "(length xs = 0) = (xs = [])"; | 
| 3860 | 97 | by (induct_tac "xs" 1); | 
| 5316 | 98 | by Auto_tac; | 
| 3860 | 99 | qed "length_0_conv"; | 
| 100 | AddIffs [length_0_conv]; | |
| 101 | ||
| 4935 | 102 | Goal "(0 = length xs) = (xs = [])"; | 
| 3860 | 103 | by (induct_tac "xs" 1); | 
| 5316 | 104 | by Auto_tac; | 
| 3860 | 105 | qed "zero_length_conv"; | 
| 106 | AddIffs [zero_length_conv]; | |
| 107 | ||
| 4935 | 108 | Goal "(0 < length xs) = (xs ~= [])"; | 
| 3860 | 109 | by (induct_tac "xs" 1); | 
| 5316 | 110 | by Auto_tac; | 
| 3860 | 111 | qed "length_greater_0_conv"; | 
| 112 | AddIffs [length_greater_0_conv]; | |
| 113 | ||
| 5296 | 114 | Goal "(length xs = Suc n) = (? y ys. xs = y#ys & length ys = n)"; | 
| 115 | by (induct_tac "xs" 1); | |
| 116 | by (Auto_tac); | |
| 117 | qed "length_Suc_conv"; | |
| 118 | AddIffs [length_Suc_conv]; | |
| 119 | ||
| 923 | 120 | (** @ - append **) | 
| 121 | ||
| 3467 | 122 | section "@ - append"; | 
| 123 | ||
| 4935 | 124 | 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 | 125 | by (induct_tac "xs" 1); | 
| 5316 | 126 | by Auto_tac; | 
| 923 | 127 | qed "append_assoc"; | 
| 2512 | 128 | Addsimps [append_assoc]; | 
| 923 | 129 | |
| 4935 | 130 | Goal "xs @ [] = xs"; | 
| 3040 
7d48671753da
Introduced a generic "induct_tac" which picks up the right induction scheme
 nipkow parents: 
3011diff
changeset | 131 | by (induct_tac "xs" 1); | 
| 5316 | 132 | by Auto_tac; | 
| 923 | 133 | qed "append_Nil2"; | 
| 2512 | 134 | Addsimps [append_Nil2]; | 
| 923 | 135 | |
| 4935 | 136 | Goal "(xs@ys = []) = (xs=[] & ys=[])"; | 
| 3040 
7d48671753da
Introduced a generic "induct_tac" which picks up the right induction scheme
 nipkow parents: 
3011diff
changeset | 137 | by (induct_tac "xs" 1); | 
| 5316 | 138 | by Auto_tac; | 
| 2608 | 139 | qed "append_is_Nil_conv"; | 
| 140 | AddIffs [append_is_Nil_conv]; | |
| 141 | ||
| 4935 | 142 | Goal "([] = xs@ys) = (xs=[] & ys=[])"; | 
| 3040 
7d48671753da
Introduced a generic "induct_tac" which picks up the right induction scheme
 nipkow parents: 
3011diff
changeset | 143 | by (induct_tac "xs" 1); | 
| 5316 | 144 | by Auto_tac; | 
| 2608 | 145 | qed "Nil_is_append_conv"; | 
| 146 | AddIffs [Nil_is_append_conv]; | |
| 923 | 147 | |
| 4935 | 148 | Goal "(xs @ ys = xs) = (ys=[])"; | 
| 3574 | 149 | by (induct_tac "xs" 1); | 
| 5316 | 150 | by Auto_tac; | 
| 3574 | 151 | qed "append_self_conv"; | 
| 152 | ||
| 4935 | 153 | Goal "(xs = xs @ ys) = (ys=[])"; | 
| 3574 | 154 | by (induct_tac "xs" 1); | 
| 5316 | 155 | by Auto_tac; | 
| 3574 | 156 | qed "self_append_conv"; | 
| 157 | AddIffs [append_self_conv,self_append_conv]; | |
| 158 | ||
| 4935 | 159 | Goal "!ys. length xs = length ys | length us = length vs \ | 
| 3860 | 160 | \ --> (xs@us = ys@vs) = (xs=ys & us=vs)"; | 
| 4423 | 161 | by (induct_tac "xs" 1); | 
| 162 | by (rtac allI 1); | |
| 163 | by (exhaust_tac "ys" 1); | |
| 164 | by (Asm_simp_tac 1); | |
| 165 | by (fast_tac (claset() addIs [less_add_Suc2] addss simpset() | |
| 3860 | 166 | addEs [less_not_refl2 RSN (2,rev_notE)]) 1); | 
| 4423 | 167 | by (rtac allI 1); | 
| 168 | by (exhaust_tac "ys" 1); | |
| 5296 | 169 | by (fast_tac (claset() addIs [less_add_Suc2] | 
| 170 | addss (simpset() delsimps [length_Suc_conv]) | |
| 5355 | 171 | addEs [(less_not_refl3) RSN (2,rev_notE)]) 1); | 
| 4423 | 172 | by (Asm_simp_tac 1); | 
| 3860 | 173 | qed_spec_mp "append_eq_append_conv"; | 
| 174 | Addsimps [append_eq_append_conv]; | |
| 175 | ||
| 4935 | 176 | Goal "(xs @ ys = xs @ zs) = (ys=zs)"; | 
| 3896 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 nipkow parents: 
3860diff
changeset | 177 | by (Simp_tac 1); | 
| 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 nipkow parents: 
3860diff
changeset | 178 | qed "same_append_eq"; | 
| 3860 | 179 | |
| 4935 | 180 | 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 | 181 | by (Simp_tac 1); | 
| 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 nipkow parents: 
3860diff
changeset | 182 | qed "append1_eq_conv"; | 
| 2608 | 183 | |
| 4935 | 184 | Goal "(ys @ xs = zs @ xs) = (ys=zs)"; | 
| 3896 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 nipkow parents: 
3860diff
changeset | 185 | by (Simp_tac 1); | 
| 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 nipkow parents: 
3860diff
changeset | 186 | qed "append_same_eq"; | 
| 2608 | 187 | |
| 3896 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 nipkow parents: 
3860diff
changeset | 188 | AddSIs | 
| 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 nipkow parents: 
3860diff
changeset | 189 | [same_append_eq RS iffD2, append1_eq_conv RS iffD2, append_same_eq RS iffD2]; | 
| 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 nipkow parents: 
3860diff
changeset | 190 | AddSDs | 
| 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 nipkow parents: 
3860diff
changeset | 191 | [same_append_eq RS iffD1, append1_eq_conv RS iffD1, append_same_eq RS iffD1]; | 
| 3571 | 192 | |
| 4935 | 193 | Goal "(xs @ ys = ys) = (xs=[])"; | 
| 5132 | 194 | by (cut_inst_tac [("zs","[]")] append_same_eq 1);
 | 
| 5316 | 195 | by Auto_tac; | 
| 4647 | 196 | qed "append_self_conv2"; | 
| 197 | ||
| 4935 | 198 | Goal "(ys = xs @ ys) = (xs=[])"; | 
| 5132 | 199 | by (simp_tac (simpset() addsimps | 
| 4647 | 200 |      [simplify (simpset()) (read_instantiate[("ys","[]")]append_same_eq)]) 1);
 | 
| 5132 | 201 | by (Blast_tac 1); | 
| 4647 | 202 | qed "self_append_conv2"; | 
| 203 | AddIffs [append_self_conv2,self_append_conv2]; | |
| 204 | ||
| 4935 | 205 | Goal "xs ~= [] --> hd xs # tl xs = xs"; | 
| 3457 | 206 | by (induct_tac "xs" 1); | 
| 5316 | 207 | by Auto_tac; | 
| 2608 | 208 | qed_spec_mp "hd_Cons_tl"; | 
| 209 | Addsimps [hd_Cons_tl]; | |
| 923 | 210 | |
| 4935 | 211 | 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 | 212 | by (induct_tac "xs" 1); | 
| 5316 | 213 | by Auto_tac; | 
| 1327 
6c29cfab679c
added new arithmetic lemmas and the functions take and drop.
 nipkow parents: 
1301diff
changeset | 214 | qed "hd_append"; | 
| 923 | 215 | |
| 5043 | 216 | Goal "xs ~= [] ==> hd(xs @ ys) = hd xs"; | 
| 4089 | 217 | by (asm_simp_tac (simpset() addsimps [hd_append] | 
| 5183 | 218 | addsplits [list.split]) 1); | 
| 3571 | 219 | qed "hd_append2"; | 
| 220 | Addsimps [hd_append2]; | |
| 221 | ||
| 4935 | 222 | Goal "tl(xs@ys) = (case xs of [] => tl(ys) | z#zs => zs@ys)"; | 
| 5183 | 223 | by (simp_tac (simpset() addsplits [list.split]) 1); | 
| 2608 | 224 | qed "tl_append"; | 
| 225 | ||
| 5043 | 226 | Goal "xs ~= [] ==> tl(xs @ ys) = (tl xs) @ ys"; | 
| 4089 | 227 | by (asm_simp_tac (simpset() addsimps [tl_append] | 
| 5183 | 228 | addsplits [list.split]) 1); | 
| 3571 | 229 | qed "tl_append2"; | 
| 230 | Addsimps [tl_append2]; | |
| 231 | ||
| 5272 | 232 | (* trivial rules for solving @-equations automatically *) | 
| 233 | ||
| 234 | Goal "xs = ys ==> xs = [] @ ys"; | |
| 5318 | 235 | by (Asm_simp_tac 1); | 
| 5272 | 236 | qed "eq_Nil_appendI"; | 
| 237 | ||
| 238 | Goal "[| x#xs1 = ys; xs = xs1 @ zs |] ==> x#xs = ys@zs"; | |
| 5318 | 239 | by (dtac sym 1); | 
| 240 | by (Asm_simp_tac 1); | |
| 5272 | 241 | qed "Cons_eq_appendI"; | 
| 242 | ||
| 243 | Goal "[| xs@xs1 = zs; ys = xs1 @ us |] ==> xs@ys = zs@us"; | |
| 5318 | 244 | by (dtac sym 1); | 
| 245 | by (Asm_simp_tac 1); | |
| 5272 | 246 | qed "append_eq_appendI"; | 
| 247 | ||
| 4830 | 248 | |
| 5427 | 249 | (*** | 
| 250 | Simplification procedure for all list equalities. | |
| 251 | Currently only tries to rearranges @ to see if | |
| 252 | - both lists end in a singleton list, | |
| 253 | - or both lists end in the same list. | |
| 254 | ***) | |
| 255 | local | |
| 256 | ||
| 257 | val list_eq_pattern = | |
| 258 |   read_cterm (sign_of List.thy) ("(xs::'a list) = ys",HOLogic.boolT);
 | |
| 259 | ||
| 260 | fun last (cons as Const("List.list.op #",_) $ _ $ xs) =
 | |
| 261 |       (case xs of Const("List.list.[]",_) => cons | _ => last xs)
 | |
| 262 |   | last (Const("List.op @",_) $ _ $ ys) = last ys
 | |
| 263 | | last t = t; | |
| 264 | ||
| 265 | fun list1 (Const("List.list.op #",_) $ _ $ Const("List.list.[]",_)) = true
 | |
| 266 | | list1 _ = false; | |
| 267 | ||
| 268 | fun butlast ((cons as Const("List.list.op #",_) $ x) $ xs) =
 | |
| 269 |       (case xs of Const("List.list.[]",_) => xs | _ => cons $ butlast xs)
 | |
| 270 |   | butlast ((app as Const("List.op @",_) $ xs) $ ys) = app $ butlast ys
 | |
| 271 |   | butlast xs = Const("List.list.[]",fastype_of xs);
 | |
| 272 | ||
| 273 | val rearr_tac = | |
| 274 | simp_tac (HOL_basic_ss addsimps [append_assoc,append_Nil,append_Cons]); | |
| 275 | ||
| 276 | fun list_eq sg _ (F as (eq as Const(_,eqT)) $ lhs $ rhs) = | |
| 277 | let | |
| 278 | val lastl = last lhs and lastr = last rhs | |
| 279 | fun rearr conv = | |
| 280 | let val lhs1 = butlast lhs and rhs1 = butlast rhs | |
| 281 | val Type(_,listT::_) = eqT | |
| 282 | val appT = [listT,listT] ---> listT | |
| 283 |           val app = Const("List.op @",appT)
 | |
| 284 | val F2 = eq $ (app$lhs1$lastl) $ (app$rhs1$lastr) | |
| 285 | val ct = cterm_of sg (HOLogic.mk_Trueprop(HOLogic.mk_eq(F,F2))) | |
| 286 | val thm = prove_goalw_cterm [] ct (K [rearr_tac 1]) | |
| 287 | handle ERROR => | |
| 288 |             error("The error(s) above occurred while trying to prove " ^
 | |
| 289 | string_of_cterm ct) | |
| 290 | in Some((conv RS (thm RS trans)) RS eq_reflection) end | |
| 291 | ||
| 292 | in if list1 lastl andalso list1 lastr | |
| 293 | then rearr append1_eq_conv | |
| 294 | else | |
| 295 | if lastl aconv lastr | |
| 296 | then rearr append_same_eq | |
| 297 | else None | |
| 298 | end; | |
| 299 | in | |
| 300 | val list_eq_simproc = mk_simproc "list_eq" [list_eq_pattern] list_eq; | |
| 301 | end; | |
| 302 | ||
| 303 | Addsimprocs [list_eq_simproc]; | |
| 304 | ||
| 305 | ||
| 2608 | 306 | (** map **) | 
| 307 | ||
| 3467 | 308 | section "map"; | 
| 309 | ||
| 5278 | 310 | Goal "(!x. x : set xs --> f x = g x) --> map f xs = map g xs"; | 
| 3457 | 311 | by (induct_tac "xs" 1); | 
| 5316 | 312 | by Auto_tac; | 
| 2608 | 313 | bind_thm("map_ext", impI RS (allI RS (result() RS mp)));
 | 
| 314 | ||
| 4935 | 315 | Goal "map (%x. x) = (%xs. xs)"; | 
| 2608 | 316 | by (rtac ext 1); | 
| 3040 
7d48671753da
Introduced a generic "induct_tac" which picks up the right induction scheme
 nipkow parents: 
3011diff
changeset | 317 | by (induct_tac "xs" 1); | 
| 5316 | 318 | by Auto_tac; | 
| 2608 | 319 | qed "map_ident"; | 
| 320 | Addsimps[map_ident]; | |
| 321 | ||
| 4935 | 322 | 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 | 323 | by (induct_tac "xs" 1); | 
| 5316 | 324 | by Auto_tac; | 
| 2608 | 325 | qed "map_append"; | 
| 326 | Addsimps[map_append]; | |
| 327 | ||
| 4935 | 328 | 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 | 329 | by (induct_tac "xs" 1); | 
| 5316 | 330 | by Auto_tac; | 
| 2608 | 331 | qed "map_compose"; | 
| 332 | Addsimps[map_compose]; | |
| 333 | ||
| 4935 | 334 | 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 | 335 | by (induct_tac "xs" 1); | 
| 5316 | 336 | by Auto_tac; | 
| 2608 | 337 | qed "rev_map"; | 
| 338 | ||
| 3589 
244daa75f890
Added function `replicate' and lemmas map_cong and set_replicate.
 nipkow parents: 
3586diff
changeset | 339 | (* a congruence rule for map: *) | 
| 5278 | 340 | Goal "(xs=ys) --> (!x. x : set ys --> f x = g x) --> map f xs = map g ys"; | 
| 4423 | 341 | by (rtac impI 1); | 
| 342 | by (hyp_subst_tac 1); | |
| 343 | by (induct_tac "ys" 1); | |
| 5316 | 344 | by Auto_tac; | 
| 3589 
244daa75f890
Added function `replicate' and lemmas map_cong and set_replicate.
 nipkow parents: 
3586diff
changeset | 345 | val lemma = result(); | 
| 
244daa75f890
Added function `replicate' and lemmas map_cong and set_replicate.
 nipkow parents: 
3586diff
changeset | 346 | bind_thm("map_cong",impI RSN (2,allI RSN (2,lemma RS mp RS mp)));
 | 
| 
244daa75f890
Added function `replicate' and lemmas map_cong and set_replicate.
 nipkow parents: 
3586diff
changeset | 347 | |
| 4935 | 348 | Goal "(map f xs = []) = (xs = [])"; | 
| 4423 | 349 | by (induct_tac "xs" 1); | 
| 5316 | 350 | by Auto_tac; | 
| 3860 | 351 | qed "map_is_Nil_conv"; | 
| 352 | AddIffs [map_is_Nil_conv]; | |
| 353 | ||
| 4935 | 354 | Goal "([] = map f xs) = (xs = [])"; | 
| 4423 | 355 | by (induct_tac "xs" 1); | 
| 5316 | 356 | by Auto_tac; | 
| 3860 | 357 | qed "Nil_is_map_conv"; | 
| 358 | AddIffs [Nil_is_map_conv]; | |
| 359 | ||
| 360 | ||
| 1169 | 361 | (** rev **) | 
| 362 | ||
| 3467 | 363 | section "rev"; | 
| 364 | ||
| 4935 | 365 | 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 | 366 | by (induct_tac "xs" 1); | 
| 5316 | 367 | by Auto_tac; | 
| 1169 | 368 | qed "rev_append"; | 
| 2512 | 369 | Addsimps[rev_append]; | 
| 1169 | 370 | |
| 4935 | 371 | Goal "rev(rev l) = l"; | 
| 3040 
7d48671753da
Introduced a generic "induct_tac" which picks up the right induction scheme
 nipkow parents: 
3011diff
changeset | 372 | by (induct_tac "l" 1); | 
| 5316 | 373 | by Auto_tac; | 
| 1169 | 374 | qed "rev_rev_ident"; | 
| 2512 | 375 | Addsimps[rev_rev_ident]; | 
| 1169 | 376 | |
| 4935 | 377 | Goal "(rev xs = []) = (xs = [])"; | 
| 4423 | 378 | by (induct_tac "xs" 1); | 
| 5316 | 379 | by Auto_tac; | 
| 3860 | 380 | qed "rev_is_Nil_conv"; | 
| 381 | AddIffs [rev_is_Nil_conv]; | |
| 382 | ||
| 4935 | 383 | Goal "([] = rev xs) = (xs = [])"; | 
| 4423 | 384 | by (induct_tac "xs" 1); | 
| 5316 | 385 | by Auto_tac; | 
| 3860 | 386 | qed "Nil_is_rev_conv"; | 
| 387 | AddIffs [Nil_is_rev_conv]; | |
| 388 | ||
| 4935 | 389 | val prems = Goal "[| P []; !!x xs. P xs ==> P(xs@[x]) |] ==> P xs"; | 
| 5132 | 390 | by (stac (rev_rev_ident RS sym) 1); | 
| 4935 | 391 | br(read_instantiate [("P","%xs. ?P(rev xs)")]list.induct)1;
 | 
| 5132 | 392 | by (ALLGOALS Simp_tac); | 
| 393 | by (resolve_tac prems 1); | |
| 394 | by (eresolve_tac prems 1); | |
| 4935 | 395 | qed "rev_induct"; | 
| 396 | ||
| 5272 | 397 | fun rev_induct_tac xs = res_inst_tac [("xs",xs)] rev_induct;
 | 
| 398 | ||
| 4935 | 399 | Goal "(xs = [] --> P) --> (!ys y. xs = ys@[y] --> P) --> P"; | 
| 5132 | 400 | by (res_inst_tac [("xs","xs")] rev_induct 1);
 | 
| 5316 | 401 | by Auto_tac; | 
| 4935 | 402 | bind_thm ("rev_exhaust",
 | 
| 403 | impI RSN (2,allI RSN (2,allI RSN (2,impI RS (result() RS mp RS mp))))); | |
| 404 | ||
| 2608 | 405 | |
| 3465 | 406 | (** set **) | 
| 1812 | 407 | |
| 3467 | 408 | section "set"; | 
| 409 | ||
| 5296 | 410 | qed_goal "finite_set" thy "finite (set xs)" | 
| 411 | (K [induct_tac "xs" 1, Auto_tac]); | |
| 412 | Addsimps[finite_set]; | |
| 413 | AddSIs[finite_set]; | |
| 414 | ||
| 4935 | 415 | 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 | 416 | by (induct_tac "xs" 1); | 
| 5316 | 417 | by Auto_tac; | 
| 3647 
a64c8fbcd98f
Renamed theorems of the form set_of_list_XXX to set_XXX
 paulson parents: 
3589diff
changeset | 418 | qed "set_append"; | 
| 
a64c8fbcd98f
Renamed theorems of the form set_of_list_XXX to set_XXX
 paulson parents: 
3589diff
changeset | 419 | Addsimps[set_append]; | 
| 1812 | 420 | |
| 4935 | 421 | Goal "set l <= set (x#l)"; | 
| 5316 | 422 | by Auto_tac; | 
| 3647 
a64c8fbcd98f
Renamed theorems of the form set_of_list_XXX to set_XXX
 paulson parents: 
3589diff
changeset | 423 | qed "set_subset_Cons"; | 
| 1936 | 424 | |
| 4935 | 425 | Goal "(set xs = {}) = (xs = [])";
 | 
| 3457 | 426 | by (induct_tac "xs" 1); | 
| 5316 | 427 | by Auto_tac; | 
| 3647 
a64c8fbcd98f
Renamed theorems of the form set_of_list_XXX to set_XXX
 paulson parents: 
3589diff
changeset | 428 | qed "set_empty"; | 
| 
a64c8fbcd98f
Renamed theorems of the form set_of_list_XXX to set_XXX
 paulson parents: 
3589diff
changeset | 429 | Addsimps [set_empty]; | 
| 2608 | 430 | |
| 4935 | 431 | Goal "set(rev xs) = set(xs)"; | 
| 3457 | 432 | by (induct_tac "xs" 1); | 
| 5316 | 433 | by Auto_tac; | 
| 3647 
a64c8fbcd98f
Renamed theorems of the form set_of_list_XXX to set_XXX
 paulson parents: 
3589diff
changeset | 434 | qed "set_rev"; | 
| 
a64c8fbcd98f
Renamed theorems of the form set_of_list_XXX to set_XXX
 paulson parents: 
3589diff
changeset | 435 | Addsimps [set_rev]; | 
| 2608 | 436 | |
| 4935 | 437 | Goal "set(map f xs) = f``(set xs)"; | 
| 3457 | 438 | by (induct_tac "xs" 1); | 
| 5316 | 439 | by Auto_tac; | 
| 3647 
a64c8fbcd98f
Renamed theorems of the form set_of_list_XXX to set_XXX
 paulson parents: 
3589diff
changeset | 440 | qed "set_map"; | 
| 
a64c8fbcd98f
Renamed theorems of the form set_of_list_XXX to set_XXX
 paulson parents: 
3589diff
changeset | 441 | Addsimps [set_map]; | 
| 2608 | 442 | |
| 5443 
e2459d18ff47
changed constants mem and list_all to mere translations
 oheimb parents: 
5427diff
changeset | 443 | Goal "(x : set (filter P xs)) = (x : set xs & P x)"; | 
| 4605 | 444 | by (induct_tac "xs" 1); | 
| 5316 | 445 | by Auto_tac; | 
| 4605 | 446 | qed "in_set_filter"; | 
| 447 | Addsimps [in_set_filter]; | |
| 448 | ||
| 5272 | 449 | Goal "(x : set xs) = (? ys zs. xs = ys@x#zs)"; | 
| 5318 | 450 | by (induct_tac "xs" 1); | 
| 451 | by (Simp_tac 1); | |
| 452 | by (Asm_simp_tac 1); | |
| 453 | by (rtac iffI 1); | |
| 454 | by (blast_tac (claset() addIs [eq_Nil_appendI,Cons_eq_appendI]) 1); | |
| 455 | by (REPEAT(etac exE 1)); | |
| 456 | by (exhaust_tac "ys" 1); | |
| 5316 | 457 | by Auto_tac; | 
| 5272 | 458 | qed "in_set_conv_decomp"; | 
| 459 | ||
| 460 | (* eliminate `lists' in favour of `set' *) | |
| 461 | ||
| 462 | Goal "(xs : lists A) = (!x : set xs. x : A)"; | |
| 5318 | 463 | by (induct_tac "xs" 1); | 
| 5316 | 464 | by Auto_tac; | 
| 5272 | 465 | qed "in_lists_conv_set"; | 
| 466 | ||
| 467 | bind_thm("in_listsD",in_lists_conv_set RS iffD1);
 | |
| 468 | AddSDs [in_listsD]; | |
| 469 | bind_thm("in_listsI",in_lists_conv_set RS iffD2);
 | |
| 470 | AddSIs [in_listsI]; | |
| 1812 | 471 | |
| 5518 | 472 | (** mem **) | 
| 473 | ||
| 474 | section "mem"; | |
| 475 | ||
| 476 | Goal "(x mem xs) = (x: set xs)"; | |
| 477 | by (induct_tac "xs" 1); | |
| 478 | by Auto_tac; | |
| 479 | qed "set_mem_eq"; | |
| 480 | ||
| 481 | ||
| 923 | 482 | (** list_all **) | 
| 483 | ||
| 3467 | 484 | section "list_all"; | 
| 485 | ||
| 5518 | 486 | Goal "list_all P xs = (!x:set xs. P x)"; | 
| 487 | by (induct_tac "xs" 1); | |
| 488 | by Auto_tac; | |
| 489 | qed "list_all_conv"; | |
| 490 | ||
| 5443 
e2459d18ff47
changed constants mem and list_all to mere translations
 oheimb parents: 
5427diff
changeset | 491 | 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 | 492 | by (induct_tac "xs" 1); | 
| 5316 | 493 | by Auto_tac; | 
| 2512 | 494 | qed "list_all_append"; | 
| 495 | Addsimps [list_all_append]; | |
| 923 | 496 | |
| 497 | ||
| 2608 | 498 | (** filter **) | 
| 923 | 499 | |
| 3467 | 500 | section "filter"; | 
| 501 | ||
| 4935 | 502 | Goal "filter P (xs@ys) = filter P xs @ filter P ys"; | 
| 3457 | 503 | by (induct_tac "xs" 1); | 
| 5316 | 504 | by Auto_tac; | 
| 2608 | 505 | qed "filter_append"; | 
| 506 | Addsimps [filter_append]; | |
| 507 | ||
| 4935 | 508 | Goal "filter (%x. True) xs = xs"; | 
| 4605 | 509 | by (induct_tac "xs" 1); | 
| 5316 | 510 | by Auto_tac; | 
| 4605 | 511 | qed "filter_True"; | 
| 512 | Addsimps [filter_True]; | |
| 513 | ||
| 4935 | 514 | Goal "filter (%x. False) xs = []"; | 
| 4605 | 515 | by (induct_tac "xs" 1); | 
| 5316 | 516 | by Auto_tac; | 
| 4605 | 517 | qed "filter_False"; | 
| 518 | Addsimps [filter_False]; | |
| 519 | ||
| 4935 | 520 | Goal "length (filter P xs) <= length xs"; | 
| 3457 | 521 | by (induct_tac "xs" 1); | 
| 5316 | 522 | by Auto_tac; | 
| 4605 | 523 | qed "length_filter"; | 
| 5443 
e2459d18ff47
changed constants mem and list_all to mere translations
 oheimb parents: 
5427diff
changeset | 524 | Addsimps[length_filter]; | 
| 2608 | 525 | |
| 5443 
e2459d18ff47
changed constants mem and list_all to mere translations
 oheimb parents: 
5427diff
changeset | 526 | Goal "set (filter P xs) <= set xs"; | 
| 
e2459d18ff47
changed constants mem and list_all to mere translations
 oheimb parents: 
5427diff
changeset | 527 | by Auto_tac; | 
| 
e2459d18ff47
changed constants mem and list_all to mere translations
 oheimb parents: 
5427diff
changeset | 528 | qed "filter_is_subset"; | 
| 
e2459d18ff47
changed constants mem and list_all to mere translations
 oheimb parents: 
5427diff
changeset | 529 | Addsimps [filter_is_subset]; | 
| 
e2459d18ff47
changed constants mem and list_all to mere translations
 oheimb parents: 
5427diff
changeset | 530 | |
| 2608 | 531 | |
| 3467 | 532 | section "concat"; | 
| 533 | ||
| 4935 | 534 | 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 | 535 | by (induct_tac "xs" 1); | 
| 5316 | 536 | by Auto_tac; | 
| 2608 | 537 | qed"concat_append"; | 
| 538 | Addsimps [concat_append]; | |
| 2512 | 539 | |
| 4935 | 540 | Goal "(concat xss = []) = (!xs:set xss. xs=[])"; | 
| 4423 | 541 | by (induct_tac "xss" 1); | 
| 5316 | 542 | by Auto_tac; | 
| 3896 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 nipkow parents: 
3860diff
changeset | 543 | qed "concat_eq_Nil_conv"; | 
| 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 nipkow parents: 
3860diff
changeset | 544 | AddIffs [concat_eq_Nil_conv]; | 
| 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 nipkow parents: 
3860diff
changeset | 545 | |
| 4935 | 546 | Goal "([] = concat xss) = (!xs:set xss. xs=[])"; | 
| 4423 | 547 | by (induct_tac "xss" 1); | 
| 5316 | 548 | by Auto_tac; | 
| 3896 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 nipkow parents: 
3860diff
changeset | 549 | qed "Nil_eq_concat_conv"; | 
| 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 nipkow parents: 
3860diff
changeset | 550 | AddIffs [Nil_eq_concat_conv]; | 
| 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 nipkow parents: 
3860diff
changeset | 551 | |
| 4935 | 552 | Goal "set(concat xs) = Union(set `` set xs)"; | 
| 3467 | 553 | by (induct_tac "xs" 1); | 
| 5316 | 554 | by Auto_tac; | 
| 3647 
a64c8fbcd98f
Renamed theorems of the form set_of_list_XXX to set_XXX
 paulson parents: 
3589diff
changeset | 555 | qed"set_concat"; | 
| 
a64c8fbcd98f
Renamed theorems of the form set_of_list_XXX to set_XXX
 paulson parents: 
3589diff
changeset | 556 | Addsimps [set_concat]; | 
| 3467 | 557 | |
| 4935 | 558 | Goal "map f (concat xs) = concat (map (map f) xs)"; | 
| 3467 | 559 | by (induct_tac "xs" 1); | 
| 5316 | 560 | by Auto_tac; | 
| 3467 | 561 | qed "map_concat"; | 
| 562 | ||
| 4935 | 563 | Goal "filter p (concat xs) = concat (map (filter p) xs)"; | 
| 3467 | 564 | by (induct_tac "xs" 1); | 
| 5316 | 565 | by Auto_tac; | 
| 3467 | 566 | qed"filter_concat"; | 
| 567 | ||
| 4935 | 568 | Goal "rev(concat xs) = concat (map rev (rev xs))"; | 
| 3467 | 569 | by (induct_tac "xs" 1); | 
| 5316 | 570 | by Auto_tac; | 
| 2608 | 571 | qed "rev_concat"; | 
| 923 | 572 | |
| 573 | (** nth **) | |
| 574 | ||
| 3467 | 575 | section "nth"; | 
| 576 | ||
| 5278 | 577 | Goal "!xs. (xs@ys)!n = (if n < length xs then xs!n else ys!(n - length xs))"; | 
| 5183 | 578 | by (induct_tac "n" 1); | 
| 3457 | 579 | by (Asm_simp_tac 1); | 
| 580 | by (rtac allI 1); | |
| 581 | by (exhaust_tac "xs" 1); | |
| 5316 | 582 | by Auto_tac; | 
| 2608 | 583 | qed_spec_mp "nth_append"; | 
| 584 | ||
| 4935 | 585 | 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 | 586 | by (induct_tac "xs" 1); | 
| 1301 | 587 | (* case [] *) | 
| 588 | by (Asm_full_simp_tac 1); | |
| 589 | (* case x#xl *) | |
| 590 | by (rtac allI 1); | |
| 5183 | 591 | by (induct_tac "n" 1); | 
| 5316 | 592 | by Auto_tac; | 
| 1485 
240cc98b94a7
Added qed_spec_mp to avoid renaming of bound vars in 'th RS spec'
 nipkow parents: 
1465diff
changeset | 593 | qed_spec_mp "nth_map"; | 
| 1301 | 594 | Addsimps [nth_map]; | 
| 595 | ||
| 5518 | 596 | Goal "!n. n < length xs --> Ball (set xs) P --> P(xs!n)"; | 
| 3040 
7d48671753da
Introduced a generic "induct_tac" which picks up the right induction scheme
 nipkow parents: 
3011diff
changeset | 597 | by (induct_tac "xs" 1); | 
| 1301 | 598 | (* case [] *) | 
| 599 | by (Simp_tac 1); | |
| 600 | (* case x#xl *) | |
| 601 | by (rtac allI 1); | |
| 5183 | 602 | by (induct_tac "n" 1); | 
| 5316 | 603 | by Auto_tac; | 
| 5518 | 604 | qed_spec_mp "list_ball_nth"; | 
| 1301 | 605 | |
| 5518 | 606 | Goal "!n. n < length xs --> xs!n : set xs"; | 
| 3040 
7d48671753da
Introduced a generic "induct_tac" which picks up the right induction scheme
 nipkow parents: 
3011diff
changeset | 607 | by (induct_tac "xs" 1); | 
| 1301 | 608 | (* case [] *) | 
| 609 | by (Simp_tac 1); | |
| 610 | (* case x#xl *) | |
| 611 | by (rtac allI 1); | |
| 5183 | 612 | by (induct_tac "n" 1); | 
| 1301 | 613 | (* case 0 *) | 
| 614 | by (Asm_full_simp_tac 1); | |
| 615 | (* case Suc x *) | |
| 4686 | 616 | by (Asm_full_simp_tac 1); | 
| 1485 
240cc98b94a7
Added qed_spec_mp to avoid renaming of bound vars in 'th RS spec'
 nipkow parents: 
1465diff
changeset | 617 | qed_spec_mp "nth_mem"; | 
| 1301 | 618 | Addsimps [nth_mem]; | 
| 619 | ||
| 5518 | 620 | |
| 5077 
71043526295f
* HOL/List: new function list_update written xs[i:=v] that updates the i-th
 nipkow parents: 
5043diff
changeset | 621 | (** list update **) | 
| 
71043526295f
* HOL/List: new function list_update written xs[i:=v] that updates the i-th
 nipkow parents: 
5043diff
changeset | 622 | |
| 
71043526295f
* HOL/List: new function list_update written xs[i:=v] that updates the i-th
 nipkow parents: 
5043diff
changeset | 623 | section "list update"; | 
| 
71043526295f
* HOL/List: new function list_update written xs[i:=v] that updates the i-th
 nipkow parents: 
5043diff
changeset | 624 | |
| 
71043526295f
* HOL/List: new function list_update written xs[i:=v] that updates the i-th
 nipkow parents: 
5043diff
changeset | 625 | 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 | 626 | 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 | 627 | by (Simp_tac 1); | 
| 5183 | 628 | 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 | 629 | 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 | 630 | Addsimps [length_list_update]; | 
| 
71043526295f
* HOL/List: new function list_update written xs[i:=v] that updates the i-th
 nipkow parents: 
5043diff
changeset | 631 | |
| 
71043526295f
* HOL/List: new function list_update written xs[i:=v] that updates the i-th
 nipkow parents: 
5043diff
changeset | 632 | |
| 3896 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 nipkow parents: 
3860diff
changeset | 633 | (** last & butlast **) | 
| 1327 
6c29cfab679c
added new arithmetic lemmas and the functions take and drop.
 nipkow parents: 
1301diff
changeset | 634 | |
| 4935 | 635 | Goal "last(xs@[x]) = x"; | 
| 4423 | 636 | by (induct_tac "xs" 1); | 
| 5316 | 637 | by Auto_tac; | 
| 3896 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 nipkow parents: 
3860diff
changeset | 638 | qed "last_snoc"; | 
| 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 nipkow parents: 
3860diff
changeset | 639 | Addsimps [last_snoc]; | 
| 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 nipkow parents: 
3860diff
changeset | 640 | |
| 4935 | 641 | Goal "butlast(xs@[x]) = xs"; | 
| 4423 | 642 | by (induct_tac "xs" 1); | 
| 5316 | 643 | by Auto_tac; | 
| 3896 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 nipkow parents: 
3860diff
changeset | 644 | qed "butlast_snoc"; | 
| 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 nipkow parents: 
3860diff
changeset | 645 | Addsimps [butlast_snoc]; | 
| 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 nipkow parents: 
3860diff
changeset | 646 | |
| 4935 | 647 | Goal "length(butlast xs) = length xs - 1"; | 
| 648 | by (res_inst_tac [("xs","xs")] rev_induct 1);
 | |
| 5316 | 649 | by Auto_tac; | 
| 4643 | 650 | qed "length_butlast"; | 
| 651 | Addsimps [length_butlast]; | |
| 652 | ||
| 5278 | 653 | Goal "!ys. butlast (xs@ys) = (if ys=[] then butlast xs else xs@butlast ys)"; | 
| 4423 | 654 | by (induct_tac "xs" 1); | 
| 5316 | 655 | by Auto_tac; | 
| 3896 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 nipkow parents: 
3860diff
changeset | 656 | qed_spec_mp "butlast_append"; | 
| 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 nipkow parents: 
3860diff
changeset | 657 | |
| 4935 | 658 | Goal "x:set(butlast xs) --> x:set xs"; | 
| 4423 | 659 | by (induct_tac "xs" 1); | 
| 5316 | 660 | by Auto_tac; | 
| 3896 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 nipkow parents: 
3860diff
changeset | 661 | qed_spec_mp "in_set_butlastD"; | 
| 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 nipkow parents: 
3860diff
changeset | 662 | |
| 5448 
40a09282ba14
in_set_butlast_appendI supersedes in_set_butlast_appendI1,2
 paulson parents: 
5443diff
changeset | 663 | 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 | 664 | by (auto_tac (claset() addDs [in_set_butlastD], | 
| 
40a09282ba14
in_set_butlast_appendI supersedes in_set_butlast_appendI1,2
 paulson parents: 
5443diff
changeset | 665 | simpset() addsimps [butlast_append])); | 
| 
40a09282ba14
in_set_butlast_appendI supersedes in_set_butlast_appendI1,2
 paulson parents: 
5443diff
changeset | 666 | qed "in_set_butlast_appendI"; | 
| 3902 | 667 | |
| 2608 | 668 | (** take & drop **) | 
| 669 | section "take & drop"; | |
| 1327 
6c29cfab679c
added new arithmetic lemmas and the functions take and drop.
 nipkow parents: 
1301diff
changeset | 670 | |
| 4935 | 671 | Goal "take 0 xs = []"; | 
| 3040 
7d48671753da
Introduced a generic "induct_tac" which picks up the right induction scheme
 nipkow parents: 
3011diff
changeset | 672 | by (induct_tac "xs" 1); | 
| 5316 | 673 | by Auto_tac; | 
| 1327 
6c29cfab679c
added new arithmetic lemmas and the functions take and drop.
 nipkow parents: 
1301diff
changeset | 674 | qed "take_0"; | 
| 
6c29cfab679c
added new arithmetic lemmas and the functions take and drop.
 nipkow parents: 
1301diff
changeset | 675 | |
| 4935 | 676 | Goal "drop 0 xs = xs"; | 
| 3040 
7d48671753da
Introduced a generic "induct_tac" which picks up the right induction scheme
 nipkow parents: 
3011diff
changeset | 677 | by (induct_tac "xs" 1); | 
| 5316 | 678 | by Auto_tac; | 
| 2608 | 679 | qed "drop_0"; | 
| 680 | ||
| 4935 | 681 | Goal "take (Suc n) (x#xs) = x # take n xs"; | 
| 1552 | 682 | by (Simp_tac 1); | 
| 1419 
a6a034a47a71
defined take/drop by induction over list rather than nat.
 nipkow parents: 
1327diff
changeset | 683 | qed "take_Suc_Cons"; | 
| 1327 
6c29cfab679c
added new arithmetic lemmas and the functions take and drop.
 nipkow parents: 
1301diff
changeset | 684 | |
| 4935 | 685 | Goal "drop (Suc n) (x#xs) = drop n xs"; | 
| 2608 | 686 | by (Simp_tac 1); | 
| 687 | qed "drop_Suc_Cons"; | |
| 688 | ||
| 689 | Delsimps [take_Cons,drop_Cons]; | |
| 690 | Addsimps [take_0,take_Suc_Cons,drop_0,drop_Suc_Cons]; | |
| 691 | ||
| 4935 | 692 | Goal "!xs. length(take n xs) = min (length xs) n"; | 
| 5183 | 693 | by (induct_tac "n" 1); | 
| 5316 | 694 | by Auto_tac; | 
| 3457 | 695 | by (exhaust_tac "xs" 1); | 
| 5316 | 696 | by Auto_tac; | 
| 2608 | 697 | qed_spec_mp "length_take"; | 
| 698 | Addsimps [length_take]; | |
| 923 | 699 | |
| 4935 | 700 | Goal "!xs. length(drop n xs) = (length xs - n)"; | 
| 5183 | 701 | by (induct_tac "n" 1); | 
| 5316 | 702 | by Auto_tac; | 
| 3457 | 703 | by (exhaust_tac "xs" 1); | 
| 5316 | 704 | by Auto_tac; | 
| 2608 | 705 | qed_spec_mp "length_drop"; | 
| 706 | Addsimps [length_drop]; | |
| 707 | ||
| 4935 | 708 | Goal "!xs. length xs <= n --> take n xs = xs"; | 
| 5183 | 709 | by (induct_tac "n" 1); | 
| 5316 | 710 | by Auto_tac; | 
| 3457 | 711 | by (exhaust_tac "xs" 1); | 
| 5316 | 712 | by Auto_tac; | 
| 2608 | 713 | qed_spec_mp "take_all"; | 
| 923 | 714 | |
| 4935 | 715 | Goal "!xs. length xs <= n --> drop n xs = []"; | 
| 5183 | 716 | by (induct_tac "n" 1); | 
| 5316 | 717 | by Auto_tac; | 
| 3457 | 718 | by (exhaust_tac "xs" 1); | 
| 5316 | 719 | by Auto_tac; | 
| 2608 | 720 | qed_spec_mp "drop_all"; | 
| 721 | ||
| 5278 | 722 | Goal "!xs. take n (xs @ ys) = (take n xs @ take (n - length xs) ys)"; | 
| 5183 | 723 | by (induct_tac "n" 1); | 
| 5316 | 724 | by Auto_tac; | 
| 3457 | 725 | by (exhaust_tac "xs" 1); | 
| 5316 | 726 | by Auto_tac; | 
| 2608 | 727 | qed_spec_mp "take_append"; | 
| 728 | Addsimps [take_append]; | |
| 729 | ||
| 4935 | 730 | Goal "!xs. drop n (xs@ys) = drop n xs @ drop (n - length xs) ys"; | 
| 5183 | 731 | by (induct_tac "n" 1); | 
| 5316 | 732 | by Auto_tac; | 
| 3457 | 733 | by (exhaust_tac "xs" 1); | 
| 5316 | 734 | by Auto_tac; | 
| 2608 | 735 | qed_spec_mp "drop_append"; | 
| 736 | Addsimps [drop_append]; | |
| 737 | ||
| 4935 | 738 | Goal "!xs n. take n (take m xs) = take (min n m) xs"; | 
| 5183 | 739 | by (induct_tac "m" 1); | 
| 5316 | 740 | by Auto_tac; | 
| 3457 | 741 | by (exhaust_tac "xs" 1); | 
| 5316 | 742 | by Auto_tac; | 
| 5183 | 743 | by (exhaust_tac "na" 1); | 
| 5316 | 744 | by Auto_tac; | 
| 2608 | 745 | qed_spec_mp "take_take"; | 
| 746 | ||
| 4935 | 747 | Goal "!xs. drop n (drop m xs) = drop (n + m) xs"; | 
| 5183 | 748 | by (induct_tac "m" 1); | 
| 5316 | 749 | by Auto_tac; | 
| 3457 | 750 | by (exhaust_tac "xs" 1); | 
| 5316 | 751 | by Auto_tac; | 
| 2608 | 752 | qed_spec_mp "drop_drop"; | 
| 923 | 753 | |
| 4935 | 754 | Goal "!xs n. take n (drop m xs) = drop m (take (n + m) xs)"; | 
| 5183 | 755 | by (induct_tac "m" 1); | 
| 5316 | 756 | by Auto_tac; | 
| 3457 | 757 | by (exhaust_tac "xs" 1); | 
| 5316 | 758 | by Auto_tac; | 
| 2608 | 759 | qed_spec_mp "take_drop"; | 
| 760 | ||
| 4935 | 761 | Goal "!xs. take n (map f xs) = map f (take n xs)"; | 
| 5183 | 762 | by (induct_tac "n" 1); | 
| 5316 | 763 | by Auto_tac; | 
| 3457 | 764 | by (exhaust_tac "xs" 1); | 
| 5316 | 765 | by Auto_tac; | 
| 2608 | 766 | qed_spec_mp "take_map"; | 
| 767 | ||
| 4935 | 768 | Goal "!xs. drop n (map f xs) = map f (drop n xs)"; | 
| 5183 | 769 | by (induct_tac "n" 1); | 
| 5316 | 770 | by Auto_tac; | 
| 3457 | 771 | by (exhaust_tac "xs" 1); | 
| 5316 | 772 | by Auto_tac; | 
| 2608 | 773 | qed_spec_mp "drop_map"; | 
| 774 | ||
| 4935 | 775 | Goal "!n i. i < n --> (take n xs)!i = xs!i"; | 
| 3457 | 776 | by (induct_tac "xs" 1); | 
| 5316 | 777 | by Auto_tac; | 
| 3457 | 778 | by (exhaust_tac "n" 1); | 
| 779 | by (Blast_tac 1); | |
| 780 | by (exhaust_tac "i" 1); | |
| 5316 | 781 | by Auto_tac; | 
| 2608 | 782 | qed_spec_mp "nth_take"; | 
| 783 | Addsimps [nth_take]; | |
| 923 | 784 | |
| 4935 | 785 | Goal "!xs i. n + i <= length xs --> (drop n xs)!i = xs!(n+i)"; | 
| 5183 | 786 | by (induct_tac "n" 1); | 
| 5316 | 787 | by Auto_tac; | 
| 3457 | 788 | by (exhaust_tac "xs" 1); | 
| 5316 | 789 | by Auto_tac; | 
| 2608 | 790 | qed_spec_mp "nth_drop"; | 
| 791 | Addsimps [nth_drop]; | |
| 792 | ||
| 793 | (** takeWhile & dropWhile **) | |
| 794 | ||
| 3467 | 795 | section "takeWhile & dropWhile"; | 
| 796 | ||
| 4935 | 797 | Goal "takeWhile P xs @ dropWhile P xs = xs"; | 
| 3586 | 798 | by (induct_tac "xs" 1); | 
| 5316 | 799 | by Auto_tac; | 
| 3586 | 800 | qed "takeWhile_dropWhile_id"; | 
| 801 | Addsimps [takeWhile_dropWhile_id]; | |
| 802 | ||
| 4935 | 803 | Goal "x:set xs & ~P(x) --> takeWhile P (xs @ ys) = takeWhile P xs"; | 
| 3457 | 804 | by (induct_tac "xs" 1); | 
| 5316 | 805 | by Auto_tac; | 
| 2608 | 806 | bind_thm("takeWhile_append1", conjI RS (result() RS mp));
 | 
| 807 | Addsimps [takeWhile_append1]; | |
| 923 | 808 | |
| 4935 | 809 | Goal "(!x:set xs. P(x)) --> takeWhile P (xs @ ys) = xs @ takeWhile P ys"; | 
| 3457 | 810 | by (induct_tac "xs" 1); | 
| 5316 | 811 | by Auto_tac; | 
| 2608 | 812 | bind_thm("takeWhile_append2", ballI RS (result() RS mp));
 | 
| 813 | Addsimps [takeWhile_append2]; | |
| 1169 | 814 | |
| 4935 | 815 | Goal "x:set xs & ~P(x) --> dropWhile P (xs @ ys) = (dropWhile P xs)@ys"; | 
| 3457 | 816 | by (induct_tac "xs" 1); | 
| 5316 | 817 | by Auto_tac; | 
| 2608 | 818 | bind_thm("dropWhile_append1", conjI RS (result() RS mp));
 | 
| 819 | Addsimps [dropWhile_append1]; | |
| 820 | ||
| 4935 | 821 | Goal "(!x:set xs. P(x)) --> dropWhile P (xs @ ys) = dropWhile P ys"; | 
| 3457 | 822 | by (induct_tac "xs" 1); | 
| 5316 | 823 | by Auto_tac; | 
| 2608 | 824 | bind_thm("dropWhile_append2", ballI RS (result() RS mp));
 | 
| 825 | Addsimps [dropWhile_append2]; | |
| 826 | ||
| 4935 | 827 | Goal "x:set(takeWhile P xs) --> x:set xs & P x"; | 
| 3457 | 828 | by (induct_tac "xs" 1); | 
| 5316 | 829 | by Auto_tac; | 
| 3647 
a64c8fbcd98f
Renamed theorems of the form set_of_list_XXX to set_XXX
 paulson parents: 
3589diff
changeset | 830 | qed_spec_mp"set_take_whileD"; | 
| 2608 | 831 | |
| 4132 | 832 | qed_goal "zip_Nil_Nil" thy "zip [] [] = []" (K [Simp_tac 1]); | 
| 833 | qed_goal "zip_Cons_Cons" thy "zip (x#xs) (y#ys) = (x,y)#zip xs ys" | |
| 834 | (K [Simp_tac 1]); | |
| 4605 | 835 | |
| 5272 | 836 | |
| 837 | (** foldl **) | |
| 838 | section "foldl"; | |
| 839 | ||
| 840 | Goal "!a. foldl f a (xs @ ys) = foldl f (foldl f a xs) ys"; | |
| 5318 | 841 | by (induct_tac "xs" 1); | 
| 5316 | 842 | by Auto_tac; | 
| 5272 | 843 | qed_spec_mp "foldl_append"; | 
| 844 | Addsimps [foldl_append]; | |
| 845 | ||
| 846 | (* Note: `n <= foldl op+ n ns' looks simpler, but is more difficult to use | |
| 847 | because it requires an additional transitivity step | |
| 848 | *) | |
| 849 | Goal "!n::nat. m <= n --> m <= foldl op+ n ns"; | |
| 5318 | 850 | by (induct_tac "ns" 1); | 
| 851 | by (Simp_tac 1); | |
| 852 | by (Asm_full_simp_tac 1); | |
| 853 | by (blast_tac (claset() addIs [trans_le_add1]) 1); | |
| 5272 | 854 | qed_spec_mp "start_le_sum"; | 
| 855 | ||
| 856 | Goal "n : set ns ==> n <= foldl op+ 0 ns"; | |
| 5318 | 857 | by (auto_tac (claset() addIs [start_le_sum], | 
| 5272 | 858 | simpset() addsimps [in_set_conv_decomp])); | 
| 859 | qed "elem_le_sum"; | |
| 860 | ||
| 861 | Goal "!m. (foldl op+ m ns = 0) = (m=0 & (!n : set ns. n=0))"; | |
| 5318 | 862 | by (induct_tac "ns" 1); | 
| 5316 | 863 | by Auto_tac; | 
| 5272 | 864 | qed_spec_mp "sum_eq_0_conv"; | 
| 865 | AddIffs [sum_eq_0_conv]; | |
| 866 | ||
| 5425 | 867 | (** upto **) | 
| 868 | ||
| 5427 | 869 | (* Does not terminate! *) | 
| 870 | Goal "[i..j(] = (if i<j then i#[Suc i..j(] else [])"; | |
| 871 | by(induct_tac "j" 1); | |
| 872 | by Auto_tac; | |
| 873 | by(REPEAT(trans_tac 1)); | |
| 874 | qed "upt_rec"; | |
| 5425 | 875 | |
| 5427 | 876 | Goal "j<=i ==> [i..j(] = []"; | 
| 877 | by(stac upt_rec 1); | |
| 878 | by(asm_simp_tac (simpset() addSolver cut_trans_tac) 1); | |
| 879 | qed "upt_conv_Nil"; | |
| 880 | Addsimps [upt_conv_Nil]; | |
| 881 | ||
| 882 | Goal "i<=j ==> [i..(Suc j)(] = [i..j(]@[j]"; | |
| 883 | by (Asm_simp_tac 1); | |
| 884 | qed "upt_Suc"; | |
| 885 | ||
| 886 | Goal "i<j ==> [i..j(] = i#[Suc i..j(]"; | |
| 5425 | 887 | br trans 1; | 
| 5427 | 888 | by(stac upt_rec 1); | 
| 889 | br refl 2; | |
| 890 | by (Asm_simp_tac 1); | |
| 891 | qed "upt_conv_Cons"; | |
| 892 | ||
| 893 | Goal "length [i..j(] = j-i"; | |
| 894 | by(induct_tac "j" 1); | |
| 895 | by (Simp_tac 1); | |
| 896 | by(asm_simp_tac (simpset() addsimps [Suc_diff_le] addSolver cut_trans_tac) 1); | |
| 897 | qed "length_upt"; | |
| 898 | Addsimps [length_upt]; | |
| 5425 | 899 | |
| 5427 | 900 | Goal "i+k < j --> [i..j(] ! k = i+k"; | 
| 901 | by(induct_tac "j" 1); | |
| 902 | by(Simp_tac 1); | |
| 5537 | 903 | by(asm_simp_tac (simpset() addsimps [nth_append,less_diff_conv]@add_ac | 
| 5427 | 904 | addSolver cut_trans_tac) 1); | 
| 905 | br conjI 1; | |
| 906 | by(Clarify_tac 1); | |
| 907 | bd add_lessD1 1; | |
| 908 | by(trans_tac 1); | |
| 909 | by(Clarify_tac 1); | |
| 910 | br conjI 1; | |
| 911 | by(Clarify_tac 1); | |
| 912 | by(subgoal_tac "n=i+k" 1); | |
| 913 | by(Asm_full_simp_tac 1); | |
| 914 | by(trans_tac 1); | |
| 915 | by(Clarify_tac 1); | |
| 916 | by(subgoal_tac "n=i+k" 1); | |
| 917 | by(Asm_full_simp_tac 1); | |
| 918 | by(trans_tac 1); | |
| 919 | qed_spec_mp "nth_upt"; | |
| 920 | Addsimps [nth_upt]; | |
| 5425 | 921 | |
| 5272 | 922 | |
| 4605 | 923 | (** nodups & remdups **) | 
| 924 | section "nodups & remdups"; | |
| 925 | ||
| 4935 | 926 | Goal "set(remdups xs) = set xs"; | 
| 4605 | 927 | by (induct_tac "xs" 1); | 
| 928 | by (Simp_tac 1); | |
| 4686 | 929 | by (asm_full_simp_tac (simpset() addsimps [insert_absorb]) 1); | 
| 4605 | 930 | qed "set_remdups"; | 
| 931 | Addsimps [set_remdups]; | |
| 932 | ||
| 4935 | 933 | Goal "nodups(remdups xs)"; | 
| 4605 | 934 | by (induct_tac "xs" 1); | 
| 5316 | 935 | by Auto_tac; | 
| 4605 | 936 | qed "nodups_remdups"; | 
| 937 | ||
| 4935 | 938 | Goal "nodups xs --> nodups (filter P xs)"; | 
| 4605 | 939 | by (induct_tac "xs" 1); | 
| 5316 | 940 | by Auto_tac; | 
| 4605 | 941 | qed_spec_mp "nodups_filter"; | 
| 942 | ||
| 3589 
244daa75f890
Added function `replicate' and lemmas map_cong and set_replicate.
 nipkow parents: 
3586diff
changeset | 943 | (** replicate **) | 
| 
244daa75f890
Added function `replicate' and lemmas map_cong and set_replicate.
 nipkow parents: 
3586diff
changeset | 944 | section "replicate"; | 
| 
244daa75f890
Added function `replicate' and lemmas map_cong and set_replicate.
 nipkow parents: 
3586diff
changeset | 945 | |
| 4935 | 946 | Goal "set(replicate (Suc n) x) = {x}";
 | 
| 4423 | 947 | by (induct_tac "n" 1); | 
| 5316 | 948 | by Auto_tac; | 
| 3589 
244daa75f890
Added function `replicate' and lemmas map_cong and set_replicate.
 nipkow parents: 
3586diff
changeset | 949 | val lemma = result(); | 
| 
244daa75f890
Added function `replicate' and lemmas map_cong and set_replicate.
 nipkow parents: 
3586diff
changeset | 950 | |
| 5043 | 951 | Goal "n ~= 0 ==> set(replicate n x) = {x}";
 | 
| 4423 | 952 | 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 | 953 | qed "set_replicate"; | 
| 
244daa75f890
Added function `replicate' and lemmas map_cong and set_replicate.
 nipkow parents: 
3586diff
changeset | 954 | Addsimps [set_replicate]; | 
| 5162 | 955 | |
| 956 | ||
| 5281 | 957 | (*** Lexcicographic orderings on lists ***) | 
| 958 | section"Lexcicographic orderings on lists"; | |
| 959 | ||
| 960 | Goal "wf r ==> wf(lexn r n)"; | |
| 5318 | 961 | by (induct_tac "n" 1); | 
| 962 | by (Simp_tac 1); | |
| 963 | by (Simp_tac 1); | |
| 964 | by (rtac wf_subset 1); | |
| 965 | by (rtac Int_lower1 2); | |
| 966 | by (rtac wf_prod_fun_image 1); | |
| 967 | by (rtac injI 2); | |
| 968 | by (Auto_tac); | |
| 5281 | 969 | qed "wf_lexn"; | 
| 970 | ||
| 971 | Goal "!xs ys. (xs,ys) : lexn r n --> length xs = n & length ys = n"; | |
| 5318 | 972 | by (induct_tac "n" 1); | 
| 973 | by (Auto_tac); | |
| 5281 | 974 | qed_spec_mp "lexn_length"; | 
| 975 | ||
| 976 | Goalw [lex_def] "wf r ==> wf(lex r)"; | |
| 5318 | 977 | by (rtac wf_UN 1); | 
| 978 | by (blast_tac (claset() addIs [wf_lexn]) 1); | |
| 979 | by (Clarify_tac 1); | |
| 980 | by (rename_tac "m n" 1); | |
| 981 | by (subgoal_tac "m ~= n" 1); | |
| 982 | by (Blast_tac 2); | |
| 983 | by (blast_tac (claset() addDs [lexn_length,not_sym]) 1); | |
| 5281 | 984 | qed "wf_lex"; | 
| 985 | AddSIs [wf_lex]; | |
| 986 | ||
| 987 | Goal | |
| 988 | "lexn r n = \ | |
| 989 | \ {(xs,ys). length xs = n & length ys = n & \
 | |
| 990 | \ (? xys x y xs' ys'. xs= xys @ x#xs' & ys= xys @ y#ys' & (x,y):r)}"; | |
| 5318 | 991 | by (induct_tac "n" 1); | 
| 992 | by (Simp_tac 1); | |
| 993 | by (Blast_tac 1); | |
| 994 | by (asm_full_simp_tac (simpset() delsimps [length_Suc_conv] | |
| 5296 | 995 | addsimps [lex_prod_def]) 1); | 
| 5318 | 996 | by (auto_tac (claset(), simpset() delsimps [length_Suc_conv])); | 
| 997 | by (Blast_tac 1); | |
| 998 | by (rename_tac "a xys x xs' y ys'" 1); | |
| 999 |  by (res_inst_tac [("x","a#xys")] exI 1);
 | |
| 1000 | by (Simp_tac 1); | |
| 1001 | by (exhaust_tac "xys" 1); | |
| 1002 | by (ALLGOALS (asm_full_simp_tac (simpset() delsimps [length_Suc_conv]))); | |
| 1003 | by (Blast_tac 1); | |
| 5281 | 1004 | qed "lexn_conv"; | 
| 1005 | ||
| 1006 | Goalw [lex_def] | |
| 1007 | "lex r = \ | |
| 1008 | \ {(xs,ys). length xs = length ys & \
 | |
| 1009 | \ (? xys x y xs' ys'. xs= xys @ x#xs' & ys= xys @ y#ys' & (x,y):r)}"; | |
| 5318 | 1010 | by (force_tac (claset(), simpset() delsimps [length_Suc_conv] addsimps [lexn_conv]) 1); | 
| 5281 | 1011 | qed "lex_conv"; | 
| 1012 | ||
| 1013 | Goalw [lexico_def] "wf r ==> wf(lexico r)"; | |
| 5318 | 1014 | by (Blast_tac 1); | 
| 5281 | 1015 | qed "wf_lexico"; | 
| 1016 | AddSIs [wf_lexico]; | |
| 1017 | ||
| 1018 | Goalw | |
| 1019 | [lexico_def,diag_def,lex_prod_def,measure_def,inv_image_def] | |
| 1020 | "lexico r = {(xs,ys). length xs < length ys | \
 | |
| 1021 | \ length xs = length ys & (xs,ys) : lex r}"; | |
| 5318 | 1022 | by (Simp_tac 1); | 
| 5281 | 1023 | qed "lexico_conv"; | 
| 1024 | ||
| 5283 | 1025 | Goal "([],ys) ~: lex r"; | 
| 5318 | 1026 | by (simp_tac (simpset() addsimps [lex_conv]) 1); | 
| 5283 | 1027 | qed "Nil_notin_lex"; | 
| 1028 | ||
| 1029 | Goal "(xs,[]) ~: lex r"; | |
| 5318 | 1030 | by (simp_tac (simpset() addsimps [lex_conv]) 1); | 
| 5283 | 1031 | qed "Nil2_notin_lex"; | 
| 1032 | ||
| 1033 | AddIffs [Nil_notin_lex,Nil2_notin_lex]; | |
| 1034 | ||
| 1035 | Goal "((x#xs,y#ys) : lex r) = \ | |
| 1036 | \ ((x,y) : r & length xs = length ys | x=y & (xs,ys) : lex r)"; | |
| 5318 | 1037 | by (simp_tac (simpset() addsimps [lex_conv]) 1); | 
| 1038 | by (rtac iffI 1); | |
| 1039 | by (blast_tac (claset() addIs [Cons_eq_appendI]) 2); | |
| 1040 | by (REPEAT(eresolve_tac [conjE, exE] 1)); | |
| 1041 | by (exhaust_tac "xys" 1); | |
| 1042 | by (Asm_full_simp_tac 1); | |
| 1043 | by (Asm_full_simp_tac 1); | |
| 1044 | by (Blast_tac 1); | |
| 5283 | 1045 | qed "Cons_in_lex"; | 
| 1046 | AddIffs [Cons_in_lex]; |