src/HOL/Induct/SList.ML
author paulson
Wed, 30 Jan 2002 12:22:59 +0100
changeset 12861 7ec4807b53cf
parent 12486 0ed8bdd883e0
permissions -rw-r--r--
mu-syntax for the LEAST operator
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
     1
(*  Title: 	HOL/ex/SList.ML
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
     2
    ID:         SList.ML,v 1.2 1994/12/14 10:17:48 clasohm Exp
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
     3
    Author: 	B. Wolff,  based on a version of Lawrence C Paulson, 
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
     4
		Cambridge University Computer Laboratory
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
     5
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
     6
Definition of type 'a list by a least fixed point
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
     7
*)
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
     8
5977
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 5535
diff changeset
     9
Goalw [List_def] "x : list (range Leaf) ==> x : List";
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 5535
diff changeset
    10
by (Asm_simp_tac 1);
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 5535
diff changeset
    11
qed "ListI";
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 5535
diff changeset
    12
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 5535
diff changeset
    13
Goalw [List_def] "x : List ==> x : list (range Leaf)";
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 5535
diff changeset
    14
by (Asm_simp_tac 1);
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 5535
diff changeset
    15
qed "ListD";
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 5535
diff changeset
    16
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    17
val list_con_defs = [NIL_def, CONS_def];
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    18
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
    19
Goal "list(A) = usum {Numb(0)} (uprod A (list(A)))";
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    20
let val rew = rewrite_rule list_con_defs in  
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
    21
by (fast_tac ((claset()) addSIs (equalityI :: map rew list.intrs)
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    22
                      addEs [rew list.elim]) 1)
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    23
end;
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    24
qed "list_unfold";
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    25
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    26
(*This justifies using list in other recursive type definitions*)
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
    27
Goalw list.defs "!!A B. A<=B ==> list(A) <= list(B)";
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    28
by (rtac lfp_mono 1);
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    29
by (REPEAT (ares_tac basic_monos 1));
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    30
qed "list_mono";
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    31
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    32
(*Type checking -- list creates well-founded sets*)
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4831
diff changeset
    33
Goalw (list_con_defs @ list.defs) "list(sexp) <= sexp";
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    34
by (rtac lfp_lowerbound 1);
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 4033
diff changeset
    35
by (fast_tac (claset() addIs sexp.intrs@[sexp_In0I,sexp_In1I]) 1);
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    36
qed "list_sexp";
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    37
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    38
(* A <= sexp ==> list(A) <= sexp *)
5977
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 5535
diff changeset
    39
bind_thm ("list_subset_sexp", [list_mono, list_sexp] MRS subset_trans);
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    40
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
    41
fun List_simp thm = (asm_full_simplify (HOL_ss addsimps [List_def]) thm)
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
    42
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    43
(*Induction for the type 'a list *)
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
    44
val prems = Goalw [Nil_def, Cons_def]
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    45
    "[| P(Nil);   \
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    46
\       !!x xs. P(xs) ==> P(x # xs) |]  ==> P(l)";
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
    47
by (rtac (Rep_List_inverse RS subst) 1);  
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
    48
			 (*types force good instantiation*)
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
    49
by (rtac ((List_simp Rep_List)  RS list.induct) 1);
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    50
by (REPEAT (ares_tac prems 1
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
    51
            ORELSE eresolve_tac [rangeE, ssubst, 
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
    52
                          (List_simp Abs_List_inverse) RS subst] 1));
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
    53
qed "list_induct";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
    54
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
    55
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
    56
(*** Isomorphisms ***)
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
    57
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
    58
Goal "inj_on Abs_List (list(range Leaf))";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
    59
by (rtac inj_on_inverseI 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
    60
by (etac (List_simp Abs_List_inverse) 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
    61
qed "inj_on_Abs_list";
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    62
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    63
(** Distinctness of constructors **)
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    64
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4831
diff changeset
    65
Goalw list_con_defs "CONS M N ~= NIL";
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    66
by (rtac In1_not_In0 1);
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    67
qed "CONS_not_NIL";
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
    68
val NIL_not_CONS = CONS_not_NIL RS not_sym;
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    69
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    70
bind_thm ("CONS_neq_NIL", (CONS_not_NIL RS notE));
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    71
val NIL_neq_CONS = sym RS CONS_neq_NIL;
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    72
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4831
diff changeset
    73
Goalw [Nil_def,Cons_def] "x # xs ~= Nil";
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
    74
by (rtac (CONS_not_NIL RS (inj_on_Abs_list RS inj_on_contraD)) 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
    75
by (REPEAT (resolve_tac (list.intrs @ [rangeI,(List_simp Rep_List)])1));
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    76
qed "Cons_not_Nil";
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    77
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
    78
bind_thm ("Nil_not_Cons", (Cons_not_Nil RS not_sym));
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
    79
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
    80
bind_thm ("Cons_neq_Nil", (Cons_not_Nil RS notE));
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
    81
val Nil_neq_Cons = sym RS Cons_neq_Nil;
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    82
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    83
(** Injectiveness of CONS and Cons **)
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    84
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
    85
Goalw [CONS_def] "(CONS K M)=(CONS L N) = (K=L & M=N)";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
    86
by (fast_tac (HOL_cs addSEs [Scons_inject, make_elim In1_inject]) 1);
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    87
qed "CONS_CONS_eq";
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    88
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    89
(*For reasoning about abstract list constructors*)
5977
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 5535
diff changeset
    90
AddIs [Rep_List RS ListD, ListI];
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 5535
diff changeset
    91
AddIs list.intrs;
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    92
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    93
AddIffs [CONS_not_NIL, NIL_not_CONS, CONS_CONS_eq];
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    94
11500
a84130c7e6ab Updated proofs to take advantage of additional theorems proved by "typedef"
paulson
parents: 9747
diff changeset
    95
AddSDs [Leaf_inject];
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    96
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4831
diff changeset
    97
Goalw [Cons_def] "(x#xs=y#ys) = (x=y & xs=ys)";
11500
a84130c7e6ab Updated proofs to take advantage of additional theorems proved by "typedef"
paulson
parents: 9747
diff changeset
    98
by (stac (thm "Abs_List_inject") 1);
a84130c7e6ab Updated proofs to take advantage of additional theorems proved by "typedef"
paulson
parents: 9747
diff changeset
    99
by (auto_tac (claset(), simpset() addsimps [thm "Rep_List_inject"])); 
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   100
qed "Cons_Cons_eq";
5977
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 5535
diff changeset
   101
bind_thm ("Cons_inject2", Cons_Cons_eq RS iffD1 RS conjE);
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   102
5977
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 5535
diff changeset
   103
Goal "CONS M N: list(A) ==> M: A & N: list(A)";
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 5535
diff changeset
   104
by (etac setup_induction 1);
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   105
by (etac list.induct 1);
5977
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 5535
diff changeset
   106
by (ALLGOALS Fast_tac);
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   107
qed "CONS_D";
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   108
5977
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 5535
diff changeset
   109
Goalw [CONS_def,In1_def] "CONS M N: sexp ==> M: sexp & N: sexp";
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 4033
diff changeset
   110
by (fast_tac (claset() addSDs [Scons_D]) 1);
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   111
qed "sexp_CONS_D";
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   112
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   113
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   114
(*Reasoning about constructors and their freeness*)
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   115
Addsimps list.intrs;
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   116
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   117
AddIffs [Cons_not_Nil, Nil_not_Cons, Cons_Cons_eq];
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   118
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5069
diff changeset
   119
Goal "N: list(A) ==> !M. N ~= CONS M N";
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   120
by (etac list.induct 1);
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   121
by (ALLGOALS Asm_simp_tac);
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   122
qed "not_CONS_self";
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   123
11500
a84130c7e6ab Updated proofs to take advantage of additional theorems proved by "typedef"
paulson
parents: 9747
diff changeset
   124
Goal "ALL x. l ~= x#l";
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   125
by (induct_thm_tac list_induct "l" 1);
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   126
by (ALLGOALS Asm_simp_tac);
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   127
qed "not_Cons_self2";
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   128
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   129
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4831
diff changeset
   130
Goal "(xs ~= []) = (? y ys. xs = y#ys)";
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   131
by (induct_thm_tac list_induct "xs" 1);
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   132
by (Simp_tac 1);
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   133
by (Asm_simp_tac 1);
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   134
by (REPEAT(resolve_tac [exI,refl,conjI] 1));
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   135
qed "neq_Nil_conv2";
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   136
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   137
(** Conversion rules for List_case: case analysis operator **)
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   138
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4831
diff changeset
   139
Goalw [List_case_def,NIL_def] "List_case c h NIL = c";
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   140
by (rtac Case_In0 1);
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   141
qed "List_case_NIL";
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   142
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4831
diff changeset
   143
Goalw [List_case_def,CONS_def]  "List_case c h (CONS M N) = h M N";
4521
c7f56322a84b Tidied by adding more default simprules
paulson
parents: 4089
diff changeset
   144
by (Simp_tac 1);
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   145
qed "List_case_CONS";
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   146
4521
c7f56322a84b Tidied by adding more default simprules
paulson
parents: 4089
diff changeset
   147
Addsimps [List_case_NIL, List_case_CONS];
c7f56322a84b Tidied by adding more default simprules
paulson
parents: 4089
diff changeset
   148
c7f56322a84b Tidied by adding more default simprules
paulson
parents: 4089
diff changeset
   149
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   150
(*** List_rec -- by wf recursion on pred_sexp ***)
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   151
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   152
(* The trancl(pred_sexp) is essential because pred_sexp_CONS_I1,2 would not
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   153
   hold if pred_sexp^+ were changed to pred_sexp. *)
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   154
5278
a903b66822e2 even more tidying of Goal commands
paulson
parents: 5223
diff changeset
   155
Goal "(%M. List_rec M c d) = wfrec (trancl pred_sexp) \
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   156
                           \     (%g. List_case c (%x y. d x y (g y)))";
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   157
by (simp_tac (HOL_ss addsimps [List_rec_def]) 1);
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   158
val List_rec_unfold = standard 
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   159
  ((wf_pred_sexp RS wf_trancl) RS ((result() RS eq_reflection) RS def_wfrec));
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   160
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   161
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   162
(*---------------------------------------------------------------------------
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   163
 * Old:
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   164
 * val List_rec_unfold = [List_rec_def,wf_pred_sexp RS wf_trancl] MRS def_wfrec
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   165
 *                     |> standard;
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   166
 *---------------------------------------------------------------------------*)
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   167
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   168
(** pred_sexp lemmas **)
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   169
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4831
diff changeset
   170
Goalw [CONS_def,In1_def]
5148
74919e8f221c More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5143
diff changeset
   171
    "[| M: sexp;  N: sexp |] ==> (M, CONS M N) : pred_sexp^+";
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   172
by (Asm_simp_tac 1);
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   173
qed "pred_sexp_CONS_I1";
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   174
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4831
diff changeset
   175
Goalw [CONS_def,In1_def]
5148
74919e8f221c More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5143
diff changeset
   176
    "[| M: sexp;  N: sexp |] ==> (N, CONS M N) : pred_sexp^+";
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   177
by (Asm_simp_tac 1);
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   178
qed "pred_sexp_CONS_I2";
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   179
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   180
Goal
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   181
    "(CONS M1 M2, N) : pred_sexp^+ ==> \
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   182
\    (M1,N) : pred_sexp^+ & (M2,N) : pred_sexp^+";
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   183
by (ftac (pred_sexp_subset_Sigma RS trancl_subset_Sigma RS subsetD) 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   184
by (blast_tac (claset() addSDs [sexp_CONS_D] 
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   185
                        addIs  [pred_sexp_CONS_I1, pred_sexp_CONS_I2,
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   186
                                trans_trancl RS transD]) 1); 
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   187
qed "pred_sexp_CONS_D";
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   188
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   189
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   190
(** Conversion rules for List_rec **)
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   191
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4831
diff changeset
   192
Goal "List_rec NIL c h = c";
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   193
by (rtac (List_rec_unfold RS trans) 1);
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   194
by (simp_tac (HOL_ss addsimps [List_case_NIL]) 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   195
qed "List_rec_NIL"; 
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   196
Addsimps [List_rec_NIL];
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   197
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5069
diff changeset
   198
Goal "[| M: sexp;  N: sexp |] ==> \
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   199
\    List_rec (CONS M N) c h = h M N (List_rec N c h)";
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   200
by (rtac (List_rec_unfold RS trans) 1);
4521
c7f56322a84b Tidied by adding more default simprules
paulson
parents: 4089
diff changeset
   201
by (asm_simp_tac (simpset() addsimps [pred_sexp_CONS_I2]) 1);
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   202
qed "List_rec_CONS";
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   203
Addsimps [List_rec_CONS];
4521
c7f56322a84b Tidied by adding more default simprules
paulson
parents: 4089
diff changeset
   204
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   205
(*** list_rec -- by List_rec ***)
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   206
5977
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 5535
diff changeset
   207
val Rep_List_in_sexp =
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 5535
diff changeset
   208
    [range_Leaf_subset_sexp RS list_subset_sexp, Rep_List RS ListD] 
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 5535
diff changeset
   209
    MRS subsetD;
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   210
11500
a84130c7e6ab Updated proofs to take advantage of additional theorems proved by "typedef"
paulson
parents: 9747
diff changeset
   211
val list_rec_simps = [ListI RS Abs_List_inverse, Rep_List_inverse,
a84130c7e6ab Updated proofs to take advantage of additional theorems proved by "typedef"
paulson
parents: 9747
diff changeset
   212
		      Rep_List RS ListD, rangeI, inj_Leaf, inv_f_f,
a84130c7e6ab Updated proofs to take advantage of additional theorems proved by "typedef"
paulson
parents: 9747
diff changeset
   213
		      sexp.LeafI, Rep_List_in_sexp];
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   214
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   215
11500
a84130c7e6ab Updated proofs to take advantage of additional theorems proved by "typedef"
paulson
parents: 9747
diff changeset
   216
Goal "list_rec Nil c h = c";
a84130c7e6ab Updated proofs to take advantage of additional theorems proved by "typedef"
paulson
parents: 9747
diff changeset
   217
by (simp_tac (simpset() addsimps list_rec_simps@ [list_rec_def, Nil_def]) 1);
a84130c7e6ab Updated proofs to take advantage of additional theorems proved by "typedef"
paulson
parents: 9747
diff changeset
   218
qed "list_rec_Nil";
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   219
Addsimps [list_rec_Nil];
11500
a84130c7e6ab Updated proofs to take advantage of additional theorems proved by "typedef"
paulson
parents: 9747
diff changeset
   220
a84130c7e6ab Updated proofs to take advantage of additional theorems proved by "typedef"
paulson
parents: 9747
diff changeset
   221
Goal "list_rec (a#l) c h = h a l (list_rec l c h)";
a84130c7e6ab Updated proofs to take advantage of additional theorems proved by "typedef"
paulson
parents: 9747
diff changeset
   222
by (simp_tac (simpset() addsimps list_rec_simps@ [list_rec_def,Cons_def]) 1);
a84130c7e6ab Updated proofs to take advantage of additional theorems proved by "typedef"
paulson
parents: 9747
diff changeset
   223
qed "list_rec_Cons";
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   224
Addsimps [list_rec_Cons];
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   225
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   226
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   227
(*Type checking.  Useful?*)
5977
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 5535
diff changeset
   228
val major::A_subset_sexp::prems = 
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 5535
diff changeset
   229
Goal "[| M: list(A);     \
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 5535
diff changeset
   230
\        A<=sexp;        \
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 5535
diff changeset
   231
\        c: C(NIL);      \
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 5535
diff changeset
   232
\        !!x y r. [| x: A;  y: list(A);  r: C(y) |] ==> h x y r: C(CONS x y) \
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 5535
diff changeset
   233
\     |] ==> List_rec M c h : C(M :: 'a item)";
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   234
val sexp_ListA_I = A_subset_sexp RS list_subset_sexp RS subsetD;
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   235
val sexp_A_I = A_subset_sexp RS subsetD;
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   236
by (rtac (major RS list.induct) 1);
5535
678999604ee9 deleted needless parentheses
paulson
parents: 5278
diff changeset
   237
by (ALLGOALS(asm_simp_tac (simpset() addsimps [sexp_A_I,sexp_ListA_I]@prems)));
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   238
qed "List_rec_type";
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   239
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   240
(** Generalized map functionals **)
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   241
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4831
diff changeset
   242
Goalw [Rep_map_def] "Rep_map f Nil = NIL";
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   243
by (rtac list_rec_Nil 1);
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   244
qed "Rep_map_Nil";
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   245
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   246
Goalw [Rep_map_def]
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   247
    "Rep_map f(x#xs) = CONS(f x)(Rep_map f xs)";
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   248
by (rtac list_rec_Cons 1);
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   249
qed "Rep_map_Cons";
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   250
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   251
Goalw [Rep_map_def] "!!f. (!!x. f(x): A) ==> Rep_map f xs: list(A)";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   252
by (rtac list_induct 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   253
by Auto_tac; 
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   254
qed "Rep_map_type";
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   255
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4831
diff changeset
   256
Goalw [Abs_map_def] "Abs_map g NIL = Nil";
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   257
by (rtac List_rec_NIL 1);
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   258
qed "Abs_map_NIL";
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   259
5977
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 5535
diff changeset
   260
Goalw [Abs_map_def]
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 5535
diff changeset
   261
    "[| M: sexp;  N: sexp |] ==> Abs_map g (CONS M N) = g(M) # Abs_map g N";
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 5535
diff changeset
   262
by (REPEAT (ares_tac [List_rec_CONS] 1));
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   263
qed "Abs_map_CONS";
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   264
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   265
(*These 2 rules ease the use of primitive recursion.  NOTE USE OF == *)
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   266
val [rew] = goal thy
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   267
    "[| !!xs. f(xs) == list_rec xs c h  |] ==> f []  = c";
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   268
by (rewtac rew);
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   269
by (rtac list_rec_Nil 1);
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   270
qed "def_list_rec_Nil";
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   271
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   272
val [rew] = goal thy
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   273
    "[| !!xs. f(xs) == list_rec xs c h  |] ==> f(x#xs) = h x xs (f xs)";
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   274
by (rewtac rew);
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   275
by (rtac list_rec_Cons 1);
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   276
qed "def_list_rec_Cons";
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   277
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   278
Addsimps [Rep_map_Nil, Rep_map_Cons, Abs_map_NIL, Abs_map_CONS];
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   279
5977
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 5535
diff changeset
   280
val [major,A_subset_sexp,minor] = 
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 5535
diff changeset
   281
Goal "[| M: list(A);  A<=sexp;  !!z. z: A ==> f(g(z)) = z |] \
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 5535
diff changeset
   282
\     ==> Rep_map f (Abs_map g M) = M";
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   283
by (rtac (major RS list.induct) 1);
5977
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 5535
diff changeset
   284
by (ALLGOALS 
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 5535
diff changeset
   285
    (asm_simp_tac (simpset() addsimps [sexp_A_I,sexp_ListA_I,minor])));
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   286
qed "Abs_map_inverse";
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   287
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   288
(*Rep_map_inverse is obtained via Abs_Rep_map and map_ident*)
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   289
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   290
(** list_case **)
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   291
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   292
(* setting up rewrite sets *)
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   293
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   294
fun list_recs def =
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   295
      [standard (def RS def_list_rec_Nil),
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   296
       standard (def RS def_list_rec_Cons)];
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   297
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   298
val [list_case_Nil,list_case_Cons] = list_recs list_case_def;
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   299
Addsimps [list_case_Nil,list_case_Cons];
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   300
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   301
(*FIXME??
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   302
val slist_ss = (simpset()) addsimps
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   303
	  [Cons_not_Nil, Nil_not_Cons, Cons_Cons_eq,
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   304
	   list_rec_Nil, list_rec_Cons,
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   305
	   slist_case_Nil,slist_case_Cons];
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   306
*)
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   307
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   308
(** list_case **)
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   309
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   310
Goal
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   311
 "P(list_case a f xs) = ((xs=[] --> P a ) & (!y ys. xs=y#ys --> P(f y ys)))";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   312
by (induct_thm_tac list_induct "xs" 1);
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   313
by (ALLGOALS Asm_simp_tac);
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   314
qed "expand_list_case";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   315
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   316
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   317
(**** Function definitions ****)
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   318
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   319
fun list_recs def =
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   320
      [standard (def RS def_list_rec_Nil),
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   321
       standard (def RS def_list_rec_Cons)];
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   322
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   323
(*** Unfolding the basic combinators ***)
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   324
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   325
val [null_Nil,null_Cons] = list_recs null_def;
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   326
val [_,hd_Cons] = list_recs hd_def;
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   327
val [_,tl_Cons] = list_recs tl_def;
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   328
val [ttl_Nil,ttl_Cons] = list_recs ttl_def;
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   329
val [append_Nil,append_Cons] = list_recs append_def;
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   330
val [mem_Nil, mem_Cons] = list_recs mem_def;
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   331
val [map_Nil,map_Cons] = list_recs map_def;
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   332
val [filter_Nil,filter_Cons] = list_recs filter_def;
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   333
val [list_all_Nil,list_all_Cons] = list_recs list_all_def;
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   334
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   335
store_thm("hd_Cons",hd_Cons);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   336
store_thm("tl_Cons",tl_Cons);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   337
store_thm("ttl_Nil" ,ttl_Nil);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   338
store_thm("ttl_Cons" ,ttl_Cons);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   339
store_thm("append_Nil", append_Nil);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   340
store_thm("append_Cons", append_Cons);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   341
store_thm("mem_Nil" ,mem_Nil);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   342
store_thm("mem_Cons" ,mem_Cons);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   343
store_thm("map_Nil", map_Nil);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   344
store_thm("map_Cons", map_Cons);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   345
store_thm("filter_Nil", filter_Nil);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   346
store_thm("filter_Cons", filter_Cons);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   347
store_thm("list_all_Nil", list_all_Nil);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   348
store_thm("list_all_Cons", list_all_Cons);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   349
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   350
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   351
Addsimps
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   352
  [null_Nil, null_Cons, hd_Cons, tl_Cons, ttl_Nil, ttl_Cons,
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   353
   mem_Nil, mem_Cons,
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   354
   append_Nil, append_Cons,
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   355
   map_Nil, map_Cons,
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   356
   list_all_Nil, list_all_Cons,
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   357
   filter_Nil, filter_Cons];
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   358
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   359
(** nth **)
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   360
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   361
val [rew] = goal Nat.thy
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   362
    "[| !!n. f == nat_rec c h |] ==> f(0) = c";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   363
by (rewtac rew);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   364
by (rtac nat_rec_0 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   365
qed "def_nat_rec_0_eta";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   366
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   367
val [rew] = goal Nat.thy
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   368
    "[| !!n. f == nat_rec c h |] ==> f(Suc(n)) = h n (f n)";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   369
by (rewtac rew);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   370
by (rtac nat_rec_Suc 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   371
qed "def_nat_rec_Suc_eta";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   372
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   373
fun nat_recs_eta def =
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   374
      [standard (def RS def_nat_rec_0_eta),
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   375
       standard (def RS def_nat_rec_Suc_eta)];
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   376
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   377
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   378
val [nth_0,nth_Suc] = nat_recs_eta nth_def; 
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   379
store_thm("nth_0",nth_0);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   380
store_thm("nth_Suc",nth_Suc);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   381
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   382
Addsimps [nth_0,nth_Suc];
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   383
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   384
(** length **)
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   385
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   386
Goalw [length_def] "length([]) = 0";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   387
by (ALLGOALS Asm_simp_tac);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   388
qed "length_Nil";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   389
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   390
Goalw [length_def] "length(a#xs) = Suc(length(xs))";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   391
by (ALLGOALS Asm_simp_tac);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   392
qed "length_Cons";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   393
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   394
Addsimps [length_Nil,length_Cons];
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   395
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   396
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   397
(** @ - append **)
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   398
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   399
Goal "(xs@ys)@zs = xs@(ys@zs)";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   400
by (induct_thm_tac list_induct "xs" 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   401
by (ALLGOALS Asm_simp_tac);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   402
by (ALLGOALS Asm_simp_tac);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   403
qed "append_assoc";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   404
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   405
Goal "xs @ [] = xs";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   406
by (induct_thm_tac list_induct "xs" 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   407
by (ALLGOALS Asm_simp_tac);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   408
qed "append_Nil2";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   409
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   410
(** mem **)
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   411
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   412
Goal "x mem (xs@ys) = (x mem xs | x mem ys)";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   413
by (induct_thm_tac list_induct "xs" 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   414
by (ALLGOALS Asm_simp_tac);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   415
qed "mem_append";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   416
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   417
Goal "x mem [x:xs. P x ] = (x mem xs & P(x))";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   418
by (induct_thm_tac list_induct "xs" 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   419
by (ALLGOALS Asm_simp_tac);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   420
qed "mem_filter";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   421
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   422
(** list_all **)
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   423
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   424
Goal "(Alls x:xs. True) = True";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   425
by (induct_thm_tac list_induct "xs" 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   426
by (ALLGOALS Asm_simp_tac);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   427
qed "list_all_True";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   428
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   429
Goal "list_all p (xs@ys) = ((list_all p xs) & (list_all p ys))";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   430
by (induct_thm_tac list_induct "xs" 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   431
by (ALLGOALS Asm_simp_tac);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   432
qed "list_all_conj";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   433
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   434
Goal "(Alls x:xs. P(x)) = (!x. x mem xs --> P(x))";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   435
by (induct_thm_tac list_induct "xs" 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   436
by (ALLGOALS Asm_simp_tac);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   437
by (fast_tac HOL_cs 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   438
qed "list_all_mem_conv";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   439
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   440
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   441
Goal "(! n. P n) = (P 0 & (! n. P (Suc n)))";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   442
by (Auto_tac);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   443
by (induct_tac "n" 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   444
by (Auto_tac);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   445
qed "nat_case_dist";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   446
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   447
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   448
val [] = Goal "(Alls u:A. P u) = (!n. n < length A --> P(nth n A))";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   449
by (induct_thm_tac list_induct "A" 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   450
by (ALLGOALS Asm_simp_tac);
12486
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12169
diff changeset
   451
by (rtac trans 1);
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12169
diff changeset
   452
by (rtac (nat_case_dist RS sym) 2);
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   453
by (ALLGOALS Asm_simp_tac);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   454
qed "alls_P_eq_P_nth";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   455
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   456
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   457
Goal "[| !x. P x --> Q x;  (Alls x:xs. P(x)) |] ==> (Alls x:xs. Q(x))";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   458
by (asm_full_simp_tac (simpset() addsimps [list_all_mem_conv]) 1); 
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   459
qed "list_all_imp";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   460
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   461
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   462
(** The functional "map" and the generalized functionals **)
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   463
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   464
val prems = 
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   465
Goal "(!!x. f(x): sexp) ==> \
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   466
\       Abs_map g (Rep_map f xs) = map (%t. g(f(t))) xs";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   467
by (induct_thm_tac list_induct "xs" 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   468
by (ALLGOALS (asm_simp_tac(simpset() addsimps
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   469
			   (prems@[Rep_map_type, list_sexp RS subsetD]))));
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   470
qed "Abs_Rep_map";
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   471
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   472
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   473
(** Additional mapping lemmas **)
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   474
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   475
Goal "map(%x. x)(xs) = xs";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   476
by (induct_thm_tac list_induct "xs" 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   477
by (ALLGOALS Asm_simp_tac);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   478
qed "map_ident";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   479
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   480
Goal "map f (xs@ys) = map f xs  @ map f ys";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   481
by (induct_thm_tac list_induct "xs" 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   482
by (ALLGOALS Asm_simp_tac);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   483
qed "map_append";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   484
Addsimps[map_append];
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   485
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   486
Goalw [o_def] "map(f o g)(xs) = map f (map g xs)";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   487
by (induct_thm_tac list_induct "xs" 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   488
by (ALLGOALS Asm_simp_tac);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   489
qed "map_compose";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   490
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   491
Addsimps
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   492
  [mem_append, mem_filter, append_assoc, append_Nil2, map_ident,
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   493
   list_all_True, list_all_conj];
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   494
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   495
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   496
Goal
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   497
"x mem (map f q) --> (? y. y mem q & x = f y)";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   498
by (induct_thm_tac list_induct "q" 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   499
by (ALLGOALS Asm_simp_tac);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   500
by (case_tac "f xa = x" 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   501
by (ALLGOALS Asm_simp_tac);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   502
by (res_inst_tac [("x","xa")] exI 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   503
by (ALLGOALS Asm_simp_tac);
12486
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12169
diff changeset
   504
by (rtac impI 1);
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   505
by (rotate_tac 1 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   506
by (ALLGOALS  Asm_full_simp_tac);
12486
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12169
diff changeset
   507
by (etac exE 1); by (etac conjE 1);
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   508
by (res_inst_tac [("x","y")] exI 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   509
by (asm_simp_tac (HOL_ss addsimps [if_cancel]) 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   510
qed "mem_map_aux1";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   511
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   512
Goal
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   513
"(? y. y mem q & x = f y) --> x mem (map f q)";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   514
by (induct_thm_tac list_induct "q" 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   515
by (Asm_simp_tac 1);
12486
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12169
diff changeset
   516
by (rtac impI 1);
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12169
diff changeset
   517
by (etac exE 1);
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12169
diff changeset
   518
by (etac conjE 1);
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   519
by (ALLGOALS Asm_simp_tac);
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   520
by (case_tac "xa = y" 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   521
by (rotate_tac 2 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   522
by (asm_full_simp_tac (HOL_ss addsimps [if_cancel]) 1);
12486
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12169
diff changeset
   523
by (etac impE 1);
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   524
by (asm_simp_tac (HOL_ss addsimps [if_cancel]) 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   525
by (case_tac "f xa = f y" 2);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   526
by (res_inst_tac [("x","y")] exI 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   527
by (asm_simp_tac (HOL_ss addsimps [if_cancel]) 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   528
by (Auto_tac);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   529
qed "mem_map_aux2";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   530
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   531
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   532
Goal
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   533
"x mem (map f q) = (? y. y mem q & x = f y)";
12486
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12169
diff changeset
   534
by (rtac iffI 1);
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12169
diff changeset
   535
by (rtac (mem_map_aux1 RS mp) 1);
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12169
diff changeset
   536
by (rtac (mem_map_aux2 RS mp) 2);
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   537
by (ALLGOALS atac);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   538
qed "mem_map";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   539
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   540
Goal "A ~= [] --> hd(A @ B) = hd(A)";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   541
by (induct_thm_tac list_induct "A" 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   542
by Auto_tac; 
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   543
qed_spec_mp "hd_append";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   544
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   545
Goal "A ~= [] --> tl(A @ B) = tl(A) @ B";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   546
by (induct_thm_tac list_induct "A" 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   547
by Auto_tac; 
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   548
qed_spec_mp "tl_append";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   549
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   550
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   551
(* ********************************************************************* *)
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   552
(* More ...         					                 *)
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   553
(* ********************************************************************* *)
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   554
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   555
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   556
(** take **)
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   557
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   558
Goal "take [] (Suc x) = []";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   559
by (asm_simp_tac  (simpset()) 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   560
qed "take_Suc1";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   561
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   562
Goal "take(a#xs)(Suc x) = a#take xs x";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   563
by (asm_simp_tac (simpset()) 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   564
qed "take_Suc2";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   565
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   566
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   567
(** drop **)
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   568
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   569
Goalw [drop_def] "drop xs 0 = xs";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   570
by (asm_simp_tac  (simpset()) 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   571
qed "drop_0";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   572
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   573
Goalw [drop_def] "drop [] (Suc x) = []";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   574
by (induct_tac "x" 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   575
by (ALLGOALS (asm_full_simp_tac ((simpset()) addsimps [ttl_Nil]) ));
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   576
qed "drop_Suc1";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   577
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   578
Goalw [drop_def] "drop(a#xs)(Suc x) = drop xs x";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   579
by (asm_simp_tac (simpset()) 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   580
qed "drop_Suc2";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   581
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   582
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   583
(** copy **)
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   584
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   585
Goalw [copy_def] "copy x 0 = []";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   586
by (asm_simp_tac (simpset()) 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   587
qed "copy_0";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   588
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   589
Goalw [copy_def] "copy x (Suc y) = x # copy x y";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   590
by (asm_simp_tac (simpset()) 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   591
qed "copy_Suc";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   592
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   593
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   594
(** fold **)
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   595
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   596
Goalw [foldl_def] "foldl f a [] = a";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   597
by (ALLGOALS Asm_simp_tac);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   598
qed "foldl_Nil";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   599
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   600
Goalw [foldl_def] "foldl f a(x#xs) = foldl f (f a x) xs";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   601
by (induct_thm_tac list_induct "xs" 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   602
by (ALLGOALS Asm_full_simp_tac);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   603
qed "foldl_Cons";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   604
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   605
Goalw [foldr_def] "foldr f a [] = a";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   606
by (ALLGOALS Asm_simp_tac);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   607
qed "foldr_Nil";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   608
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   609
Goalw [foldr_def] "foldr f z(x#xs)  = f x (foldr f z xs)";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   610
by (ALLGOALS Asm_simp_tac);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   611
qed "foldr_Cons";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   612
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   613
Addsimps  			
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   614
		[length_Nil,length_Cons,
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   615
		 take_0, take_Suc1,take_Suc2,
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   616
		 drop_0,drop_Suc1,drop_Suc2,copy_0,copy_Suc,
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   617
		 foldl_Nil,foldl_Cons,foldr_Nil,foldr_Cons];
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   618
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   619
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   620
(** flat **)
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   621
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   622
Goalw [flat_def] 
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   623
"flat [] = []";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   624
by (ALLGOALS Asm_simp_tac);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   625
qed "flat_Nil";
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   626
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   627
Goalw [flat_def] 
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   628
"flat (x # xs) = x @ flat xs";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   629
by (ALLGOALS Asm_simp_tac);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   630
qed "flat_Cons";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   631
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   632
Addsimps  [flat_Nil,flat_Cons];			
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   633
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   634
(** rev **)
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   635
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   636
Goalw [rev_def] 
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   637
"rev [] = []";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   638
by (ALLGOALS Asm_simp_tac);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   639
qed "rev_Nil";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   640
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   641
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   642
Goalw [rev_def] 
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   643
"rev (x # xs) = rev xs @ [x]";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   644
by (ALLGOALS Asm_simp_tac);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   645
qed "rev_Cons";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   646
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   647
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   648
Addsimps [rev_Nil,rev_Cons];			
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   649
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   650
(** zip **)
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   651
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   652
Goalw [zipWith_def] 
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   653
"zipWith f (a#as,b#bs)   = f(a,b) # zipWith f (as,bs)";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   654
by (ALLGOALS Asm_simp_tac);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   655
qed "zipWith_Cons_Cons";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   656
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   657
Goalw [zipWith_def] 
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   658
"zipWith f ([],[])      = []";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   659
by (ALLGOALS Asm_simp_tac);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   660
qed "zipWith_Nil_Nil";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   661
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   662
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   663
Goalw [zipWith_def] 
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   664
"zipWith f (x,[])  = []";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   665
by (induct_thm_tac list_induct "x" 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   666
by (ALLGOALS Asm_simp_tac);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   667
qed "zipWith_Cons_Nil";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   668
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   669
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   670
Goalw [zipWith_def] 
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   671
"zipWith f ([],x) = []";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   672
by (induct_thm_tac list_induct "x" 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   673
by (ALLGOALS Asm_simp_tac);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   674
qed "zipWith_Nil_Cons";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   675
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   676
Goalw [unzip_def] "unzip [] = ([],[])";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   677
by (ALLGOALS Asm_simp_tac);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   678
qed "unzip_Nil";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   679
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   680
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   681
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   682
(** SOME LIST THEOREMS **)
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   683
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   684
(* SQUIGGOL LEMMAS *)
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   685
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   686
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   687
Goalw [o_def] "map(f o g) = ((map f) o (map g))";
12486
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12169
diff changeset
   688
by (rtac ext 1);
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   689
by (simp_tac (HOL_ss addsimps [map_compose RS sym,o_def]) 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   690
qed "map_compose_ext";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   691
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   692
Goal "map f (flat S) = flat(map (map f) S)";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   693
by (induct_thm_tac list_induct "S" 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   694
by (ALLGOALS Asm_simp_tac);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   695
qed "map_flat";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   696
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   697
Goal "(Alls u:xs. f(u) = g(u)) --> map f xs = map g xs";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   698
by (induct_thm_tac list_induct "xs" 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   699
by (ALLGOALS Asm_simp_tac);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   700
qed "list_all_map_eq";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   701
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   702
Goal "filter p (map f xs) = map f (filter(p o f)(xs))";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   703
by (induct_thm_tac list_induct "xs" 1);
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   704
by (ALLGOALS Asm_simp_tac);
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   705
qed "filter_map_d";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   706
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   707
Goal "filter p (filter q xs) = filter(%x. p x & q x) xs";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   708
by (induct_thm_tac list_induct "xs" 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   709
by (ALLGOALS Asm_simp_tac);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   710
qed "filter_compose";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   711
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   712
(* "filter(p, filter(q ,xs)) = filter(q, filter(p ,xs))",
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   713
   "filter(p, filter(p ,xs)) = filter(p,xs)" BIRD's thms.*)
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   714
 
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   715
Goal "ALL B. filter p (A @ B) = (filter p A @ filter p B)";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   716
by (induct_thm_tac list_induct "A" 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   717
by (ALLGOALS Asm_simp_tac);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   718
qed_spec_mp "filter_append";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   719
Addsimps [filter_append];
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   720
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   721
(* inits(xs) == map(fst,splits(xs)), 
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   722
 
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   723
   splits([]) = []
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   724
   splits(a # xs) = <[],xs> @ map(%x. <a # fst(x),snd(x)>, splits(xs))
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   725
   (x @ y = z) = <x,y> mem splits(z)
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   726
   x mem xs & y mem ys = <x,y> mem diag(xs,ys) *)
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   727
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   728
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   729
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   730
Goal "length(xs@ys) = length(xs)+length(ys)";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   731
by (induct_thm_tac list_induct "xs" 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   732
by (ALLGOALS Asm_simp_tac);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   733
qed "length_append";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   734
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   735
Goal "length(map f xs) = length(xs)";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   736
by (induct_thm_tac list_induct "xs" 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   737
by (ALLGOALS Asm_simp_tac);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   738
qed "length_map";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   739
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   740
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   741
Goal "take [] n = []";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   742
by (induct_tac "n" 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   743
by (ALLGOALS Asm_simp_tac);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   744
qed "take_Nil";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   745
Addsimps [take_Nil];
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   746
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   747
Goal "ALL n. take (take xs n) n = take xs n";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   748
by (induct_thm_tac list_induct "xs" 1); 
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   749
by (ALLGOALS Asm_simp_tac);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   750
by (rtac allI 1); 
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   751
by (induct_tac "n" 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   752
by Auto_tac; 
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   753
qed "take_take_eq";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   754
Addsimps [take_take_eq];
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   755
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   756
Goal "ALL n. take (take xs(Suc(n+m))) n = take xs n";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   757
by (induct_thm_tac list_induct "xs" 1); 
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   758
by (ALLGOALS Asm_simp_tac);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   759
by (rtac allI 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   760
by (induct_tac "n" 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   761
by Auto_tac; 
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   762
qed_spec_mp "take_take_Suc_eq1";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   763
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   764
Delsimps [take_Suc];
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   765
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   766
Goal "take (take xs (n+m)) n = take xs n";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   767
by (induct_tac "m" 1); 
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   768
by (ALLGOALS (asm_simp_tac (simpset() addsimps [take_take_Suc_eq1])));
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   769
qed "take_take_1";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   770
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   771
Goal "ALL n. take (take xs n)(Suc(n+m)) = take xs n";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   772
by (induct_thm_tac list_induct "xs" 1); 
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   773
by (ALLGOALS Asm_simp_tac);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   774
by (rtac allI 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   775
by (induct_tac "n" 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   776
by Auto_tac; 
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   777
qed_spec_mp "take_take_Suc_eq2";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   778
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   779
Goal "take(take xs n)(n+m) = take xs n";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   780
by (induct_tac "m" 1); 
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   781
by (ALLGOALS (asm_simp_tac (simpset() addsimps [take_take_Suc_eq2])));
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   782
qed "take_take_2";
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   783
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   784
(* length(take(xs,n)) = min(n, length(xs)) *)
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   785
(* length(drop(xs,n)) = length(xs) - n *)
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   786
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   787
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   788
Goal "drop  [] n  = []";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   789
by (induct_tac "n" 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   790
by (ALLGOALS(asm_full_simp_tac (simpset())));
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   791
qed "drop_Nil";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   792
Addsimps [drop_Nil];
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   793
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   794
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   795
qed_goal "drop_drop" SList.thy "drop (drop xs m) n = drop xs(m+n)"
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   796
(fn _=>[res_inst_tac [("x","xs")] allE 1,
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   797
      	atac 2,
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   798
	induct_tac "m" 1,
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   799
	ALLGOALS(asm_full_simp_tac (simpset() 
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   800
				addsimps [drop_Nil])),
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   801
	rtac allI 1,
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   802
	induct_thm_tac list_induct "x" 1,
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   803
	ALLGOALS(asm_full_simp_tac (simpset() 
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   804
				addsimps [drop_Nil]))]);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   805
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   806
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   807
qed_goal "take_drop" SList.thy  "(take xs n) @ (drop xs n) = xs"
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   808
(fn _=>[res_inst_tac [("x","xs")] allE 1,
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   809
      	atac 2,
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   810
	induct_tac "n" 1,
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   811
	ALLGOALS(asm_full_simp_tac (simpset())),
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   812
	rtac allI 1,
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   813
	induct_thm_tac list_induct "x" 1,
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   814
	ALLGOALS(asm_full_simp_tac (simpset()
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   815
					addsimps [drop_Nil,take_Nil] ))]);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   816
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   817
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   818
qed_goal "copy_copy" SList.thy  "copy x n @ copy x m = copy x(n+m)"
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   819
(fn _=>[induct_tac "n" 1,
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   820
	ALLGOALS(asm_full_simp_tac (simpset()))]);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   821
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   822
qed_goal "length_copy"  SList.thy  "length(copy x n)  = n"
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   823
(fn _=>[induct_tac "n" 1,
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   824
	ALLGOALS(asm_full_simp_tac (simpset()))]);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   825
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   826
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   827
Goal "!xs. length(take xs n) = min (length xs) n";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   828
by (induct_tac "n" 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   829
 by Auto_tac;
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   830
by (induct_thm_tac list_induct "xs" 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   831
 by Auto_tac;
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   832
qed_spec_mp "length_take";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   833
Addsimps [length_take];
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   834
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   835
Goal "length(take A k) + length(drop A k)=length(A)";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   836
by (simp_tac (HOL_ss addsimps [length_append RS sym, take_drop]) 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   837
qed "length_take_drop";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   838
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   839
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   840
Goal "ALL A. length(A) = n --> take(A@B) n = A";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   841
by (induct_tac "n" 1);
12486
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12169
diff changeset
   842
by (rtac allI 1);
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12169
diff changeset
   843
by (rtac allI 2);
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   844
by (induct_thm_tac list_induct "A" 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   845
by (induct_thm_tac list_induct "A" 3);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   846
by (ALLGOALS  Asm_full_simp_tac);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   847
qed_spec_mp "take_append";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   848
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   849
Goal "ALL A. length(A) = n --> take(A@B) (n+k) = A@take B k";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   850
by (induct_tac "n" 1);
12486
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12169
diff changeset
   851
by (rtac allI 1);
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12169
diff changeset
   852
by (rtac allI 2);
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   853
by (induct_thm_tac list_induct "A" 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   854
by (induct_thm_tac list_induct "A" 3);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   855
by (ALLGOALS  Asm_full_simp_tac);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   856
qed_spec_mp "take_append2";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   857
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   858
Goal "ALL n. take (map f A) n = map f (take A n)";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   859
by (induct_thm_tac list_induct "A" 1);
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   860
by (ALLGOALS Asm_simp_tac);
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   861
by (rtac allI 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   862
by (induct_tac "n" 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   863
by (ALLGOALS Asm_simp_tac);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   864
qed_spec_mp "take_map";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   865
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   866
Goal "ALL A. length(A) = n --> drop(A@B)n = B";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   867
by (induct_tac "n" 1);
12486
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12169
diff changeset
   868
by (rtac allI 1);
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12169
diff changeset
   869
by (rtac allI 2);
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   870
by (induct_thm_tac list_induct "A" 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   871
by (induct_thm_tac list_induct "A" 3);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   872
by (ALLGOALS  Asm_full_simp_tac);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   873
qed_spec_mp "drop_append";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   874
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   875
Goal "ALL A. length(A) = n --> drop(A@B)(n+k) = drop B k";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   876
by (induct_tac "n" 1);
12486
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12169
diff changeset
   877
by (rtac allI 1);
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12169
diff changeset
   878
by (rtac allI 2);
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   879
by (induct_thm_tac list_induct "A" 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   880
by (induct_thm_tac list_induct "A" 3);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   881
by (ALLGOALS  Asm_full_simp_tac);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   882
qed_spec_mp "drop_append2";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   883
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   884
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   885
Goal "ALL A. length(A) = n --> drop A n = []";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   886
by (induct_tac "n" 1);
12486
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12169
diff changeset
   887
by (rtac allI 1);
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12169
diff changeset
   888
by (rtac allI 2);
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   889
by (induct_thm_tac list_induct "A" 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   890
by (induct_thm_tac list_induct "A" 3);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   891
by Auto_tac; 
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   892
qed_spec_mp "drop_all";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   893
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   894
Goal "ALL n. drop (map f A) n = map f (drop A n)";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   895
by (induct_thm_tac list_induct "A" 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   896
by (ALLGOALS Asm_simp_tac);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   897
by (rtac allI 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   898
by (induct_tac "n" 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   899
by (ALLGOALS Asm_simp_tac);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   900
qed_spec_mp "drop_map";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   901
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   902
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   903
Goal "ALL A. length(A) = n --> take A n = A";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   904
by (induct_tac "n" 1);
12486
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12169
diff changeset
   905
by (rtac allI 1);
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12169
diff changeset
   906
by (rtac allI 2);
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   907
by (induct_thm_tac list_induct "A" 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   908
by (induct_thm_tac list_induct "A" 3);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   909
by (ALLGOALS (simp_tac (simpset())));
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   910
by (asm_simp_tac ((simpset()) addsimps [le_eq_less_or_eq]) 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   911
qed_spec_mp "take_all";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   912
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   913
Goal "foldl f a [b] = f a b";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   914
by (ALLGOALS Asm_simp_tac);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   915
qed "foldl_single";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   916
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   917
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   918
Goal "ALL a. foldl f a (A @ B) = foldl f (foldl f a A) B";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   919
by (induct_thm_tac list_induct "A" 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   920
by (ALLGOALS Asm_simp_tac);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   921
qed_spec_mp "foldl_append";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   922
Addsimps [foldl_append];
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   923
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   924
Goal "ALL e. foldl f e (map g S) = foldl (%x y. f x (g y)) e S";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   925
by (induct_thm_tac list_induct "S" 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   926
by (ALLGOALS Asm_simp_tac);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   927
qed_spec_mp "foldl_map";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   928
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   929
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   930
qed_goal "foldl_neutr_distr" SList.thy 
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   931
"[| !a. f a e = a; !a. f e a = a;          \
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   932
\   !a b c. f a (f b c) = f(f a b) c |]   \
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   933
\ ==> foldl f y A = f y (foldl f e A)"
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   934
(fn [r_neutr,l_neutr,assoc] =>
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   935
	[res_inst_tac [("x","y")] spec 1,
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   936
	induct_thm_tac list_induct "A" 1,
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   937
	ALLGOALS strip_tac,
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   938
	ALLGOALS(simp_tac (simpset() addsimps [r_neutr,l_neutr])),
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   939
	etac all_dupE 1,
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   940
	rtac trans 1,
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   941
	atac 1,
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   942
	simp_tac (HOL_ss addsimps [assoc RS spec RS spec RS spec RS sym])1,
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   943
	res_inst_tac [("f","%c. f xa c")] arg_cong 1,
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   944
	rtac sym 1,
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   945
	etac allE 1,
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   946
	atac 1]);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   947
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   948
Goal 
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   949
"[| !a. f a e = a; !a. f e a = a;          \
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   950
\   !a b c. f a (f b c) = f(f a b) c |]   \
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   951
\ ==> foldl f e (A @ B) = f(foldl f e A)(foldl f e B)";
12486
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12169
diff changeset
   952
by (rtac trans 1);
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12169
diff changeset
   953
by (rtac foldl_append 1);
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12169
diff changeset
   954
by (rtac (foldl_neutr_distr) 1);
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   955
by Auto_tac; 
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   956
qed "foldl_append_sym";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   957
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   958
Goal "ALL a. foldr f a (A @ B) = foldr f (foldr f a B) A";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   959
by (induct_thm_tac list_induct "A" 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   960
by (ALLGOALS Asm_simp_tac);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   961
qed_spec_mp "foldr_append";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   962
Addsimps [foldr_append];
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   963
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   964
Goalw [o_def] "ALL e. foldr f e (map g S) = foldr (f o g) e S";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   965
by (induct_thm_tac list_induct "S" 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   966
by (ALLGOALS Asm_simp_tac);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   967
qed_spec_mp "foldr_map";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   968
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   969
Goal "foldr op Un {} S = (UN X: {t. t mem S}.X)";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   970
by (induct_thm_tac list_induct "S" 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   971
by Auto_tac; 
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   972
qed "foldr_Un_eq_UN";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   973
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   974
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   975
Goal "[| !a. f e a = a; !a b c. f a (f b c) = f(f a b) c |]   \
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   976
\ ==> foldr f y S = f (foldr f e S) y";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   977
by (induct_thm_tac list_induct "S" 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   978
by Auto_tac; 
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   979
qed "foldr_neutr_distr";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   980
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   981
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   982
 Goal 
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   983
"[| !a. f e a = a; !a b c. f a (f b c) = f(f a b) c |] ==> \
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   984
\ foldr f e (A @ B) = f (foldr f e A) (foldr f e B)";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   985
by Auto_tac;
12486
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12169
diff changeset
   986
by (rtac foldr_neutr_distr 1);
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   987
by Auto_tac; 
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   988
qed "foldr_append2";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   989
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   990
Goal 
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   991
"[| !a. f e a = a; !a b c. f a (f b c) = f(f a b) c |] ==> \
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   992
\ foldr f e (flat S) = (foldr f e)(map (foldr f e) S)";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   993
by (induct_thm_tac list_induct "S" 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   994
by (ALLGOALS(asm_simp_tac (simpset() delsimps [foldr_append] 
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   995
                                 addsimps [foldr_append2])));
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   996
qed "foldr_flat";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   997
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   998
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
   999
Goal "(Alls x:map f xs .P(x)) = (Alls x:xs.(P o f)(x))";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1000
by (induct_thm_tac list_induct "xs" 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1001
by Auto_tac; 
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1002
qed "list_all_map";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1003
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1004
Goal 
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1005
"(Alls x:xs. P(x)&Q(x)) = ((Alls x:xs. P(x))&(Alls x:xs. Q(x)))";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1006
by (induct_thm_tac list_induct "xs" 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1007
by Auto_tac; 
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1008
qed "list_all_and";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1009
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
  1010
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1011
(* necessary to circumvent Bug in rewriter *)
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1012
val [pre] = Goal 
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1013
"(!!x. PQ(x) = (P(x) & Q(x))) ==> \
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1014
\ (Alls x:xs. PQ(x)) = ((Alls x:xs. P(x))&(Alls x:xs. Q(x)))";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1015
by (simp_tac (HOL_ss addsimps [pre]) 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1016
by (rtac list_all_and 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1017
qed "list_all_and_R";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1018
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1019
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1020
Goal "ALL i. i < length(A)  --> nth i (map f A) = f(nth i A)";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1021
by (induct_thm_tac list_induct "A" 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1022
by (ALLGOALS(asm_simp_tac (simpset() delsimps [less_Suc_eq])));
12486
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12169
diff changeset
  1023
by (rtac allI 1);
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1024
by (induct_tac "i" 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1025
by (ALLGOALS(asm_simp_tac (simpset() addsimps [nth_0,nth_Suc])));
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1026
qed_spec_mp "nth_map";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1027
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1028
Goal "ALL i. i < length(A)  --> nth i(A@B) = nth i A";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1029
by (induct_thm_tac list_induct "A" 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1030
by (ALLGOALS(asm_simp_tac (simpset() delsimps [less_Suc_eq])));
12486
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12169
diff changeset
  1031
by (rtac allI 1);
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1032
by (induct_tac "i" 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1033
by (ALLGOALS(asm_simp_tac (simpset() addsimps [nth_0,nth_Suc])));
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1034
qed_spec_mp "nth_app_cancel_right";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1035
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1036
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1037
Goal "ALL n. n = length(A) --> nth(n+i)(A@B) = nth i B";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1038
by (induct_thm_tac list_induct "A" 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1039
by (ALLGOALS Asm_simp_tac);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1040
qed_spec_mp "nth_app_cancel_left";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1041
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1042
(** flat **)
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1043
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1044
Goal  "flat(xs@ys) = flat(xs) @ flat(ys)";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1045
by (induct_thm_tac list_induct "xs" 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1046
by Auto_tac; 
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1047
qed "flat_append";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1048
Addsimps [flat_append];
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1049
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1050
Goal "filter p (flat S) = flat(map (filter p) S)";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1051
by (induct_thm_tac list_induct "S" 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1052
by Auto_tac;
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1053
qed "filter_flat";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1054
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1055
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1056
(** rev **)
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1057
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1058
Goal "rev(xs@ys) = rev(ys) @ rev(xs)";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1059
by (induct_thm_tac list_induct "xs" 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1060
by Auto_tac; 
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1061
qed "rev_append";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1062
Addsimps[rev_append];
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1063
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1064
Goal "rev(rev l) = l";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1065
by (induct_thm_tac list_induct "l" 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1066
by Auto_tac; 
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1067
qed "rev_rev_ident";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1068
Addsimps[rev_rev_ident];
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1069
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1070
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1071
Goal "rev(flat ls) = flat (map rev (rev ls))";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1072
by (induct_thm_tac list_induct "ls" 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1073
by Auto_tac;
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1074
qed "rev_flat";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1075
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1076
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1077
Goal "rev(map f l) = map f (rev l)";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1078
by (induct_thm_tac list_induct "l" 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1079
by Auto_tac; 
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1080
qed "rev_map_distrib";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1081
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1082
Goal "foldl f b (rev l) = foldr (%x y. f y x) b l";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1083
by (induct_thm_tac list_induct "l" 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1084
by Auto_tac; 
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1085
qed "foldl_rev";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1086
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1087
Goal "foldr f b (rev l) = foldl (%x y. f y x) b l";
12486
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12169
diff changeset
  1088
by (rtac sym 1);
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12169
diff changeset
  1089
by (rtac trans 1);
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12169
diff changeset
  1090
by (rtac foldl_rev 2);
12169
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1091
by (simp_tac (HOL_ss addsimps [rev_rev_ident]) 1);
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1092
qed "foldr_rev";
d4ed9802082a new SList theory from Bu Wolff
paulson
parents: 11500
diff changeset
  1093