| author | kleing | 
| Tue, 15 Feb 2000 21:02:55 +0100 | |
| changeset 8247 | 635339ef2dca | 
| parent 8144 | c4b5cbfb90dd | 
| child 8254 | 84a5fe44520f | 
| 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  | 
|
| 6141 | 35  | 
val listsE = lists.mk_cases "x#l : lists A";  | 
| 3468 | 36  | 
AddSEs [listsE];  | 
37  | 
AddSIs lists.intrs;  | 
|
38  | 
||
| 5043 | 39  | 
Goal "l: lists A ==> l: lists B --> l: lists (A Int B)";  | 
| 3468 | 40  | 
by (etac lists.induct 1);  | 
41  | 
by (ALLGOALS Blast_tac);  | 
|
42  | 
qed_spec_mp "lists_IntI";  | 
|
43  | 
||
| 4935 | 44  | 
Goal "lists (A Int B) = lists A Int lists B";  | 
| 4423 | 45  | 
by (rtac (mono_Int RS equalityI) 1);  | 
| 4089 | 46  | 
by (simp_tac (simpset() addsimps [mono_def, lists_mono]) 1);  | 
47  | 
by (blast_tac (claset() addSIs [lists_IntI]) 1);  | 
|
| 3468 | 48  | 
qed "lists_Int_eq";  | 
49  | 
Addsimps [lists_Int_eq];  | 
|
50  | 
||
| 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  | 
||
| 7028 | 90  | 
Goal "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);  | 
|
| 
6813
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
116  | 
by Auto_tac;  | 
| 5296 | 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 =  | 
|
| 6394 | 254  | 
  Thm.read_cterm (Theory.sign_of List.thy) ("(xs::'a list) = ys",HOLogic.boolT);
 | 
| 5427 | 255  | 
|
| 7224 | 256  | 
fun last (cons as Const("List.list.Cons",_) $ _ $ xs) =
 | 
257  | 
      (case xs of Const("List.list.Nil",_) => cons | _ => last xs)
 | 
|
| 5427 | 258  | 
  | last (Const("List.op @",_) $ _ $ ys) = last ys
 | 
259  | 
| last t = t;  | 
|
260  | 
||
| 7224 | 261  | 
fun list1 (Const("List.list.Cons",_) $ _ $ Const("List.list.Nil",_)) = true
 | 
| 5427 | 262  | 
| list1 _ = false;  | 
263  | 
||
| 7224 | 264  | 
fun butlast ((cons as Const("List.list.Cons",_) $ x) $ xs) =
 | 
265  | 
      (case xs of Const("List.list.Nil",_) => xs | _ => cons $ butlast xs)
 | 
|
| 5427 | 266  | 
  | butlast ((app as Const("List.op @",_) $ xs) $ ys) = app $ butlast ys
 | 
| 7224 | 267  | 
  | butlast xs = Const("List.list.Nil",fastype_of xs);
 | 
| 5427 | 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: *)  | 
| 6451 | 336  | 
Goal "xs=ys ==> (!x. x : set ys --> f x = g x) --> map f xs = map g ys";  | 
| 4423 | 337  | 
by (hyp_subst_tac 1);  | 
338  | 
by (induct_tac "ys" 1);  | 
|
| 5316 | 339  | 
by Auto_tac;  | 
| 6451 | 340  | 
bind_thm("map_cong", impI RSN (2,allI RSN (2, result() RS mp)));
 | 
| 
3589
 
244daa75f890
Added function `replicate' and lemmas map_cong and set_replicate.
 
nipkow 
parents: 
3586 
diff
changeset
 | 
341  | 
|
| 4935 | 342  | 
Goal "(map f xs = []) = (xs = [])";  | 
| 8009 | 343  | 
by (exhaust_tac "xs" 1);  | 
| 5316 | 344  | 
by Auto_tac;  | 
| 3860 | 345  | 
qed "map_is_Nil_conv";  | 
346  | 
AddIffs [map_is_Nil_conv];  | 
|
347  | 
||
| 4935 | 348  | 
Goal "([] = map f xs) = (xs = [])";  | 
| 8009 | 349  | 
by (exhaust_tac "xs" 1);  | 
| 5316 | 350  | 
by Auto_tac;  | 
| 3860 | 351  | 
qed "Nil_is_map_conv";  | 
352  | 
AddIffs [Nil_is_map_conv];  | 
|
353  | 
||
| 8009 | 354  | 
Goal "(map f xs = y#ys) = (? x xs'. xs = x#xs' & f x = y & map f xs' = ys)";  | 
355  | 
by (exhaust_tac "xs" 1);  | 
|
356  | 
by (ALLGOALS Asm_simp_tac);  | 
|
357  | 
qed "map_eq_Cons";  | 
|
358  | 
||
359  | 
Goal "!xs. map f xs = map f ys --> (!x y. f x = f y --> x=y) --> xs=ys";  | 
|
360  | 
by (induct_tac "ys" 1);  | 
|
361  | 
by (Asm_simp_tac 1);  | 
|
362  | 
by (fast_tac (claset() addss (simpset() addsimps [map_eq_Cons])) 1);  | 
|
363  | 
qed_spec_mp "map_injective";  | 
|
364  | 
||
365  | 
Goal "inj f ==> inj (map f)";  | 
|
| 8064 | 366  | 
by (blast_tac (claset() addDs [map_injective,injD] addIs [injI]) 1);  | 
| 8009 | 367  | 
qed "inj_mapI";  | 
368  | 
||
369  | 
Goalw [inj_on_def] "inj (map f) ==> inj f";  | 
|
| 8064 | 370  | 
by (Clarify_tac 1);  | 
371  | 
by (eres_inst_tac [("x","[x]")] ballE 1);
 | 
|
372  | 
 by (eres_inst_tac [("x","[y]")] ballE 1);
 | 
|
373  | 
by (Asm_full_simp_tac 1);  | 
|
374  | 
by (Blast_tac 1);  | 
|
375  | 
by (Blast_tac 1);  | 
|
| 8009 | 376  | 
qed "inj_mapD";  | 
377  | 
||
378  | 
Goal "inj (map f) = inj f";  | 
|
| 8064 | 379  | 
by (blast_tac (claset() addDs [inj_mapD] addIs [inj_mapI]) 1);  | 
| 8009 | 380  | 
qed "inj_map";  | 
| 3860 | 381  | 
|
| 1169 | 382  | 
(** rev **)  | 
383  | 
||
| 3467 | 384  | 
section "rev";  | 
385  | 
||
| 4935 | 386  | 
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
 | 
387  | 
by (induct_tac "xs" 1);  | 
| 5316 | 388  | 
by Auto_tac;  | 
| 1169 | 389  | 
qed "rev_append";  | 
| 2512 | 390  | 
Addsimps[rev_append];  | 
| 1169 | 391  | 
|
| 4935 | 392  | 
Goal "rev(rev l) = l";  | 
| 
3040
 
7d48671753da
Introduced a generic "induct_tac" which picks up the right induction scheme
 
nipkow 
parents: 
3011 
diff
changeset
 | 
393  | 
by (induct_tac "l" 1);  | 
| 5316 | 394  | 
by Auto_tac;  | 
| 1169 | 395  | 
qed "rev_rev_ident";  | 
| 2512 | 396  | 
Addsimps[rev_rev_ident];  | 
| 1169 | 397  | 
|
| 4935 | 398  | 
Goal "(rev xs = []) = (xs = [])";  | 
| 4423 | 399  | 
by (induct_tac "xs" 1);  | 
| 5316 | 400  | 
by Auto_tac;  | 
| 3860 | 401  | 
qed "rev_is_Nil_conv";  | 
402  | 
AddIffs [rev_is_Nil_conv];  | 
|
403  | 
||
| 4935 | 404  | 
Goal "([] = rev xs) = (xs = [])";  | 
| 4423 | 405  | 
by (induct_tac "xs" 1);  | 
| 5316 | 406  | 
by Auto_tac;  | 
| 3860 | 407  | 
qed "Nil_is_rev_conv";  | 
408  | 
AddIffs [Nil_is_rev_conv];  | 
|
409  | 
||
| 6820 | 410  | 
Goal "!ys. (rev xs = rev ys) = (xs = ys)";  | 
| 6831 | 411  | 
by (induct_tac "xs" 1);  | 
| 6820 | 412  | 
by (Force_tac 1);  | 
| 6831 | 413  | 
by (rtac allI 1);  | 
414  | 
by (exhaust_tac "ys" 1);  | 
|
| 6820 | 415  | 
by (Asm_simp_tac 1);  | 
416  | 
by (Force_tac 1);  | 
|
417  | 
qed_spec_mp "rev_is_rev_conv";  | 
|
418  | 
AddIffs [rev_is_rev_conv];  | 
|
419  | 
||
| 4935 | 420  | 
val prems = Goal "[| P []; !!x xs. P xs ==> P(xs@[x]) |] ==> P xs";  | 
| 5132 | 421  | 
by (stac (rev_rev_ident RS sym) 1);  | 
| 6162 | 422  | 
by (res_inst_tac [("list", "rev xs")] list.induct 1);
 | 
| 5132 | 423  | 
by (ALLGOALS Simp_tac);  | 
424  | 
by (resolve_tac prems 1);  | 
|
425  | 
by (eresolve_tac prems 1);  | 
|
| 4935 | 426  | 
qed "rev_induct";  | 
427  | 
||
| 5272 | 428  | 
fun rev_induct_tac xs = res_inst_tac [("xs",xs)] rev_induct;
 | 
429  | 
||
| 4935 | 430  | 
Goal "(xs = [] --> P) --> (!ys y. xs = ys@[y] --> P) --> P";  | 
| 5132 | 431  | 
by (res_inst_tac [("xs","xs")] rev_induct 1);
 | 
| 5316 | 432  | 
by Auto_tac;  | 
| 4935 | 433  | 
bind_thm ("rev_exhaust",
 | 
434  | 
impI RSN (2,allI RSN (2,allI RSN (2,impI RS (result() RS mp RS mp)))));  | 
|
435  | 
||
| 2608 | 436  | 
|
| 3465 | 437  | 
(** set **)  | 
| 1812 | 438  | 
|
| 3467 | 439  | 
section "set";  | 
440  | 
||
| 7032 | 441  | 
Goal "finite (set xs)";  | 
442  | 
by (induct_tac "xs" 1);  | 
|
443  | 
by Auto_tac;  | 
|
444  | 
qed "finite_set";  | 
|
445  | 
AddIffs [finite_set];  | 
|
| 5296 | 446  | 
|
| 4935 | 447  | 
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
 | 
448  | 
by (induct_tac "xs" 1);  | 
| 5316 | 449  | 
by Auto_tac;  | 
| 
3647
 
a64c8fbcd98f
Renamed theorems of the form set_of_list_XXX to set_XXX
 
paulson 
parents: 
3589 
diff
changeset
 | 
450  | 
qed "set_append";  | 
| 
 
a64c8fbcd98f
Renamed theorems of the form set_of_list_XXX to set_XXX
 
paulson 
parents: 
3589 
diff
changeset
 | 
451  | 
Addsimps[set_append];  | 
| 1812 | 452  | 
|
| 4935 | 453  | 
Goal "set l <= set (x#l)";  | 
| 5316 | 454  | 
by Auto_tac;  | 
| 
3647
 
a64c8fbcd98f
Renamed theorems of the form set_of_list_XXX to set_XXX
 
paulson 
parents: 
3589 
diff
changeset
 | 
455  | 
qed "set_subset_Cons";  | 
| 1936 | 456  | 
|
| 4935 | 457  | 
Goal "(set xs = {}) = (xs = [])";
 | 
| 3457 | 458  | 
by (induct_tac "xs" 1);  | 
| 5316 | 459  | 
by Auto_tac;  | 
| 
3647
 
a64c8fbcd98f
Renamed theorems of the form set_of_list_XXX to set_XXX
 
paulson 
parents: 
3589 
diff
changeset
 | 
460  | 
qed "set_empty";  | 
| 
 
a64c8fbcd98f
Renamed theorems of the form set_of_list_XXX to set_XXX
 
paulson 
parents: 
3589 
diff
changeset
 | 
461  | 
Addsimps [set_empty];  | 
| 2608 | 462  | 
|
| 4935 | 463  | 
Goal "set(rev xs) = set(xs)";  | 
| 3457 | 464  | 
by (induct_tac "xs" 1);  | 
| 5316 | 465  | 
by Auto_tac;  | 
| 
3647
 
a64c8fbcd98f
Renamed theorems of the form set_of_list_XXX to set_XXX
 
paulson 
parents: 
3589 
diff
changeset
 | 
466  | 
qed "set_rev";  | 
| 
 
a64c8fbcd98f
Renamed theorems of the form set_of_list_XXX to set_XXX
 
paulson 
parents: 
3589 
diff
changeset
 | 
467  | 
Addsimps [set_rev];  | 
| 2608 | 468  | 
|
| 4935 | 469  | 
Goal "set(map f xs) = f``(set xs)";  | 
| 3457 | 470  | 
by (induct_tac "xs" 1);  | 
| 5316 | 471  | 
by Auto_tac;  | 
| 
3647
 
a64c8fbcd98f
Renamed theorems of the form set_of_list_XXX to set_XXX
 
paulson 
parents: 
3589 
diff
changeset
 | 
472  | 
qed "set_map";  | 
| 
 
a64c8fbcd98f
Renamed theorems of the form set_of_list_XXX to set_XXX
 
paulson 
parents: 
3589 
diff
changeset
 | 
473  | 
Addsimps [set_map];  | 
| 2608 | 474  | 
|
| 6433 | 475  | 
Goal "set(filter P xs) = {x. x : set xs & P x}";
 | 
| 
6813
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
476  | 
by (induct_tac "xs" 1);  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
477  | 
by Auto_tac;  | 
| 6433 | 478  | 
qed "set_filter";  | 
479  | 
Addsimps [set_filter];  | 
|
| 8009 | 480  | 
|
| 6433 | 481  | 
Goal "set[i..j(] = {k. i <= k & k < j}";
 | 
| 
6813
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
482  | 
by (induct_tac "j" 1);  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
483  | 
by Auto_tac;  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
484  | 
by (arith_tac 1);  | 
| 6433 | 485  | 
qed "set_upt";  | 
486  | 
Addsimps [set_upt];  | 
|
487  | 
||
488  | 
Goal "!i < size xs. set(xs[i := x]) <= insert x (set xs)";  | 
|
| 
6813
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
489  | 
by (induct_tac "xs" 1);  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
490  | 
by (Simp_tac 1);  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
491  | 
by (asm_simp_tac (simpset() addsplits [nat.split]) 1);  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
492  | 
by (Blast_tac 1);  | 
| 6433 | 493  | 
qed_spec_mp "set_list_update_subset";  | 
| 4605 | 494  | 
|
| 5272 | 495  | 
Goal "(x : set xs) = (? ys zs. xs = ys@x#zs)";  | 
| 5318 | 496  | 
by (induct_tac "xs" 1);  | 
497  | 
by (Simp_tac 1);  | 
|
498  | 
by (Asm_simp_tac 1);  | 
|
499  | 
by (rtac iffI 1);  | 
|
500  | 
by (blast_tac (claset() addIs [eq_Nil_appendI,Cons_eq_appendI]) 1);  | 
|
501  | 
by (REPEAT(etac exE 1));  | 
|
502  | 
by (exhaust_tac "ys" 1);  | 
|
| 5316 | 503  | 
by Auto_tac;  | 
| 5272 | 504  | 
qed "in_set_conv_decomp";  | 
505  | 
||
| 8009 | 506  | 
|
| 5272 | 507  | 
(* eliminate `lists' in favour of `set' *)  | 
508  | 
||
509  | 
Goal "(xs : lists A) = (!x : set xs. x : A)";  | 
|
| 5318 | 510  | 
by (induct_tac "xs" 1);  | 
| 5316 | 511  | 
by Auto_tac;  | 
| 5272 | 512  | 
qed "in_lists_conv_set";  | 
513  | 
||
514  | 
bind_thm("in_listsD",in_lists_conv_set RS iffD1);
 | 
|
515  | 
AddSDs [in_listsD];  | 
|
516  | 
bind_thm("in_listsI",in_lists_conv_set RS iffD2);
 | 
|
517  | 
AddSIs [in_listsI];  | 
|
| 1812 | 518  | 
|
| 5518 | 519  | 
(** mem **)  | 
520  | 
||
521  | 
section "mem";  | 
|
522  | 
||
523  | 
Goal "(x mem xs) = (x: set xs)";  | 
|
524  | 
by (induct_tac "xs" 1);  | 
|
525  | 
by Auto_tac;  | 
|
526  | 
qed "set_mem_eq";  | 
|
527  | 
||
528  | 
||
| 923 | 529  | 
(** list_all **)  | 
530  | 
||
| 3467 | 531  | 
section "list_all";  | 
532  | 
||
| 5518 | 533  | 
Goal "list_all P xs = (!x:set xs. P x)";  | 
534  | 
by (induct_tac "xs" 1);  | 
|
535  | 
by Auto_tac;  | 
|
536  | 
qed "list_all_conv";  | 
|
537  | 
||
| 
5443
 
e2459d18ff47
changed constants mem and list_all to mere translations
 
oheimb 
parents: 
5427 
diff
changeset
 | 
538  | 
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
 | 
539  | 
by (induct_tac "xs" 1);  | 
| 5316 | 540  | 
by Auto_tac;  | 
| 2512 | 541  | 
qed "list_all_append";  | 
542  | 
Addsimps [list_all_append];  | 
|
| 923 | 543  | 
|
544  | 
||
| 2608 | 545  | 
(** filter **)  | 
| 923 | 546  | 
|
| 3467 | 547  | 
section "filter";  | 
548  | 
||
| 4935 | 549  | 
Goal "filter P (xs@ys) = filter P xs @ filter P ys";  | 
| 3457 | 550  | 
by (induct_tac "xs" 1);  | 
| 5316 | 551  | 
by Auto_tac;  | 
| 2608 | 552  | 
qed "filter_append";  | 
553  | 
Addsimps [filter_append];  | 
|
554  | 
||
| 4935 | 555  | 
Goal "filter (%x. True) xs = xs";  | 
| 4605 | 556  | 
by (induct_tac "xs" 1);  | 
| 5316 | 557  | 
by Auto_tac;  | 
| 4605 | 558  | 
qed "filter_True";  | 
559  | 
Addsimps [filter_True];  | 
|
560  | 
||
| 4935 | 561  | 
Goal "filter (%x. False) xs = []";  | 
| 4605 | 562  | 
by (induct_tac "xs" 1);  | 
| 5316 | 563  | 
by Auto_tac;  | 
| 4605 | 564  | 
qed "filter_False";  | 
565  | 
Addsimps [filter_False];  | 
|
566  | 
||
| 4935 | 567  | 
Goal "length (filter P xs) <= length xs";  | 
| 3457 | 568  | 
by (induct_tac "xs" 1);  | 
| 5316 | 569  | 
by Auto_tac;  | 
| 4605 | 570  | 
qed "length_filter";  | 
| 
5443
 
e2459d18ff47
changed constants mem and list_all to mere translations
 
oheimb 
parents: 
5427 
diff
changeset
 | 
571  | 
Addsimps[length_filter];  | 
| 2608 | 572  | 
|
| 
5443
 
e2459d18ff47
changed constants mem and list_all to mere translations
 
oheimb 
parents: 
5427 
diff
changeset
 | 
573  | 
Goal "set (filter P xs) <= set xs";  | 
| 
 
e2459d18ff47
changed constants mem and list_all to mere translations
 
oheimb 
parents: 
5427 
diff
changeset
 | 
574  | 
by Auto_tac;  | 
| 
 
e2459d18ff47
changed constants mem and list_all to mere translations
 
oheimb 
parents: 
5427 
diff
changeset
 | 
575  | 
qed "filter_is_subset";  | 
| 
 
e2459d18ff47
changed constants mem and list_all to mere translations
 
oheimb 
parents: 
5427 
diff
changeset
 | 
576  | 
Addsimps [filter_is_subset];  | 
| 
 
e2459d18ff47
changed constants mem and list_all to mere translations
 
oheimb 
parents: 
5427 
diff
changeset
 | 
577  | 
|
| 2608 | 578  | 
|
| 3467 | 579  | 
section "concat";  | 
580  | 
||
| 4935 | 581  | 
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
 | 
582  | 
by (induct_tac "xs" 1);  | 
| 5316 | 583  | 
by Auto_tac;  | 
| 2608 | 584  | 
qed"concat_append";  | 
585  | 
Addsimps [concat_append];  | 
|
| 2512 | 586  | 
|
| 4935 | 587  | 
Goal "(concat xss = []) = (!xs:set xss. xs=[])";  | 
| 4423 | 588  | 
by (induct_tac "xss" 1);  | 
| 5316 | 589  | 
by Auto_tac;  | 
| 
3896
 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 
nipkow 
parents: 
3860 
diff
changeset
 | 
590  | 
qed "concat_eq_Nil_conv";  | 
| 
 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 
nipkow 
parents: 
3860 
diff
changeset
 | 
591  | 
AddIffs [concat_eq_Nil_conv];  | 
| 
 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 
nipkow 
parents: 
3860 
diff
changeset
 | 
592  | 
|
| 4935 | 593  | 
Goal "([] = concat xss) = (!xs:set xss. xs=[])";  | 
| 4423 | 594  | 
by (induct_tac "xss" 1);  | 
| 5316 | 595  | 
by Auto_tac;  | 
| 
3896
 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 
nipkow 
parents: 
3860 
diff
changeset
 | 
596  | 
qed "Nil_eq_concat_conv";  | 
| 
 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 
nipkow 
parents: 
3860 
diff
changeset
 | 
597  | 
AddIffs [Nil_eq_concat_conv];  | 
| 
 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 
nipkow 
parents: 
3860 
diff
changeset
 | 
598  | 
|
| 4935 | 599  | 
Goal "set(concat xs) = Union(set `` set xs)";  | 
| 3467 | 600  | 
by (induct_tac "xs" 1);  | 
| 5316 | 601  | 
by Auto_tac;  | 
| 
3647
 
a64c8fbcd98f
Renamed theorems of the form set_of_list_XXX to set_XXX
 
paulson 
parents: 
3589 
diff
changeset
 | 
602  | 
qed"set_concat";  | 
| 
 
a64c8fbcd98f
Renamed theorems of the form set_of_list_XXX to set_XXX
 
paulson 
parents: 
3589 
diff
changeset
 | 
603  | 
Addsimps [set_concat];  | 
| 3467 | 604  | 
|
| 4935 | 605  | 
Goal "map f (concat xs) = concat (map (map f) xs)";  | 
| 3467 | 606  | 
by (induct_tac "xs" 1);  | 
| 5316 | 607  | 
by Auto_tac;  | 
| 3467 | 608  | 
qed "map_concat";  | 
609  | 
||
| 4935 | 610  | 
Goal "filter p (concat xs) = concat (map (filter p) xs)";  | 
| 3467 | 611  | 
by (induct_tac "xs" 1);  | 
| 5316 | 612  | 
by Auto_tac;  | 
| 3467 | 613  | 
qed"filter_concat";  | 
614  | 
||
| 4935 | 615  | 
Goal "rev(concat xs) = concat (map rev (rev xs))";  | 
| 3467 | 616  | 
by (induct_tac "xs" 1);  | 
| 5316 | 617  | 
by Auto_tac;  | 
| 2608 | 618  | 
qed "rev_concat";  | 
| 923 | 619  | 
|
620  | 
(** nth **)  | 
|
621  | 
||
| 3467 | 622  | 
section "nth";  | 
623  | 
||
| 6408 | 624  | 
Goal "(x#xs)!0 = x";  | 
625  | 
by Auto_tac;  | 
|
626  | 
qed "nth_Cons_0";  | 
|
627  | 
Addsimps [nth_Cons_0];  | 
|
| 5644 | 628  | 
|
| 6408 | 629  | 
Goal "(x#xs)!(Suc n) = xs!n";  | 
630  | 
by Auto_tac;  | 
|
631  | 
qed "nth_Cons_Suc";  | 
|
632  | 
Addsimps [nth_Cons_Suc];  | 
|
633  | 
||
634  | 
Delsimps (thms "nth.simps");  | 
|
635  | 
||
636  | 
Goal "!n. (xs@ys)!n = (if n < length xs then xs!n else ys!(n - length xs))";  | 
|
637  | 
by (induct_tac "xs" 1);  | 
|
| 3457 | 638  | 
by (Asm_simp_tac 1);  | 
639  | 
by (rtac allI 1);  | 
|
| 6408 | 640  | 
by (exhaust_tac "n" 1);  | 
| 5316 | 641  | 
by Auto_tac;  | 
| 2608 | 642  | 
qed_spec_mp "nth_append";  | 
643  | 
||
| 4935 | 644  | 
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
 | 
645  | 
by (induct_tac "xs" 1);  | 
| 8118 | 646  | 
by (Asm_full_simp_tac 1);  | 
| 1301 | 647  | 
by (rtac allI 1);  | 
| 5183 | 648  | 
by (induct_tac "n" 1);  | 
| 5316 | 649  | 
by Auto_tac;  | 
| 
1485
 
240cc98b94a7
Added qed_spec_mp to avoid renaming of bound vars in 'th RS spec'
 
nipkow 
parents: 
1465 
diff
changeset
 | 
650  | 
qed_spec_mp "nth_map";  | 
| 1301 | 651  | 
Addsimps [nth_map];  | 
652  | 
||
| 8118 | 653  | 
Goal "set xs = {xs!i |i. i < length xs}";
 | 
| 
3040
 
7d48671753da
Introduced a generic "induct_tac" which picks up the right induction scheme
 
nipkow 
parents: 
3011 
diff
changeset
 | 
654  | 
by (induct_tac "xs" 1);  | 
| 8118 | 655  | 
by (Simp_tac 1);  | 
656  | 
by(Asm_simp_tac 1);  | 
|
657  | 
by(Safe_tac);  | 
|
658  | 
  by(res_inst_tac [("x","0")] exI 1);
 | 
|
659  | 
by (Simp_tac 1);  | 
|
660  | 
 by(res_inst_tac [("x","Suc i")] exI 1);
 | 
|
661  | 
by(Asm_simp_tac 1);  | 
|
662  | 
by(exhaust_tac "i" 1);  | 
|
663  | 
by(Asm_full_simp_tac 1);  | 
|
664  | 
by(rename_tac "j" 1);  | 
|
665  | 
 by(res_inst_tac [("x","j")] exI 1);
 | 
|
666  | 
by(Asm_simp_tac 1);  | 
|
667  | 
qed "set_conv_nth";  | 
|
668  | 
||
669  | 
Goal "n < length xs ==> Ball (set xs) P --> P(xs!n)";  | 
|
670  | 
by (simp_tac (simpset() addsimps [set_conv_nth]) 1);  | 
|
671  | 
by(Blast_tac 1);  | 
|
| 5518 | 672  | 
qed_spec_mp "list_ball_nth";  | 
| 1301 | 673  | 
|
| 8118 | 674  | 
Goal "n < length xs ==> xs!n : set xs";  | 
675  | 
by (simp_tac (simpset() addsimps [set_conv_nth]) 1);  | 
|
676  | 
by(Blast_tac 1);  | 
|
| 
1485
 
240cc98b94a7
Added qed_spec_mp to avoid renaming of bound vars in 'th RS spec'
 
nipkow 
parents: 
1465 
diff
changeset
 | 
677  | 
qed_spec_mp "nth_mem";  | 
| 1301 | 678  | 
Addsimps [nth_mem];  | 
679  | 
||
| 8009 | 680  | 
Goal "(!i. i < length xs --> P(xs!i)) --> (!x : set xs. P x)";  | 
| 8118 | 681  | 
by (simp_tac (simpset() addsimps [set_conv_nth]) 1);  | 
682  | 
by(Blast_tac 1);  | 
|
| 8009 | 683  | 
qed_spec_mp "all_nth_imp_all_set";  | 
684  | 
||
685  | 
Goal "(!x : set xs. P x) = (!i. i<length xs --> P (xs ! i))";  | 
|
| 8118 | 686  | 
by (simp_tac (simpset() addsimps [set_conv_nth]) 1);  | 
687  | 
by(Blast_tac 1);  | 
|
| 8009 | 688  | 
qed_spec_mp "all_set_conv_all_nth";  | 
689  | 
||
690  | 
||
| 
5077
 
71043526295f
* HOL/List: new function list_update written xs[i:=v] that updates the i-th
 
nipkow 
parents: 
5043 
diff
changeset
 | 
691  | 
(** list update **)  | 
| 
 
71043526295f
* HOL/List: new function list_update written xs[i:=v] that updates the i-th
 
nipkow 
parents: 
5043 
diff
changeset
 | 
692  | 
|
| 
 
71043526295f
* HOL/List: new function list_update written xs[i:=v] that updates the i-th
 
nipkow 
parents: 
5043 
diff
changeset
 | 
693  | 
section "list update";  | 
| 
 
71043526295f
* HOL/List: new function list_update written xs[i:=v] that updates the i-th
 
nipkow 
parents: 
5043 
diff
changeset
 | 
694  | 
|
| 
 
71043526295f
* HOL/List: new function list_update written xs[i:=v] that updates the i-th
 
nipkow 
parents: 
5043 
diff
changeset
 | 
695  | 
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
 | 
696  | 
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
 | 
697  | 
by (Simp_tac 1);  | 
| 5183 | 698  | 
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
 | 
699  | 
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
 | 
700  | 
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
 | 
701  | 
|
| 5644 | 702  | 
Goal "!i j. i < length xs --> (xs[i:=x])!j = (if i=j then x else xs!j)";  | 
| 6162 | 703  | 
by (induct_tac "xs" 1);  | 
704  | 
by (Simp_tac 1);  | 
|
705  | 
by (auto_tac (claset(), simpset() addsimps [nth_Cons] addsplits [nat.split]));  | 
|
| 5644 | 706  | 
qed_spec_mp "nth_list_update";  | 
707  | 
||
| 8144 | 708  | 
Goal "i < length xs ==> (xs[i:=x])!i = x";  | 
709  | 
by (asm_simp_tac (simpset() addsimps [nth_list_update]) 1);  | 
|
710  | 
qed "nth_list_update_eq";  | 
|
711  | 
Addsimps [nth_list_update_eq];  | 
|
712  | 
||
713  | 
Goal "!i j. i ~= j --> xs[i:=x]!j = xs!j";  | 
|
714  | 
by (induct_tac "xs" 1);  | 
|
715  | 
by (Simp_tac 1);  | 
|
716  | 
by (auto_tac (claset(), simpset() addsimps [nth_Cons] addsplits [nat.split]));  | 
|
717  | 
qed_spec_mp "nth_list_update_neq";  | 
|
718  | 
Addsimps [nth_list_update_neq];  | 
|
719  | 
||
| 6433 | 720  | 
Goal "!i. i < size xs --> xs[i:=x, i:=y] = xs[i:=y]";  | 
| 
6813
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
721  | 
by (induct_tac "xs" 1);  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
722  | 
by (Simp_tac 1);  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
723  | 
by (asm_simp_tac (simpset() addsplits [nat.split]) 1);  | 
| 6433 | 724  | 
qed_spec_mp "list_update_overwrite";  | 
725  | 
Addsimps [list_update_overwrite];  | 
|
726  | 
||
727  | 
Goal "!i < length xs. (xs[i := x] = xs) = (xs!i = x)";  | 
|
| 
6813
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
728  | 
by (induct_tac "xs" 1);  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
729  | 
by (Simp_tac 1);  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
730  | 
by (simp_tac (simpset() addsplits [nat.split]) 1);  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
731  | 
by (Blast_tac 1);  | 
| 6433 | 732  | 
qed_spec_mp "list_update_same_conv";  | 
733  | 
||
| 8009 | 734  | 
Goal "!i xy xs. length xs = length ys --> \  | 
735  | 
\ (zip xs ys)[i:=xy] = zip (xs[i:=fst xy]) (ys[i:=snd xy])";  | 
|
736  | 
by (induct_tac "ys" 1);  | 
|
737  | 
by Auto_tac;  | 
|
738  | 
by (exhaust_tac "xs" 1);  | 
|
739  | 
by (auto_tac (claset(), simpset() addsplits [nat.split]));  | 
|
740  | 
qed_spec_mp "update_zip";  | 
|
741  | 
||
742  | 
Goal "!i. set(xs[i:=x]) <= insert x (set xs)";  | 
|
743  | 
by (induct_tac "xs" 1);  | 
|
744  | 
by (asm_full_simp_tac (simpset() addsimps []) 1);  | 
|
745  | 
by (asm_full_simp_tac (simpset() addsplits [nat.split]) 1);  | 
|
746  | 
by (Fast_tac 1);  | 
|
747  | 
qed_spec_mp "set_update_subset";  | 
|
748  | 
||
| 
5077
 
71043526295f
* HOL/List: new function list_update written xs[i:=v] that updates the i-th
 
nipkow 
parents: 
5043 
diff
changeset
 | 
749  | 
|
| 
3896
 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 
nipkow 
parents: 
3860 
diff
changeset
 | 
750  | 
(** last & butlast **)  | 
| 
1327
 
6c29cfab679c
added new arithmetic lemmas and the functions take and drop.
 
nipkow 
parents: 
1301 
diff
changeset
 | 
751  | 
|
| 5644 | 752  | 
section "last / butlast";  | 
753  | 
||
| 4935 | 754  | 
Goal "last(xs@[x]) = x";  | 
| 4423 | 755  | 
by (induct_tac "xs" 1);  | 
| 5316 | 756  | 
by Auto_tac;  | 
| 
3896
 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 
nipkow 
parents: 
3860 
diff
changeset
 | 
757  | 
qed "last_snoc";  | 
| 
 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 
nipkow 
parents: 
3860 
diff
changeset
 | 
758  | 
Addsimps [last_snoc];  | 
| 
 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 
nipkow 
parents: 
3860 
diff
changeset
 | 
759  | 
|
| 4935 | 760  | 
Goal "butlast(xs@[x]) = xs";  | 
| 4423 | 761  | 
by (induct_tac "xs" 1);  | 
| 5316 | 762  | 
by Auto_tac;  | 
| 
3896
 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 
nipkow 
parents: 
3860 
diff
changeset
 | 
763  | 
qed "butlast_snoc";  | 
| 
 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 
nipkow 
parents: 
3860 
diff
changeset
 | 
764  | 
Addsimps [butlast_snoc];  | 
| 
 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 
nipkow 
parents: 
3860 
diff
changeset
 | 
765  | 
|
| 4935 | 766  | 
Goal "length(butlast xs) = length xs - 1";  | 
767  | 
by (res_inst_tac [("xs","xs")] rev_induct 1);
 | 
|
| 5316 | 768  | 
by Auto_tac;  | 
| 4643 | 769  | 
qed "length_butlast";  | 
770  | 
Addsimps [length_butlast];  | 
|
771  | 
||
| 5278 | 772  | 
Goal "!ys. butlast (xs@ys) = (if ys=[] then butlast xs else xs@butlast ys)";  | 
| 4423 | 773  | 
by (induct_tac "xs" 1);  | 
| 5316 | 774  | 
by Auto_tac;  | 
| 
3896
 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 
nipkow 
parents: 
3860 
diff
changeset
 | 
775  | 
qed_spec_mp "butlast_append";  | 
| 
 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 
nipkow 
parents: 
3860 
diff
changeset
 | 
776  | 
|
| 8118 | 777  | 
Goal "xs ~= [] --> butlast xs @ [last xs] = xs";  | 
778  | 
by(induct_tac "xs" 1);  | 
|
779  | 
by(ALLGOALS Asm_simp_tac);  | 
|
780  | 
qed_spec_mp "append_butlast_last_id";  | 
|
781  | 
Addsimps [append_butlast_last_id];  | 
|
782  | 
||
| 4935 | 783  | 
Goal "x:set(butlast xs) --> x:set xs";  | 
| 4423 | 784  | 
by (induct_tac "xs" 1);  | 
| 5316 | 785  | 
by Auto_tac;  | 
| 
3896
 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 
nipkow 
parents: 
3860 
diff
changeset
 | 
786  | 
qed_spec_mp "in_set_butlastD";  | 
| 
 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 
nipkow 
parents: 
3860 
diff
changeset
 | 
787  | 
|
| 
5448
 
40a09282ba14
in_set_butlast_appendI supersedes in_set_butlast_appendI1,2
 
paulson 
parents: 
5443 
diff
changeset
 | 
788  | 
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
 | 
789  | 
by (auto_tac (claset() addDs [in_set_butlastD],  | 
| 
 
40a09282ba14
in_set_butlast_appendI supersedes in_set_butlast_appendI1,2
 
paulson 
parents: 
5443 
diff
changeset
 | 
790  | 
simpset() addsimps [butlast_append]));  | 
| 
 
40a09282ba14
in_set_butlast_appendI supersedes in_set_butlast_appendI1,2
 
paulson 
parents: 
5443 
diff
changeset
 | 
791  | 
qed "in_set_butlast_appendI";  | 
| 3902 | 792  | 
|
| 2608 | 793  | 
(** take & drop **)  | 
794  | 
section "take & drop";  | 
|
| 
1327
 
6c29cfab679c
added new arithmetic lemmas and the functions take and drop.
 
nipkow 
parents: 
1301 
diff
changeset
 | 
795  | 
|
| 4935 | 796  | 
Goal "take 0 xs = []";  | 
| 
3040
 
7d48671753da
Introduced a generic "induct_tac" which picks up the right induction scheme
 
nipkow 
parents: 
3011 
diff
changeset
 | 
797  | 
by (induct_tac "xs" 1);  | 
| 5316 | 798  | 
by Auto_tac;  | 
| 
1327
 
6c29cfab679c
added new arithmetic lemmas and the functions take and drop.
 
nipkow 
parents: 
1301 
diff
changeset
 | 
799  | 
qed "take_0";  | 
| 
 
6c29cfab679c
added new arithmetic lemmas and the functions take and drop.
 
nipkow 
parents: 
1301 
diff
changeset
 | 
800  | 
|
| 4935 | 801  | 
Goal "drop 0 xs = xs";  | 
| 
3040
 
7d48671753da
Introduced a generic "induct_tac" which picks up the right induction scheme
 
nipkow 
parents: 
3011 
diff
changeset
 | 
802  | 
by (induct_tac "xs" 1);  | 
| 5316 | 803  | 
by Auto_tac;  | 
| 2608 | 804  | 
qed "drop_0";  | 
805  | 
||
| 4935 | 806  | 
Goal "take (Suc n) (x#xs) = x # take n xs";  | 
| 1552 | 807  | 
by (Simp_tac 1);  | 
| 
1419
 
a6a034a47a71
defined take/drop by induction over list rather than nat.
 
nipkow 
parents: 
1327 
diff
changeset
 | 
808  | 
qed "take_Suc_Cons";  | 
| 
1327
 
6c29cfab679c
added new arithmetic lemmas and the functions take and drop.
 
nipkow 
parents: 
1301 
diff
changeset
 | 
809  | 
|
| 4935 | 810  | 
Goal "drop (Suc n) (x#xs) = drop n xs";  | 
| 2608 | 811  | 
by (Simp_tac 1);  | 
812  | 
qed "drop_Suc_Cons";  | 
|
813  | 
||
814  | 
Delsimps [take_Cons,drop_Cons];  | 
|
815  | 
Addsimps [take_0,take_Suc_Cons,drop_0,drop_Suc_Cons];  | 
|
816  | 
||
| 4935 | 817  | 
Goal "!xs. length(take n xs) = min (length xs) n";  | 
| 5183 | 818  | 
by (induct_tac "n" 1);  | 
| 5316 | 819  | 
by Auto_tac;  | 
| 3457 | 820  | 
by (exhaust_tac "xs" 1);  | 
| 5316 | 821  | 
by Auto_tac;  | 
| 2608 | 822  | 
qed_spec_mp "length_take";  | 
823  | 
Addsimps [length_take];  | 
|
| 923 | 824  | 
|
| 4935 | 825  | 
Goal "!xs. length(drop n xs) = (length xs - n)";  | 
| 5183 | 826  | 
by (induct_tac "n" 1);  | 
| 5316 | 827  | 
by Auto_tac;  | 
| 3457 | 828  | 
by (exhaust_tac "xs" 1);  | 
| 5316 | 829  | 
by Auto_tac;  | 
| 2608 | 830  | 
qed_spec_mp "length_drop";  | 
831  | 
Addsimps [length_drop];  | 
|
832  | 
||
| 4935 | 833  | 
Goal "!xs. length xs <= n --> take n xs = xs";  | 
| 5183 | 834  | 
by (induct_tac "n" 1);  | 
| 5316 | 835  | 
by Auto_tac;  | 
| 3457 | 836  | 
by (exhaust_tac "xs" 1);  | 
| 5316 | 837  | 
by Auto_tac;  | 
| 2608 | 838  | 
qed_spec_mp "take_all";  | 
| 7246 | 839  | 
Addsimps [take_all];  | 
| 923 | 840  | 
|
| 4935 | 841  | 
Goal "!xs. length xs <= n --> drop n xs = []";  | 
| 5183 | 842  | 
by (induct_tac "n" 1);  | 
| 5316 | 843  | 
by Auto_tac;  | 
| 3457 | 844  | 
by (exhaust_tac "xs" 1);  | 
| 5316 | 845  | 
by Auto_tac;  | 
| 2608 | 846  | 
qed_spec_mp "drop_all";  | 
| 7246 | 847  | 
Addsimps [drop_all];  | 
| 2608 | 848  | 
|
| 5278 | 849  | 
Goal "!xs. take n (xs @ ys) = (take n xs @ take (n - length xs) ys)";  | 
| 5183 | 850  | 
by (induct_tac "n" 1);  | 
| 5316 | 851  | 
by Auto_tac;  | 
| 3457 | 852  | 
by (exhaust_tac "xs" 1);  | 
| 5316 | 853  | 
by Auto_tac;  | 
| 2608 | 854  | 
qed_spec_mp "take_append";  | 
855  | 
Addsimps [take_append];  | 
|
856  | 
||
| 4935 | 857  | 
Goal "!xs. drop n (xs@ys) = drop n xs @ drop (n - length xs) ys";  | 
| 5183 | 858  | 
by (induct_tac "n" 1);  | 
| 5316 | 859  | 
by Auto_tac;  | 
| 3457 | 860  | 
by (exhaust_tac "xs" 1);  | 
| 5316 | 861  | 
by Auto_tac;  | 
| 2608 | 862  | 
qed_spec_mp "drop_append";  | 
863  | 
Addsimps [drop_append];  | 
|
864  | 
||
| 4935 | 865  | 
Goal "!xs n. take n (take m xs) = take (min n m) xs";  | 
| 5183 | 866  | 
by (induct_tac "m" 1);  | 
| 5316 | 867  | 
by Auto_tac;  | 
| 3457 | 868  | 
by (exhaust_tac "xs" 1);  | 
| 5316 | 869  | 
by Auto_tac;  | 
| 5183 | 870  | 
by (exhaust_tac "na" 1);  | 
| 5316 | 871  | 
by Auto_tac;  | 
| 2608 | 872  | 
qed_spec_mp "take_take";  | 
| 7570 | 873  | 
Addsimps [take_take];  | 
| 2608 | 874  | 
|
| 4935 | 875  | 
Goal "!xs. drop n (drop m xs) = drop (n + m) xs";  | 
| 5183 | 876  | 
by (induct_tac "m" 1);  | 
| 5316 | 877  | 
by Auto_tac;  | 
| 3457 | 878  | 
by (exhaust_tac "xs" 1);  | 
| 5316 | 879  | 
by Auto_tac;  | 
| 2608 | 880  | 
qed_spec_mp "drop_drop";  | 
| 7570 | 881  | 
Addsimps [drop_drop];  | 
| 923 | 882  | 
|
| 4935 | 883  | 
Goal "!xs n. take n (drop m xs) = drop m (take (n + m) xs)";  | 
| 5183 | 884  | 
by (induct_tac "m" 1);  | 
| 5316 | 885  | 
by Auto_tac;  | 
| 3457 | 886  | 
by (exhaust_tac "xs" 1);  | 
| 5316 | 887  | 
by Auto_tac;  | 
| 2608 | 888  | 
qed_spec_mp "take_drop";  | 
889  | 
||
| 
6813
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
890  | 
Goal "!xs. take n xs @ drop n xs = xs";  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
891  | 
by (induct_tac "n" 1);  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
892  | 
by Auto_tac;  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
893  | 
by (exhaust_tac "xs" 1);  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
894  | 
by Auto_tac;  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
895  | 
qed_spec_mp "append_take_drop_id";  | 
| 8118 | 896  | 
Addsimps [append_take_drop_id];  | 
| 
6813
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
897  | 
|
| 4935 | 898  | 
Goal "!xs. take n (map f xs) = map f (take n xs)";  | 
| 5183 | 899  | 
by (induct_tac "n" 1);  | 
| 5316 | 900  | 
by Auto_tac;  | 
| 3457 | 901  | 
by (exhaust_tac "xs" 1);  | 
| 5316 | 902  | 
by Auto_tac;  | 
| 2608 | 903  | 
qed_spec_mp "take_map";  | 
904  | 
||
| 4935 | 905  | 
Goal "!xs. drop n (map f xs) = map f (drop n xs)";  | 
| 5183 | 906  | 
by (induct_tac "n" 1);  | 
| 5316 | 907  | 
by Auto_tac;  | 
| 3457 | 908  | 
by (exhaust_tac "xs" 1);  | 
| 5316 | 909  | 
by Auto_tac;  | 
| 2608 | 910  | 
qed_spec_mp "drop_map";  | 
911  | 
||
| 4935 | 912  | 
Goal "!n i. i < n --> (take n xs)!i = xs!i";  | 
| 3457 | 913  | 
by (induct_tac "xs" 1);  | 
| 5316 | 914  | 
by Auto_tac;  | 
| 3457 | 915  | 
by (exhaust_tac "n" 1);  | 
916  | 
by (Blast_tac 1);  | 
|
917  | 
by (exhaust_tac "i" 1);  | 
|
| 5316 | 918  | 
by Auto_tac;  | 
| 2608 | 919  | 
qed_spec_mp "nth_take";  | 
920  | 
Addsimps [nth_take];  | 
|
| 923 | 921  | 
|
| 4935 | 922  | 
Goal "!xs i. n + i <= length xs --> (drop n xs)!i = xs!(n+i)";  | 
| 5183 | 923  | 
by (induct_tac "n" 1);  | 
| 5316 | 924  | 
by Auto_tac;  | 
| 3457 | 925  | 
by (exhaust_tac "xs" 1);  | 
| 5316 | 926  | 
by Auto_tac;  | 
| 2608 | 927  | 
qed_spec_mp "nth_drop";  | 
928  | 
Addsimps [nth_drop];  | 
|
929  | 
||
| 8118 | 930  | 
|
931  | 
Goal  | 
|
932  | 
"!zs. (xs@ys = zs) = (xs = take (length xs) zs & ys = drop (length xs) zs)";  | 
|
933  | 
by(induct_tac "xs" 1);  | 
|
934  | 
by(Simp_tac 1);  | 
|
935  | 
by(Asm_full_simp_tac 1);  | 
|
936  | 
by(Clarify_tac 1);  | 
|
937  | 
by(exhaust_tac "zs" 1);  | 
|
938  | 
by(Auto_tac);  | 
|
939  | 
qed_spec_mp "append_eq_conv_conj";  | 
|
940  | 
||
| 2608 | 941  | 
(** takeWhile & dropWhile **)  | 
942  | 
||
| 3467 | 943  | 
section "takeWhile & dropWhile";  | 
944  | 
||
| 4935 | 945  | 
Goal "takeWhile P xs @ dropWhile P xs = xs";  | 
| 3586 | 946  | 
by (induct_tac "xs" 1);  | 
| 5316 | 947  | 
by Auto_tac;  | 
| 3586 | 948  | 
qed "takeWhile_dropWhile_id";  | 
949  | 
Addsimps [takeWhile_dropWhile_id];  | 
|
950  | 
||
| 4935 | 951  | 
Goal "x:set xs & ~P(x) --> takeWhile P (xs @ ys) = takeWhile P xs";  | 
| 3457 | 952  | 
by (induct_tac "xs" 1);  | 
| 5316 | 953  | 
by Auto_tac;  | 
| 2608 | 954  | 
bind_thm("takeWhile_append1", conjI RS (result() RS mp));
 | 
955  | 
Addsimps [takeWhile_append1];  | 
|
| 923 | 956  | 
|
| 4935 | 957  | 
Goal "(!x:set xs. P(x)) --> takeWhile P (xs @ ys) = xs @ takeWhile P ys";  | 
| 3457 | 958  | 
by (induct_tac "xs" 1);  | 
| 5316 | 959  | 
by Auto_tac;  | 
| 2608 | 960  | 
bind_thm("takeWhile_append2", ballI RS (result() RS mp));
 | 
961  | 
Addsimps [takeWhile_append2];  | 
|
| 1169 | 962  | 
|
| 4935 | 963  | 
Goal "x:set xs & ~P(x) --> dropWhile P (xs @ ys) = (dropWhile P xs)@ys";  | 
| 3457 | 964  | 
by (induct_tac "xs" 1);  | 
| 5316 | 965  | 
by Auto_tac;  | 
| 2608 | 966  | 
bind_thm("dropWhile_append1", conjI RS (result() RS mp));
 | 
967  | 
Addsimps [dropWhile_append1];  | 
|
968  | 
||
| 4935 | 969  | 
Goal "(!x:set xs. P(x)) --> dropWhile P (xs @ ys) = dropWhile P ys";  | 
| 3457 | 970  | 
by (induct_tac "xs" 1);  | 
| 5316 | 971  | 
by Auto_tac;  | 
| 2608 | 972  | 
bind_thm("dropWhile_append2", ballI RS (result() RS mp));
 | 
973  | 
Addsimps [dropWhile_append2];  | 
|
974  | 
||
| 4935 | 975  | 
Goal "x:set(takeWhile P xs) --> x:set xs & P x";  | 
| 3457 | 976  | 
by (induct_tac "xs" 1);  | 
| 5316 | 977  | 
by Auto_tac;  | 
| 
3647
 
a64c8fbcd98f
Renamed theorems of the form set_of_list_XXX to set_XXX
 
paulson 
parents: 
3589 
diff
changeset
 | 
978  | 
qed_spec_mp"set_take_whileD";  | 
| 2608 | 979  | 
|
| 6306 | 980  | 
(** zip **)  | 
981  | 
section "zip";  | 
|
982  | 
||
983  | 
Goal "zip [] ys = []";  | 
|
| 
6813
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
984  | 
by (induct_tac "ys" 1);  | 
| 6306 | 985  | 
by Auto_tac;  | 
986  | 
qed "zip_Nil";  | 
|
987  | 
Addsimps [zip_Nil];  | 
|
988  | 
||
989  | 
Goal "zip (x#xs) (y#ys) = (x,y)#zip xs ys";  | 
|
| 
6813
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
990  | 
by (Simp_tac 1);  | 
| 6306 | 991  | 
qed "zip_Cons_Cons";  | 
992  | 
Addsimps [zip_Cons_Cons];  | 
|
993  | 
||
994  | 
Delsimps(tl (thms"zip.simps"));  | 
|
| 4605 | 995  | 
|
| 8118 | 996  | 
Goal "!xs. length (zip xs ys) = min (length xs) (length ys)";  | 
| 8009 | 997  | 
by (induct_tac "ys" 1);  | 
998  | 
by (Simp_tac 1);  | 
|
999  | 
by (Clarify_tac 1);  | 
|
1000  | 
by (exhaust_tac "xs" 1);  | 
|
| 8064 | 1001  | 
by (Auto_tac);  | 
| 8009 | 1002  | 
qed_spec_mp "length_zip";  | 
1003  | 
Addsimps [length_zip];  | 
|
1004  | 
||
1005  | 
Goal  | 
|
| 8118 | 1006  | 
"!xs. zip (xs@ys) zs = \  | 
1007  | 
\ zip xs (take (length xs) zs) @ zip ys (drop (length xs) zs)";  | 
|
1008  | 
by(induct_tac "zs" 1);  | 
|
1009  | 
by(Simp_tac 1);  | 
|
| 8064 | 1010  | 
by (Clarify_tac 1);  | 
| 8118 | 1011  | 
by(exhaust_tac "xs" 1);  | 
1012  | 
by(Asm_simp_tac 1);  | 
|
1013  | 
by(Asm_simp_tac 1);  | 
|
1014  | 
qed_spec_mp "zip_append1";  | 
|
1015  | 
||
1016  | 
Goal  | 
|
1017  | 
"!ys. zip xs (ys@zs) = \  | 
|
1018  | 
\ zip (take (length ys) xs) ys @ zip (drop (length ys) xs) zs";  | 
|
1019  | 
by(induct_tac "xs" 1);  | 
|
1020  | 
by(Simp_tac 1);  | 
|
1021  | 
by (Clarify_tac 1);  | 
|
1022  | 
by(exhaust_tac "ys" 1);  | 
|
1023  | 
by(Asm_simp_tac 1);  | 
|
1024  | 
by(Asm_simp_tac 1);  | 
|
1025  | 
qed_spec_mp "zip_append2";  | 
|
1026  | 
||
1027  | 
Goal  | 
|
1028  | 
"[| length xs = length us; length ys = length vs |] ==> \  | 
|
1029  | 
\ zip (xs@ys) (us@vs) = zip xs us @ zip ys vs";  | 
|
1030  | 
by(asm_simp_tac (simpset() addsimps [zip_append1]) 1);  | 
|
| 8009 | 1031  | 
qed_spec_mp "zip_append";  | 
| 8118 | 1032  | 
Addsimps [zip_append];  | 
| 8009 | 1033  | 
|
1034  | 
Goal "!xs. length xs = length ys --> zip (rev xs) (rev ys) = rev (zip xs ys)";  | 
|
| 8064 | 1035  | 
by (induct_tac "ys" 1);  | 
1036  | 
by (Asm_full_simp_tac 1);  | 
|
1037  | 
by (Asm_full_simp_tac 1);  | 
|
1038  | 
by (Clarify_tac 1);  | 
|
1039  | 
by (exhaust_tac "xs" 1);  | 
|
1040  | 
by (Auto_tac);  | 
|
| 8009 | 1041  | 
qed_spec_mp "zip_rev";  | 
1042  | 
||
| 8115 | 1043  | 
|
1044  | 
Goal  | 
|
| 8009 | 1045  | 
"!i xs. i < length xs --> i < length ys --> (zip xs ys)!i = (xs!i, ys!i)";  | 
1046  | 
by (induct_tac "ys" 1);  | 
|
1047  | 
by (Simp_tac 1);  | 
|
1048  | 
by (Clarify_tac 1);  | 
|
| 8064 | 1049  | 
by (exhaust_tac "xs" 1);  | 
1050  | 
by (Auto_tac);  | 
|
| 8009 | 1051  | 
by (asm_full_simp_tac (simpset() addsimps (thms"nth.simps") addsplits [nat.split]) 1);  | 
1052  | 
qed_spec_mp "nth_zip";  | 
|
1053  | 
Addsimps [nth_zip];  | 
|
1054  | 
||
| 8118 | 1055  | 
Goal "set(zip xs ys) = {(xs!i,ys!i) |i. i < min (length xs) (length ys)}";
 | 
1056  | 
by (simp_tac (simpset() addsimps [set_conv_nth]addcongs [rev_conj_cong]) 1);  | 
|
1057  | 
qed_spec_mp "set_zip";  | 
|
1058  | 
||
| 8009 | 1059  | 
Goal  | 
1060  | 
"length xs = length ys ==> zip (xs[i:=x]) (ys[i:=y]) = (zip xs ys)[i:=(x,y)]";  | 
|
| 8064 | 1061  | 
by (rtac sym 1);  | 
1062  | 
by (asm_simp_tac (simpset() addsimps [update_zip]) 1);  | 
|
| 8009 | 1063  | 
qed_spec_mp "zip_update";  | 
1064  | 
||
1065  | 
Goal "!j. zip (replicate i x) (replicate j y) = replicate (min i j) (x,y)";  | 
|
1066  | 
by (induct_tac "i" 1);  | 
|
| 8064 | 1067  | 
by (Auto_tac);  | 
| 8009 | 1068  | 
by (exhaust_tac "j" 1);  | 
| 8064 | 1069  | 
by (Auto_tac);  | 
| 8009 | 1070  | 
qed "zip_replicate";  | 
1071  | 
Addsimps [zip_replicate];  | 
|
1072  | 
||
| 8115 | 1073  | 
(** list_all2 **)  | 
1074  | 
section "list_all2";  | 
|
1075  | 
||
1076  | 
Goalw [list_all2_def] "list_all2 P xs ys ==> length xs = length ys";  | 
|
1077  | 
by(Asm_simp_tac 1);  | 
|
1078  | 
qed "list_all2_lengthD";  | 
|
1079  | 
||
1080  | 
Goalw [list_all2_def] "list_all2 P [] ys = (ys=[])";  | 
|
1081  | 
by (Simp_tac 1);  | 
|
1082  | 
qed "list_all2_Nil";  | 
|
1083  | 
AddIffs [list_all2_Nil];  | 
|
1084  | 
||
1085  | 
Goalw [list_all2_def] "list_all2 P xs [] = (xs=[])";  | 
|
1086  | 
by (Simp_tac 1);  | 
|
1087  | 
qed "list_all2_Nil2";  | 
|
1088  | 
AddIffs [list_all2_Nil2];  | 
|
1089  | 
||
1090  | 
Goalw [list_all2_def]  | 
|
1091  | 
"list_all2 P (x#xs) (y#ys) = (P x y & list_all2 P xs ys)";  | 
|
1092  | 
by (Auto_tac);  | 
|
1093  | 
qed "list_all2_Cons";  | 
|
1094  | 
AddIffs[list_all2_Cons];  | 
|
1095  | 
||
1096  | 
Goalw [list_all2_def]  | 
|
| 8118 | 1097  | 
"list_all2 P (x#xs) ys = (? z zs. ys = z#zs & P x z & list_all2 P xs zs)";  | 
1098  | 
by(exhaust_tac "ys" 1);  | 
|
1099  | 
by(Auto_tac);  | 
|
1100  | 
qed "list_all2_Cons1";  | 
|
1101  | 
||
1102  | 
Goalw [list_all2_def]  | 
|
1103  | 
"list_all2 P xs (y#ys) = (? z zs. xs = z#zs & P z y & list_all2 P zs ys)";  | 
|
1104  | 
by(exhaust_tac "xs" 1);  | 
|
1105  | 
by(Auto_tac);  | 
|
1106  | 
qed "list_all2_Cons2";  | 
|
1107  | 
||
1108  | 
Goalw [list_all2_def]  | 
|
1109  | 
"list_all2 P (xs@ys) zs = \  | 
|
1110  | 
\ (EX us vs. zs = us@vs & length us = length xs & length vs = length ys & \  | 
|
1111  | 
\ list_all2 P xs us & list_all2 P ys vs)";  | 
|
1112  | 
by(simp_tac (simpset() addsimps [zip_append1]) 1);  | 
|
1113  | 
br iffI 1;  | 
|
1114  | 
 by(res_inst_tac [("x","take (length xs) zs")] exI 1);
 | 
|
1115  | 
 by(res_inst_tac [("x","drop (length xs) zs")] exI 1);
 | 
|
1116  | 
by(asm_full_simp_tac (simpset() addsimps [min_def,eq_sym_conv]) 1);  | 
|
1117  | 
by (Clarify_tac 1);  | 
|
1118  | 
by(asm_full_simp_tac (simpset() addsimps [ball_Un]) 1);  | 
|
1119  | 
qed "list_all2_append1";  | 
|
1120  | 
||
1121  | 
Goalw [list_all2_def]  | 
|
1122  | 
"list_all2 P xs (ys@zs) = \  | 
|
1123  | 
\ (EX us vs. xs = us@vs & length us = length ys & length vs = length zs & \  | 
|
1124  | 
\ list_all2 P us ys & list_all2 P vs zs)";  | 
|
1125  | 
by(simp_tac (simpset() addsimps [zip_append2]) 1);  | 
|
1126  | 
br iffI 1;  | 
|
1127  | 
 by(res_inst_tac [("x","take (length ys) xs")] exI 1);
 | 
|
1128  | 
 by(res_inst_tac [("x","drop (length ys) xs")] exI 1);
 | 
|
1129  | 
by(asm_full_simp_tac (simpset() addsimps [min_def,eq_sym_conv]) 1);  | 
|
1130  | 
by (Clarify_tac 1);  | 
|
1131  | 
by(asm_full_simp_tac (simpset() addsimps [ball_Un]) 1);  | 
|
1132  | 
qed "list_all2_append2";  | 
|
1133  | 
||
1134  | 
Goalw [list_all2_def]  | 
|
| 8115 | 1135  | 
"list_all2 P xs ys = \  | 
1136  | 
\ (length xs = length ys & (!i<length xs. P (xs!i) (ys!i)))";  | 
|
1137  | 
by(force_tac (claset(), simpset() addsimps [set_zip]) 1);  | 
|
1138  | 
qed "list_all2_conv_all_nth";  | 
|
| 5272 | 1139  | 
|
1140  | 
(** foldl **)  | 
|
1141  | 
section "foldl";  | 
|
1142  | 
||
1143  | 
Goal "!a. foldl f a (xs @ ys) = foldl f (foldl f a xs) ys";  | 
|
| 5318 | 1144  | 
by (induct_tac "xs" 1);  | 
| 5316 | 1145  | 
by Auto_tac;  | 
| 5272 | 1146  | 
qed_spec_mp "foldl_append";  | 
1147  | 
Addsimps [foldl_append];  | 
|
1148  | 
||
1149  | 
(* Note: `n <= foldl op+ n ns' looks simpler, but is more difficult to use  | 
|
1150  | 
because it requires an additional transitivity step  | 
|
1151  | 
*)  | 
|
1152  | 
Goal "!n::nat. m <= n --> m <= foldl op+ n ns";  | 
|
| 5318 | 1153  | 
by (induct_tac "ns" 1);  | 
| 6058 | 1154  | 
by Auto_tac;  | 
| 5272 | 1155  | 
qed_spec_mp "start_le_sum";  | 
1156  | 
||
1157  | 
Goal "n : set ns ==> n <= foldl op+ 0 ns";  | 
|
| 
5758
 
27a2b36efd95
corrected auto_tac (applications of unsafe wrappers)
 
oheimb 
parents: 
5644 
diff
changeset
 | 
1158  | 
by (force_tac (claset() addIs [start_le_sum],  | 
| 
 
27a2b36efd95
corrected auto_tac (applications of unsafe wrappers)
 
oheimb 
parents: 
5644 
diff
changeset
 | 
1159  | 
simpset() addsimps [in_set_conv_decomp]) 1);  | 
| 5272 | 1160  | 
qed "elem_le_sum";  | 
1161  | 
||
1162  | 
Goal "!m. (foldl op+ m ns = 0) = (m=0 & (!n : set ns. n=0))";  | 
|
| 5318 | 1163  | 
by (induct_tac "ns" 1);  | 
| 5316 | 1164  | 
by Auto_tac;  | 
| 5272 | 1165  | 
qed_spec_mp "sum_eq_0_conv";  | 
1166  | 
AddIffs [sum_eq_0_conv];  | 
|
1167  | 
||
| 5425 | 1168  | 
(** upto **)  | 
1169  | 
||
| 5427 | 1170  | 
(* Does not terminate! *)  | 
1171  | 
Goal "[i..j(] = (if i<j then i#[Suc i..j(] else [])";  | 
|
| 6162 | 1172  | 
by (induct_tac "j" 1);  | 
| 5427 | 1173  | 
by Auto_tac;  | 
1174  | 
qed "upt_rec";  | 
|
| 5425 | 1175  | 
|
| 5427 | 1176  | 
Goal "j<=i ==> [i..j(] = []";  | 
| 6162 | 1177  | 
by (stac upt_rec 1);  | 
1178  | 
by (Asm_simp_tac 1);  | 
|
| 5427 | 1179  | 
qed "upt_conv_Nil";  | 
1180  | 
Addsimps [upt_conv_Nil];  | 
|
1181  | 
||
1182  | 
Goal "i<=j ==> [i..(Suc j)(] = [i..j(]@[j]";  | 
|
1183  | 
by (Asm_simp_tac 1);  | 
|
1184  | 
qed "upt_Suc";  | 
|
1185  | 
||
1186  | 
Goal "i<j ==> [i..j(] = i#[Suc i..j(]";  | 
|
| 6162 | 1187  | 
by (rtac trans 1);  | 
1188  | 
by (stac upt_rec 1);  | 
|
1189  | 
by (rtac refl 2);  | 
|
| 5427 | 1190  | 
by (Asm_simp_tac 1);  | 
1191  | 
qed "upt_conv_Cons";  | 
|
1192  | 
||
1193  | 
Goal "length [i..j(] = j-i";  | 
|
| 6162 | 1194  | 
by (induct_tac "j" 1);  | 
| 5427 | 1195  | 
by (Simp_tac 1);  | 
| 6162 | 1196  | 
by (asm_simp_tac (simpset() addsimps [Suc_diff_le]) 1);  | 
| 5427 | 1197  | 
qed "length_upt";  | 
1198  | 
Addsimps [length_upt];  | 
|
| 5425 | 1199  | 
|
| 5427 | 1200  | 
Goal "i+k < j --> [i..j(] ! k = i+k";  | 
| 6162 | 1201  | 
by (induct_tac "j" 1);  | 
1202  | 
by (Simp_tac 1);  | 
|
1203  | 
by (asm_simp_tac (simpset() addsimps [nth_append,less_diff_conv]@add_ac) 1);  | 
|
1204  | 
by (Clarify_tac 1);  | 
|
1205  | 
by (subgoal_tac "n=i+k" 1);  | 
|
1206  | 
by (Asm_simp_tac 2);  | 
|
1207  | 
by (Asm_simp_tac 1);  | 
|
| 5427 | 1208  | 
qed_spec_mp "nth_upt";  | 
1209  | 
Addsimps [nth_upt];  | 
|
| 5425 | 1210  | 
|
| 6433 | 1211  | 
Goal "!i. i+m <= n --> take m [i..n(] = [i..i+m(]";  | 
| 
6813
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1212  | 
by (induct_tac "m" 1);  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1213  | 
by (Simp_tac 1);  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1214  | 
by (Clarify_tac 1);  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1215  | 
by (stac upt_rec 1);  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1216  | 
by (rtac sym 1);  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1217  | 
by (stac upt_rec 1);  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1218  | 
by (asm_simp_tac (simpset() delsimps (thms"upt.simps")) 1);  | 
| 6433 | 1219  | 
qed_spec_mp "take_upt";  | 
1220  | 
Addsimps [take_upt];  | 
|
1221  | 
||
1222  | 
Goal "!m i. i < n-m --> (map f [m..n(]) ! i = f(m+i)";  | 
|
| 
6813
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1223  | 
by (induct_tac "n" 1);  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1224  | 
by (Simp_tac 1);  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1225  | 
by (Clarify_tac 1);  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1226  | 
by (subgoal_tac "m < Suc n" 1);  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1227  | 
by (arith_tac 2);  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1228  | 
by (stac upt_rec 1);  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1229  | 
by (asm_simp_tac (simpset() delsplits [split_if]) 1);  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1230  | 
by (split_tac [split_if] 1);  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1231  | 
by (rtac conjI 1);  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1232  | 
by (simp_tac (simpset() addsimps [nth_Cons] addsplits [nat.split]) 1);  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1233  | 
by (simp_tac (simpset() addsimps [nth_append] addsplits [nat.split]) 1);  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1234  | 
by (Clarify_tac 1);  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1235  | 
by (rtac conjI 1);  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1236  | 
by (Clarify_tac 1);  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1237  | 
by (subgoal_tac "Suc(m+nat) < n" 1);  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1238  | 
by (arith_tac 2);  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1239  | 
by (Asm_simp_tac 1);  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1240  | 
by (Clarify_tac 1);  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1241  | 
by (subgoal_tac "n = Suc(m+nat)" 1);  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1242  | 
by (arith_tac 2);  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1243  | 
by (Asm_simp_tac 1);  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1244  | 
by (simp_tac (simpset() addsimps [nth_Cons] addsplits [nat.split]) 1);  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1245  | 
by (arith_tac 1);  | 
| 6433 | 1246  | 
qed_spec_mp "nth_map_upt";  | 
1247  | 
||
| 
6813
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1248  | 
Goal "ALL xs ys. k <= length xs --> k <= length ys --> \  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1249  | 
\ (ALL i. i < k --> xs!i = ys!i) \  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1250  | 
\ --> take k xs = take k ys";  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1251  | 
by (induct_tac "k" 1);  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1252  | 
by (ALLGOALS (asm_simp_tac (simpset() addsimps [less_Suc_eq_0_disj,  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1253  | 
all_conj_distrib])));  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1254  | 
by (Clarify_tac 1);  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1255  | 
(*Both lists must be non-empty*)  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1256  | 
by (exhaust_tac "xs" 1);  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1257  | 
by (exhaust_tac "ys" 2);  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1258  | 
by (ALLGOALS Clarify_tac);  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1259  | 
(*prenexing's needed, not miniscoping*)  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1260  | 
by (ALLGOALS (full_simp_tac (simpset() addsimps (all_simps RL [sym])  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1261  | 
delsimps (all_simps))));  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1262  | 
by (Blast_tac 1);  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1263  | 
qed_spec_mp "nth_take_lemma";  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1264  | 
|
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1265  | 
Goal "[| length xs = length ys; \  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1266  | 
\ ALL i. i < length xs --> xs!i = ys!i |] \  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1267  | 
\ ==> xs = ys";  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1268  | 
by (forward_tac [[le_refl, eq_imp_le] MRS nth_take_lemma] 1);  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1269  | 
by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [take_all])));  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1270  | 
qed_spec_mp "nth_equalityI";  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1271  | 
|
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1272  | 
(*The famous take-lemma*)  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1273  | 
Goal "(ALL i. take i xs = take i ys) ==> xs = ys";  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1274  | 
by (dres_inst_tac [("x", "max (length xs) (length ys)")] spec 1);
 | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1275  | 
by (full_simp_tac (simpset() addsimps [le_max_iff_disj, take_all]) 1);  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1276  | 
qed_spec_mp "take_equalityI";  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1277  | 
|
| 5272 | 1278  | 
|
| 4605 | 1279  | 
(** nodups & remdups **)  | 
1280  | 
section "nodups & remdups";  | 
|
1281  | 
||
| 4935 | 1282  | 
Goal "set(remdups xs) = set xs";  | 
| 4605 | 1283  | 
by (induct_tac "xs" 1);  | 
1284  | 
by (Simp_tac 1);  | 
|
| 4686 | 1285  | 
by (asm_full_simp_tac (simpset() addsimps [insert_absorb]) 1);  | 
| 4605 | 1286  | 
qed "set_remdups";  | 
1287  | 
Addsimps [set_remdups];  | 
|
1288  | 
||
| 4935 | 1289  | 
Goal "nodups(remdups xs)";  | 
| 4605 | 1290  | 
by (induct_tac "xs" 1);  | 
| 5316 | 1291  | 
by Auto_tac;  | 
| 4605 | 1292  | 
qed "nodups_remdups";  | 
1293  | 
||
| 4935 | 1294  | 
Goal "nodups xs --> nodups (filter P xs)";  | 
| 4605 | 1295  | 
by (induct_tac "xs" 1);  | 
| 5316 | 1296  | 
by Auto_tac;  | 
| 4605 | 1297  | 
qed_spec_mp "nodups_filter";  | 
1298  | 
||
| 
3589
 
244daa75f890
Added function `replicate' and lemmas map_cong and set_replicate.
 
nipkow 
parents: 
3586 
diff
changeset
 | 
1299  | 
(** replicate **)  | 
| 
 
244daa75f890
Added function `replicate' and lemmas map_cong and set_replicate.
 
nipkow 
parents: 
3586 
diff
changeset
 | 
1300  | 
section "replicate";  | 
| 
 
244daa75f890
Added function `replicate' and lemmas map_cong and set_replicate.
 
nipkow 
parents: 
3586 
diff
changeset
 | 
1301  | 
|
| 6794 | 1302  | 
Goal "length(replicate n x) = n";  | 
| 
6813
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1303  | 
by (induct_tac "n" 1);  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1304  | 
by Auto_tac;  | 
| 6794 | 1305  | 
qed "length_replicate";  | 
1306  | 
Addsimps [length_replicate];  | 
|
1307  | 
||
1308  | 
Goal "map f (replicate n x) = replicate n (f x)";  | 
|
1309  | 
by (induct_tac "n" 1);  | 
|
| 
6813
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1310  | 
by Auto_tac;  | 
| 6794 | 1311  | 
qed "map_replicate";  | 
1312  | 
Addsimps [map_replicate];  | 
|
1313  | 
||
1314  | 
Goal "(replicate n x) @ (x#xs) = x # replicate n x @ xs";  | 
|
1315  | 
by (induct_tac "n" 1);  | 
|
| 
6813
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1316  | 
by Auto_tac;  | 
| 6794 | 1317  | 
qed "replicate_app_Cons_same";  | 
1318  | 
||
1319  | 
Goal "rev(replicate n x) = replicate n x";  | 
|
1320  | 
by (induct_tac "n" 1);  | 
|
| 
6813
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1321  | 
by (Simp_tac 1);  | 
| 6794 | 1322  | 
by (asm_simp_tac (simpset() addsimps [replicate_app_Cons_same]) 1);  | 
1323  | 
qed "rev_replicate";  | 
|
1324  | 
Addsimps [rev_replicate];  | 
|
1325  | 
||
| 8009 | 1326  | 
Goal "replicate (n+m) x = replicate n x @ replicate m x";  | 
1327  | 
by (induct_tac "n" 1);  | 
|
1328  | 
by Auto_tac;  | 
|
1329  | 
qed "replicate_add";  | 
|
1330  | 
||
| 6794 | 1331  | 
Goal"n ~= 0 --> hd(replicate n x) = x";  | 
1332  | 
by (induct_tac "n" 1);  | 
|
| 
6813
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1333  | 
by Auto_tac;  | 
| 6794 | 1334  | 
qed_spec_mp "hd_replicate";  | 
1335  | 
Addsimps [hd_replicate];  | 
|
1336  | 
||
1337  | 
Goal "n ~= 0 --> tl(replicate n x) = replicate (n-1) x";  | 
|
1338  | 
by (induct_tac "n" 1);  | 
|
| 
6813
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1339  | 
by Auto_tac;  | 
| 6794 | 1340  | 
qed_spec_mp "tl_replicate";  | 
1341  | 
Addsimps [tl_replicate];  | 
|
1342  | 
||
1343  | 
Goal "n ~= 0 --> last(replicate n x) = x";  | 
|
1344  | 
by (induct_tac "n" 1);  | 
|
| 
6813
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1345  | 
by Auto_tac;  | 
| 6794 | 1346  | 
qed_spec_mp "last_replicate";  | 
1347  | 
Addsimps [last_replicate];  | 
|
1348  | 
||
1349  | 
Goal "!i. i<n --> (replicate n x)!i = x";  | 
|
| 
6813
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1350  | 
by (induct_tac "n" 1);  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1351  | 
by (Simp_tac 1);  | 
| 
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1352  | 
by (asm_simp_tac (simpset() addsimps [nth_Cons] addsplits [nat.split]) 1);  | 
| 6794 | 1353  | 
qed_spec_mp "nth_replicate";  | 
1354  | 
Addsimps [nth_replicate];  | 
|
1355  | 
||
| 4935 | 1356  | 
Goal "set(replicate (Suc n) x) = {x}";
 | 
| 4423 | 1357  | 
by (induct_tac "n" 1);  | 
| 5316 | 1358  | 
by Auto_tac;  | 
| 
3589
 
244daa75f890
Added function `replicate' and lemmas map_cong and set_replicate.
 
nipkow 
parents: 
3586 
diff
changeset
 | 
1359  | 
val lemma = result();  | 
| 
 
244daa75f890
Added function `replicate' and lemmas map_cong and set_replicate.
 
nipkow 
parents: 
3586 
diff
changeset
 | 
1360  | 
|
| 5043 | 1361  | 
Goal "n ~= 0 ==> set(replicate n x) = {x}";
 | 
| 4423 | 1362  | 
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
 | 
1363  | 
qed "set_replicate";  | 
| 
 
244daa75f890
Added function `replicate' and lemmas map_cong and set_replicate.
 
nipkow 
parents: 
3586 
diff
changeset
 | 
1364  | 
Addsimps [set_replicate];  | 
| 5162 | 1365  | 
|
| 8009 | 1366  | 
Goal "set(replicate n x) = (if n=0 then {} else {x})";
 | 
| 8064 | 1367  | 
by (Auto_tac);  | 
| 8009 | 1368  | 
qed "set_replicate_conv_if";  | 
1369  | 
||
1370  | 
Goal "x : set(replicate n y) --> x=y";  | 
|
| 8064 | 1371  | 
by (asm_simp_tac (simpset() addsimps [set_replicate_conv_if]) 1);  | 
| 8009 | 1372  | 
qed_spec_mp "in_set_replicateD";  | 
1373  | 
||
| 5162 | 1374  | 
|
| 5281 | 1375  | 
(*** Lexcicographic orderings on lists ***)  | 
1376  | 
section"Lexcicographic orderings on lists";  | 
|
1377  | 
||
1378  | 
Goal "wf r ==> wf(lexn r n)";  | 
|
| 5318 | 1379  | 
by (induct_tac "n" 1);  | 
1380  | 
by (Simp_tac 1);  | 
|
1381  | 
by (Simp_tac 1);  | 
|
1382  | 
by (rtac wf_subset 1);  | 
|
1383  | 
by (rtac Int_lower1 2);  | 
|
1384  | 
by (rtac wf_prod_fun_image 1);  | 
|
1385  | 
by (rtac injI 2);  | 
|
| 
6813
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1386  | 
by Auto_tac;  | 
| 5281 | 1387  | 
qed "wf_lexn";  | 
1388  | 
||
1389  | 
Goal "!xs ys. (xs,ys) : lexn r n --> length xs = n & length ys = n";  | 
|
| 5318 | 1390  | 
by (induct_tac "n" 1);  | 
| 
6813
 
bf90f86502b2
many new lemmas about take & drop, incl the famous take-lemma
 
paulson 
parents: 
6794 
diff
changeset
 | 
1391  | 
by Auto_tac;  | 
| 5281 | 1392  | 
qed_spec_mp "lexn_length";  | 
1393  | 
||
1394  | 
Goalw [lex_def] "wf r ==> wf(lex r)";  | 
|
| 5318 | 1395  | 
by (rtac wf_UN 1);  | 
1396  | 
by (blast_tac (claset() addIs [wf_lexn]) 1);  | 
|
1397  | 
by (Clarify_tac 1);  | 
|
1398  | 
by (rename_tac "m n" 1);  | 
|
1399  | 
by (subgoal_tac "m ~= n" 1);  | 
|
1400  | 
by (Blast_tac 2);  | 
|
1401  | 
by (blast_tac (claset() addDs [lexn_length,not_sym]) 1);  | 
|
| 5281 | 1402  | 
qed "wf_lex";  | 
1403  | 
AddSIs [wf_lex];  | 
|
1404  | 
||
1405  | 
Goal  | 
|
1406  | 
"lexn r n = \  | 
|
1407  | 
\ {(xs,ys). length xs = n & length ys = n & \
 | 
|
1408  | 
\ (? xys x y xs' ys'. xs= xys @ x#xs' & ys= xys @ y#ys' & (x,y):r)}";  | 
|
| 5318 | 1409  | 
by (induct_tac "n" 1);  | 
1410  | 
by (Simp_tac 1);  | 
|
1411  | 
by (Blast_tac 1);  | 
|
| 5641 | 1412  | 
by (asm_full_simp_tac (simpset()  | 
| 5296 | 1413  | 
addsimps [lex_prod_def]) 1);  | 
| 5641 | 1414  | 
by (auto_tac (claset(), simpset()));  | 
| 5318 | 1415  | 
by (Blast_tac 1);  | 
1416  | 
by (rename_tac "a xys x xs' y ys'" 1);  | 
|
1417  | 
 by (res_inst_tac [("x","a#xys")] exI 1);
 | 
|
1418  | 
by (Simp_tac 1);  | 
|
1419  | 
by (exhaust_tac "xys" 1);  | 
|
| 5641 | 1420  | 
by (ALLGOALS (asm_full_simp_tac (simpset())));  | 
| 5318 | 1421  | 
by (Blast_tac 1);  | 
| 5281 | 1422  | 
qed "lexn_conv";  | 
1423  | 
||
1424  | 
Goalw [lex_def]  | 
|
1425  | 
"lex r = \  | 
|
1426  | 
\ {(xs,ys). length xs = length ys & \
 | 
|
1427  | 
\ (? xys x y xs' ys'. xs= xys @ x#xs' & ys= xys @ y#ys' & (x,y):r)}";  | 
|
| 5641 | 1428  | 
by (force_tac (claset(), simpset() addsimps [lexn_conv]) 1);  | 
| 5281 | 1429  | 
qed "lex_conv";  | 
1430  | 
||
1431  | 
Goalw [lexico_def] "wf r ==> wf(lexico r)";  | 
|
| 5318 | 1432  | 
by (Blast_tac 1);  | 
| 5281 | 1433  | 
qed "wf_lexico";  | 
1434  | 
AddSIs [wf_lexico];  | 
|
1435  | 
||
1436  | 
Goalw  | 
|
1437  | 
[lexico_def,diag_def,lex_prod_def,measure_def,inv_image_def]  | 
|
1438  | 
"lexico r = {(xs,ys). length xs < length ys | \
 | 
|
1439  | 
\ length xs = length ys & (xs,ys) : lex r}";  | 
|
| 5318 | 1440  | 
by (Simp_tac 1);  | 
| 5281 | 1441  | 
qed "lexico_conv";  | 
1442  | 
||
| 5283 | 1443  | 
Goal "([],ys) ~: lex r";  | 
| 5318 | 1444  | 
by (simp_tac (simpset() addsimps [lex_conv]) 1);  | 
| 5283 | 1445  | 
qed "Nil_notin_lex";  | 
1446  | 
||
1447  | 
Goal "(xs,[]) ~: lex r";  | 
|
| 5318 | 1448  | 
by (simp_tac (simpset() addsimps [lex_conv]) 1);  | 
| 5283 | 1449  | 
qed "Nil2_notin_lex";  | 
1450  | 
||
1451  | 
AddIffs [Nil_notin_lex,Nil2_notin_lex];  | 
|
1452  | 
||
1453  | 
Goal "((x#xs,y#ys) : lex r) = \  | 
|
1454  | 
\ ((x,y) : r & length xs = length ys | x=y & (xs,ys) : lex r)";  | 
|
| 5318 | 1455  | 
by (simp_tac (simpset() addsimps [lex_conv]) 1);  | 
1456  | 
by (rtac iffI 1);  | 
|
1457  | 
by (blast_tac (claset() addIs [Cons_eq_appendI]) 2);  | 
|
1458  | 
by (REPEAT(eresolve_tac [conjE, exE] 1));  | 
|
1459  | 
by (exhaust_tac "xys" 1);  | 
|
1460  | 
by (Asm_full_simp_tac 1);  | 
|
1461  | 
by (Asm_full_simp_tac 1);  | 
|
1462  | 
by (Blast_tac 1);  | 
|
| 5283 | 1463  | 
qed "Cons_in_lex";  | 
1464  | 
AddIffs [Cons_in_lex];  | 
|
| 7032 | 1465  | 
|
1466  | 
||
1467  | 
(*** Versions of some theorems above using binary numerals ***)  | 
|
1468  | 
||
1469  | 
AddIffs (map (rename_numerals thy)  | 
|
1470  | 
[length_0_conv, zero_length_conv, length_greater_0_conv,  | 
|
1471  | 
sum_eq_0_conv]);  | 
|
1472  | 
||
1473  | 
Goal "take n (x#xs) = (if n = #0 then [] else x # take (n-#1) xs)";  | 
|
1474  | 
by (exhaust_tac "n" 1);  | 
|
1475  | 
by (ALLGOALS  | 
|
1476  | 
(asm_simp_tac (simpset() addsimps [numeral_0_eq_0, numeral_1_eq_1])));  | 
|
1477  | 
qed "take_Cons'";  | 
|
1478  | 
||
1479  | 
Goal "drop n (x#xs) = (if n = #0 then x#xs else drop (n-#1) xs)";  | 
|
1480  | 
by (exhaust_tac "n" 1);  | 
|
1481  | 
by (ALLGOALS  | 
|
1482  | 
(asm_simp_tac (simpset() addsimps [numeral_0_eq_0, numeral_1_eq_1])));  | 
|
1483  | 
qed "drop_Cons'";  | 
|
1484  | 
||
1485  | 
Goal "(x#xs)!n = (if n = #0 then x else xs!(n-#1))";  | 
|
1486  | 
by (exhaust_tac "n" 1);  | 
|
1487  | 
by (ALLGOALS  | 
|
1488  | 
(asm_simp_tac (simpset() addsimps [numeral_0_eq_0, numeral_1_eq_1])));  | 
|
1489  | 
qed "nth_Cons'";  | 
|
1490  | 
||
1491  | 
Addsimps (map (inst "n" "number_of ?v") [take_Cons', drop_Cons', nth_Cons']);  | 
|
1492  |