1 (* Title: HOL/llist |
|
2 ID: $Id$ |
|
3 Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
4 Copyright 1993 University of Cambridge |
|
5 |
|
6 SHOULD LListD_Fun_CONS_I, etc., be equations (for rewriting)? |
|
7 *) |
|
8 |
|
9 open LList; |
|
10 |
|
11 (** Simplification **) |
|
12 |
|
13 val llist_ss = univ_ss addcongs [split_weak_cong, sum_case_weak_cong] |
|
14 setloop split_tac [expand_split, expand_sum_case]; |
|
15 |
|
16 (*For adding _eqI rules to a simpset; we must remove Pair_eq because |
|
17 it may turn an instance of reflexivity into a conjunction!*) |
|
18 fun add_eqI ss = ss addsimps [range_eqI, image_eqI] |
|
19 delsimps [Pair_eq]; |
|
20 |
|
21 |
|
22 (*This justifies using llist in other recursive type definitions*) |
|
23 goalw LList.thy llist.defs "!!A B. A<=B ==> llist(A) <= llist(B)"; |
|
24 by (rtac gfp_mono 1); |
|
25 by (REPEAT (ares_tac basic_monos 1)); |
|
26 qed "llist_mono"; |
|
27 |
|
28 |
|
29 goal LList.thy "llist(A) = {Numb(0)} <+> (A <*> llist(A))"; |
|
30 let val rew = rewrite_rule [NIL_def, CONS_def] in |
|
31 by (fast_tac (univ_cs addSIs (equalityI :: map rew llist.intrs) |
|
32 addEs [rew llist.elim]) 1) |
|
33 end; |
|
34 qed "llist_unfold"; |
|
35 |
|
36 |
|
37 (*** Type checking by coinduction, using list_Fun |
|
38 THE COINDUCTIVE DEFINITION PACKAGE COULD DO THIS! |
|
39 ***) |
|
40 |
|
41 goalw LList.thy [list_Fun_def] |
|
42 "!!M. [| M : X; X <= list_Fun(A, X Un llist(A)) |] ==> M : llist(A)"; |
|
43 by (etac llist.coinduct 1); |
|
44 by (etac (subsetD RS CollectD) 1); |
|
45 by (assume_tac 1); |
|
46 qed "llist_coinduct"; |
|
47 |
|
48 goalw LList.thy [list_Fun_def, NIL_def] "NIL: list_Fun(A,X)"; |
|
49 by (fast_tac set_cs 1); |
|
50 qed "list_Fun_NIL_I"; |
|
51 |
|
52 goalw LList.thy [list_Fun_def,CONS_def] |
|
53 "!!M N. [| M: A; N: X |] ==> CONS(M,N) : list_Fun(A,X)"; |
|
54 by (fast_tac set_cs 1); |
|
55 qed "list_Fun_CONS_I"; |
|
56 |
|
57 (*Utilise the "strong" part, i.e. gfp(f)*) |
|
58 goalw LList.thy (llist.defs @ [list_Fun_def]) |
|
59 "!!M N. M: llist(A) ==> M : list_Fun(A, X Un llist(A))"; |
|
60 by (etac (llist.mono RS gfp_fun_UnI2) 1); |
|
61 qed "list_Fun_llist_I"; |
|
62 |
|
63 (*** LList_corec satisfies the desired recurion equation ***) |
|
64 |
|
65 (*A continuity result?*) |
|
66 goalw LList.thy [CONS_def] "CONS(M, UN x.f(x)) = (UN x. CONS(M, f(x)))"; |
|
67 by (simp_tac (univ_ss addsimps [In1_UN1, Scons_UN1_y]) 1); |
|
68 qed "CONS_UN1"; |
|
69 |
|
70 (*UNUSED; obsolete? |
|
71 goal Prod.thy "split(p, %x y.UN z.f(x,y,z)) = (UN z. split(p, %x y.f(x,y,z)))"; |
|
72 by (simp_tac (prod_ss setloop (split_tac [expand_split])) 1); |
|
73 qed "split_UN1"; |
|
74 |
|
75 goal Sum.thy "sum_case(s,f,%y.UN z.g(y,z)) = (UN z.sum_case(s,f,%y. g(y,z)))"; |
|
76 by (simp_tac (sum_ss setloop (split_tac [expand_sum_case])) 1); |
|
77 qed "sum_case2_UN1"; |
|
78 *) |
|
79 |
|
80 val prems = goalw LList.thy [CONS_def] |
|
81 "[| M<=M'; N<=N' |] ==> CONS(M,N) <= CONS(M',N')"; |
|
82 by (REPEAT (resolve_tac ([In1_mono,Scons_mono]@prems) 1)); |
|
83 qed "CONS_mono"; |
|
84 |
|
85 val corec_fun_simps = [LList_corec_fun_def RS def_nat_rec_0, |
|
86 LList_corec_fun_def RS def_nat_rec_Suc]; |
|
87 val corec_fun_ss = llist_ss addsimps corec_fun_simps; |
|
88 |
|
89 (** The directions of the equality are proved separately **) |
|
90 |
|
91 goalw LList.thy [LList_corec_def] |
|
92 "LList_corec(a,f) <= sum_case(%u.NIL, \ |
|
93 \ split(%z w. CONS(z, LList_corec(w,f))), f(a))"; |
|
94 by (rtac UN1_least 1); |
|
95 by (res_inst_tac [("n","k")] natE 1); |
|
96 by (ALLGOALS (asm_simp_tac corec_fun_ss)); |
|
97 by (REPEAT (resolve_tac [allI, impI, subset_refl RS CONS_mono, UN1_upper] 1)); |
|
98 qed "LList_corec_subset1"; |
|
99 |
|
100 goalw LList.thy [LList_corec_def] |
|
101 "sum_case(%u.NIL, split(%z w. CONS(z, LList_corec(w,f))), f(a)) <= \ |
|
102 \ LList_corec(a,f)"; |
|
103 by (simp_tac (corec_fun_ss addsimps [CONS_UN1]) 1); |
|
104 by (safe_tac set_cs); |
|
105 by (ALLGOALS (res_inst_tac [("x","Suc(?k)")] UN1_I THEN' |
|
106 asm_simp_tac corec_fun_ss)); |
|
107 qed "LList_corec_subset2"; |
|
108 |
|
109 (*the recursion equation for LList_corec -- NOT SUITABLE FOR REWRITING!*) |
|
110 goal LList.thy |
|
111 "LList_corec(a,f) = sum_case(%u. NIL, \ |
|
112 \ split(%z w. CONS(z, LList_corec(w,f))), f(a))"; |
|
113 by (REPEAT (resolve_tac [equalityI, LList_corec_subset1, |
|
114 LList_corec_subset2] 1)); |
|
115 qed "LList_corec"; |
|
116 |
|
117 (*definitional version of same*) |
|
118 val [rew] = goal LList.thy |
|
119 "[| !!x. h(x) == LList_corec(x,f) |] ==> \ |
|
120 \ h(a) = sum_case(%u.NIL, split(%z w. CONS(z, h(w))), f(a))"; |
|
121 by (rewtac rew); |
|
122 by (rtac LList_corec 1); |
|
123 qed "def_LList_corec"; |
|
124 |
|
125 (*A typical use of co-induction to show membership in the gfp. |
|
126 Bisimulation is range(%x. LList_corec(x,f)) *) |
|
127 goal LList.thy "LList_corec(a,f) : llist({u.True})"; |
|
128 by (res_inst_tac [("X", "range(%x.LList_corec(x,?g))")] llist_coinduct 1); |
|
129 by (rtac rangeI 1); |
|
130 by (safe_tac set_cs); |
|
131 by (stac LList_corec 1); |
|
132 by (simp_tac (llist_ss addsimps [list_Fun_NIL_I, list_Fun_CONS_I, CollectI] |
|
133 |> add_eqI) 1); |
|
134 qed "LList_corec_type"; |
|
135 |
|
136 (*Lemma for the proof of llist_corec*) |
|
137 goal LList.thy |
|
138 "LList_corec(a, %z.sum_case(Inl, split(%v w.Inr(<Leaf(v),w>)), f(z))) : \ |
|
139 \ llist(range(Leaf))"; |
|
140 by (res_inst_tac [("X", "range(%x.LList_corec(x,?g))")] llist_coinduct 1); |
|
141 by (rtac rangeI 1); |
|
142 by (safe_tac set_cs); |
|
143 by (stac LList_corec 1); |
|
144 by (asm_simp_tac (llist_ss addsimps [list_Fun_NIL_I]) 1); |
|
145 by (fast_tac (set_cs addSIs [list_Fun_CONS_I]) 1); |
|
146 qed "LList_corec_type2"; |
|
147 |
|
148 |
|
149 (**** llist equality as a gfp; the bisimulation principle ****) |
|
150 |
|
151 (*This theorem is actually used, unlike the many similar ones in ZF*) |
|
152 goal LList.thy "LListD(r) = diag({Numb(0)}) <++> (r <**> LListD(r))"; |
|
153 let val rew = rewrite_rule [NIL_def, CONS_def] in |
|
154 by (fast_tac (univ_cs addSIs (equalityI :: map rew LListD.intrs) |
|
155 addEs [rew LListD.elim]) 1) |
|
156 end; |
|
157 qed "LListD_unfold"; |
|
158 |
|
159 goal LList.thy "!M N. <M,N> : LListD(diag(A)) --> ntrunc(k,M) = ntrunc(k,N)"; |
|
160 by (res_inst_tac [("n", "k")] less_induct 1); |
|
161 by (safe_tac set_cs); |
|
162 by (etac LListD.elim 1); |
|
163 by (safe_tac (prod_cs addSEs [diagE])); |
|
164 by (res_inst_tac [("n", "n")] natE 1); |
|
165 by (asm_simp_tac (univ_ss addsimps [ntrunc_0]) 1); |
|
166 by (rename_tac "n'" 1); |
|
167 by (res_inst_tac [("n", "n'")] natE 1); |
|
168 by (asm_simp_tac (univ_ss addsimps [CONS_def, ntrunc_one_In1]) 1); |
|
169 by (asm_simp_tac (univ_ss addsimps [CONS_def, ntrunc_In1, ntrunc_Scons]) 1); |
|
170 qed "LListD_implies_ntrunc_equality"; |
|
171 |
|
172 (*The domain of the LListD relation*) |
|
173 goalw LList.thy (llist.defs @ [NIL_def, CONS_def]) |
|
174 "fst``LListD(diag(A)) <= llist(A)"; |
|
175 by (rtac gfp_upperbound 1); |
|
176 (*avoids unfolding LListD on the rhs*) |
|
177 by (res_inst_tac [("P", "%x. fst``x <= ?B")] (LListD_unfold RS ssubst) 1); |
|
178 by (simp_tac fst_image_ss 1); |
|
179 by (fast_tac univ_cs 1); |
|
180 qed "fst_image_LListD"; |
|
181 |
|
182 (*This inclusion justifies the use of coinduction to show M=N*) |
|
183 goal LList.thy "LListD(diag(A)) <= diag(llist(A))"; |
|
184 by (rtac subsetI 1); |
|
185 by (res_inst_tac [("p","x")] PairE 1); |
|
186 by (safe_tac HOL_cs); |
|
187 by (rtac diag_eqI 1); |
|
188 by (rtac (LListD_implies_ntrunc_equality RS spec RS spec RS mp RS |
|
189 ntrunc_equality) 1); |
|
190 by (assume_tac 1); |
|
191 by (etac (fst_imageI RS (fst_image_LListD RS subsetD)) 1); |
|
192 qed "LListD_subset_diag"; |
|
193 |
|
194 (** Coinduction, using LListD_Fun |
|
195 THE COINDUCTIVE DEFINITION PACKAGE COULD DO THIS! |
|
196 **) |
|
197 |
|
198 goalw LList.thy [LListD_Fun_def] |
|
199 "!!M. [| M : X; X <= LListD_Fun(r, X Un LListD(r)) |] ==> M : LListD(r)"; |
|
200 by (etac LListD.coinduct 1); |
|
201 by (etac (subsetD RS CollectD) 1); |
|
202 by (assume_tac 1); |
|
203 qed "LListD_coinduct"; |
|
204 |
|
205 goalw LList.thy [LListD_Fun_def,NIL_def] "<NIL,NIL> : LListD_Fun(r,s)"; |
|
206 by (fast_tac set_cs 1); |
|
207 qed "LListD_Fun_NIL_I"; |
|
208 |
|
209 goalw LList.thy [LListD_Fun_def,CONS_def] |
|
210 "!!x. [| x:A; <M,N>:s |] ==> <CONS(x,M), CONS(x,N)> : LListD_Fun(diag(A),s)"; |
|
211 by (fast_tac univ_cs 1); |
|
212 qed "LListD_Fun_CONS_I"; |
|
213 |
|
214 (*Utilise the "strong" part, i.e. gfp(f)*) |
|
215 goalw LList.thy (LListD.defs @ [LListD_Fun_def]) |
|
216 "!!M N. M: LListD(r) ==> M : LListD_Fun(r, X Un LListD(r))"; |
|
217 by (etac (LListD.mono RS gfp_fun_UnI2) 1); |
|
218 qed "LListD_Fun_LListD_I"; |
|
219 |
|
220 |
|
221 (*This converse inclusion helps to strengthen LList_equalityI*) |
|
222 goal LList.thy "diag(llist(A)) <= LListD(diag(A))"; |
|
223 by (rtac subsetI 1); |
|
224 by (etac LListD_coinduct 1); |
|
225 by (rtac subsetI 1); |
|
226 by (etac diagE 1); |
|
227 by (etac ssubst 1); |
|
228 by (eresolve_tac [llist.elim] 1); |
|
229 by (ALLGOALS |
|
230 (asm_simp_tac (llist_ss addsimps [diagI, LListD_Fun_NIL_I, |
|
231 LListD_Fun_CONS_I]))); |
|
232 qed "diag_subset_LListD"; |
|
233 |
|
234 goal LList.thy "LListD(diag(A)) = diag(llist(A))"; |
|
235 by (REPEAT (resolve_tac [equalityI, LListD_subset_diag, |
|
236 diag_subset_LListD] 1)); |
|
237 qed "LListD_eq_diag"; |
|
238 |
|
239 goal LList.thy |
|
240 "!!M N. M: llist(A) ==> <M,M> : LListD_Fun(diag(A), X Un diag(llist(A)))"; |
|
241 by (rtac (LListD_eq_diag RS subst) 1); |
|
242 by (rtac LListD_Fun_LListD_I 1); |
|
243 by (asm_simp_tac (HOL_ss addsimps [LListD_eq_diag, diagI]) 1); |
|
244 qed "LListD_Fun_diag_I"; |
|
245 |
|
246 |
|
247 (** To show two LLists are equal, exhibit a bisimulation! |
|
248 [also admits true equality] |
|
249 Replace "A" by some particular set, like {x.True}??? *) |
|
250 goal LList.thy |
|
251 "!!r. [| <M,N> : r; r <= LListD_Fun(diag(A), r Un diag(llist(A))) \ |
|
252 \ |] ==> M=N"; |
|
253 by (rtac (LListD_subset_diag RS subsetD RS diagE) 1); |
|
254 by (etac LListD_coinduct 1); |
|
255 by (asm_simp_tac (HOL_ss addsimps [LListD_eq_diag]) 1); |
|
256 by (safe_tac prod_cs); |
|
257 qed "LList_equalityI"; |
|
258 |
|
259 |
|
260 (*** Finality of llist(A): Uniqueness of functions defined by corecursion ***) |
|
261 |
|
262 (*abstract proof using a bisimulation*) |
|
263 val [prem1,prem2] = goal LList.thy |
|
264 "[| !!x. h1(x) = sum_case(%u.NIL, split(%z w. CONS(z,h1(w))), f(x)); \ |
|
265 \ !!x. h2(x) = sum_case(%u.NIL, split(%z w. CONS(z,h2(w))), f(x)) |]\ |
|
266 \ ==> h1=h2"; |
|
267 by (rtac ext 1); |
|
268 (*next step avoids an unknown (and flexflex pair) in simplification*) |
|
269 by (res_inst_tac [("A", "{u.True}"), |
|
270 ("r", "range(%u. <h1(u),h2(u)>)")] LList_equalityI 1); |
|
271 by (rtac rangeI 1); |
|
272 by (safe_tac set_cs); |
|
273 by (stac prem1 1); |
|
274 by (stac prem2 1); |
|
275 by (simp_tac (llist_ss addsimps [LListD_Fun_NIL_I, |
|
276 CollectI RS LListD_Fun_CONS_I] |
|
277 |> add_eqI) 1); |
|
278 qed "LList_corec_unique"; |
|
279 |
|
280 val [prem] = goal LList.thy |
|
281 "[| !!x. h(x) = sum_case(%u.NIL, split(%z w. CONS(z,h(w))), f(x)) |] \ |
|
282 \ ==> h = (%x.LList_corec(x,f))"; |
|
283 by (rtac (LList_corec RS (prem RS LList_corec_unique)) 1); |
|
284 qed "equals_LList_corec"; |
|
285 |
|
286 |
|
287 (** Obsolete LList_corec_unique proof: complete induction, not coinduction **) |
|
288 |
|
289 goalw LList.thy [CONS_def] "ntrunc(Suc(0), CONS(M,N)) = {}"; |
|
290 by (rtac ntrunc_one_In1 1); |
|
291 qed "ntrunc_one_CONS"; |
|
292 |
|
293 goalw LList.thy [CONS_def] |
|
294 "ntrunc(Suc(Suc(k)), CONS(M,N)) = CONS (ntrunc(k,M), ntrunc(k,N))"; |
|
295 by (simp_tac (HOL_ss addsimps [ntrunc_Scons,ntrunc_In1]) 1); |
|
296 qed "ntrunc_CONS"; |
|
297 |
|
298 val [prem1,prem2] = goal LList.thy |
|
299 "[| !!x. h1(x) = sum_case(%u.NIL, split(%z w. CONS(z,h1(w))), f(x)); \ |
|
300 \ !!x. h2(x) = sum_case(%u.NIL, split(%z w. CONS(z,h2(w))), f(x)) |]\ |
|
301 \ ==> h1=h2"; |
|
302 by (rtac (ntrunc_equality RS ext) 1); |
|
303 by (res_inst_tac [("x", "x")] spec 1); |
|
304 by (res_inst_tac [("n", "k")] less_induct 1); |
|
305 by (rtac allI 1); |
|
306 by (stac prem1 1); |
|
307 by (stac prem2 1); |
|
308 by (simp_tac (sum_ss setloop (split_tac [expand_split,expand_sum_case])) 1); |
|
309 by (strip_tac 1); |
|
310 by (res_inst_tac [("n", "n")] natE 1); |
|
311 by (res_inst_tac [("n", "xc")] natE 2); |
|
312 by (ALLGOALS(asm_simp_tac(nat_ss addsimps |
|
313 [ntrunc_0,ntrunc_one_CONS,ntrunc_CONS]))); |
|
314 result(); |
|
315 |
|
316 |
|
317 (*** Lconst -- defined directly using lfp, but equivalent to a LList_corec ***) |
|
318 |
|
319 goal LList.thy "mono(CONS(M))"; |
|
320 by (REPEAT (ares_tac [monoI, subset_refl, CONS_mono] 1)); |
|
321 qed "Lconst_fun_mono"; |
|
322 |
|
323 (* Lconst(M) = CONS(M,Lconst(M)) *) |
|
324 bind_thm ("Lconst", (Lconst_fun_mono RS (Lconst_def RS def_lfp_Tarski))); |
|
325 |
|
326 (*A typical use of co-induction to show membership in the gfp. |
|
327 The containing set is simply the singleton {Lconst(M)}. *) |
|
328 goal LList.thy "!!M A. M:A ==> Lconst(M): llist(A)"; |
|
329 by (rtac (singletonI RS llist_coinduct) 1); |
|
330 by (safe_tac set_cs); |
|
331 by (res_inst_tac [("P", "%u. u: ?A")] (Lconst RS ssubst) 1); |
|
332 by (REPEAT (ares_tac [list_Fun_CONS_I, singletonI, UnI1] 1)); |
|
333 qed "Lconst_type"; |
|
334 |
|
335 goal LList.thy "Lconst(M) = LList_corec(M, %x.Inr(<x,x>))"; |
|
336 by (rtac (equals_LList_corec RS fun_cong) 1); |
|
337 by (simp_tac sum_ss 1); |
|
338 by (rtac Lconst 1); |
|
339 qed "Lconst_eq_LList_corec"; |
|
340 |
|
341 (*Thus we could have used gfp in the definition of Lconst*) |
|
342 goal LList.thy "gfp(%N. CONS(M, N)) = LList_corec(M, %x.Inr(<x,x>))"; |
|
343 by (rtac (equals_LList_corec RS fun_cong) 1); |
|
344 by (simp_tac sum_ss 1); |
|
345 by (rtac (Lconst_fun_mono RS gfp_Tarski) 1); |
|
346 qed "gfp_Lconst_eq_LList_corec"; |
|
347 |
|
348 |
|
349 (*** Isomorphisms ***) |
|
350 |
|
351 goal LList.thy "inj(Rep_llist)"; |
|
352 by (rtac inj_inverseI 1); |
|
353 by (rtac Rep_llist_inverse 1); |
|
354 qed "inj_Rep_llist"; |
|
355 |
|
356 goal LList.thy "inj_onto(Abs_llist,llist(range(Leaf)))"; |
|
357 by (rtac inj_onto_inverseI 1); |
|
358 by (etac Abs_llist_inverse 1); |
|
359 qed "inj_onto_Abs_llist"; |
|
360 |
|
361 (** Distinctness of constructors **) |
|
362 |
|
363 goalw LList.thy [LNil_def,LCons_def] "~ LCons(x,xs) = LNil"; |
|
364 by (rtac (CONS_not_NIL RS (inj_onto_Abs_llist RS inj_onto_contraD)) 1); |
|
365 by (REPEAT (resolve_tac (llist.intrs @ [rangeI, Rep_llist]) 1)); |
|
366 qed "LCons_not_LNil"; |
|
367 |
|
368 bind_thm ("LNil_not_LCons", (LCons_not_LNil RS not_sym)); |
|
369 |
|
370 bind_thm ("LCons_neq_LNil", (LCons_not_LNil RS notE)); |
|
371 val LNil_neq_LCons = sym RS LCons_neq_LNil; |
|
372 |
|
373 (** llist constructors **) |
|
374 |
|
375 goalw LList.thy [LNil_def] |
|
376 "Rep_llist(LNil) = NIL"; |
|
377 by (rtac (llist.NIL_I RS Abs_llist_inverse) 1); |
|
378 qed "Rep_llist_LNil"; |
|
379 |
|
380 goalw LList.thy [LCons_def] |
|
381 "Rep_llist(LCons(x,l)) = CONS(Leaf(x),Rep_llist(l))"; |
|
382 by (REPEAT (resolve_tac [llist.CONS_I RS Abs_llist_inverse, |
|
383 rangeI, Rep_llist] 1)); |
|
384 qed "Rep_llist_LCons"; |
|
385 |
|
386 (** Injectiveness of CONS and LCons **) |
|
387 |
|
388 goalw LList.thy [CONS_def] "(CONS(M,N)=CONS(M',N')) = (M=M' & N=N')"; |
|
389 by (fast_tac (HOL_cs addSEs [Scons_inject, make_elim In1_inject]) 1); |
|
390 qed "CONS_CONS_eq"; |
|
391 |
|
392 bind_thm ("CONS_inject", (CONS_CONS_eq RS iffD1 RS conjE)); |
|
393 |
|
394 |
|
395 (*For reasoning about abstract llist constructors*) |
|
396 val llist_cs = set_cs addIs [Rep_llist]@llist.intrs |
|
397 addSEs [CONS_neq_NIL,NIL_neq_CONS,CONS_inject] |
|
398 addSDs [inj_onto_Abs_llist RS inj_ontoD, |
|
399 inj_Rep_llist RS injD, Leaf_inject]; |
|
400 |
|
401 goalw LList.thy [LCons_def] "(LCons(x,xs)=LCons(y,ys)) = (x=y & xs=ys)"; |
|
402 by (fast_tac llist_cs 1); |
|
403 qed "LCons_LCons_eq"; |
|
404 bind_thm ("LCons_inject", (LCons_LCons_eq RS iffD1 RS conjE)); |
|
405 |
|
406 val [major] = goal LList.thy "CONS(M,N): llist(A) ==> M: A & N: llist(A)"; |
|
407 by (rtac (major RS llist.elim) 1); |
|
408 by (etac CONS_neq_NIL 1); |
|
409 by (fast_tac llist_cs 1); |
|
410 qed "CONS_D"; |
|
411 |
|
412 |
|
413 (****** Reasoning about llist(A) ******) |
|
414 |
|
415 (*Don't use llist_ss, as it does case splits!*) |
|
416 val List_case_ss = univ_ss addsimps [List_case_NIL, List_case_CONS]; |
|
417 |
|
418 (*A special case of list_equality for functions over lazy lists*) |
|
419 val [Mlist,gMlist,NILcase,CONScase] = goal LList.thy |
|
420 "[| M: llist(A); g(NIL): llist(A); \ |
|
421 \ f(NIL)=g(NIL); \ |
|
422 \ !!x l. [| x:A; l: llist(A) |] ==> \ |
|
423 \ <f(CONS(x,l)),g(CONS(x,l))> : \ |
|
424 \ LListD_Fun(diag(A), (%u.<f(u),g(u)>)``llist(A) Un \ |
|
425 \ diag(llist(A))) \ |
|
426 \ |] ==> f(M) = g(M)"; |
|
427 by (rtac LList_equalityI 1); |
|
428 by (rtac (Mlist RS imageI) 1); |
|
429 by (rtac subsetI 1); |
|
430 by (etac imageE 1); |
|
431 by (etac ssubst 1); |
|
432 by (etac llist.elim 1); |
|
433 by (etac ssubst 1); |
|
434 by (stac NILcase 1); |
|
435 by (rtac (gMlist RS LListD_Fun_diag_I) 1); |
|
436 by (etac ssubst 1); |
|
437 by (REPEAT (ares_tac [CONScase] 1)); |
|
438 qed "LList_fun_equalityI"; |
|
439 |
|
440 |
|
441 (*** The functional "Lmap" ***) |
|
442 |
|
443 goal LList.thy "Lmap(f,NIL) = NIL"; |
|
444 by (rtac (Lmap_def RS def_LList_corec RS trans) 1); |
|
445 by (simp_tac List_case_ss 1); |
|
446 qed "Lmap_NIL"; |
|
447 |
|
448 goal LList.thy "Lmap(f, CONS(M,N)) = CONS(f(M), Lmap(f,N))"; |
|
449 by (rtac (Lmap_def RS def_LList_corec RS trans) 1); |
|
450 by (simp_tac List_case_ss 1); |
|
451 qed "Lmap_CONS"; |
|
452 |
|
453 (*Another type-checking proof by coinduction*) |
|
454 val [major,minor] = goal LList.thy |
|
455 "[| M: llist(A); !!x. x:A ==> f(x):B |] ==> Lmap(f,M): llist(B)"; |
|
456 by (rtac (major RS imageI RS llist_coinduct) 1); |
|
457 by (safe_tac set_cs); |
|
458 by (etac llist.elim 1); |
|
459 by (ALLGOALS (asm_simp_tac (HOL_ss addsimps [Lmap_NIL,Lmap_CONS]))); |
|
460 by (REPEAT (ares_tac [list_Fun_NIL_I, list_Fun_CONS_I, |
|
461 minor, imageI, UnI1] 1)); |
|
462 qed "Lmap_type"; |
|
463 |
|
464 (*This type checking rule synthesises a sufficiently large set for f*) |
|
465 val [major] = goal LList.thy "M: llist(A) ==> Lmap(f,M): llist(f``A)"; |
|
466 by (rtac (major RS Lmap_type) 1); |
|
467 by (etac imageI 1); |
|
468 qed "Lmap_type2"; |
|
469 |
|
470 (** Two easy results about Lmap **) |
|
471 |
|
472 val [prem] = goalw LList.thy [o_def] |
|
473 "M: llist(A) ==> Lmap(f o g, M) = Lmap(f, Lmap(g, M))"; |
|
474 by (rtac (prem RS imageI RS LList_equalityI) 1); |
|
475 by (safe_tac set_cs); |
|
476 by (etac llist.elim 1); |
|
477 by (ALLGOALS (asm_simp_tac (HOL_ss addsimps [Lmap_NIL,Lmap_CONS]))); |
|
478 by (REPEAT (ares_tac [LListD_Fun_NIL_I, imageI, UnI1, |
|
479 rangeI RS LListD_Fun_CONS_I] 1)); |
|
480 qed "Lmap_compose"; |
|
481 |
|
482 val [prem] = goal LList.thy "M: llist(A) ==> Lmap(%x.x, M) = M"; |
|
483 by (rtac (prem RS imageI RS LList_equalityI) 1); |
|
484 by (safe_tac set_cs); |
|
485 by (etac llist.elim 1); |
|
486 by (ALLGOALS (asm_simp_tac (HOL_ss addsimps [Lmap_NIL,Lmap_CONS]))); |
|
487 by (REPEAT (ares_tac [LListD_Fun_NIL_I, imageI RS UnI1, |
|
488 rangeI RS LListD_Fun_CONS_I] 1)); |
|
489 qed "Lmap_ident"; |
|
490 |
|
491 |
|
492 (*** Lappend -- its two arguments cause some complications! ***) |
|
493 |
|
494 goalw LList.thy [Lappend_def] "Lappend(NIL,NIL) = NIL"; |
|
495 by (rtac (LList_corec RS trans) 1); |
|
496 by (simp_tac List_case_ss 1); |
|
497 qed "Lappend_NIL_NIL"; |
|
498 |
|
499 goalw LList.thy [Lappend_def] |
|
500 "Lappend(NIL,CONS(N,N')) = CONS(N, Lappend(NIL,N'))"; |
|
501 by (rtac (LList_corec RS trans) 1); |
|
502 by (simp_tac List_case_ss 1); |
|
503 qed "Lappend_NIL_CONS"; |
|
504 |
|
505 goalw LList.thy [Lappend_def] |
|
506 "Lappend(CONS(M,M'), N) = CONS(M, Lappend(M',N))"; |
|
507 by (rtac (LList_corec RS trans) 1); |
|
508 by (simp_tac List_case_ss 1); |
|
509 qed "Lappend_CONS"; |
|
510 |
|
511 val Lappend_ss = |
|
512 List_case_ss addsimps [llist.NIL_I, Lappend_NIL_NIL, Lappend_NIL_CONS, |
|
513 Lappend_CONS, LListD_Fun_CONS_I] |
|
514 |> add_eqI; |
|
515 |
|
516 goal LList.thy "!!M. M: llist(A) ==> Lappend(NIL,M) = M"; |
|
517 by (etac LList_fun_equalityI 1); |
|
518 by (ALLGOALS (asm_simp_tac Lappend_ss)); |
|
519 qed "Lappend_NIL"; |
|
520 |
|
521 goal LList.thy "!!M. M: llist(A) ==> Lappend(M,NIL) = M"; |
|
522 by (etac LList_fun_equalityI 1); |
|
523 by (ALLGOALS (asm_simp_tac Lappend_ss)); |
|
524 qed "Lappend_NIL2"; |
|
525 |
|
526 (** Alternative type-checking proofs for Lappend **) |
|
527 |
|
528 (*weak co-induction: bisimulation and case analysis on both variables*) |
|
529 goal LList.thy |
|
530 "!!M N. [| M: llist(A); N: llist(A) |] ==> Lappend(M,N): llist(A)"; |
|
531 by (res_inst_tac |
|
532 [("X", "UN u:llist(A). UN v: llist(A). {Lappend(u,v)}")] llist_coinduct 1); |
|
533 by (fast_tac set_cs 1); |
|
534 by (safe_tac set_cs); |
|
535 by (eres_inst_tac [("a", "u")] llist.elim 1); |
|
536 by (eres_inst_tac [("a", "v")] llist.elim 1); |
|
537 by (ALLGOALS |
|
538 (asm_simp_tac Lappend_ss THEN' |
|
539 fast_tac (set_cs addSIs [llist.NIL_I, list_Fun_NIL_I, list_Fun_CONS_I]))); |
|
540 qed "Lappend_type"; |
|
541 |
|
542 (*strong co-induction: bisimulation and case analysis on one variable*) |
|
543 goal LList.thy |
|
544 "!!M N. [| M: llist(A); N: llist(A) |] ==> Lappend(M,N): llist(A)"; |
|
545 by (res_inst_tac [("X", "(%u.Lappend(u,N))``llist(A)")] llist_coinduct 1); |
|
546 by (etac imageI 1); |
|
547 by (rtac subsetI 1); |
|
548 by (etac imageE 1); |
|
549 by (eres_inst_tac [("a", "u")] llist.elim 1); |
|
550 by (asm_simp_tac (Lappend_ss addsimps [Lappend_NIL, list_Fun_llist_I]) 1); |
|
551 by (asm_simp_tac Lappend_ss 1); |
|
552 by (fast_tac (set_cs addSIs [list_Fun_CONS_I]) 1); |
|
553 qed "Lappend_type"; |
|
554 |
|
555 (**** Lazy lists as the type 'a llist -- strongly typed versions of above ****) |
|
556 |
|
557 (** llist_case: case analysis for 'a llist **) |
|
558 |
|
559 val Rep_llist_simps = |
|
560 [List_case_NIL, List_case_CONS, |
|
561 Abs_llist_inverse, Rep_llist_inverse, |
|
562 Rep_llist, rangeI, inj_Leaf, Inv_f_f] |
|
563 @ llist.intrs; |
|
564 val Rep_llist_ss = llist_ss addsimps Rep_llist_simps; |
|
565 |
|
566 goalw LList.thy [llist_case_def,LNil_def] "llist_case(c, d, LNil) = c"; |
|
567 by (simp_tac Rep_llist_ss 1); |
|
568 qed "llist_case_LNil"; |
|
569 |
|
570 goalw LList.thy [llist_case_def,LCons_def] |
|
571 "llist_case(c, d, LCons(M,N)) = d(M,N)"; |
|
572 by (simp_tac Rep_llist_ss 1); |
|
573 qed "llist_case_LCons"; |
|
574 |
|
575 (*Elimination is case analysis, not induction.*) |
|
576 val [prem1,prem2] = goalw LList.thy [NIL_def,CONS_def] |
|
577 "[| l=LNil ==> P; !!x l'. l=LCons(x,l') ==> P \ |
|
578 \ |] ==> P"; |
|
579 by (rtac (Rep_llist RS llist.elim) 1); |
|
580 by (rtac (inj_Rep_llist RS injD RS prem1) 1); |
|
581 by (stac Rep_llist_LNil 1); |
|
582 by (assume_tac 1); |
|
583 by (etac rangeE 1); |
|
584 by (rtac (inj_Rep_llist RS injD RS prem2) 1); |
|
585 by (asm_simp_tac (HOL_ss addsimps [Rep_llist_LCons]) 1); |
|
586 by (etac (Abs_llist_inverse RS ssubst) 1); |
|
587 by (rtac refl 1); |
|
588 qed "llistE"; |
|
589 |
|
590 (** llist_corec: corecursion for 'a llist **) |
|
591 |
|
592 goalw LList.thy [llist_corec_def,LNil_def,LCons_def] |
|
593 "llist_corec(a,f) = sum_case(%u. LNil, \ |
|
594 \ split(%z w. LCons(z, llist_corec(w,f))), f(a))"; |
|
595 by (stac LList_corec 1); |
|
596 by (res_inst_tac [("s","f(a)")] sumE 1); |
|
597 by (asm_simp_tac (llist_ss addsimps [LList_corec_type2,Abs_llist_inverse]) 1); |
|
598 by (res_inst_tac [("p","y")] PairE 1); |
|
599 by (asm_simp_tac (llist_ss addsimps [LList_corec_type2,Abs_llist_inverse]) 1); |
|
600 (*FIXME: correct case splits usd to be found automatically: |
|
601 by (ASM_SIMP_TAC(llist_ss addsimps [LList_corec_type2,Abs_llist_inverse]) 1);*) |
|
602 qed "llist_corec"; |
|
603 |
|
604 (*definitional version of same*) |
|
605 val [rew] = goal LList.thy |
|
606 "[| !!x. h(x) == llist_corec(x,f) |] ==> \ |
|
607 \ h(a) = sum_case(%u.LNil, split(%z w. LCons(z, h(w))), f(a))"; |
|
608 by (rewtac rew); |
|
609 by (rtac llist_corec 1); |
|
610 qed "def_llist_corec"; |
|
611 |
|
612 (**** Proofs about type 'a llist functions ****) |
|
613 |
|
614 (*** Deriving llist_equalityI -- llist equality is a bisimulation ***) |
|
615 |
|
616 goalw LList.thy [LListD_Fun_def] |
|
617 "!!r A. r <= Sigma(llist(A), %x.llist(A)) ==> \ |
|
618 \ LListD_Fun(diag(A),r) <= Sigma(llist(A), %x.llist(A))"; |
|
619 by (stac llist_unfold 1); |
|
620 by (simp_tac (HOL_ss addsimps [NIL_def, CONS_def]) 1); |
|
621 by (fast_tac univ_cs 1); |
|
622 qed "LListD_Fun_subset_Sigma_llist"; |
|
623 |
|
624 goal LList.thy |
|
625 "prod_fun(Rep_llist,Rep_llist) `` r <= \ |
|
626 \ Sigma(llist(range(Leaf)), %x.llist(range(Leaf)))"; |
|
627 by (fast_tac (prod_cs addIs [Rep_llist]) 1); |
|
628 qed "subset_Sigma_llist"; |
|
629 |
|
630 val [prem] = goal LList.thy |
|
631 "r <= Sigma(llist(range(Leaf)), %x.llist(range(Leaf))) ==> \ |
|
632 \ prod_fun(Rep_llist o Abs_llist, Rep_llist o Abs_llist) `` r <= r"; |
|
633 by (safe_tac prod_cs); |
|
634 by (rtac (prem RS subsetD RS SigmaE2) 1); |
|
635 by (assume_tac 1); |
|
636 by (asm_simp_tac (HOL_ss addsimps [o_def,prod_fun,Abs_llist_inverse]) 1); |
|
637 qed "prod_fun_lemma"; |
|
638 |
|
639 goal LList.thy |
|
640 "prod_fun(Rep_llist, Rep_llist) `` range(%x. <x, x>) = \ |
|
641 \ diag(llist(range(Leaf)))"; |
|
642 by (rtac equalityI 1); |
|
643 by (fast_tac (univ_cs addIs [Rep_llist]) 1); |
|
644 by (fast_tac (univ_cs addSEs [Abs_llist_inverse RS subst]) 1); |
|
645 qed "prod_fun_range_eq_diag"; |
|
646 |
|
647 (** To show two llists are equal, exhibit a bisimulation! |
|
648 [also admits true equality] **) |
|
649 val [prem1,prem2] = goalw LList.thy [llistD_Fun_def] |
|
650 "[| <l1,l2> : r; r <= llistD_Fun(r Un range(%x.<x,x>)) |] ==> l1=l2"; |
|
651 by (rtac (inj_Rep_llist RS injD) 1); |
|
652 by (res_inst_tac [("r", "prod_fun(Rep_llist,Rep_llist)``r"), |
|
653 ("A", "range(Leaf)")] |
|
654 LList_equalityI 1); |
|
655 by (rtac (prem1 RS prod_fun_imageI) 1); |
|
656 by (rtac (prem2 RS image_mono RS subset_trans) 1); |
|
657 by (rtac (image_compose RS subst) 1); |
|
658 by (rtac (prod_fun_compose RS subst) 1); |
|
659 by (rtac (image_Un RS ssubst) 1); |
|
660 by (stac prod_fun_range_eq_diag 1); |
|
661 by (rtac (LListD_Fun_subset_Sigma_llist RS prod_fun_lemma) 1); |
|
662 by (rtac (subset_Sigma_llist RS Un_least) 1); |
|
663 by (rtac diag_subset_Sigma 1); |
|
664 qed "llist_equalityI"; |
|
665 |
|
666 (** Rules to prove the 2nd premise of llist_equalityI **) |
|
667 goalw LList.thy [llistD_Fun_def,LNil_def] "<LNil,LNil> : llistD_Fun(r)"; |
|
668 by (rtac (LListD_Fun_NIL_I RS prod_fun_imageI) 1); |
|
669 qed "llistD_Fun_LNil_I"; |
|
670 |
|
671 val [prem] = goalw LList.thy [llistD_Fun_def,LCons_def] |
|
672 "<l1,l2>:r ==> <LCons(x,l1), LCons(x,l2)> : llistD_Fun(r)"; |
|
673 by (rtac (rangeI RS LListD_Fun_CONS_I RS prod_fun_imageI) 1); |
|
674 by (rtac (prem RS prod_fun_imageI) 1); |
|
675 qed "llistD_Fun_LCons_I"; |
|
676 |
|
677 (*Utilise the "strong" part, i.e. gfp(f)*) |
|
678 goalw LList.thy [llistD_Fun_def] |
|
679 "!!l. <l,l> : llistD_Fun(r Un range(%x.<x,x>))"; |
|
680 by (rtac (Rep_llist_inverse RS subst) 1); |
|
681 by (rtac prod_fun_imageI 1); |
|
682 by (rtac (image_Un RS ssubst) 1); |
|
683 by (stac prod_fun_range_eq_diag 1); |
|
684 by (rtac (Rep_llist RS LListD_Fun_diag_I) 1); |
|
685 qed "llistD_Fun_range_I"; |
|
686 |
|
687 (*A special case of list_equality for functions over lazy lists*) |
|
688 val [prem1,prem2] = goal LList.thy |
|
689 "[| f(LNil)=g(LNil); \ |
|
690 \ !!x l. <f(LCons(x,l)),g(LCons(x,l))> : \ |
|
691 \ llistD_Fun(range(%u. <f(u),g(u)>) Un range(%v. <v,v>)) \ |
|
692 \ |] ==> f(l) = (g(l :: 'a llist) :: 'b llist)"; |
|
693 by (res_inst_tac [("r", "range(%u. <f(u),g(u)>)")] llist_equalityI 1); |
|
694 by (rtac rangeI 1); |
|
695 by (rtac subsetI 1); |
|
696 by (etac rangeE 1); |
|
697 by (etac ssubst 1); |
|
698 by (res_inst_tac [("l", "u")] llistE 1); |
|
699 by (etac ssubst 1); |
|
700 by (stac prem1 1); |
|
701 by (rtac llistD_Fun_range_I 1); |
|
702 by (etac ssubst 1); |
|
703 by (rtac prem2 1); |
|
704 qed "llist_fun_equalityI"; |
|
705 |
|
706 (*simpset for llist bisimulations*) |
|
707 val llistD_simps = [llist_case_LNil, llist_case_LCons, |
|
708 llistD_Fun_LNil_I, llistD_Fun_LCons_I]; |
|
709 (*Don't use llist_ss, as it does case splits!*) |
|
710 val llistD_ss = univ_ss addsimps llistD_simps |> add_eqI; |
|
711 |
|
712 |
|
713 (*** The functional "lmap" ***) |
|
714 |
|
715 goal LList.thy "lmap(f,LNil) = LNil"; |
|
716 by (rtac (lmap_def RS def_llist_corec RS trans) 1); |
|
717 by (simp_tac llistD_ss 1); |
|
718 qed "lmap_LNil"; |
|
719 |
|
720 goal LList.thy "lmap(f, LCons(M,N)) = LCons(f(M), lmap(f,N))"; |
|
721 by (rtac (lmap_def RS def_llist_corec RS trans) 1); |
|
722 by (simp_tac llistD_ss 1); |
|
723 qed "lmap_LCons"; |
|
724 |
|
725 |
|
726 (** Two easy results about lmap **) |
|
727 |
|
728 goal LList.thy "lmap(f o g, l) = lmap(f, lmap(g, l))"; |
|
729 by (res_inst_tac [("l","l")] llist_fun_equalityI 1); |
|
730 by (ALLGOALS (simp_tac (llistD_ss addsimps [lmap_LNil, lmap_LCons]))); |
|
731 qed "lmap_compose"; |
|
732 |
|
733 goal LList.thy "lmap(%x.x, l) = l"; |
|
734 by (res_inst_tac [("l","l")] llist_fun_equalityI 1); |
|
735 by (ALLGOALS (simp_tac (llistD_ss addsimps [lmap_LNil, lmap_LCons]))); |
|
736 qed "lmap_ident"; |
|
737 |
|
738 |
|
739 (*** iterates -- llist_fun_equalityI cannot be used! ***) |
|
740 |
|
741 goal LList.thy "iterates(f,x) = LCons(x, iterates(f,f(x)))"; |
|
742 by (rtac (iterates_def RS def_llist_corec RS trans) 1); |
|
743 by (simp_tac sum_ss 1); |
|
744 qed "iterates"; |
|
745 |
|
746 goal LList.thy "lmap(f, iterates(f,x)) = iterates(f,f(x))"; |
|
747 by (res_inst_tac [("r", "range(%u.<lmap(f,iterates(f,u)),iterates(f,f(u))>)")] |
|
748 llist_equalityI 1); |
|
749 by (rtac rangeI 1); |
|
750 by (safe_tac set_cs); |
|
751 by (res_inst_tac [("x1", "f(u)")] (iterates RS ssubst) 1); |
|
752 by (res_inst_tac [("x1", "u")] (iterates RS ssubst) 1); |
|
753 by (simp_tac (llistD_ss addsimps [lmap_LCons]) 1); |
|
754 qed "lmap_iterates"; |
|
755 |
|
756 goal LList.thy "iterates(f,x) = LCons(x, lmap(f, iterates(f,x)))"; |
|
757 by (rtac (lmap_iterates RS ssubst) 1); |
|
758 by (rtac iterates 1); |
|
759 qed "iterates_lmap"; |
|
760 |
|
761 (*** A rather complex proof about iterates -- cf Andy Pitts ***) |
|
762 |
|
763 (** Two lemmas about natrec(n,x,%m.g), which is essentially (g^n)(x) **) |
|
764 |
|
765 goal LList.thy |
|
766 "nat_rec(n, LCons(b, l), %m. lmap(f)) = \ |
|
767 \ LCons(nat_rec(n, b, %m. f), nat_rec(n, l, %m. lmap(f)))"; |
|
768 by (nat_ind_tac "n" 1); |
|
769 by (ALLGOALS (asm_simp_tac (nat_ss addsimps [lmap_LCons]))); |
|
770 qed "fun_power_lmap"; |
|
771 |
|
772 goal Nat.thy "nat_rec(n, g(x), %m. g) = nat_rec(Suc(n), x, %m. g)"; |
|
773 by (nat_ind_tac "n" 1); |
|
774 by (ALLGOALS (asm_simp_tac nat_ss)); |
|
775 qed "fun_power_Suc"; |
|
776 |
|
777 val Pair_cong = read_instantiate_sg (sign_of Prod.thy) |
|
778 [("f","Pair")] (standard(refl RS cong RS cong)); |
|
779 |
|
780 (*The bisimulation consists of {<lmap(f)^n (h(u)), lmap(f)^n (iterates(f,u))>} |
|
781 for all u and all n::nat.*) |
|
782 val [prem] = goal LList.thy |
|
783 "(!!x. h(x) = LCons(x, lmap(f,h(x)))) ==> h = iterates(f)"; |
|
784 by (rtac ext 1); |
|
785 by (res_inst_tac [("r", |
|
786 "UN u. range(%n. <nat_rec(n, h(u), %m y.lmap(f,y)), \ |
|
787 \ nat_rec(n, iterates(f,u), %m y.lmap(f,y))>)")] |
|
788 llist_equalityI 1); |
|
789 by (REPEAT (resolve_tac [UN1_I, range_eqI, Pair_cong, nat_rec_0 RS sym] 1)); |
|
790 by (safe_tac set_cs); |
|
791 by (stac iterates 1); |
|
792 by (stac prem 1); |
|
793 by (stac fun_power_lmap 1); |
|
794 by (stac fun_power_lmap 1); |
|
795 by (rtac llistD_Fun_LCons_I 1); |
|
796 by (rtac (lmap_iterates RS subst) 1); |
|
797 by (stac fun_power_Suc 1); |
|
798 by (stac fun_power_Suc 1); |
|
799 by (rtac (UN1_I RS UnI1) 1); |
|
800 by (rtac rangeI 1); |
|
801 qed "iterates_equality"; |
|
802 |
|
803 |
|
804 (*** lappend -- its two arguments cause some complications! ***) |
|
805 |
|
806 goalw LList.thy [lappend_def] "lappend(LNil,LNil) = LNil"; |
|
807 by (rtac (llist_corec RS trans) 1); |
|
808 by (simp_tac llistD_ss 1); |
|
809 qed "lappend_LNil_LNil"; |
|
810 |
|
811 goalw LList.thy [lappend_def] |
|
812 "lappend(LNil,LCons(l,l')) = LCons(l, lappend(LNil,l'))"; |
|
813 by (rtac (llist_corec RS trans) 1); |
|
814 by (simp_tac llistD_ss 1); |
|
815 qed "lappend_LNil_LCons"; |
|
816 |
|
817 goalw LList.thy [lappend_def] |
|
818 "lappend(LCons(l,l'), N) = LCons(l, lappend(l',N))"; |
|
819 by (rtac (llist_corec RS trans) 1); |
|
820 by (simp_tac llistD_ss 1); |
|
821 qed "lappend_LCons"; |
|
822 |
|
823 goal LList.thy "lappend(LNil,l) = l"; |
|
824 by (res_inst_tac [("l","l")] llist_fun_equalityI 1); |
|
825 by (ALLGOALS |
|
826 (simp_tac (llistD_ss addsimps [lappend_LNil_LNil, lappend_LNil_LCons]))); |
|
827 qed "lappend_LNil"; |
|
828 |
|
829 goal LList.thy "lappend(l,LNil) = l"; |
|
830 by (res_inst_tac [("l","l")] llist_fun_equalityI 1); |
|
831 by (ALLGOALS |
|
832 (simp_tac (llistD_ss addsimps [lappend_LNil_LNil, lappend_LCons]))); |
|
833 qed "lappend_LNil2"; |
|
834 |
|
835 (*The infinite first argument blocks the second*) |
|
836 goal LList.thy "lappend(iterates(f,x), N) = iterates(f,x)"; |
|
837 by (res_inst_tac [("r", "range(%u.<lappend(iterates(f,u),N),iterates(f,u)>)")] |
|
838 llist_equalityI 1); |
|
839 by (rtac rangeI 1); |
|
840 by (safe_tac set_cs); |
|
841 by (stac iterates 1); |
|
842 by (simp_tac (llistD_ss addsimps [lappend_LCons]) 1); |
|
843 qed "lappend_iterates"; |
|
844 |
|
845 (** Two proofs that lmap distributes over lappend **) |
|
846 |
|
847 (*Long proof requiring case analysis on both both arguments*) |
|
848 goal LList.thy "lmap(f, lappend(l,n)) = lappend(lmap(f,l), lmap(f,n))"; |
|
849 by (res_inst_tac |
|
850 [("r", |
|
851 "UN n. range(%l.<lmap(f,lappend(l,n)), lappend(lmap(f,l),lmap(f,n))>)")] |
|
852 llist_equalityI 1); |
|
853 by (rtac UN1_I 1); |
|
854 by (rtac rangeI 1); |
|
855 by (safe_tac set_cs); |
|
856 by (res_inst_tac [("l", "l")] llistE 1); |
|
857 by (res_inst_tac [("l", "n")] llistE 1); |
|
858 by (ALLGOALS (asm_simp_tac (llistD_ss addsimps |
|
859 [lappend_LNil_LNil,lappend_LCons,lappend_LNil_LCons, |
|
860 lmap_LNil,lmap_LCons]))); |
|
861 by (REPEAT_SOME (ares_tac [llistD_Fun_LCons_I, UN1_I RS UnI1, rangeI])); |
|
862 by (rtac range_eqI 1); |
|
863 by (rtac (refl RS Pair_cong) 1); |
|
864 by (stac lmap_LNil 1); |
|
865 by (rtac refl 1); |
|
866 qed "lmap_lappend_distrib"; |
|
867 |
|
868 (*Shorter proof of theorem above using llist_equalityI as strong coinduction*) |
|
869 goal LList.thy "lmap(f, lappend(l,n)) = lappend(lmap(f,l), lmap(f,n))"; |
|
870 by (res_inst_tac [("l","l")] llist_fun_equalityI 1); |
|
871 by (simp_tac (llistD_ss addsimps [lappend_LNil, lmap_LNil])1); |
|
872 by (simp_tac (llistD_ss addsimps [lappend_LCons, lmap_LCons]) 1); |
|
873 qed "lmap_lappend_distrib"; |
|
874 |
|
875 (*Without strong coinduction, three case analyses might be needed*) |
|
876 goal LList.thy "lappend(lappend(l1,l2) ,l3) = lappend(l1, lappend(l2,l3))"; |
|
877 by (res_inst_tac [("l","l1")] llist_fun_equalityI 1); |
|
878 by (simp_tac (llistD_ss addsimps [lappend_LNil])1); |
|
879 by (simp_tac (llistD_ss addsimps [lappend_LCons]) 1); |
|
880 qed "lappend_assoc"; |
|