| author | paulson |
| Wed, 30 Jan 2002 12:22:59 +0100 | |
| changeset 12861 | 7ec4807b53cf |
| parent 12486 | 0ed8bdd883e0 |
| permissions | -rw-r--r-- |
| 12169 | 1 |
(* Title: HOL/ex/SList.ML |
2 |
ID: SList.ML,v 1.2 1994/12/14 10:17:48 clasohm Exp |
|
3 |
Author: B. Wolff, based on a version of Lawrence C Paulson, |
|
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 | 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 | 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 | 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 | 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 | 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 | 41 |
fun List_simp thm = (asm_full_simplify (HOL_ss addsimps [List_def]) thm) |
42 |
||
|
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
43 |
(*Induction for the type 'a list *) |
| 12169 | 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 | 47 |
by (rtac (Rep_List_inverse RS subst) 1); |
48 |
(*types force good instantiation*) |
|
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 | 51 |
ORELSE eresolve_tac [rangeE, ssubst, |
52 |
(List_simp Abs_List_inverse) RS subst] 1)); |
|
53 |
qed "list_induct"; |
|
54 |
||
55 |
||
56 |
(*** Isomorphisms ***) |
|
57 |
||
58 |
Goal "inj_on Abs_List (list(range Leaf))"; |
|
59 |
by (rtac inj_on_inverseI 1); |
|
60 |
by (etac (List_simp Abs_List_inverse) 1); |
|
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 | 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 | 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 | 73 |
Goalw [Nil_def,Cons_def] "x # xs ~= Nil"; |
| 12169 | 74 |
by (rtac (CONS_not_NIL RS (inj_on_Abs_list RS inj_on_contraD)) 1); |
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 | 78 |
bind_thm ("Nil_not_Cons", (Cons_not_Nil RS not_sym));
|
79 |
||
80 |
bind_thm ("Cons_neq_Nil", (Cons_not_Nil RS notE));
|
|
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 | 85 |
Goalw [CONS_def] "(CONS K M)=(CONS L N) = (K=L & M=N)"; |
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 | 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 | 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 | 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 | 130 |
Goal "(xs ~= []) = (? y ys. xs = y#ys)"; |
| 12169 | 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 | 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 | 143 |
Goalw [List_case_def,CONS_def] "List_case c h (CONS M N) = h M N"; |
| 4521 | 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 | 147 |
Addsimps [List_case_NIL, List_case_CONS]; |
148 |
||
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 | 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 | 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 | 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 | 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 | 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 | 183 |
by (ftac (pred_sexp_subset_Sigma RS trancl_subset_Sigma RS subsetD) 1); |
184 |
by (blast_tac (claset() addSDs [sexp_CONS_D] |
|
185 |
addIs [pred_sexp_CONS_I1, pred_sexp_CONS_I2, |
|
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 | 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 | 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 | 194 |
by (simp_tac (HOL_ss addsimps [List_case_NIL]) 1); |
195 |
qed "List_rec_NIL"; |
|
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 | 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 | 203 |
Addsimps [List_rec_CONS]; |
| 4521 | 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 | 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 | 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 | 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 | 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 | 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 | 246 |
Goalw [Rep_map_def] |
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 | 251 |
Goalw [Rep_map_def] "!!f. (!!x. f(x): A) ==> Rep_map f xs: list(A)"; |
252 |
by (rtac list_induct 1); |
|
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 | 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 | 266 |
val [rew] = goal thy |
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 | 272 |
val [rew] = goal thy |
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 | 292 |
(* setting up rewrite sets *) |
293 |
||
294 |
fun list_recs def = |
|
295 |
[standard (def RS def_list_rec_Nil), |
|
296 |
standard (def RS def_list_rec_Cons)]; |
|
297 |
||
298 |
val [list_case_Nil,list_case_Cons] = list_recs list_case_def; |
|
299 |
Addsimps [list_case_Nil,list_case_Cons]; |
|
300 |
||
301 |
(*FIXME?? |
|
302 |
val slist_ss = (simpset()) addsimps |
|
303 |
[Cons_not_Nil, Nil_not_Cons, Cons_Cons_eq, |
|
304 |
list_rec_Nil, list_rec_Cons, |
|
305 |
slist_case_Nil,slist_case_Cons]; |
|
306 |
*) |
|
307 |
||
308 |
(** list_case **) |
|
309 |
||
310 |
Goal |
|
311 |
"P(list_case a f xs) = ((xs=[] --> P a ) & (!y ys. xs=y#ys --> P(f y ys)))"; |
|
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 | 314 |
qed "expand_list_case"; |
315 |
||
316 |
||
317 |
(**** Function definitions ****) |
|
318 |
||
319 |
fun list_recs def = |
|
320 |
[standard (def RS def_list_rec_Nil), |
|
321 |
standard (def RS def_list_rec_Cons)]; |
|
322 |
||
323 |
(*** Unfolding the basic combinators ***) |
|
324 |
||
325 |
val [null_Nil,null_Cons] = list_recs null_def; |
|
326 |
val [_,hd_Cons] = list_recs hd_def; |
|
327 |
val [_,tl_Cons] = list_recs tl_def; |
|
328 |
val [ttl_Nil,ttl_Cons] = list_recs ttl_def; |
|
329 |
val [append_Nil,append_Cons] = list_recs append_def; |
|
330 |
val [mem_Nil, mem_Cons] = list_recs mem_def; |
|
331 |
val [map_Nil,map_Cons] = list_recs map_def; |
|
332 |
val [filter_Nil,filter_Cons] = list_recs filter_def; |
|
333 |
val [list_all_Nil,list_all_Cons] = list_recs list_all_def; |
|
334 |
||
335 |
store_thm("hd_Cons",hd_Cons);
|
|
336 |
store_thm("tl_Cons",tl_Cons);
|
|
337 |
store_thm("ttl_Nil" ,ttl_Nil);
|
|
338 |
store_thm("ttl_Cons" ,ttl_Cons);
|
|
339 |
store_thm("append_Nil", append_Nil);
|
|
340 |
store_thm("append_Cons", append_Cons);
|
|
341 |
store_thm("mem_Nil" ,mem_Nil);
|
|
342 |
store_thm("mem_Cons" ,mem_Cons);
|
|
343 |
store_thm("map_Nil", map_Nil);
|
|
344 |
store_thm("map_Cons", map_Cons);
|
|
345 |
store_thm("filter_Nil", filter_Nil);
|
|
346 |
store_thm("filter_Cons", filter_Cons);
|
|
347 |
store_thm("list_all_Nil", list_all_Nil);
|
|
348 |
store_thm("list_all_Cons", list_all_Cons);
|
|
349 |
||
350 |
||
351 |
Addsimps |
|
352 |
[null_Nil, null_Cons, hd_Cons, tl_Cons, ttl_Nil, ttl_Cons, |
|
353 |
mem_Nil, mem_Cons, |
|
354 |
append_Nil, append_Cons, |
|
355 |
map_Nil, map_Cons, |
|
356 |
list_all_Nil, list_all_Cons, |
|
357 |
filter_Nil, filter_Cons]; |
|
358 |
||
359 |
(** nth **) |
|
360 |
||
361 |
val [rew] = goal Nat.thy |
|
362 |
"[| !!n. f == nat_rec c h |] ==> f(0) = c"; |
|
363 |
by (rewtac rew); |
|
364 |
by (rtac nat_rec_0 1); |
|
365 |
qed "def_nat_rec_0_eta"; |
|
366 |
||
367 |
val [rew] = goal Nat.thy |
|
368 |
"[| !!n. f == nat_rec c h |] ==> f(Suc(n)) = h n (f n)"; |
|
369 |
by (rewtac rew); |
|
370 |
by (rtac nat_rec_Suc 1); |
|
371 |
qed "def_nat_rec_Suc_eta"; |
|
372 |
||
373 |
fun nat_recs_eta def = |
|
374 |
[standard (def RS def_nat_rec_0_eta), |
|
375 |
standard (def RS def_nat_rec_Suc_eta)]; |
|
376 |
||
377 |
||
378 |
val [nth_0,nth_Suc] = nat_recs_eta nth_def; |
|
379 |
store_thm("nth_0",nth_0);
|
|
380 |
store_thm("nth_Suc",nth_Suc);
|
|
381 |
||
382 |
Addsimps [nth_0,nth_Suc]; |
|
383 |
||
384 |
(** length **) |
|
385 |
||
386 |
Goalw [length_def] "length([]) = 0"; |
|
387 |
by (ALLGOALS Asm_simp_tac); |
|
388 |
qed "length_Nil"; |
|
389 |
||
390 |
Goalw [length_def] "length(a#xs) = Suc(length(xs))"; |
|
391 |
by (ALLGOALS Asm_simp_tac); |
|
392 |
qed "length_Cons"; |
|
393 |
||
394 |
Addsimps [length_Nil,length_Cons]; |
|
395 |
||
396 |
||
397 |
(** @ - append **) |
|
398 |
||
399 |
Goal "(xs@ys)@zs = xs@(ys@zs)"; |
|
400 |
by (induct_thm_tac list_induct "xs" 1); |
|
401 |
by (ALLGOALS Asm_simp_tac); |
|
402 |
by (ALLGOALS Asm_simp_tac); |
|
403 |
qed "append_assoc"; |
|
404 |
||
405 |
Goal "xs @ [] = xs"; |
|
406 |
by (induct_thm_tac list_induct "xs" 1); |
|
407 |
by (ALLGOALS Asm_simp_tac); |
|
408 |
qed "append_Nil2"; |
|
409 |
||
410 |
(** mem **) |
|
411 |
||
412 |
Goal "x mem (xs@ys) = (x mem xs | x mem ys)"; |
|
413 |
by (induct_thm_tac list_induct "xs" 1); |
|
414 |
by (ALLGOALS Asm_simp_tac); |
|
415 |
qed "mem_append"; |
|
416 |
||
417 |
Goal "x mem [x:xs. P x ] = (x mem xs & P(x))"; |
|
418 |
by (induct_thm_tac list_induct "xs" 1); |
|
419 |
by (ALLGOALS Asm_simp_tac); |
|
420 |
qed "mem_filter"; |
|
421 |
||
422 |
(** list_all **) |
|
423 |
||
424 |
Goal "(Alls x:xs. True) = True"; |
|
425 |
by (induct_thm_tac list_induct "xs" 1); |
|
426 |
by (ALLGOALS Asm_simp_tac); |
|
427 |
qed "list_all_True"; |
|
428 |
||
429 |
Goal "list_all p (xs@ys) = ((list_all p xs) & (list_all p ys))"; |
|
430 |
by (induct_thm_tac list_induct "xs" 1); |
|
431 |
by (ALLGOALS Asm_simp_tac); |
|
432 |
qed "list_all_conj"; |
|
433 |
||
434 |
Goal "(Alls x:xs. P(x)) = (!x. x mem xs --> P(x))"; |
|
435 |
by (induct_thm_tac list_induct "xs" 1); |
|
436 |
by (ALLGOALS Asm_simp_tac); |
|
437 |
by (fast_tac HOL_cs 1); |
|
438 |
qed "list_all_mem_conv"; |
|
439 |
||
440 |
||
441 |
Goal "(! n. P n) = (P 0 & (! n. P (Suc n)))"; |
|
442 |
by (Auto_tac); |
|
443 |
by (induct_tac "n" 1); |
|
444 |
by (Auto_tac); |
|
445 |
qed "nat_case_dist"; |
|
446 |
||
447 |
||
448 |
val [] = Goal "(Alls u:A. P u) = (!n. n < length A --> P(nth n A))"; |
|
449 |
by (induct_thm_tac list_induct "A" 1); |
|
450 |
by (ALLGOALS Asm_simp_tac); |
|
| 12486 | 451 |
by (rtac trans 1); |
452 |
by (rtac (nat_case_dist RS sym) 2); |
|
| 12169 | 453 |
by (ALLGOALS Asm_simp_tac); |
454 |
qed "alls_P_eq_P_nth"; |
|
455 |
||
456 |
||
457 |
Goal "[| !x. P x --> Q x; (Alls x:xs. P(x)) |] ==> (Alls x:xs. Q(x))"; |
|
458 |
by (asm_full_simp_tac (simpset() addsimps [list_all_mem_conv]) 1); |
|
459 |
qed "list_all_imp"; |
|
460 |
||
461 |
||
462 |
(** The functional "map" and the generalized functionals **) |
|
463 |
||
464 |
val prems = |
|
465 |
Goal "(!!x. f(x): sexp) ==> \ |
|
466 |
\ Abs_map g (Rep_map f xs) = map (%t. g(f(t))) xs"; |
|
467 |
by (induct_thm_tac list_induct "xs" 1); |
|
468 |
by (ALLGOALS (asm_simp_tac(simpset() addsimps |
|
469 |
(prems@[Rep_map_type, list_sexp RS subsetD])))); |
|
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 | 475 |
Goal "map(%x. x)(xs) = xs"; |
476 |
by (induct_thm_tac list_induct "xs" 1); |
|
477 |
by (ALLGOALS Asm_simp_tac); |
|
478 |
qed "map_ident"; |
|
479 |
||
480 |
Goal "map f (xs@ys) = map f xs @ map f ys"; |
|
481 |
by (induct_thm_tac list_induct "xs" 1); |
|
482 |
by (ALLGOALS Asm_simp_tac); |
|
483 |
qed "map_append"; |
|
484 |
Addsimps[map_append]; |
|
485 |
||
486 |
Goalw [o_def] "map(f o g)(xs) = map f (map g xs)"; |
|
487 |
by (induct_thm_tac list_induct "xs" 1); |
|
488 |
by (ALLGOALS Asm_simp_tac); |
|
489 |
qed "map_compose"; |
|
490 |
||
491 |
Addsimps |
|
492 |
[mem_append, mem_filter, append_assoc, append_Nil2, map_ident, |
|
493 |
list_all_True, list_all_conj]; |
|
494 |
||
495 |
||
496 |
Goal |
|
497 |
"x mem (map f q) --> (? y. y mem q & x = f y)"; |
|
498 |
by (induct_thm_tac list_induct "q" 1); |
|
499 |
by (ALLGOALS Asm_simp_tac); |
|
500 |
by (case_tac "f xa = x" 1); |
|
501 |
by (ALLGOALS Asm_simp_tac); |
|
502 |
by (res_inst_tac [("x","xa")] exI 1);
|
|
503 |
by (ALLGOALS Asm_simp_tac); |
|
| 12486 | 504 |
by (rtac impI 1); |
| 12169 | 505 |
by (rotate_tac 1 1); |
506 |
by (ALLGOALS Asm_full_simp_tac); |
|
| 12486 | 507 |
by (etac exE 1); by (etac conjE 1); |
| 12169 | 508 |
by (res_inst_tac [("x","y")] exI 1);
|
509 |
by (asm_simp_tac (HOL_ss addsimps [if_cancel]) 1); |
|
510 |
qed "mem_map_aux1"; |
|
511 |
||
512 |
Goal |
|
513 |
"(? y. y mem q & x = f y) --> x mem (map f q)"; |
|
514 |
by (induct_thm_tac list_induct "q" 1); |
|
515 |
by (Asm_simp_tac 1); |
|
| 12486 | 516 |
by (rtac impI 1); |
517 |
by (etac exE 1); |
|
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 | 520 |
by (case_tac "xa = y" 1); |
521 |
by (rotate_tac 2 1); |
|
522 |
by (asm_full_simp_tac (HOL_ss addsimps [if_cancel]) 1); |
|
| 12486 | 523 |
by (etac impE 1); |
| 12169 | 524 |
by (asm_simp_tac (HOL_ss addsimps [if_cancel]) 1); |
525 |
by (case_tac "f xa = f y" 2); |
|
526 |
by (res_inst_tac [("x","y")] exI 1);
|
|
527 |
by (asm_simp_tac (HOL_ss addsimps [if_cancel]) 1); |
|
528 |
by (Auto_tac); |
|
529 |
qed "mem_map_aux2"; |
|
530 |
||
531 |
||
532 |
Goal |
|
533 |
"x mem (map f q) = (? y. y mem q & x = f y)"; |
|
| 12486 | 534 |
by (rtac iffI 1); |
535 |
by (rtac (mem_map_aux1 RS mp) 1); |
|
536 |
by (rtac (mem_map_aux2 RS mp) 2); |
|
| 12169 | 537 |
by (ALLGOALS atac); |
538 |
qed "mem_map"; |
|
539 |
||
540 |
Goal "A ~= [] --> hd(A @ B) = hd(A)"; |
|
541 |
by (induct_thm_tac list_induct "A" 1); |
|
542 |
by Auto_tac; |
|
543 |
qed_spec_mp "hd_append"; |
|
544 |
||
545 |
Goal "A ~= [] --> tl(A @ B) = tl(A) @ B"; |
|
546 |
by (induct_thm_tac list_induct "A" 1); |
|
547 |
by Auto_tac; |
|
548 |
qed_spec_mp "tl_append"; |
|
549 |
||
550 |
||
551 |
(* ********************************************************************* *) |
|
552 |
(* More ... *) |
|
553 |
(* ********************************************************************* *) |
|
554 |
||
555 |
||
556 |
(** take **) |
|
557 |
||
558 |
Goal "take [] (Suc x) = []"; |
|
559 |
by (asm_simp_tac (simpset()) 1); |
|
560 |
qed "take_Suc1"; |
|
561 |
||
562 |
Goal "take(a#xs)(Suc x) = a#take xs x"; |
|
563 |
by (asm_simp_tac (simpset()) 1); |
|
564 |
qed "take_Suc2"; |
|
565 |
||
566 |
||
567 |
(** drop **) |
|
568 |
||
569 |
Goalw [drop_def] "drop xs 0 = xs"; |
|
570 |
by (asm_simp_tac (simpset()) 1); |
|
571 |
qed "drop_0"; |
|
572 |
||
573 |
Goalw [drop_def] "drop [] (Suc x) = []"; |
|
574 |
by (induct_tac "x" 1); |
|
575 |
by (ALLGOALS (asm_full_simp_tac ((simpset()) addsimps [ttl_Nil]) )); |
|
576 |
qed "drop_Suc1"; |
|
577 |
||
578 |
Goalw [drop_def] "drop(a#xs)(Suc x) = drop xs x"; |
|
579 |
by (asm_simp_tac (simpset()) 1); |
|
580 |
qed "drop_Suc2"; |
|
581 |
||
582 |
||
583 |
(** copy **) |
|
584 |
||
585 |
Goalw [copy_def] "copy x 0 = []"; |
|
586 |
by (asm_simp_tac (simpset()) 1); |
|
587 |
qed "copy_0"; |
|
588 |
||
589 |
Goalw [copy_def] "copy x (Suc y) = x # copy x y"; |
|
590 |
by (asm_simp_tac (simpset()) 1); |
|
591 |
qed "copy_Suc"; |
|
592 |
||
593 |
||
594 |
(** fold **) |
|
595 |
||
596 |
Goalw [foldl_def] "foldl f a [] = a"; |
|
597 |
by (ALLGOALS Asm_simp_tac); |
|
598 |
qed "foldl_Nil"; |
|
599 |
||
600 |
Goalw [foldl_def] "foldl f a(x#xs) = foldl f (f a x) xs"; |
|
601 |
by (induct_thm_tac list_induct "xs" 1); |
|
602 |
by (ALLGOALS Asm_full_simp_tac); |
|
603 |
qed "foldl_Cons"; |
|
604 |
||
605 |
Goalw [foldr_def] "foldr f a [] = a"; |
|
606 |
by (ALLGOALS Asm_simp_tac); |
|
607 |
qed "foldr_Nil"; |
|
608 |
||
609 |
Goalw [foldr_def] "foldr f z(x#xs) = f x (foldr f z xs)"; |
|
610 |
by (ALLGOALS Asm_simp_tac); |
|
611 |
qed "foldr_Cons"; |
|
612 |
||
613 |
Addsimps |
|
614 |
[length_Nil,length_Cons, |
|
615 |
take_0, take_Suc1,take_Suc2, |
|
616 |
drop_0,drop_Suc1,drop_Suc2,copy_0,copy_Suc, |
|
617 |
foldl_Nil,foldl_Cons,foldr_Nil,foldr_Cons]; |
|
618 |
||
619 |
||
620 |
(** flat **) |
|
621 |
||
622 |
Goalw [flat_def] |
|
623 |
"flat [] = []"; |
|
624 |
by (ALLGOALS Asm_simp_tac); |
|
625 |
qed "flat_Nil"; |
|
|
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
626 |
|
| 12169 | 627 |
Goalw [flat_def] |
628 |
"flat (x # xs) = x @ flat xs"; |
|
629 |
by (ALLGOALS Asm_simp_tac); |
|
630 |
qed "flat_Cons"; |
|
631 |
||
632 |
Addsimps [flat_Nil,flat_Cons]; |
|
633 |
||
634 |
(** rev **) |
|
635 |
||
636 |
Goalw [rev_def] |
|
637 |
"rev [] = []"; |
|
638 |
by (ALLGOALS Asm_simp_tac); |
|
639 |
qed "rev_Nil"; |
|
640 |
||
641 |
||
642 |
Goalw [rev_def] |
|
643 |
"rev (x # xs) = rev xs @ [x]"; |
|
644 |
by (ALLGOALS Asm_simp_tac); |
|
645 |
qed "rev_Cons"; |
|
646 |
||
647 |
||
648 |
Addsimps [rev_Nil,rev_Cons]; |
|
649 |
||
650 |
(** zip **) |
|
651 |
||
652 |
Goalw [zipWith_def] |
|
653 |
"zipWith f (a#as,b#bs) = f(a,b) # zipWith f (as,bs)"; |
|
654 |
by (ALLGOALS Asm_simp_tac); |
|
655 |
qed "zipWith_Cons_Cons"; |
|
656 |
||
657 |
Goalw [zipWith_def] |
|
658 |
"zipWith f ([],[]) = []"; |
|
659 |
by (ALLGOALS Asm_simp_tac); |
|
660 |
qed "zipWith_Nil_Nil"; |
|
661 |
||
662 |
||
663 |
Goalw [zipWith_def] |
|
664 |
"zipWith f (x,[]) = []"; |
|
665 |
by (induct_thm_tac list_induct "x" 1); |
|
666 |
by (ALLGOALS Asm_simp_tac); |
|
667 |
qed "zipWith_Cons_Nil"; |
|
668 |
||
669 |
||
670 |
Goalw [zipWith_def] |
|
671 |
"zipWith f ([],x) = []"; |
|
672 |
by (induct_thm_tac list_induct "x" 1); |
|
673 |
by (ALLGOALS Asm_simp_tac); |
|
674 |
qed "zipWith_Nil_Cons"; |
|
675 |
||
676 |
Goalw [unzip_def] "unzip [] = ([],[])"; |
|
677 |
by (ALLGOALS Asm_simp_tac); |
|
678 |
qed "unzip_Nil"; |
|
679 |
||
680 |
||
681 |
||
682 |
(** SOME LIST THEOREMS **) |
|
683 |
||
684 |
(* SQUIGGOL LEMMAS *) |
|
685 |
||
686 |
||
687 |
Goalw [o_def] "map(f o g) = ((map f) o (map g))"; |
|
| 12486 | 688 |
by (rtac ext 1); |
| 12169 | 689 |
by (simp_tac (HOL_ss addsimps [map_compose RS sym,o_def]) 1); |
690 |
qed "map_compose_ext"; |
|
691 |
||
692 |
Goal "map f (flat S) = flat(map (map f) S)"; |
|
693 |
by (induct_thm_tac list_induct "S" 1); |
|
694 |
by (ALLGOALS Asm_simp_tac); |
|
695 |
qed "map_flat"; |
|
696 |
||
697 |
Goal "(Alls u:xs. f(u) = g(u)) --> map f xs = map g xs"; |
|
698 |
by (induct_thm_tac list_induct "xs" 1); |
|
699 |
by (ALLGOALS Asm_simp_tac); |
|
700 |
qed "list_all_map_eq"; |
|
701 |
||
702 |
Goal "filter p (map f xs) = map f (filter(p o f)(xs))"; |
|
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 | 705 |
qed "filter_map_d"; |
706 |
||
707 |
Goal "filter p (filter q xs) = filter(%x. p x & q x) xs"; |
|
708 |
by (induct_thm_tac list_induct "xs" 1); |
|
709 |
by (ALLGOALS Asm_simp_tac); |
|
710 |
qed "filter_compose"; |
|
711 |
||
712 |
(* "filter(p, filter(q ,xs)) = filter(q, filter(p ,xs))", |
|
713 |
"filter(p, filter(p ,xs)) = filter(p,xs)" BIRD's thms.*) |
|
714 |
||
715 |
Goal "ALL B. filter p (A @ B) = (filter p A @ filter p B)"; |
|
716 |
by (induct_thm_tac list_induct "A" 1); |
|
717 |
by (ALLGOALS Asm_simp_tac); |
|
718 |
qed_spec_mp "filter_append"; |
|
719 |
Addsimps [filter_append]; |
|
720 |
||
721 |
(* inits(xs) == map(fst,splits(xs)), |
|
722 |
||
723 |
splits([]) = [] |
|
724 |
splits(a # xs) = <[],xs> @ map(%x. <a # fst(x),snd(x)>, splits(xs)) |
|
725 |
(x @ y = z) = <x,y> mem splits(z) |
|
726 |
x mem xs & y mem ys = <x,y> mem diag(xs,ys) *) |
|
727 |
||
728 |
||
729 |
||
730 |
Goal "length(xs@ys) = length(xs)+length(ys)"; |
|
731 |
by (induct_thm_tac list_induct "xs" 1); |
|
732 |
by (ALLGOALS Asm_simp_tac); |
|
733 |
qed "length_append"; |
|
734 |
||
735 |
Goal "length(map f xs) = length(xs)"; |
|
736 |
by (induct_thm_tac list_induct "xs" 1); |
|
737 |
by (ALLGOALS Asm_simp_tac); |
|
738 |
qed "length_map"; |
|
739 |
||
740 |
||
741 |
Goal "take [] n = []"; |
|
742 |
by (induct_tac "n" 1); |
|
743 |
by (ALLGOALS Asm_simp_tac); |
|
744 |
qed "take_Nil"; |
|
745 |
Addsimps [take_Nil]; |
|
746 |
||
747 |
Goal "ALL n. take (take xs n) n = take xs n"; |
|
748 |
by (induct_thm_tac list_induct "xs" 1); |
|
749 |
by (ALLGOALS Asm_simp_tac); |
|
750 |
by (rtac allI 1); |
|
751 |
by (induct_tac "n" 1); |
|
752 |
by Auto_tac; |
|
753 |
qed "take_take_eq"; |
|
754 |
Addsimps [take_take_eq]; |
|
755 |
||
756 |
Goal "ALL n. take (take xs(Suc(n+m))) n = take xs n"; |
|
757 |
by (induct_thm_tac list_induct "xs" 1); |
|
758 |
by (ALLGOALS Asm_simp_tac); |
|
759 |
by (rtac allI 1); |
|
760 |
by (induct_tac "n" 1); |
|
761 |
by Auto_tac; |
|
762 |
qed_spec_mp "take_take_Suc_eq1"; |
|
763 |
||
764 |
Delsimps [take_Suc]; |
|
765 |
||
766 |
Goal "take (take xs (n+m)) n = take xs n"; |
|
767 |
by (induct_tac "m" 1); |
|
768 |
by (ALLGOALS (asm_simp_tac (simpset() addsimps [take_take_Suc_eq1]))); |
|
769 |
qed "take_take_1"; |
|
770 |
||
771 |
Goal "ALL n. take (take xs n)(Suc(n+m)) = take xs n"; |
|
772 |
by (induct_thm_tac list_induct "xs" 1); |
|
773 |
by (ALLGOALS Asm_simp_tac); |
|
774 |
by (rtac allI 1); |
|
775 |
by (induct_tac "n" 1); |
|
776 |
by Auto_tac; |
|
777 |
qed_spec_mp "take_take_Suc_eq2"; |
|
778 |
||
779 |
Goal "take(take xs n)(n+m) = take xs n"; |
|
780 |
by (induct_tac "m" 1); |
|
781 |
by (ALLGOALS (asm_simp_tac (simpset() addsimps [take_take_Suc_eq2]))); |
|
782 |
qed "take_take_2"; |
|
|
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
783 |
|
| 12169 | 784 |
(* length(take(xs,n)) = min(n, length(xs)) *) |
785 |
(* length(drop(xs,n)) = length(xs) - n *) |
|
786 |
||
787 |
||
788 |
Goal "drop [] n = []"; |
|
789 |
by (induct_tac "n" 1); |
|
790 |
by (ALLGOALS(asm_full_simp_tac (simpset()))); |
|
791 |
qed "drop_Nil"; |
|
792 |
Addsimps [drop_Nil]; |
|
793 |
||
794 |
||
795 |
qed_goal "drop_drop" SList.thy "drop (drop xs m) n = drop xs(m+n)" |
|
796 |
(fn _=>[res_inst_tac [("x","xs")] allE 1,
|
|
797 |
atac 2, |
|
798 |
induct_tac "m" 1, |
|
799 |
ALLGOALS(asm_full_simp_tac (simpset() |
|
800 |
addsimps [drop_Nil])), |
|
801 |
rtac allI 1, |
|
802 |
induct_thm_tac list_induct "x" 1, |
|
803 |
ALLGOALS(asm_full_simp_tac (simpset() |
|
804 |
addsimps [drop_Nil]))]); |
|
805 |
||
806 |
||
807 |
qed_goal "take_drop" SList.thy "(take xs n) @ (drop xs n) = xs" |
|
808 |
(fn _=>[res_inst_tac [("x","xs")] allE 1,
|
|
809 |
atac 2, |
|
810 |
induct_tac "n" 1, |
|
811 |
ALLGOALS(asm_full_simp_tac (simpset())), |
|
812 |
rtac allI 1, |
|
813 |
induct_thm_tac list_induct "x" 1, |
|
814 |
ALLGOALS(asm_full_simp_tac (simpset() |
|
815 |
addsimps [drop_Nil,take_Nil] ))]); |
|
816 |
||
817 |
||
818 |
qed_goal "copy_copy" SList.thy "copy x n @ copy x m = copy x(n+m)" |
|
819 |
(fn _=>[induct_tac "n" 1, |
|
820 |
ALLGOALS(asm_full_simp_tac (simpset()))]); |
|
821 |
||
822 |
qed_goal "length_copy" SList.thy "length(copy x n) = n" |
|
823 |
(fn _=>[induct_tac "n" 1, |
|
824 |
ALLGOALS(asm_full_simp_tac (simpset()))]); |
|
825 |
||
826 |
||
827 |
Goal "!xs. length(take xs n) = min (length xs) n"; |
|
828 |
by (induct_tac "n" 1); |
|
829 |
by Auto_tac; |
|
830 |
by (induct_thm_tac list_induct "xs" 1); |
|
831 |
by Auto_tac; |
|
832 |
qed_spec_mp "length_take"; |
|
833 |
Addsimps [length_take]; |
|
834 |
||
835 |
Goal "length(take A k) + length(drop A k)=length(A)"; |
|
836 |
by (simp_tac (HOL_ss addsimps [length_append RS sym, take_drop]) 1); |
|
837 |
qed "length_take_drop"; |
|
838 |
||
839 |
||
840 |
Goal "ALL A. length(A) = n --> take(A@B) n = A"; |
|
841 |
by (induct_tac "n" 1); |
|
| 12486 | 842 |
by (rtac allI 1); |
843 |
by (rtac allI 2); |
|
| 12169 | 844 |
by (induct_thm_tac list_induct "A" 1); |
845 |
by (induct_thm_tac list_induct "A" 3); |
|
846 |
by (ALLGOALS Asm_full_simp_tac); |
|
847 |
qed_spec_mp "take_append"; |
|
848 |
||
849 |
Goal "ALL A. length(A) = n --> take(A@B) (n+k) = A@take B k"; |
|
850 |
by (induct_tac "n" 1); |
|
| 12486 | 851 |
by (rtac allI 1); |
852 |
by (rtac allI 2); |
|
| 12169 | 853 |
by (induct_thm_tac list_induct "A" 1); |
854 |
by (induct_thm_tac list_induct "A" 3); |
|
855 |
by (ALLGOALS Asm_full_simp_tac); |
|
856 |
qed_spec_mp "take_append2"; |
|
857 |
||
858 |
Goal "ALL n. take (map f A) n = map f (take A n)"; |
|
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 | 861 |
by (rtac allI 1); |
862 |
by (induct_tac "n" 1); |
|
863 |
by (ALLGOALS Asm_simp_tac); |
|
864 |
qed_spec_mp "take_map"; |
|
865 |
||
866 |
Goal "ALL A. length(A) = n --> drop(A@B)n = B"; |
|
867 |
by (induct_tac "n" 1); |
|
| 12486 | 868 |
by (rtac allI 1); |
869 |
by (rtac allI 2); |
|
| 12169 | 870 |
by (induct_thm_tac list_induct "A" 1); |
871 |
by (induct_thm_tac list_induct "A" 3); |
|
872 |
by (ALLGOALS Asm_full_simp_tac); |
|
873 |
qed_spec_mp "drop_append"; |
|
874 |
||
875 |
Goal "ALL A. length(A) = n --> drop(A@B)(n+k) = drop B k"; |
|
876 |
by (induct_tac "n" 1); |
|
| 12486 | 877 |
by (rtac allI 1); |
878 |
by (rtac allI 2); |
|
| 12169 | 879 |
by (induct_thm_tac list_induct "A" 1); |
880 |
by (induct_thm_tac list_induct "A" 3); |
|
881 |
by (ALLGOALS Asm_full_simp_tac); |
|
882 |
qed_spec_mp "drop_append2"; |
|
883 |
||
884 |
||
885 |
Goal "ALL A. length(A) = n --> drop A n = []"; |
|
886 |
by (induct_tac "n" 1); |
|
| 12486 | 887 |
by (rtac allI 1); |
888 |
by (rtac allI 2); |
|
| 12169 | 889 |
by (induct_thm_tac list_induct "A" 1); |
890 |
by (induct_thm_tac list_induct "A" 3); |
|
891 |
by Auto_tac; |
|
892 |
qed_spec_mp "drop_all"; |
|
893 |
||
894 |
Goal "ALL n. drop (map f A) n = map f (drop A n)"; |
|
895 |
by (induct_thm_tac list_induct "A" 1); |
|
896 |
by (ALLGOALS Asm_simp_tac); |
|
897 |
by (rtac allI 1); |
|
898 |
by (induct_tac "n" 1); |
|
899 |
by (ALLGOALS Asm_simp_tac); |
|
900 |
qed_spec_mp "drop_map"; |
|
901 |
||
902 |
||
903 |
Goal "ALL A. length(A) = n --> take A n = A"; |
|
904 |
by (induct_tac "n" 1); |
|
| 12486 | 905 |
by (rtac allI 1); |
906 |
by (rtac allI 2); |
|
| 12169 | 907 |
by (induct_thm_tac list_induct "A" 1); |
908 |
by (induct_thm_tac list_induct "A" 3); |
|
909 |
by (ALLGOALS (simp_tac (simpset()))); |
|
910 |
by (asm_simp_tac ((simpset()) addsimps [le_eq_less_or_eq]) 1); |
|
911 |
qed_spec_mp "take_all"; |
|
912 |
||
913 |
Goal "foldl f a [b] = f a b"; |
|
914 |
by (ALLGOALS Asm_simp_tac); |
|
915 |
qed "foldl_single"; |
|
916 |
||
917 |
||
918 |
Goal "ALL a. foldl f a (A @ B) = foldl f (foldl f a A) B"; |
|
919 |
by (induct_thm_tac list_induct "A" 1); |
|
920 |
by (ALLGOALS Asm_simp_tac); |
|
921 |
qed_spec_mp "foldl_append"; |
|
922 |
Addsimps [foldl_append]; |
|
923 |
||
924 |
Goal "ALL e. foldl f e (map g S) = foldl (%x y. f x (g y)) e S"; |
|
925 |
by (induct_thm_tac list_induct "S" 1); |
|
926 |
by (ALLGOALS Asm_simp_tac); |
|
927 |
qed_spec_mp "foldl_map"; |
|
928 |
||
|
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
929 |
|
| 12169 | 930 |
qed_goal "foldl_neutr_distr" SList.thy |
931 |
"[| !a. f a e = a; !a. f e a = a; \ |
|
932 |
\ !a b c. f a (f b c) = f(f a b) c |] \ |
|
933 |
\ ==> foldl f y A = f y (foldl f e A)" |
|
934 |
(fn [r_neutr,l_neutr,assoc] => |
|
935 |
[res_inst_tac [("x","y")] spec 1,
|
|
936 |
induct_thm_tac list_induct "A" 1, |
|
937 |
ALLGOALS strip_tac, |
|
938 |
ALLGOALS(simp_tac (simpset() addsimps [r_neutr,l_neutr])), |
|
939 |
etac all_dupE 1, |
|
940 |
rtac trans 1, |
|
941 |
atac 1, |
|
942 |
simp_tac (HOL_ss addsimps [assoc RS spec RS spec RS spec RS sym])1, |
|
943 |
res_inst_tac [("f","%c. f xa c")] arg_cong 1,
|
|
944 |
rtac sym 1, |
|
945 |
etac allE 1, |
|
946 |
atac 1]); |
|
947 |
||
948 |
Goal |
|
949 |
"[| !a. f a e = a; !a. f e a = a; \ |
|
950 |
\ !a b c. f a (f b c) = f(f a b) c |] \ |
|
951 |
\ ==> foldl f e (A @ B) = f(foldl f e A)(foldl f e B)"; |
|
| 12486 | 952 |
by (rtac trans 1); |
953 |
by (rtac foldl_append 1); |
|
954 |
by (rtac (foldl_neutr_distr) 1); |
|
| 12169 | 955 |
by Auto_tac; |
956 |
qed "foldl_append_sym"; |
|
957 |
||
958 |
Goal "ALL a. foldr f a (A @ B) = foldr f (foldr f a B) A"; |
|
959 |
by (induct_thm_tac list_induct "A" 1); |
|
960 |
by (ALLGOALS Asm_simp_tac); |
|
961 |
qed_spec_mp "foldr_append"; |
|
962 |
Addsimps [foldr_append]; |
|
963 |
||
964 |
Goalw [o_def] "ALL e. foldr f e (map g S) = foldr (f o g) e S"; |
|
965 |
by (induct_thm_tac list_induct "S" 1); |
|
966 |
by (ALLGOALS Asm_simp_tac); |
|
967 |
qed_spec_mp "foldr_map"; |
|
968 |
||
969 |
Goal "foldr op Un {} S = (UN X: {t. t mem S}.X)";
|
|
970 |
by (induct_thm_tac list_induct "S" 1); |
|
971 |
by Auto_tac; |
|
972 |
qed "foldr_Un_eq_UN"; |
|
973 |
||
974 |
||
975 |
Goal "[| !a. f e a = a; !a b c. f a (f b c) = f(f a b) c |] \ |
|
976 |
\ ==> foldr f y S = f (foldr f e S) y"; |
|
977 |
by (induct_thm_tac list_induct "S" 1); |
|
978 |
by Auto_tac; |
|
979 |
qed "foldr_neutr_distr"; |
|
980 |
||
981 |
||
982 |
Goal |
|
983 |
"[| !a. f e a = a; !a b c. f a (f b c) = f(f a b) c |] ==> \ |
|
984 |
\ foldr f e (A @ B) = f (foldr f e A) (foldr f e B)"; |
|
985 |
by Auto_tac; |
|
| 12486 | 986 |
by (rtac foldr_neutr_distr 1); |
| 12169 | 987 |
by Auto_tac; |
988 |
qed "foldr_append2"; |
|
989 |
||
990 |
Goal |
|
991 |
"[| !a. f e a = a; !a b c. f a (f b c) = f(f a b) c |] ==> \ |
|
992 |
\ foldr f e (flat S) = (foldr f e)(map (foldr f e) S)"; |
|
993 |
by (induct_thm_tac list_induct "S" 1); |
|
994 |
by (ALLGOALS(asm_simp_tac (simpset() delsimps [foldr_append] |
|
995 |
addsimps [foldr_append2]))); |
|
996 |
qed "foldr_flat"; |
|
997 |
||
998 |
||
999 |
Goal "(Alls x:map f xs .P(x)) = (Alls x:xs.(P o f)(x))"; |
|
1000 |
by (induct_thm_tac list_induct "xs" 1); |
|
1001 |
by Auto_tac; |
|
1002 |
qed "list_all_map"; |
|
1003 |
||
1004 |
Goal |
|
1005 |
"(Alls x:xs. P(x)&Q(x)) = ((Alls x:xs. P(x))&(Alls x:xs. Q(x)))"; |
|
1006 |
by (induct_thm_tac list_induct "xs" 1); |
|
1007 |
by Auto_tac; |
|
1008 |
qed "list_all_and"; |
|
1009 |
||
|
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
1010 |
|
| 12169 | 1011 |
(* necessary to circumvent Bug in rewriter *) |
1012 |
val [pre] = Goal |
|
1013 |
"(!!x. PQ(x) = (P(x) & Q(x))) ==> \ |
|
1014 |
\ (Alls x:xs. PQ(x)) = ((Alls x:xs. P(x))&(Alls x:xs. Q(x)))"; |
|
1015 |
by (simp_tac (HOL_ss addsimps [pre]) 1); |
|
1016 |
by (rtac list_all_and 1); |
|
1017 |
qed "list_all_and_R"; |
|
1018 |
||
1019 |
||
1020 |
Goal "ALL i. i < length(A) --> nth i (map f A) = f(nth i A)"; |
|
1021 |
by (induct_thm_tac list_induct "A" 1); |
|
1022 |
by (ALLGOALS(asm_simp_tac (simpset() delsimps [less_Suc_eq]))); |
|
| 12486 | 1023 |
by (rtac allI 1); |
| 12169 | 1024 |
by (induct_tac "i" 1); |
1025 |
by (ALLGOALS(asm_simp_tac (simpset() addsimps [nth_0,nth_Suc]))); |
|
1026 |
qed_spec_mp "nth_map"; |
|
1027 |
||
1028 |
Goal "ALL i. i < length(A) --> nth i(A@B) = nth i A"; |
|
1029 |
by (induct_thm_tac list_induct "A" 1); |
|
1030 |
by (ALLGOALS(asm_simp_tac (simpset() delsimps [less_Suc_eq]))); |
|
| 12486 | 1031 |
by (rtac allI 1); |
| 12169 | 1032 |
by (induct_tac "i" 1); |
1033 |
by (ALLGOALS(asm_simp_tac (simpset() addsimps [nth_0,nth_Suc]))); |
|
1034 |
qed_spec_mp "nth_app_cancel_right"; |
|
1035 |
||
1036 |
||
1037 |
Goal "ALL n. n = length(A) --> nth(n+i)(A@B) = nth i B"; |
|
1038 |
by (induct_thm_tac list_induct "A" 1); |
|
1039 |
by (ALLGOALS Asm_simp_tac); |
|
1040 |
qed_spec_mp "nth_app_cancel_left"; |
|
1041 |
||
1042 |
(** flat **) |
|
1043 |
||
1044 |
Goal "flat(xs@ys) = flat(xs) @ flat(ys)"; |
|
1045 |
by (induct_thm_tac list_induct "xs" 1); |
|
1046 |
by Auto_tac; |
|
1047 |
qed "flat_append"; |
|
1048 |
Addsimps [flat_append]; |
|
1049 |
||
1050 |
Goal "filter p (flat S) = flat(map (filter p) S)"; |
|
1051 |
by (induct_thm_tac list_induct "S" 1); |
|
1052 |
by Auto_tac; |
|
1053 |
qed "filter_flat"; |
|
1054 |
||
1055 |
||
1056 |
(** rev **) |
|
1057 |
||
1058 |
Goal "rev(xs@ys) = rev(ys) @ rev(xs)"; |
|
1059 |
by (induct_thm_tac list_induct "xs" 1); |
|
1060 |
by Auto_tac; |
|
1061 |
qed "rev_append"; |
|
1062 |
Addsimps[rev_append]; |
|
1063 |
||
1064 |
Goal "rev(rev l) = l"; |
|
1065 |
by (induct_thm_tac list_induct "l" 1); |
|
1066 |
by Auto_tac; |
|
1067 |
qed "rev_rev_ident"; |
|
1068 |
Addsimps[rev_rev_ident]; |
|
1069 |
||
1070 |
||
1071 |
Goal "rev(flat ls) = flat (map rev (rev ls))"; |
|
1072 |
by (induct_thm_tac list_induct "ls" 1); |
|
1073 |
by Auto_tac; |
|
1074 |
qed "rev_flat"; |
|
1075 |
||
1076 |
||
1077 |
Goal "rev(map f l) = map f (rev l)"; |
|
1078 |
by (induct_thm_tac list_induct "l" 1); |
|
1079 |
by Auto_tac; |
|
1080 |
qed "rev_map_distrib"; |
|
1081 |
||
1082 |
Goal "foldl f b (rev l) = foldr (%x y. f y x) b l"; |
|
1083 |
by (induct_thm_tac list_induct "l" 1); |
|
1084 |
by Auto_tac; |
|
1085 |
qed "foldl_rev"; |
|
1086 |
||
1087 |
Goal "foldr f b (rev l) = foldl (%x y. f y x) b l"; |
|
| 12486 | 1088 |
by (rtac sym 1); |
1089 |
by (rtac trans 1); |
|
1090 |
by (rtac foldl_rev 2); |
|
| 12169 | 1091 |
by (simp_tac (HOL_ss addsimps [rev_rev_ident]) 1); |
1092 |
qed "foldr_rev"; |
|
1093 |