| author | paulson | 
| Mon, 16 Nov 1998 10:39:30 +0100 | |
| changeset 5867 | 1c4806b4bf43 | 
| parent 5529 | 4a54acae6a15 | 
| child 6053 | 8a1059aa01f0 | 
| permissions | -rw-r--r-- | 
| 1461 | 1  | 
(* Title: ZF/List.ML  | 
| 0 | 2  | 
ID: $Id$  | 
| 1461 | 3  | 
Author: Lawrence C Paulson, Cambridge University Computer Laboratory  | 
| 0 | 4  | 
Copyright 1993 University of Cambridge  | 
5  | 
||
6  | 
Datatype definition of Lists  | 
|
7  | 
*)  | 
|
8  | 
||
| 516 | 9  | 
open List;  | 
| 0 | 10  | 
|
| 516 | 11  | 
(*** Aspects of the datatype definition ***)  | 
| 0 | 12  | 
|
| 2469 | 13  | 
Addsimps list.case_eqns;  | 
14  | 
||
15  | 
||
| 0 | 16  | 
(*An elimination rule, for type-checking*)  | 
| 516 | 17  | 
val ConsE = list.mk_cases list.con_defs "Cons(a,l) : list(A)";  | 
| 0 | 18  | 
|
19  | 
(*Proving freeness results*)  | 
|
| 516 | 20  | 
val Cons_iff = list.mk_free "Cons(a,l)=Cons(a',l') <-> a=a' & l=l'";  | 
21  | 
val Nil_Cons_iff = list.mk_free "~ Nil=Cons(a,l)";  | 
|
| 0 | 22  | 
|
23  | 
(*Perform induction on l, then prove the major premise using prems. *)  | 
|
24  | 
fun list_ind_tac a prems i =  | 
|
| 516 | 25  | 
    EVERY [res_inst_tac [("x",a)] list.induct i,
 | 
| 1461 | 26  | 
rename_last_tac a ["1"] (i+2),  | 
27  | 
ares_tac prems i];  | 
|
| 0 | 28  | 
|
| 5067 | 29  | 
Goal "list(A) = {0} + (A * list(A))";
 | 
| 
525
 
e62519a8497d
ZF/List, ex/Brouwer,Data,LList,Ntree,TF,Term: much simplified proof of _unfold
 
lcp 
parents: 
516 
diff
changeset
 | 
30  | 
let open list; val rew = rewrite_rule con_defs in  | 
| 4091 | 31  | 
by (blast_tac (claset() addSIs (map rew intrs) addEs [rew elim]) 1)  | 
| 
525
 
e62519a8497d
ZF/List, ex/Brouwer,Data,LList,Ntree,TF,Term: much simplified proof of _unfold
 
lcp 
parents: 
516 
diff
changeset
 | 
32  | 
end;  | 
| 760 | 33  | 
qed "list_unfold";  | 
| 435 | 34  | 
|
| 0 | 35  | 
(** Lemmas to justify using "list" in other recursive type definitions **)  | 
36  | 
||
| 5137 | 37  | 
Goalw list.defs "A<=B ==> list(A) <= list(B)";  | 
| 0 | 38  | 
by (rtac lfp_mono 1);  | 
| 516 | 39  | 
by (REPEAT (rtac list.bnd_mono 1));  | 
| 0 | 40  | 
by (REPEAT (ares_tac (univ_mono::basic_monos) 1));  | 
| 760 | 41  | 
qed "list_mono";  | 
| 0 | 42  | 
|
43  | 
(*There is a similar proof by list induction.*)  | 
|
| 5067 | 44  | 
Goalw (list.defs@list.con_defs) "list(univ(A)) <= univ(A)";  | 
| 0 | 45  | 
by (rtac lfp_lowerbound 1);  | 
46  | 
by (rtac (A_subset_univ RS univ_mono) 2);  | 
|
| 4091 | 47  | 
by (blast_tac (claset() addSIs [zero_in_univ, Inl_in_univ, Inr_in_univ,  | 
| 1461 | 48  | 
Pair_in_univ]) 1);  | 
| 760 | 49  | 
qed "list_univ";  | 
| 0 | 50  | 
|
| 
908
 
1f99a44c10cb
Updated comment about list_subset_univ and list_into_univ.
 
lcp 
parents: 
782 
diff
changeset
 | 
51  | 
(*These two theorems justify datatypes involving list(nat), list(A), ...*)  | 
| 
782
 
200a16083201
added bind_thm for theorems defined by "standard ..."
 
clasohm 
parents: 
760 
diff
changeset
 | 
52  | 
bind_thm ("list_subset_univ", ([list_mono, list_univ] MRS subset_trans));
 | 
| 0 | 53  | 
|
| 5137 | 54  | 
Goal "[| l: list(A); A <= univ(B) |] ==> l: univ(B)";  | 
| 435 | 55  | 
by (REPEAT (ares_tac [list_subset_univ RS subsetD] 1));  | 
| 760 | 56  | 
qed "list_into_univ";  | 
| 435 | 57  | 
|
| 5321 | 58  | 
val major::prems = Goal  | 
| 0 | 59  | 
"[| l: list(A); \  | 
| 
15
 
6c6d2f6e3185
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 
lcp 
parents: 
6 
diff
changeset
 | 
60  | 
\ c: C(Nil); \  | 
| 
 
6c6d2f6e3185
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 
lcp 
parents: 
6 
diff
changeset
 | 
61  | 
\ !!x y. [| x: A; y: list(A) |] ==> h(x,y): C(Cons(x,y)) \  | 
| 
 
6c6d2f6e3185
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 
lcp 
parents: 
6 
diff
changeset
 | 
62  | 
\ |] ==> list_case(c,h,l) : C(l)";  | 
| 516 | 63  | 
by (rtac (major RS list.induct) 1);  | 
| 5529 | 64  | 
by (ALLGOALS (asm_simp_tac (simpset() addsimps list.case_eqns @ prems)));  | 
| 760 | 65  | 
qed "list_case_type";  | 
| 0 | 66  | 
|
67  | 
||
68  | 
(** For recursion **)  | 
|
69  | 
||
| 5067 | 70  | 
Goalw list.con_defs "rank(a) < rank(Cons(a,l))";  | 
| 
6
 
8ce8c4d13d4d
Installation of new simplifier for ZF.  Deleted all congruence rules not
 
lcp 
parents: 
0 
diff
changeset
 | 
71  | 
by (simp_tac rank_ss 1);  | 
| 760 | 72  | 
qed "rank_Cons1";  | 
| 0 | 73  | 
|
| 5067 | 74  | 
Goalw list.con_defs "rank(l) < rank(Cons(a,l))";  | 
| 
6
 
8ce8c4d13d4d
Installation of new simplifier for ZF.  Deleted all congruence rules not
 
lcp 
parents: 
0 
diff
changeset
 | 
75  | 
by (simp_tac rank_ss 1);  | 
| 760 | 76  | 
qed "rank_Cons2";  | 
| 0 | 77  | 
|
| 516 | 78  | 
|
79  | 
(*** List functions ***)  | 
|
80  | 
||
81  | 
(** hd and tl **)  | 
|
82  | 
||
| 5067 | 83  | 
Goalw [hd_def] "hd(Cons(a,l)) = a";  | 
| 516 | 84  | 
by (resolve_tac list.case_eqns 1);  | 
| 760 | 85  | 
qed "hd_Cons";  | 
| 516 | 86  | 
|
| 5067 | 87  | 
Goalw [tl_def] "tl(Nil) = Nil";  | 
| 516 | 88  | 
by (resolve_tac list.case_eqns 1);  | 
| 760 | 89  | 
qed "tl_Nil";  | 
| 516 | 90  | 
|
| 5067 | 91  | 
Goalw [tl_def] "tl(Cons(a,l)) = l";  | 
| 516 | 92  | 
by (resolve_tac list.case_eqns 1);  | 
| 760 | 93  | 
qed "tl_Cons";  | 
| 516 | 94  | 
|
| 2469 | 95  | 
Addsimps [hd_Cons, tl_Nil, tl_Cons];  | 
96  | 
||
| 5137 | 97  | 
Goal "l: list(A) ==> tl(l) : list(A)";  | 
| 516 | 98  | 
by (etac list.elim 1);  | 
| 4091 | 99  | 
by (ALLGOALS (asm_simp_tac (simpset() addsimps list.intrs)));  | 
| 760 | 100  | 
qed "tl_type";  | 
| 516 | 101  | 
|
102  | 
(** drop **)  | 
|
103  | 
||
| 5067 | 104  | 
Goalw [drop_def] "drop(0, l) = l";  | 
| 516 | 105  | 
by (rtac rec_0 1);  | 
| 760 | 106  | 
qed "drop_0";  | 
| 516 | 107  | 
|
| 5137 | 108  | 
Goalw [drop_def] "i:nat ==> drop(i, Nil) = Nil";  | 
| 516 | 109  | 
by (etac nat_induct 1);  | 
| 2469 | 110  | 
by (ALLGOALS Asm_simp_tac);  | 
| 760 | 111  | 
qed "drop_Nil";  | 
| 516 | 112  | 
|
| 5067 | 113  | 
Goalw [drop_def]  | 
| 
5147
 
825877190618
More tidying and removal of "\!\!... from Goal commands
 
paulson 
parents: 
5137 
diff
changeset
 | 
114  | 
"i:nat ==> drop(succ(i), Cons(a,l)) = drop(i,l)";  | 
| 2493 | 115  | 
by (rtac sym 1);  | 
| 516 | 116  | 
by (etac nat_induct 1);  | 
| 2469 | 117  | 
by (Simp_tac 1);  | 
118  | 
by (Asm_simp_tac 1);  | 
|
| 760 | 119  | 
qed "drop_succ_Cons";  | 
| 516 | 120  | 
|
| 2469 | 121  | 
Addsimps [drop_0, drop_Nil, drop_succ_Cons];  | 
122  | 
||
| 5067 | 123  | 
Goalw [drop_def]  | 
| 
5147
 
825877190618
More tidying and removal of "\!\!... from Goal commands
 
paulson 
parents: 
5137 
diff
changeset
 | 
124  | 
"[| i:nat; l: list(A) |] ==> drop(i,l) : list(A)";  | 
| 516 | 125  | 
by (etac nat_induct 1);  | 
| 4091 | 126  | 
by (ALLGOALS (asm_simp_tac (simpset() addsimps [tl_type])));  | 
| 760 | 127  | 
qed "drop_type";  | 
| 516 | 128  | 
|
129  | 
(** list_rec -- by Vset recursion **)  | 
|
130  | 
||
| 5067 | 131  | 
Goal "list_rec(Nil,c,h) = c";  | 
| 516 | 132  | 
by (rtac (list_rec_def RS def_Vrec RS trans) 1);  | 
| 4091 | 133  | 
by (simp_tac (simpset() addsimps list.case_eqns) 1);  | 
| 760 | 134  | 
qed "list_rec_Nil";  | 
| 516 | 135  | 
|
| 5067 | 136  | 
Goal "list_rec(Cons(a,l), c, h) = h(a, l, list_rec(l,c,h))";  | 
| 516 | 137  | 
by (rtac (list_rec_def RS def_Vrec RS trans) 1);  | 
| 5529 | 138  | 
by (simp_tac (rank_ss addsimps rank_Cons2::list.case_eqns) 1);  | 
| 760 | 139  | 
qed "list_rec_Cons";  | 
| 516 | 140  | 
|
| 2469 | 141  | 
Addsimps [list_rec_Nil, list_rec_Cons];  | 
142  | 
||
143  | 
||
| 516 | 144  | 
(*Type checking -- proved by induction, as usual*)  | 
| 5321 | 145  | 
val prems = Goal  | 
| 516 | 146  | 
"[| l: list(A); \  | 
147  | 
\ c: C(Nil); \  | 
|
148  | 
\ !!x y r. [| x:A; y: list(A); r: C(y) |] ==> h(x,y,r): C(Cons(x,y)) \  | 
|
149  | 
\ |] ==> list_rec(l,c,h) : C(l)";  | 
|
150  | 
by (list_ind_tac "l" prems 1);  | 
|
| 4091 | 151  | 
by (ALLGOALS (asm_simp_tac (simpset() addsimps prems)));  | 
| 760 | 152  | 
qed "list_rec_type";  | 
| 516 | 153  | 
|
154  | 
(** Versions for use with definitions **)  | 
|
155  | 
||
| 5321 | 156  | 
val [rew] = Goal "[| !!l. j(l)==list_rec(l, c, h) |] ==> j(Nil) = c";  | 
| 516 | 157  | 
by (rewtac rew);  | 
158  | 
by (rtac list_rec_Nil 1);  | 
|
| 760 | 159  | 
qed "def_list_rec_Nil";  | 
| 516 | 160  | 
|
| 5321 | 161  | 
val [rew] = Goal "[| !!l. j(l)==list_rec(l, c, h) |] ==> j(Cons(a,l)) = h(a,l,j(l))";  | 
| 516 | 162  | 
by (rewtac rew);  | 
163  | 
by (rtac list_rec_Cons 1);  | 
|
| 760 | 164  | 
qed "def_list_rec_Cons";  | 
| 516 | 165  | 
|
166  | 
fun list_recs def = map standard  | 
|
| 1461 | 167  | 
([def] RL [def_list_rec_Nil, def_list_rec_Cons]);  | 
| 516 | 168  | 
|
169  | 
(** map **)  | 
|
170  | 
||
171  | 
val [map_Nil,map_Cons] = list_recs map_def;  | 
|
| 2469 | 172  | 
Addsimps [map_Nil, map_Cons];  | 
| 516 | 173  | 
|
| 5321 | 174  | 
val prems = Goalw [map_def]  | 
| 516 | 175  | 
"[| l: list(A); !!x. x: A ==> h(x): B |] ==> map(h,l) : list(B)";  | 
176  | 
by (REPEAT (ares_tac (prems @ list.intrs @ [list_rec_type]) 1));  | 
|
| 760 | 177  | 
qed "map_type";  | 
| 516 | 178  | 
|
| 5321 | 179  | 
Goal "l: list(A) ==> map(h,l) : list({h(u). u:A})";
 | 
180  | 
by (etac map_type 1);  | 
|
| 516 | 181  | 
by (etac RepFunI 1);  | 
| 760 | 182  | 
qed "map_type2";  | 
| 516 | 183  | 
|
184  | 
(** length **)  | 
|
185  | 
||
186  | 
val [length_Nil,length_Cons] = list_recs length_def;  | 
|
| 2469 | 187  | 
Addsimps [length_Nil,length_Cons];  | 
| 516 | 188  | 
|
| 5067 | 189  | 
Goalw [length_def]  | 
| 
5147
 
825877190618
More tidying and removal of "\!\!... from Goal commands
 
paulson 
parents: 
5137 
diff
changeset
 | 
190  | 
"l: list(A) ==> length(l) : nat";  | 
| 516 | 191  | 
by (REPEAT (ares_tac [list_rec_type, nat_0I, nat_succI] 1));  | 
| 760 | 192  | 
qed "length_type";  | 
| 516 | 193  | 
|
194  | 
(** app **)  | 
|
195  | 
||
196  | 
val [app_Nil,app_Cons] = list_recs app_def;  | 
|
| 2469 | 197  | 
Addsimps [app_Nil, app_Cons];  | 
| 516 | 198  | 
|
| 5067 | 199  | 
Goalw [app_def]  | 
| 
5147
 
825877190618
More tidying and removal of "\!\!... from Goal commands
 
paulson 
parents: 
5137 
diff
changeset
 | 
200  | 
"[| xs: list(A); ys: list(A) |] ==> xs@ys : list(A)";  | 
| 516 | 201  | 
by (REPEAT (ares_tac [list_rec_type, list.Cons_I] 1));  | 
| 760 | 202  | 
qed "app_type";  | 
| 516 | 203  | 
|
204  | 
(** rev **)  | 
|
205  | 
||
206  | 
val [rev_Nil,rev_Cons] = list_recs rev_def;  | 
|
| 2469 | 207  | 
Addsimps [rev_Nil,rev_Cons] ;  | 
| 516 | 208  | 
|
| 5067 | 209  | 
Goalw [rev_def]  | 
| 
5147
 
825877190618
More tidying and removal of "\!\!... from Goal commands
 
paulson 
parents: 
5137 
diff
changeset
 | 
210  | 
"xs: list(A) ==> rev(xs) : list(A)";  | 
| 516 | 211  | 
by (REPEAT (ares_tac (list.intrs @ [list_rec_type, app_type]) 1));  | 
| 760 | 212  | 
qed "rev_type";  | 
| 516 | 213  | 
|
214  | 
||
215  | 
(** flat **)  | 
|
216  | 
||
217  | 
val [flat_Nil,flat_Cons] = list_recs flat_def;  | 
|
| 2469 | 218  | 
Addsimps [flat_Nil,flat_Cons];  | 
| 516 | 219  | 
|
| 5067 | 220  | 
Goalw [flat_def]  | 
| 
5147
 
825877190618
More tidying and removal of "\!\!... from Goal commands
 
paulson 
parents: 
5137 
diff
changeset
 | 
221  | 
"ls: list(list(A)) ==> flat(ls) : list(A)";  | 
| 516 | 222  | 
by (REPEAT (ares_tac (list.intrs @ [list_rec_type, app_type]) 1));  | 
| 760 | 223  | 
qed "flat_type";  | 
| 516 | 224  | 
|
225  | 
||
| 1926 | 226  | 
(** set_of_list **)  | 
227  | 
||
228  | 
val [set_of_list_Nil,set_of_list_Cons] = list_recs set_of_list_def;  | 
|
| 2469 | 229  | 
Addsimps [set_of_list_Nil,set_of_list_Cons];  | 
| 1926 | 230  | 
|
| 5067 | 231  | 
Goalw [set_of_list_def]  | 
| 
5147
 
825877190618
More tidying and removal of "\!\!... from Goal commands
 
paulson 
parents: 
5137 
diff
changeset
 | 
232  | 
"l: list(A) ==> set_of_list(l) : Pow(A)";  | 
| 2033 | 233  | 
by (etac list_rec_type 1);  | 
| 3016 | 234  | 
by (ALLGOALS (Blast_tac));  | 
| 1926 | 235  | 
qed "set_of_list_type";  | 
236  | 
||
| 
5147
 
825877190618
More tidying and removal of "\!\!... from Goal commands
 
paulson 
parents: 
5137 
diff
changeset
 | 
237  | 
Goal "xs: list(A) ==> \  | 
| 1926 | 238  | 
\ set_of_list (xs@ys) = set_of_list(xs) Un set_of_list(ys)";  | 
239  | 
by (etac list.induct 1);  | 
|
| 4091 | 240  | 
by (ALLGOALS (asm_simp_tac (simpset() addsimps [Un_cons])));  | 
| 1926 | 241  | 
qed "set_of_list_append";  | 
242  | 
||
243  | 
||
| 516 | 244  | 
(** list_add **)  | 
245  | 
||
246  | 
val [list_add_Nil,list_add_Cons] = list_recs list_add_def;  | 
|
| 2469 | 247  | 
Addsimps [list_add_Nil,list_add_Cons];  | 
| 516 | 248  | 
|
| 5067 | 249  | 
Goalw [list_add_def]  | 
| 
5147
 
825877190618
More tidying and removal of "\!\!... from Goal commands
 
paulson 
parents: 
5137 
diff
changeset
 | 
250  | 
"xs: list(nat) ==> list_add(xs) : nat";  | 
| 516 | 251  | 
by (REPEAT (ares_tac [list_rec_type, nat_0I, add_type] 1));  | 
| 760 | 252  | 
qed "list_add_type";  | 
| 516 | 253  | 
|
254  | 
val list_typechecks =  | 
|
255  | 
list.intrs @  | 
|
256  | 
[list_rec_type, map_type, map_type2, app_type, length_type,  | 
|
257  | 
rev_type, flat_type, list_add_type];  | 
|
258  | 
||
| 4091 | 259  | 
simpset_ref() := simpset() setSolver (type_auto_tac list_typechecks);  | 
| 516 | 260  | 
|
261  | 
||
262  | 
(*** theorems about map ***)  | 
|
263  | 
||
| 5321 | 264  | 
Goal "l: list(A) ==> map(%u. u, l) = l";  | 
265  | 
by (list_ind_tac "l" [] 1);  | 
|
| 3016 | 266  | 
by (ALLGOALS Asm_simp_tac);  | 
| 760 | 267  | 
qed "map_ident";  | 
| 516 | 268  | 
|
| 5321 | 269  | 
Goal "l: list(A) ==> map(h, map(j,l)) = map(%u. h(j(u)), l)";  | 
270  | 
by (list_ind_tac "l" [] 1);  | 
|
| 3016 | 271  | 
by (ALLGOALS Asm_simp_tac);  | 
| 760 | 272  | 
qed "map_compose";  | 
| 516 | 273  | 
|
| 5321 | 274  | 
Goal "xs: list(A) ==> map(h, xs@ys) = map(h,xs) @ map(h,ys)";  | 
275  | 
by (list_ind_tac "xs" [] 1);  | 
|
| 3016 | 276  | 
by (ALLGOALS Asm_simp_tac);  | 
| 760 | 277  | 
qed "map_app_distrib";  | 
| 516 | 278  | 
|
| 5321 | 279  | 
Goal "ls: list(list(A)) ==> map(h, flat(ls)) = flat(map(map(h),ls))";  | 
280  | 
by (list_ind_tac "ls" [] 1);  | 
|
| 4091 | 281  | 
by (ALLGOALS (asm_simp_tac (simpset() addsimps [map_app_distrib])));  | 
| 760 | 282  | 
qed "map_flat";  | 
| 516 | 283  | 
|
| 5321 | 284  | 
Goal "l: list(A) ==> \  | 
| 516 | 285  | 
\ list_rec(map(h,l), c, d) = \  | 
286  | 
\ list_rec(l, c, %x xs r. d(h(x), map(h,xs), r))";  | 
|
| 5321 | 287  | 
by (list_ind_tac "l" [] 1);  | 
| 3016 | 288  | 
by (ALLGOALS Asm_simp_tac);  | 
| 760 | 289  | 
qed "list_rec_map";  | 
| 516 | 290  | 
|
291  | 
(** theorems about list(Collect(A,P)) -- used in ex/term.ML **)  | 
|
292  | 
||
293  | 
(* c : list(Collect(B,P)) ==> c : list(B) *)  | 
|
| 
782
 
200a16083201
added bind_thm for theorems defined by "standard ..."
 
clasohm 
parents: 
760 
diff
changeset
 | 
294  | 
bind_thm ("list_CollectD", (Collect_subset RS list_mono RS subsetD));
 | 
| 516 | 295  | 
|
| 5321 | 296  | 
Goal "l: list({x:A. h(x)=j(x)}) ==> map(h,l) = map(j,l)";
 | 
297  | 
by (list_ind_tac "l" [] 1);  | 
|
| 3016 | 298  | 
by (ALLGOALS Asm_simp_tac);  | 
| 760 | 299  | 
qed "map_list_Collect";  | 
| 516 | 300  | 
|
301  | 
(*** theorems about length ***)  | 
|
302  | 
||
| 5321 | 303  | 
Goal "xs: list(A) ==> length(map(h,xs)) = length(xs)";  | 
304  | 
by (list_ind_tac "xs" [] 1);  | 
|
| 3016 | 305  | 
by (ALLGOALS Asm_simp_tac);  | 
| 760 | 306  | 
qed "length_map";  | 
| 516 | 307  | 
|
| 5321 | 308  | 
Goal "xs: list(A) ==> length(xs@ys) = length(xs) #+ length(ys)";  | 
309  | 
by (list_ind_tac "xs" [] 1);  | 
|
| 3016 | 310  | 
by (ALLGOALS Asm_simp_tac);  | 
| 760 | 311  | 
qed "length_app";  | 
| 516 | 312  | 
|
313  | 
(* [| m: nat; n: nat |] ==> m #+ succ(n) = succ(n) #+ m  | 
|
314  | 
Used for rewriting below*)  | 
|
315  | 
val add_commute_succ = nat_succI RSN (2,add_commute);  | 
|
316  | 
||
| 5321 | 317  | 
Goal "xs: list(A) ==> length(rev(xs)) = length(xs)";  | 
318  | 
by (list_ind_tac "xs" [] 1);  | 
|
| 4091 | 319  | 
by (ALLGOALS (asm_simp_tac (simpset() addsimps [length_app, add_commute_succ])));  | 
| 760 | 320  | 
qed "length_rev";  | 
| 516 | 321  | 
|
| 5321 | 322  | 
Goal "ls: list(list(A)) ==> length(flat(ls)) = list_add(map(length,ls))";  | 
323  | 
by (list_ind_tac "ls" [] 1);  | 
|
| 4091 | 324  | 
by (ALLGOALS (asm_simp_tac (simpset() addsimps [length_app])));  | 
| 760 | 325  | 
qed "length_flat";  | 
| 516 | 326  | 
|
327  | 
(** Length and drop **)  | 
|
328  | 
||
329  | 
(*Lemma for the inductive step of drop_length*)  | 
|
| 
5147
 
825877190618
More tidying and removal of "\!\!... from Goal commands
 
paulson 
parents: 
5137 
diff
changeset
 | 
330  | 
Goal "xs: list(A) ==> \  | 
| 516 | 331  | 
\ ALL x. EX z zs. drop(length(xs), Cons(x,xs)) = Cons(z,zs)";  | 
332  | 
by (etac list.induct 1);  | 
|
| 2469 | 333  | 
by (ALLGOALS Asm_simp_tac);  | 
| 3016 | 334  | 
by (Blast_tac 1);  | 
| 760 | 335  | 
qed "drop_length_Cons_lemma";  | 
| 
782
 
200a16083201
added bind_thm for theorems defined by "standard ..."
 
clasohm 
parents: 
760 
diff
changeset
 | 
336  | 
bind_thm ("drop_length_Cons", (drop_length_Cons_lemma RS spec));
 | 
| 516 | 337  | 
|
| 
5147
 
825877190618
More tidying and removal of "\!\!... from Goal commands
 
paulson 
parents: 
5137 
diff
changeset
 | 
338  | 
Goal "l: list(A) ==> ALL i: length(l). EX z zs. drop(i,l) = Cons(z,zs)";  | 
| 516 | 339  | 
by (etac list.induct 1);  | 
| 2469 | 340  | 
by (ALLGOALS Asm_simp_tac);  | 
| 516 | 341  | 
by (rtac conjI 1);  | 
342  | 
by (etac drop_length_Cons 1);  | 
|
343  | 
by (rtac ballI 1);  | 
|
344  | 
by (rtac natE 1);  | 
|
345  | 
by (etac ([asm_rl, length_type, Ord_nat] MRS Ord_trans) 1);  | 
|
346  | 
by (assume_tac 1);  | 
|
| 3016 | 347  | 
by (ALLGOALS Asm_simp_tac);  | 
| 4091 | 348  | 
by (ALLGOALS (blast_tac (claset() addIs [succ_in_naturalD, length_type])));  | 
| 760 | 349  | 
qed "drop_length_lemma";  | 
| 
782
 
200a16083201
added bind_thm for theorems defined by "standard ..."
 
clasohm 
parents: 
760 
diff
changeset
 | 
350  | 
bind_thm ("drop_length", (drop_length_lemma RS bspec));
 | 
| 516 | 351  | 
|
352  | 
||
353  | 
(*** theorems about app ***)  | 
|
354  | 
||
| 5321 | 355  | 
Goal "xs: list(A) ==> xs@Nil=xs";  | 
356  | 
by (etac list.induct 1);  | 
|
| 3016 | 357  | 
by (ALLGOALS Asm_simp_tac);  | 
| 760 | 358  | 
qed "app_right_Nil";  | 
| 516 | 359  | 
|
| 5321 | 360  | 
Goal "xs: list(A) ==> (xs@ys)@zs = xs@(ys@zs)";  | 
361  | 
by (list_ind_tac "xs" [] 1);  | 
|
| 3016 | 362  | 
by (ALLGOALS Asm_simp_tac);  | 
| 760 | 363  | 
qed "app_assoc";  | 
| 516 | 364  | 
|
| 5321 | 365  | 
Goal "ls: list(list(A)) ==> flat(ls@ms) = flat(ls)@flat(ms)";  | 
366  | 
by (list_ind_tac "ls" [] 1);  | 
|
| 4091 | 367  | 
by (ALLGOALS (asm_simp_tac (simpset() addsimps [app_assoc])));  | 
| 760 | 368  | 
qed "flat_app_distrib";  | 
| 516 | 369  | 
|
370  | 
(*** theorems about rev ***)  | 
|
371  | 
||
| 5321 | 372  | 
Goal "l: list(A) ==> rev(map(h,l)) = map(h,rev(l))";  | 
373  | 
by (list_ind_tac "l" [] 1);  | 
|
| 4091 | 374  | 
by (ALLGOALS (asm_simp_tac (simpset() addsimps [map_app_distrib])));  | 
| 760 | 375  | 
qed "rev_map_distrib";  | 
| 516 | 376  | 
|
377  | 
(*Simplifier needs the premises as assumptions because rewriting will not  | 
|
378  | 
instantiate the variable ?A in the rules' typing conditions; note that  | 
|
379  | 
rev_type does not instantiate ?A. Only the premises do.  | 
|
380  | 
*)  | 
|
| 
5147
 
825877190618
More tidying and removal of "\!\!... from Goal commands
 
paulson 
parents: 
5137 
diff
changeset
 | 
381  | 
Goal "[| xs: list(A); ys: list(A) |] ==> rev(xs@ys) = rev(ys)@rev(xs)";  | 
| 516 | 382  | 
by (etac list.induct 1);  | 
| 4091 | 383  | 
by (ALLGOALS (asm_simp_tac (simpset() addsimps [app_right_Nil,app_assoc])));  | 
| 760 | 384  | 
qed "rev_app_distrib";  | 
| 516 | 385  | 
|
| 5321 | 386  | 
Goal "l: list(A) ==> rev(rev(l))=l";  | 
387  | 
by (list_ind_tac "l" [] 1);  | 
|
| 4091 | 388  | 
by (ALLGOALS (asm_simp_tac (simpset() addsimps [rev_app_distrib])));  | 
| 760 | 389  | 
qed "rev_rev_ident";  | 
| 516 | 390  | 
|
| 5321 | 391  | 
Goal "ls: list(list(A)) ==> rev(flat(ls)) = flat(map(rev,rev(ls)))";  | 
392  | 
by (list_ind_tac "ls" [] 1);  | 
|
| 4091 | 393  | 
by (ALLGOALS (asm_simp_tac (simpset() addsimps  | 
| 516 | 394  | 
[map_app_distrib, flat_app_distrib, rev_app_distrib, app_right_Nil])));  | 
| 760 | 395  | 
qed "rev_flat";  | 
| 516 | 396  | 
|
397  | 
||
398  | 
(*** theorems about list_add ***)  | 
|
399  | 
||
| 5321 | 400  | 
Goal "[| xs: list(nat); ys: list(nat) |] ==> \  | 
| 516 | 401  | 
\ list_add(xs@ys) = list_add(ys) #+ list_add(xs)";  | 
| 5321 | 402  | 
by (list_ind_tac "xs" [] 1);  | 
| 516 | 403  | 
by (ALLGOALS  | 
| 4091 | 404  | 
(asm_simp_tac (simpset() addsimps [add_0_right, add_assoc RS sym])));  | 
| 516 | 405  | 
by (rtac (add_commute RS subst_context) 1);  | 
406  | 
by (REPEAT (ares_tac [refl, list_add_type] 1));  | 
|
| 760 | 407  | 
qed "list_add_app";  | 
| 516 | 408  | 
|
| 5321 | 409  | 
Goal "l: list(nat) ==> list_add(rev(l)) = list_add(l)";  | 
410  | 
by (list_ind_tac "l" [] 1);  | 
|
| 516 | 411  | 
by (ALLGOALS  | 
| 4091 | 412  | 
(asm_simp_tac (simpset() addsimps [list_add_app, add_0_right])));  | 
| 760 | 413  | 
qed "list_add_rev";  | 
| 516 | 414  | 
|
| 5321 | 415  | 
Goal "ls: list(list(nat)) ==> list_add(flat(ls)) = list_add(map(list_add,ls))";  | 
416  | 
by (list_ind_tac "ls" [] 1);  | 
|
| 4091 | 417  | 
by (ALLGOALS (asm_simp_tac (simpset() addsimps [list_add_app])));  | 
| 516 | 418  | 
by (REPEAT (ares_tac [refl, list_add_type, map_type, add_commute] 1));  | 
| 760 | 419  | 
qed "list_add_flat";  | 
| 516 | 420  | 
|
421  | 
(** New induction rule **)  | 
|
422  | 
||
| 5321 | 423  | 
val major::prems = Goal  | 
| 516 | 424  | 
"[| l: list(A); \  | 
425  | 
\ P(Nil); \  | 
|
426  | 
\ !!x y. [| x: A; y: list(A); P(y) |] ==> P(y @ [x]) \  | 
|
427  | 
\ |] ==> P(l)";  | 
|
428  | 
by (rtac (major RS rev_rev_ident RS subst) 1);  | 
|
429  | 
by (rtac (major RS rev_type RS list.induct) 1);  | 
|
| 4091 | 430  | 
by (ALLGOALS (asm_simp_tac (simpset() addsimps prems)));  | 
| 760 | 431  | 
qed "list_append_induct";  | 
| 516 | 432  |