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