LList.ML
changeset 2 befa4e9f7c90
parent 0 7949f97df77a
child 22 17b6487e1ac7
equal deleted inserted replaced
1:142f1eb707b4 2:befa4e9f7c90
     3     Author: 	Lawrence C Paulson, Cambridge University Computer Laboratory
     3     Author: 	Lawrence C Paulson, Cambridge University Computer Laboratory
     4     Copyright   1993  University of Cambridge
     4     Copyright   1993  University of Cambridge
     5 
     5 
     6 For llist.thy.  
     6 For llist.thy.  
     7 
     7 
     8 SHOULD LListD_Fun_CONS_I, etc., be equations (for rewriting)
     8 SHOULD LListD_Fun_CONS_I, etc., be equations (for rewriting)?
       
     9 
       
    10 TOO LONG!  needs splitting up
     9 *)
    11 *)
    10 
    12 
    11 open LList;
    13 open LList;
    12 
    14 
    13 (** Simplification **)
    15 (** Simplification **)
    14 
    16 
    15 val llist_simps = [case_Inl, case_Inr];
    17 val llist_simps = [case_Inl, case_Inr];
    16 val llist_ss = univ_ss addsimps llist_simps
    18 val llist_ss = univ_ss addsimps llist_simps
    17                        setloop (split_tac [expand_split,expand_case]);
    19                        addcongs [split_weak_cong, case_weak_cong]
       
    20                        setloop (split_tac [expand_split, expand_case]);
    18 
    21 
    19 (** the llist functional **)
    22 (** the llist functional **)
    20 
    23 
    21 val LList_unfold = rewrite_rule [List_Fun_def]
    24 val LList_unfold = rewrite_rule [List_Fun_def]
    22 	 (List_Fun_mono RS (LList_def RS def_gfp_Tarski));
    25 	 (List_Fun_mono RS (LList_def RS def_gfp_Tarski));
    43 
    46 
    44 
    47 
    45 (*** Type checking by co-induction, using List_Fun ***)
    48 (*** Type checking by co-induction, using List_Fun ***)
    46 
    49 
    47 val prems = goalw LList.thy [LList_def]
    50 val prems = goalw LList.thy [LList_def]
    48     "[| M: X;  X <= List_Fun(A,X) |] ==> M: LList(A)";
    51     "[| M : X;  X <= List_Fun(A, X Un LList(A)) |] ==>  M : LList(A)";
    49 by (REPEAT (resolve_tac (prems@[coinduct]) 1));
    52 by (REPEAT (resolve_tac (prems@[List_Fun_mono RS coinduct]) 1));
    50 val LList_coinduct = result();
    53 val LList_coinduct = result();
    51 
       
    52 (*stronger version*)
       
    53 val prems = goalw LList.thy [LList_def]
       
    54     "[| M : X;  X <= List_Fun(A, X) Un LList(A) |] ==>  M : LList(A)";
       
    55 by (REPEAT (resolve_tac (prems@[coinduct2,List_Fun_mono]) 1));
       
    56 val LList_coinduct2 = result();
       
    57 
    54 
    58 (** Rules to prove the 2nd premise of LList_coinduct **)
    55 (** Rules to prove the 2nd premise of LList_coinduct **)
    59 
    56 
    60 goalw LList.thy [List_Fun_def,NIL_def] "NIL: List_Fun(A,X)";
    57 goalw LList.thy [List_Fun_def,NIL_def] "NIL: List_Fun(A,X)";
    61 by (resolve_tac [singletonI RS usum_In0I] 1);
    58 by (resolve_tac [singletonI RS usum_In0I] 1);
    64 goalw LList.thy [List_Fun_def,CONS_def]
    61 goalw LList.thy [List_Fun_def,CONS_def]
    65     "!!M N. [| M: A;  N: X |] ==> CONS(M,N) : List_Fun(A,X)";
    62     "!!M N. [| M: A;  N: X |] ==> CONS(M,N) : List_Fun(A,X)";
    66 by (REPEAT (ares_tac [uprodI RS usum_In1I] 1));
    63 by (REPEAT (ares_tac [uprodI RS usum_In1I] 1));
    67 val List_Fun_CONS_I = result();
    64 val List_Fun_CONS_I = result();
    68 
    65 
       
    66 (*Utilise the "strong" part, i.e. gfp(f)*)
       
    67 goalw LList.thy [LList_def]
       
    68     "!!M N. M: LList(A) ==> M : List_Fun(A, X Un LList(A))";
       
    69 by (etac (List_Fun_mono RS gfp_fun_UnI2) 1);
       
    70 val List_Fun_LList_I = result();
       
    71 
    69 (*** LList_corec satisfies the desired recurion equation ***)
    72 (*** LList_corec satisfies the desired recurion equation ***)
    70 
    73 
    71 (*A continuity result?*)
    74 (*A continuity result?*)
    72 goalw LList.thy [CONS_def] "CONS(M, UN x.f(x)) = (UN x. CONS(M, f(x)))";
    75 goalw LList.thy [CONS_def] "CONS(M, UN x.f(x)) = (UN x. CONS(M, f(x)))";
    73 by(simp_tac (univ_ss addsimps [In1_UN1, Scons_UN1_y]) 1);
    76 by (simp_tac (univ_ss addsimps [In1_UN1, Scons_UN1_y]) 1);
    74 val CONS_UN1 = result();
    77 val CONS_UN1 = result();
    75 
    78 
    76 goal Prod.thy "split(p, %x y.UN z.f(x,y,z)) = (UN z. split(p, %x y.f(x,y,z)))";
    79 goal Prod.thy "split(p, %x y.UN z.f(x,y,z)) = (UN z. split(p, %x y.f(x,y,z)))";
    77 by(simp_tac (pair_ss setloop (split_tac [expand_split])) 1);
    80 by (simp_tac (pair_ss setloop (split_tac [expand_split])) 1);
    78 val split_UN1 = result();
    81 val split_UN1 = result();
    79 
    82 
    80 goal Sum.thy "case(s, f, %y. UN z.g(y,z)) = (UN z. case(s, f, %y. g(y,z)))";
    83 goal Sum.thy "case(s, f, %y. UN z.g(y,z)) = (UN z. case(s, f, %y. g(y,z)))";
    81 by(simp_tac (sum_ss setloop (split_tac [expand_case])) 1);
    84 by (simp_tac (sum_ss setloop (split_tac [expand_case])) 1);
    82 val case2_UN1 = result();
    85 val case2_UN1 = result();
    83 
    86 
    84 val prems = goalw LList.thy [CONS_def]
    87 val prems = goalw LList.thy [CONS_def]
    85     "[| M<=M';  N<=N' |] ==> CONS(M,N) <= CONS(M',N')";
    88     "[| M<=M';  N<=N' |] ==> CONS(M,N) <= CONS(M',N')";
    86 by (REPEAT (resolve_tac ([In1_mono,Scons_mono]@prems) 1));
    89 by (REPEAT (resolve_tac ([In1_mono,Scons_mono]@prems) 1));
    95 goalw LList.thy [LList_corec_def]
    98 goalw LList.thy [LList_corec_def]
    96     "LList_corec(a,f) <= case(f(a), %u.NIL, \
    99     "LList_corec(a,f) <= case(f(a), %u.NIL, \
    97 \			   %v. split(v, %z w. CONS(z, LList_corec(w,f))))";
   100 \			   %v. split(v, %z w. CONS(z, LList_corec(w,f))))";
    98 by (rtac UN1_least 1);
   101 by (rtac UN1_least 1);
    99 by (nat_ind_tac "k" 1);
   102 by (nat_ind_tac "k" 1);
   100 by(ALLGOALS(simp_tac corec_fun_ss));
   103 by (ALLGOALS(simp_tac corec_fun_ss));
   101 by (REPEAT (resolve_tac [allI, impI, subset_refl RS CONS_mono, UN1_upper] 1));
   104 by (REPEAT (resolve_tac [allI, impI, subset_refl RS CONS_mono, UN1_upper] 1));
   102 val LList_corec_subset1 = result();
   105 val LList_corec_subset1 = result();
   103 
   106 
   104 goalw LList.thy [LList_corec_def]
   107 goalw LList.thy [LList_corec_def]
   105     "case(f(a), %u.NIL, %v. split(v, %z w. CONS(z, LList_corec(w,f)))) <= \
   108     "case(f(a), %u.NIL, %v. split(v, %z w. CONS(z, LList_corec(w,f)))) <= \
   131 goal LList.thy "LList_corec(a,f) : LList({u.True})";
   134 goal LList.thy "LList_corec(a,f) : LList({u.True})";
   132 by (res_inst_tac [("X", "range(%x.LList_corec(x,?g))")] LList_coinduct 1);
   135 by (res_inst_tac [("X", "range(%x.LList_corec(x,?g))")] LList_coinduct 1);
   133 by (rtac rangeI 1);
   136 by (rtac rangeI 1);
   134 by (safe_tac set_cs);
   137 by (safe_tac set_cs);
   135 by (stac LList_corec 1);
   138 by (stac LList_corec 1);
   136 by(simp_tac (llist_ss addsimps [List_Fun_NIL_I,List_Fun_CONS_I,
   139 by (simp_tac (llist_ss addsimps [List_Fun_NIL_I,List_Fun_CONS_I,
   137                                 CollectI, range_eqI]) 1);
   140                                  CollectI, range_eqI]) 1);
   138 (* 6.7 vs 3.4 !!!
   141 (* 6.7 vs 3.4 !!!
   139 by (ASM_SIMP_TAC (llist_ss addsimps [List_Fun_NIL_I,List_Fun_CONS_I,
   142 by (ASM_SIMP_TAC (llist_ss addsimps [List_Fun_NIL_I,List_Fun_CONS_I,
   140 				    CollectI, rangeI]) 1);
   143 				    CollectI, rangeI]) 1);
   141 *)
   144 *)
   142 val LList_corec_type = result();
   145 val LList_corec_type = result();
   149 by (rtac rangeI 1);
   152 by (rtac rangeI 1);
   150 by (safe_tac set_cs);
   153 by (safe_tac set_cs);
   151 by (stac LList_corec 1);
   154 by (stac LList_corec 1);
   152 (*nested "case"; requires an explicit split*)
   155 (*nested "case"; requires an explicit split*)
   153 by (res_inst_tac [("s", "f(xa)")] sumE 1);
   156 by (res_inst_tac [("s", "f(xa)")] sumE 1);
   154 by(asm_simp_tac (univ_ss addsimps (llist_simps@[List_Fun_NIL_I])) 1);
   157 by (asm_simp_tac (univ_ss addsimps (llist_simps@[List_Fun_NIL_I])) 1);
   155 by(asm_simp_tac (univ_ss addsimps (llist_simps@[List_Fun_CONS_I, range_eqI])
   158 by (asm_simp_tac (univ_ss addsimps (llist_simps@[List_Fun_CONS_I, range_eqI])
   156                      setloop (split_tac [expand_split])) 1);
   159                      setloop (split_tac [expand_split])) 1);
   157 (* FIXME: can the selection of the case split be automated?
   160 (* FIXME: can the selection of the case split be automated?
   158 by (ASM_SIMP_TAC (llist_ss addsimps [List_Fun_CONS_I, rangeI]) 1);*)
   161 by (ASM_SIMP_TAC (llist_ss addsimps [List_Fun_CONS_I, rangeI]) 1);*)
   159 val LList_corec_type2 = result();
   162 val LList_corec_type2 = result();
   160 
   163 
   161 (**** LList equality as a gfp; the bisimulation principle ****)
   164 (**** LList equality as a gfp; the bisimulation principle ****)
   162 
   165 
   163 goalw LList.thy [LListD_Fun_def] "mono(LListD_Fun(r))";
   166 goalw LList.thy [LListD_Fun_def] "mono(LListD_Fun(r))";
   164 by (REPEAT (ares_tac [monoI, subset_refl, dsum_mono, dprod_mono] 1));
   167 by (REPEAT (ares_tac [monoI, subset_refl, dsum_mono, dprod_mono] 1));
   165 val LListD_fun_mono = result();
   168 val LListD_Fun_mono = result();
   166 
   169 
   167 val LListD_unfold = rewrite_rule [LListD_Fun_def]
   170 val LListD_unfold = rewrite_rule [LListD_Fun_def]
   168 	 (LListD_fun_mono RS (LListD_def RS def_gfp_Tarski));
   171 	 (LListD_Fun_mono RS (LListD_def RS def_gfp_Tarski));
   169 
   172 
   170 goal LList.thy "!M N. <M,N> : LListD(diag(A)) --> ntrunc(k,M) = ntrunc(k,N)";
   173 goal LList.thy "!M N. <M,N> : LListD(diag(A)) --> ntrunc(k,M) = ntrunc(k,N)";
   171 by (res_inst_tac [("n", "k")] less_induct 1);
   174 by (res_inst_tac [("n", "k")] less_induct 1);
   172 by (safe_tac set_cs);
   175 by (safe_tac set_cs);
   173 by (etac (LListD_unfold RS equalityD1 RS subsetD RS dsumE) 1);
   176 by (etac (LListD_unfold RS equalityD1 RS subsetD RS dsumE) 1);
   174 by (safe_tac (set_cs addSEs [Pair_inject, dprodE, diagE]));
   177 by (safe_tac (set_cs addSEs [Pair_inject, dprodE, diagE]));
   175 by (res_inst_tac [("n", "n")] natE 1);
   178 by (res_inst_tac [("n", "n")] natE 1);
   176 by(asm_simp_tac (univ_ss addsimps [ntrunc_0]) 1);
   179 by (asm_simp_tac (univ_ss addsimps [ntrunc_0]) 1);
   177 by (res_inst_tac [("n", "xb")] natE 1);
   180 by (res_inst_tac [("n", "xb")] natE 1);
   178 by(asm_simp_tac (univ_ss addsimps [ntrunc_one_In1]) 1);
   181 by (asm_simp_tac (univ_ss addsimps [ntrunc_one_In1]) 1);
   179 by(asm_simp_tac (univ_ss addsimps [ntrunc_In1, ntrunc_Scons]) 1);
   182 by (asm_simp_tac (univ_ss addsimps [ntrunc_In1, ntrunc_Scons]) 1);
   180 val LListD_implies_ntrunc_equality = result();
   183 val LListD_implies_ntrunc_equality = result();
   181 
   184 
   182 goalw LList.thy [LList_def,List_Fun_def] "fst``LListD(diag(A)) <= LList(A)";
   185 goalw LList.thy [LList_def,List_Fun_def] "fst``LListD(diag(A)) <= LList(A)";
   183 by (rtac gfp_upperbound 1);
   186 by (rtac gfp_upperbound 1);
   184 by (res_inst_tac [("P", "%x. fst``x <= ?B")] (LListD_unfold RS ssubst) 1);
   187 by (res_inst_tac [("P", "%x. fst``x <= ?B")] (LListD_unfold RS ssubst) 1);
   185 by(simp_tac fst_image_ss 1);
   188 by (simp_tac fst_image_ss 1);
   186 val fst_image_LListD = result();
   189 val fst_image_LListD = result();
   187 
   190 
   188 (*This inclusion justifies the use of coinduction to show M=N*)
   191 (*This inclusion justifies the use of coinduction to show M=N*)
   189 goal LList.thy "LListD(diag(A)) <= diag(LList(A))";
   192 goal LList.thy "LListD(diag(A)) <= diag(LList(A))";
   190 by (rtac subsetI 1);
   193 by (rtac subsetI 1);
   213 goal LList.thy "LListD(diag(A)) = diag(LList(A))";
   216 goal LList.thy "LListD(diag(A)) = diag(LList(A))";
   214 by (REPEAT (resolve_tac [equalityI, LListD_subset_diag, 
   217 by (REPEAT (resolve_tac [equalityI, LListD_subset_diag, 
   215 			 diag_subset_LListD] 1));
   218 			 diag_subset_LListD] 1));
   216 val LListD_eq_diag = result();
   219 val LListD_eq_diag = result();
   217 
   220 
   218 (** To show two LLists are equal, exhibit a bisimulation! **)
   221 (** To show two LLists are equal, exhibit a bisimulation! 
   219 (* Replace "A" by some particular set, like {x.True}??? *)
   222       [also admits true equality]
   220 val prems = goal LList.thy 
   223    Replace "A" by some particular set, like {x.True}??? *)
   221     "[| <M,N> : r;  r <= LListD_Fun(diag(A), r) |] ==>  M=N";
   224 goal LList.thy 
       
   225     "!!r. [| <M,N> : r;  r <= LListD_Fun(diag(A), r Un diag(LList(A))) \
       
   226 \         |] ==>  M=N";
   222 by (rtac (rewrite_rule [LListD_def]
   227 by (rtac (rewrite_rule [LListD_def]
   223            (LListD_subset_diag RS subsetD RS diagE)) 1);
   228            (LListD_subset_diag RS subsetD RS diagE)) 1);
   224 by (REPEAT (resolve_tac (prems@[coinduct]) 1));
   229 by (etac (LListD_Fun_mono RS coinduct) 1);
       
   230 by (etac (rewrite_rule [LListD_def] LListD_eq_diag RS ssubst) 1);
   225 by (safe_tac (set_cs addSEs [Pair_inject]));
   231 by (safe_tac (set_cs addSEs [Pair_inject]));
   226 val LList_equalityI = result();
   232 val LList_equalityI = result();
   227 
       
   228 (*Stronger notion of bisimulation -- also admits true equality*)
       
   229 val prems = goal LList.thy 
       
   230     "[| <M,N> : r;  r <= LListD_Fun(diag(A), r) Un diag(LList(A)) |] ==>  M=N";
       
   231 by (rtac (rewrite_rule [LListD_def]
       
   232            (LListD_subset_diag RS subsetD RS diagE)) 1);
       
   233 by (rtac coinduct2 1);
       
   234 by (stac (rewrite_rule [LListD_def] LListD_eq_diag) 2);
       
   235 by (REPEAT (resolve_tac (prems@[LListD_fun_mono]) 1));
       
   236 by (safe_tac (set_cs addSEs [Pair_inject]));
       
   237 val LList_equalityI2 = result();
       
   238 
   233 
   239 (** Rules to prove the 2nd premise of LList_equalityI **)
   234 (** Rules to prove the 2nd premise of LList_equalityI **)
   240 
   235 
   241 goalw LList.thy [LListD_Fun_def,NIL_def] "<NIL,NIL> : LListD_Fun(r,s)";
   236 goalw LList.thy [LListD_Fun_def,NIL_def] "<NIL,NIL> : LListD_Fun(r,s)";
   242 by (rtac (singletonI RS diagI RS dsum_In0I) 1);
   237 by (rtac (singletonI RS diagI RS dsum_In0I) 1);
   245 val prems = goalw LList.thy [LListD_Fun_def,CONS_def]
   240 val prems = goalw LList.thy [LListD_Fun_def,CONS_def]
   246     "[| x:A;  <M,N>:s |] ==> <CONS(x,M), CONS(x,N)> : LListD_Fun(diag(A),s)";
   241     "[| x:A;  <M,N>:s |] ==> <CONS(x,M), CONS(x,N)> : LListD_Fun(diag(A),s)";
   247 by (rtac (dprodI RS dsum_In1I) 1);
   242 by (rtac (dprodI RS dsum_In1I) 1);
   248 by (REPEAT (resolve_tac (diagI::prems) 1));
   243 by (REPEAT (resolve_tac (diagI::prems) 1));
   249 val LListD_Fun_CONS_I = result();
   244 val LListD_Fun_CONS_I = result();
       
   245 
       
   246 (*Utilise the "strong" part, i.e. gfp(f)*)
       
   247 goal LList.thy 
       
   248     "!!M N. M: LList(A) ==> <M,M> : LListD_Fun(diag(A), X Un diag(LList(A)))";
       
   249 br (rewrite_rule [LListD_def] LListD_eq_diag RS subst) 1;
       
   250 br (LListD_Fun_mono RS gfp_fun_UnI2) 1;
       
   251 br (rewrite_rule [LListD_def] LListD_eq_diag RS ssubst) 1;
       
   252 be diagI 1;
       
   253 val LListD_Fun_diag_I = result();
   250 
   254 
   251 
   255 
   252 (*** Finality of LList(A): Uniqueness of functions defined by corecursion ***)
   256 (*** Finality of LList(A): Uniqueness of functions defined by corecursion ***)
   253 
   257 
   254 (*abstract proof using a bisimulation*)
   258 (*abstract proof using a bisimulation*)
   262 		  ("r", "range(%u. <h1(u),h2(u)>)")] LList_equalityI 1);
   266 		  ("r", "range(%u. <h1(u),h2(u)>)")] LList_equalityI 1);
   263 by (rtac rangeI 1);
   267 by (rtac rangeI 1);
   264 by (safe_tac set_cs);
   268 by (safe_tac set_cs);
   265 by (stac prem1 1);
   269 by (stac prem1 1);
   266 by (stac prem2 1);
   270 by (stac prem2 1);
   267 by(simp_tac (llist_ss addsimps [LListD_Fun_NIL_I, range_eqI,
   271 by (simp_tac (llist_ss addsimps [LListD_Fun_NIL_I, range_eqI,
   268 				    CollectI RS LListD_Fun_CONS_I]) 1);
   272 				 CollectI RS LListD_Fun_CONS_I]) 1);
   269 (* 9.5 vs 9.2/4.1/4.3
   273 (* 9.5 vs 9.2/4.1/4.3
   270 by (ASM_SIMP_TAC (llist_ss addsimps [LListD_Fun_NIL_I, rangeI,
   274 by (ASM_SIMP_TAC (llist_ss addsimps [LListD_Fun_NIL_I, rangeI,
   271 				    CollectI RS LListD_Fun_CONS_I]) 1);*)
   275 				    CollectI RS LListD_Fun_CONS_I]) 1);*)
   272 val LList_corec_unique = result();
   276 val LList_corec_unique = result();
   273 
   277 
   284 by (rtac ntrunc_one_In1 1);
   288 by (rtac ntrunc_one_In1 1);
   285 val ntrunc_one_CONS = result();
   289 val ntrunc_one_CONS = result();
   286 
   290 
   287 goalw LList.thy [CONS_def]
   291 goalw LList.thy [CONS_def]
   288     "ntrunc(Suc(Suc(k)), CONS(M,N)) = CONS (ntrunc(k,M), ntrunc(k,N))";
   292     "ntrunc(Suc(Suc(k)), CONS(M,N)) = CONS (ntrunc(k,M), ntrunc(k,N))";
   289 by(simp_tac (HOL_ss addsimps [ntrunc_Scons,ntrunc_In1]) 1);
   293 by (simp_tac (HOL_ss addsimps [ntrunc_Scons,ntrunc_In1]) 1);
   290 val ntrunc_CONS = result();
   294 val ntrunc_CONS = result();
   291 
   295 
   292 val [prem1,prem2] = goal LList.thy
   296 val [prem1,prem2] = goal LList.thy
   293  "[| !!x. h1(x) = case(f(x), %u.NIL, %v. split(v, %z w. CONS(z,h1(w))));  \
   297  "[| !!x. h1(x) = case(f(x), %u.NIL, %v. split(v, %z w. CONS(z,h1(w))));  \
   294 \    !!x. h2(x) = case(f(x), %u.NIL, %v. split(v, %z w. CONS(z,h2(w)))) |] \
   298 \    !!x. h2(x) = case(f(x), %u.NIL, %v. split(v, %z w. CONS(z,h2(w)))) |] \
   297 by (res_inst_tac [("x", "x")] spec 1);
   301 by (res_inst_tac [("x", "x")] spec 1);
   298 by (res_inst_tac [("n", "k")] less_induct 1);
   302 by (res_inst_tac [("n", "k")] less_induct 1);
   299 by (rtac allI 1);
   303 by (rtac allI 1);
   300 by (stac prem1 1);
   304 by (stac prem1 1);
   301 by (stac prem2 1);
   305 by (stac prem2 1);
   302 by(simp_tac (sum_ss setloop (split_tac [expand_split,expand_case])) 1);
   306 by (simp_tac (sum_ss setloop (split_tac [expand_split,expand_case])) 1);
   303 by (strip_tac 1);
   307 by (strip_tac 1);
   304 by (res_inst_tac [("n", "n")] natE 1);
   308 by (res_inst_tac [("n", "n")] natE 1);
   305 by (res_inst_tac [("n", "xc")] natE 2);
   309 by (res_inst_tac [("n", "xc")] natE 2);
   306 by(ALLGOALS(asm_simp_tac(nat_ss addsimps
   310 by (ALLGOALS(asm_simp_tac(nat_ss addsimps
   307             [ntrunc_0,ntrunc_one_CONS,ntrunc_CONS])));
   311             [ntrunc_0,ntrunc_one_CONS,ntrunc_CONS])));
   308 val LList_corec_unique = result();
   312 val LList_corec_unique = result();
   309 
   313 
   310 
   314 
   311 (*** Lconst -- defined directly using lfp, but equivalent to a LList_corec ***)
   315 (*** Lconst -- defined directly using lfp, but equivalent to a LList_corec ***)
   321   The containing set is simply the singleton {Lconst(M)}. *)
   325   The containing set is simply the singleton {Lconst(M)}. *)
   322 goal LList.thy "!!M A. M:A ==> Lconst(M): LList(A)";
   326 goal LList.thy "!!M A. M:A ==> Lconst(M): LList(A)";
   323 by (rtac (singletonI RS LList_coinduct) 1);
   327 by (rtac (singletonI RS LList_coinduct) 1);
   324 by (safe_tac set_cs);
   328 by (safe_tac set_cs);
   325 by (res_inst_tac [("P", "%u. u: ?A")] (Lconst RS ssubst) 1);
   329 by (res_inst_tac [("P", "%u. u: ?A")] (Lconst RS ssubst) 1);
   326 by (REPEAT (ares_tac [List_Fun_CONS_I, singletonI] 1));
   330 by (REPEAT (ares_tac [List_Fun_CONS_I, singletonI, UnI1] 1));
   327 val Lconst_type = result();
   331 val Lconst_type = result();
   328 
   332 
   329 goal LList.thy "Lconst(M) = LList_corec(M, %x.Inr(<x,x>))";
   333 goal LList.thy "Lconst(M) = LList_corec(M, %x.Inr(<x,x>))";
   330 by (rtac (equals_LList_corec RS fun_cong) 1);
   334 by (rtac (equals_LList_corec RS fun_cong) 1);
   331 by(simp_tac sum_ss 1);
   335 by (simp_tac sum_ss 1);
   332 by (rtac Lconst 1);
   336 by (rtac Lconst 1);
   333 val Lconst_eq_LList_corec = result();
   337 val Lconst_eq_LList_corec = result();
   334 
   338 
   335 (*Thus we could have used gfp in the definition of Lconst*)
   339 (*Thus we could have used gfp in the definition of Lconst*)
   336 goal LList.thy "gfp(%N. CONS(M, N)) = LList_corec(M, %x.Inr(<x,x>))";
   340 goal LList.thy "gfp(%N. CONS(M, N)) = LList_corec(M, %x.Inr(<x,x>))";
   337 by (rtac (equals_LList_corec RS fun_cong) 1);
   341 by (rtac (equals_LList_corec RS fun_cong) 1);
   338 by(simp_tac sum_ss 1);
   342 by (simp_tac sum_ss 1);
   339 by (rtac (Lconst_fun_mono RS gfp_Tarski) 1);
   343 by (rtac (Lconst_fun_mono RS gfp_Tarski) 1);
   340 val gfp_Lconst_eq_LList_corec = result();
   344 val gfp_Lconst_eq_LList_corec = result();
   341 
   345 
   342 
   346 
   343 (** Introduction rules for LList constructors **)
   347 (** Introduction rules for LList constructors **)
   429 val [MList,gMList,NILcase,CONScase] = goal LList.thy
   433 val [MList,gMList,NILcase,CONScase] = goal LList.thy
   430  "[| M: LList(A); g(NIL): LList(A); 				\
   434  "[| M: LList(A); g(NIL): LList(A); 				\
   431 \    f(NIL)=g(NIL);						\
   435 \    f(NIL)=g(NIL);						\
   432 \    !!x l. [| x:A;  l: LList(A) |] ==>				\
   436 \    !!x l. [| x:A;  l: LList(A) |] ==>				\
   433 \	    <f(CONS(x,l)),g(CONS(x,l))> :			\
   437 \	    <f(CONS(x,l)),g(CONS(x,l))> :			\
   434 \               LListD_Fun(diag(A), (%u.<f(u),g(u)>)``LList(A)) Un  \
   438 \               LListD_Fun(diag(A), (%u.<f(u),g(u)>)``LList(A) Un  \
   435 \               diag(LList(A))	\
   439 \                                   diag(LList(A)))		\
   436 \ |] ==> f(M) = g(M)";
   440 \ |] ==> f(M) = g(M)";
   437 by (rtac LList_equalityI2 1);
   441 by (rtac LList_equalityI 1);
   438 br (MList RS imageI) 1;
   442 br (MList RS imageI) 1;
   439 by (rtac subsetI 1);
   443 by (rtac subsetI 1);
   440 by (etac imageE 1);
   444 by (etac imageE 1);
   441 by (etac ssubst 1);
   445 by (etac ssubst 1);
   442 by (etac LListE 1);
   446 by (etac LListE 1);
   443 by (etac ssubst 1);
   447 by (etac ssubst 1);
   444 by (stac NILcase 1);
   448 by (stac NILcase 1);
   445 br (gMList RS diagI RS UnI2) 1;
   449 br (gMList RS LListD_Fun_diag_I) 1;
   446 by (etac ssubst 1);
   450 by (etac ssubst 1);
   447 by (REPEAT (ares_tac [CONScase] 1));
   451 by (REPEAT (ares_tac [CONScase] 1));
   448 val LList_fun_equalityI = result();
   452 val LList_fun_equalityI = result();
   449 
   453 
   450 
   454 
   451 (*** The functional "Lmap" ***)
   455 (*** The functional "Lmap" ***)
   452 
   456 
   453 goal LList.thy "Lmap(f,NIL) = NIL";
   457 goal LList.thy "Lmap(f,NIL) = NIL";
   454 by (rtac (Lmap_def RS def_LList_corec RS trans) 1);
   458 by (rtac (Lmap_def RS def_LList_corec RS trans) 1);
   455 by(simp_tac List_case_ss 1);
   459 by (simp_tac List_case_ss 1);
   456 val Lmap_NIL = result();
   460 val Lmap_NIL = result();
   457 
   461 
   458 goal LList.thy "Lmap(f, CONS(M,N)) = CONS(f(M), Lmap(f,N))";
   462 goal LList.thy "Lmap(f, CONS(M,N)) = CONS(f(M), Lmap(f,N))";
   459 by (rtac (Lmap_def RS def_LList_corec RS trans) 1);
   463 by (rtac (Lmap_def RS def_LList_corec RS trans) 1);
   460 by(simp_tac List_case_ss 1);
   464 by (simp_tac List_case_ss 1);
   461 val Lmap_CONS = result();
   465 val Lmap_CONS = result();
   462 
   466 
   463 (*Another type-checking proof by coinduction*)
   467 (*Another type-checking proof by coinduction*)
   464 val [major,minor] = goal LList.thy
   468 val [major,minor] = goal LList.thy
   465     "[| M: LList(A);  !!x. x:A ==> f(x):B |] ==> Lmap(f,M): LList(B)";
   469     "[| M: LList(A);  !!x. x:A ==> f(x):B |] ==> Lmap(f,M): LList(B)";
   466 by (rtac (major RS imageI RS LList_coinduct) 1);
   470 by (rtac (major RS imageI RS LList_coinduct) 1);
   467 by (safe_tac set_cs);
   471 by (safe_tac set_cs);
   468 by (etac LListE 1);
   472 by (etac LListE 1);
   469 by(ALLGOALS (asm_simp_tac (HOL_ss addsimps [Lmap_NIL,Lmap_CONS])));
   473 by (ALLGOALS (asm_simp_tac (HOL_ss addsimps [Lmap_NIL,Lmap_CONS])));
   470 by (REPEAT (ares_tac [List_Fun_NIL_I, List_Fun_CONS_I, minor, imageI] 1));
   474 by (REPEAT (ares_tac [List_Fun_NIL_I, List_Fun_CONS_I, 
       
   475 		      minor, imageI, UnI1] 1));
   471 val Lmap_type = result();
   476 val Lmap_type = result();
   472 
   477 
   473 (*This type checking rule synthesises a sufficiently large set for f*)
   478 (*This type checking rule synthesises a sufficiently large set for f*)
   474 val [major] = goal LList.thy  "M: LList(A) ==> Lmap(f,M): LList(f``A)";
   479 val [major] = goal LList.thy  "M: LList(A) ==> Lmap(f,M): LList(f``A)";
   475 by (rtac (major RS Lmap_type) 1);
   480 by (rtac (major RS Lmap_type) 1);
   482     "M: LList(A) ==> Lmap(f o g, M) = Lmap(f, Lmap(g, M))";
   487     "M: LList(A) ==> Lmap(f o g, M) = Lmap(f, Lmap(g, M))";
   483 by (rtac (prem RS imageI RS LList_equalityI) 1);
   488 by (rtac (prem RS imageI RS LList_equalityI) 1);
   484 by (stac o_def 1);
   489 by (stac o_def 1);
   485 by (safe_tac set_cs);
   490 by (safe_tac set_cs);
   486 by (etac LListE 1);
   491 by (etac LListE 1);
   487 by(ALLGOALS (asm_simp_tac (HOL_ss addsimps [Lmap_NIL,Lmap_CONS])));
   492 by (ALLGOALS (asm_simp_tac (HOL_ss addsimps [Lmap_NIL,Lmap_CONS])));
   488 by (REPEAT (ares_tac [LListD_Fun_NIL_I, imageI,
   493 by (REPEAT (ares_tac [LListD_Fun_NIL_I, imageI, UnI1,
   489 			 rangeI RS LListD_Fun_CONS_I] 1));
   494 		      rangeI RS LListD_Fun_CONS_I] 1));
   490 val Lmap_compose = result();
   495 val Lmap_compose = result();
   491 
   496 
   492 val [prem] = goal LList.thy "M: LList(A) ==> Lmap(%x.x, M) = M";
   497 val [prem] = goal LList.thy "M: LList(A) ==> Lmap(%x.x, M) = M";
   493 by (rtac (prem RS imageI RS LList_equalityI) 1);
   498 by (rtac (prem RS imageI RS LList_equalityI) 1);
   494 by (safe_tac set_cs);
   499 by (safe_tac set_cs);
   495 by (etac LListE 1);
   500 by (etac LListE 1);
   496 by(ALLGOALS (asm_simp_tac (HOL_ss addsimps [Lmap_NIL,Lmap_CONS])));
   501 by (ALLGOALS (asm_simp_tac (HOL_ss addsimps [Lmap_NIL,Lmap_CONS])));
   497 by (REPEAT (ares_tac [LListD_Fun_NIL_I, imageI,
   502 by (REPEAT (ares_tac [LListD_Fun_NIL_I, imageI RS UnI1,
   498 			 rangeI RS LListD_Fun_CONS_I] 1));
   503 		      rangeI RS LListD_Fun_CONS_I] 1));
   499 val Lmap_ident = result();
   504 val Lmap_ident = result();
   500 
   505 
   501 
   506 
   502 (*** Lappend -- its two arguments cause some complications! ***)
   507 (*** Lappend -- its two arguments cause some complications! ***)
   503 
   508 
   504 goalw LList.thy [Lappend_def] "Lappend(NIL,NIL) = NIL";
   509 goalw LList.thy [Lappend_def] "Lappend(NIL,NIL) = NIL";
   505 by (rtac (LList_corec RS trans) 1);
   510 by (rtac (LList_corec RS trans) 1);
   506 (* takes 2.4(3.4 w NORM) vs 0.9 w/o NORM terms *)
   511 (* takes 2.4(3.4 w NORM) vs 0.9 w/o NORM terms *)
   507 by(simp_tac List_case_ss 1);
   512 by (simp_tac List_case_ss 1);
   508 (*by (SIMP_TAC List_case_ss 1);*)
   513 (*by (SIMP_TAC List_case_ss 1);*)
   509 val Lappend_NIL_NIL = result();
   514 val Lappend_NIL_NIL = result();
   510 
   515 
   511 goalw LList.thy [Lappend_def]
   516 goalw LList.thy [Lappend_def]
   512     "Lappend(NIL,CONS(N,N')) = CONS(N, Lappend(NIL,N'))";
   517     "Lappend(NIL,CONS(N,N')) = CONS(N, Lappend(NIL,N'))";
   513 by (rtac (LList_corec RS trans) 1);
   518 by (rtac (LList_corec RS trans) 1);
   514 (* takes 5(7 w NORM) vs 2.1 w/o NORM terms *)
   519 (* takes 5(7 w NORM) vs 2.1 w/o NORM terms *)
   515 by(simp_tac List_case_ss 1);
   520 by (simp_tac List_case_ss 1);
   516 (*by (SIMP_TAC List_case_ss 1);*)
   521 (*by (SIMP_TAC List_case_ss 1);*)
   517 val Lappend_NIL_CONS = result();
   522 val Lappend_NIL_CONS = result();
   518 
   523 
   519 goalw LList.thy [Lappend_def]
   524 goalw LList.thy [Lappend_def]
   520     "Lappend(CONS(M,M'), N) = CONS(M, Lappend(M',N))";
   525     "Lappend(CONS(M,M'), N) = CONS(M, Lappend(M',N))";
   521 by (rtac (LList_corec RS trans) 1);
   526 by (rtac (LList_corec RS trans) 1);
   522 (* takes 4.9(6.7) vs 2.2 w/o NORM terms *)
   527 (* takes 4.9(6.7) vs 2.2 w/o NORM terms *)
   523 by(simp_tac List_case_ss 1);
   528 by (simp_tac List_case_ss 1);
   524 (*by (SIMP_TAC List_case_ss 1);*)
   529 (*by (SIMP_TAC List_case_ss 1);*)
   525 val Lappend_CONS = result();
   530 val Lappend_CONS = result();
   526 
   531 
   527 val Lappend_ss = List_case_ss addsimps
   532 val Lappend_ss = List_case_ss addsimps
   528                  [NIL_LListI, Lappend_NIL_NIL, Lappend_NIL_CONS,
   533                  [NIL_LListI, Lappend_NIL_NIL, Lappend_NIL_CONS,
   561 val Lappend_type = result();
   566 val Lappend_type = result();
   562 
   567 
   563 (*strong co-induction: bisimulation and case analysis on one variable*)
   568 (*strong co-induction: bisimulation and case analysis on one variable*)
   564 goal LList.thy
   569 goal LList.thy
   565     "!!M N. [| M: LList(A); N: LList(A) |] ==> Lappend(M,N): LList(A)";
   570     "!!M N. [| M: LList(A); N: LList(A) |] ==> Lappend(M,N): LList(A)";
   566 by (res_inst_tac [("X", "(%u.Lappend(u,N))``LList(A)")] LList_coinduct2 1);
   571 by (res_inst_tac [("X", "(%u.Lappend(u,N))``LList(A)")] LList_coinduct 1);
   567 fe imageI;
   572 be imageI 1;
   568 br subsetI 1;
   573 br subsetI 1;
   569 be imageE 1;
   574 be imageE 1;
   570 by (eres_inst_tac [("L", "u")] LListE 1);
   575 by (eres_inst_tac [("L", "u")] LListE 1);
   571 by (asm_simp_tac (Lappend_ss addsimps [Lappend_NIL]) 1);
   576 by (asm_simp_tac (Lappend_ss addsimps [Lappend_NIL, List_Fun_LList_I]) 1);
   572 by (asm_simp_tac Lappend_ss 1);
   577 by (asm_simp_tac Lappend_ss 1);
   573 by (fast_tac (set_cs addSIs [List_Fun_CONS_I]) 1);
   578 by (fast_tac (set_cs addSIs [List_Fun_CONS_I]) 1);
   574 val Lappend_type = result();
   579 val Lappend_type = result();
   575 
   580 
   576 (**** Lazy lists as the type 'a llist -- strongly typed versions of above ****)
   581 (**** Lazy lists as the type 'a llist -- strongly typed versions of above ****)
   582 		 Abs_LList_inverse, Rep_LList_inverse, NIL_LListI, CONS_LListI,
   587 		 Abs_LList_inverse, Rep_LList_inverse, NIL_LListI, CONS_LListI,
   583 		 Rep_LList, rangeI, inj_Leaf, Inv_f_f];
   588 		 Rep_LList, rangeI, inj_Leaf, Inv_f_f];
   584 val Rep_LList_ss = llist_ss addsimps Rep_LList_simps;
   589 val Rep_LList_ss = llist_ss addsimps Rep_LList_simps;
   585 
   590 
   586 goalw LList.thy [llist_case_def,LNil_def]  "llist_case(LNil, c, d) = c";
   591 goalw LList.thy [llist_case_def,LNil_def]  "llist_case(LNil, c, d) = c";
   587 by(simp_tac Rep_LList_ss 1);
   592 by (simp_tac Rep_LList_ss 1);
   588 val llist_case_LNil = result();
   593 val llist_case_LNil = result();
   589 
   594 
   590 goalw LList.thy [llist_case_def,LCons_def]
   595 goalw LList.thy [llist_case_def,LCons_def]
   591     "llist_case(LCons(M,N), c, d) = d(M,N)";
   596     "llist_case(LCons(M,N), c, d) = d(M,N)";
   592 by(simp_tac Rep_LList_ss 1);
   597 by (simp_tac Rep_LList_ss 1);
   593 val llist_case_LCons = result();
   598 val llist_case_LCons = result();
   594 
   599 
   595 (*Elimination is case analysis, not induction.*)
   600 (*Elimination is case analysis, not induction.*)
   596 val [prem1,prem2] = goalw LList.thy [NIL_def,CONS_def]
   601 val [prem1,prem2] = goalw LList.thy [NIL_def,CONS_def]
   597     "[| l=LNil ==> P;  !!x l'. l=LCons(x,l') ==> P \
   602     "[| l=LNil ==> P;  !!x l'. l=LCons(x,l') ==> P \
   600 by (rtac (inj_Rep_LList RS injD RS prem1) 1);
   605 by (rtac (inj_Rep_LList RS injD RS prem1) 1);
   601 by (stac Rep_LList_LNil 1);
   606 by (stac Rep_LList_LNil 1);
   602 by (assume_tac 1);
   607 by (assume_tac 1);
   603 by (etac rangeE 1);
   608 by (etac rangeE 1);
   604 by (rtac (inj_Rep_LList RS injD RS prem2) 1);
   609 by (rtac (inj_Rep_LList RS injD RS prem2) 1);
   605 by(asm_simp_tac (HOL_ss addsimps [Rep_LList_LCons]) 1);
   610 by (asm_simp_tac (HOL_ss addsimps [Rep_LList_LCons]) 1);
   606 by (etac (Abs_LList_inverse RS ssubst) 1);
   611 by (etac (Abs_LList_inverse RS ssubst) 1);
   607 by (rtac refl 1);
   612 by (rtac refl 1);
   608 val llistE = result();
   613 val llistE = result();
   609 
   614 
   610 (** llist_corec: corecursion for 'a llist **)
   615 (** llist_corec: corecursion for 'a llist **)
   611 
   616 
   612 goalw LList.thy [llist_corec_def,LNil_def,LCons_def]
   617 goalw LList.thy [llist_corec_def,LNil_def,LCons_def]
   613     "llist_corec(a,f) = case(f(a), %u. LNil, \
   618     "llist_corec(a,f) = case(f(a), %u. LNil, \
   614 \			     %v. split(v, %z w. LCons(z, llist_corec(w,f))))";
   619 \			     %v. split(v, %z w. LCons(z, llist_corec(w,f))))";
   615 by (stac LList_corec 1);
   620 by (stac LList_corec 1);
   616 by(res_inst_tac [("s","f(a)")] sumE 1);
   621 by (res_inst_tac [("s","f(a)")] sumE 1);
   617 by(asm_simp_tac (llist_ss addsimps [LList_corec_type2,Abs_LList_inverse]) 1);
   622 by (asm_simp_tac (llist_ss addsimps [LList_corec_type2,Abs_LList_inverse]) 1);
   618 by(res_inst_tac [("p","y")] PairE 1);
   623 by (res_inst_tac [("p","y")] PairE 1);
   619 by(asm_simp_tac (llist_ss addsimps [LList_corec_type2,Abs_LList_inverse]) 1);
   624 by (asm_simp_tac (llist_ss addsimps [LList_corec_type2,Abs_LList_inverse]) 1);
   620 (*FIXME: correct case splits usd to be found automatically:
   625 (*FIXME: correct case splits usd to be found automatically:
   621 by(ASM_SIMP_TAC(llist_ss addsimps [LList_corec_type2,Abs_LList_inverse]) 1);*)
   626 by (ASM_SIMP_TAC(llist_ss addsimps [LList_corec_type2,Abs_LList_inverse]) 1);*)
   622 val llist_corec = result();
   627 val llist_corec = result();
   623 
   628 
   624 (*definitional version of same*)
   629 (*definitional version of same*)
   625 val [rew] = goal LList.thy
   630 val [rew] = goal LList.thy
   626     "[| !!x. h(x) == llist_corec(x,f) |] ==>	\
   631     "[| !!x. h(x) == llist_corec(x,f) |] ==>	\
   651     "r <= Sigma(LList(range(Leaf)), %x.LList(range(Leaf))) ==> \
   656     "r <= Sigma(LList(range(Leaf)), %x.LList(range(Leaf))) ==> \
   652 \    prod_fun(Rep_LList o Abs_LList, Rep_LList o Abs_LList) `` r <= r";
   657 \    prod_fun(Rep_LList o Abs_LList, Rep_LList o Abs_LList) `` r <= r";
   653 by (safe_tac (set_cs addSEs [prod_fun_imageE]));
   658 by (safe_tac (set_cs addSEs [prod_fun_imageE]));
   654 by (rtac (prem RS subsetD RS SigmaE2) 1);
   659 by (rtac (prem RS subsetD RS SigmaE2) 1);
   655 by (assume_tac 1);
   660 by (assume_tac 1);
   656 by(asm_simp_tac (HOL_ss addsimps [o_def,prod_fun,Abs_LList_inverse]) 1);
   661 by (asm_simp_tac (HOL_ss addsimps [o_def,prod_fun,Abs_LList_inverse]) 1);
   657 val prod_fun_lemma = result();
   662 val prod_fun_lemma = result();
   658 
   663 
   659 (** To show two llists are equal, exhibit a bisimulation! **)
   664 goal LList.thy
       
   665     "prod_fun(Rep_LList, Rep_LList) `` range(%x. <x, x>) = \
       
   666 \    diag(LList(range(Leaf)))";
       
   667 br equalityI 1;
       
   668 by (fast_tac (set_cs addIs  [diagI, Rep_LList]
       
   669 		     addSEs [prod_fun_imageE, Pair_inject]) 1);
       
   670 by (fast_tac (set_cs addIs  [prod_fun_imageI, rangeI]
       
   671 		     addSEs [diagE, Abs_LList_inverse RS subst]) 1);
       
   672 val prod_fun_range_eq_diag = result();
       
   673 
       
   674 (** To show two llists are equal, exhibit a bisimulation! 
       
   675       [also admits true equality] **)
   660 val [prem1,prem2] = goalw LList.thy [llistD_Fun_def]
   676 val [prem1,prem2] = goalw LList.thy [llistD_Fun_def]
   661     "[| <l1,l2> : r;  r <= llistD_Fun(r) |] ==>  l1=l2";
   677     "[| <l1,l2> : r;  r <= llistD_Fun(r Un range(%x.<x,x>)) |] ==> l1=l2";
   662 by (rtac (inj_Rep_LList RS injD) 1);
   678 by (rtac (inj_Rep_LList RS injD) 1);
   663 by (res_inst_tac [("r", "prod_fun(Rep_LList,Rep_LList)``r")] 
   679 by (res_inst_tac [("r", "prod_fun(Rep_LList,Rep_LList)``r"),
       
   680 		  ("A", "range(Leaf)")] 
   664 	LList_equalityI 1);
   681 	LList_equalityI 1);
   665 by (rtac (prem1 RS prod_fun_imageI) 1);
   682 by (rtac (prem1 RS prod_fun_imageI) 1);
   666 by (rtac (prem2 RS image_mono RS subset_trans) 1);
   683 by (rtac (prem2 RS image_mono RS subset_trans) 1);
   667 by (rtac (image_compose RS subst) 1);
   684 by (rtac (image_compose RS subst) 1);
   668 by (rtac (prod_fun_compose RS subst) 1);
   685 by (rtac (prod_fun_compose RS subst) 1);
   669 by (rtac (subset_Sigma_LList RS LListD_Fun_subset_Sigma_LList RS 
   686 by (rtac (image_Un RS ssubst) 1);
   670 	  prod_fun_lemma) 1);
   687 by (stac prod_fun_range_eq_diag 1);
       
   688 by (rtac (LListD_Fun_subset_Sigma_LList RS prod_fun_lemma) 1);
       
   689 by (rtac (subset_Sigma_LList RS Un_least) 1);
       
   690 by (rtac diag_subset_Sigma 1);
   671 val llist_equalityI = result();
   691 val llist_equalityI = result();
   672 
       
   673 
       
   674 (*Stronger notion of bisimulation -- also admits true equality*)
       
   675 val [prem1,prem2] = goalw LList.thy [llistD_Fun_def]
       
   676     "[| <l1,l2> : r;  r <= llistD_Fun(r) Un range(%x.<x,x>) |] ==> l1=l2";
       
   677 by (rtac (inj_Rep_LList RS injD) 1);
       
   678 by (res_inst_tac [("r", "prod_fun(Rep_LList,Rep_LList)``r"),
       
   679 		  ("A", "range(Leaf)")] 
       
   680 	LList_equalityI2 1);
       
   681 by (rtac (prem1 RS prod_fun_imageI) 1);
       
   682 by (rtac (prem2 RS image_mono RS subset_trans) 1);
       
   683 by (rtac (image_Un RS ssubst) 1);
       
   684 by (rtac Un_least 1);
       
   685 by (rtac (image_compose RS subst) 1);
       
   686 by (rtac (prod_fun_compose RS subst) 1);
       
   687 by (rtac (subset_Sigma_LList RS LListD_Fun_subset_Sigma_LList RS 
       
   688 	  prod_fun_lemma RS subset_trans) 1);
       
   689 by (rtac Un_upper1 1);
       
   690 by (fast_tac (set_cs addSEs [prod_fun_imageE, Pair_inject] 
       
   691                      addIs [diagI,Rep_LList]) 1);
       
   692 val llist_equalityI2 = result();
       
   693 
   692 
   694 (** Rules to prove the 2nd premise of llist_equalityI **)
   693 (** Rules to prove the 2nd premise of llist_equalityI **)
   695 goalw LList.thy [llistD_Fun_def,LNil_def] "<LNil,LNil> : llistD_Fun(r)";
   694 goalw LList.thy [llistD_Fun_def,LNil_def] "<LNil,LNil> : llistD_Fun(r)";
   696 by (rtac (LListD_Fun_NIL_I RS prod_fun_imageI) 1);
   695 by (rtac (LListD_Fun_NIL_I RS prod_fun_imageI) 1);
   697 val llistD_Fun_LNil_I = result();
   696 val llistD_Fun_LNil_I = result();
   700     "<l1,l2>:r ==> <LCons(x,l1), LCons(x,l2)> : llistD_Fun(r)";
   699     "<l1,l2>:r ==> <LCons(x,l1), LCons(x,l2)> : llistD_Fun(r)";
   701 by (rtac (rangeI RS LListD_Fun_CONS_I RS prod_fun_imageI) 1);
   700 by (rtac (rangeI RS LListD_Fun_CONS_I RS prod_fun_imageI) 1);
   702 by (rtac (prem RS prod_fun_imageI) 1);
   701 by (rtac (prem RS prod_fun_imageI) 1);
   703 val llistD_Fun_LCons_I = result();
   702 val llistD_Fun_LCons_I = result();
   704 
   703 
       
   704 (*Utilise the "strong" part, i.e. gfp(f)*)
       
   705 goalw LList.thy [llistD_Fun_def]
       
   706      "!!l. <l,l> : llistD_Fun(r Un range(%x.<x,x>))";
       
   707 br (Rep_LList_inverse RS subst) 1;
       
   708 br prod_fun_imageI 1;
       
   709 by (rtac (image_Un RS ssubst) 1);
       
   710 by (stac prod_fun_range_eq_diag 1);
       
   711 br (Rep_LList RS LListD_Fun_diag_I) 1;
       
   712 val llistD_Fun_range_I = result();
   705 
   713 
   706 (*A special case of list_equality for functions over lazy lists*)
   714 (*A special case of list_equality for functions over lazy lists*)
   707 val [prem1,prem2] = goal LList.thy
   715 val [prem1,prem2] = goal LList.thy
   708     "[| f(LNil)=g(LNil);						\
   716     "[| f(LNil)=g(LNil);						\
   709 \       !!x l. <f(LCons(x,l)),g(LCons(x,l))> :				\
   717 \       !!x l. <f(LCons(x,l)),g(LCons(x,l))> :				\
   710 \              llistD_Fun(range(%u. <f(u),g(u)>)) Un range(%v. <v,v>)	\
   718 \              llistD_Fun(range(%u. <f(u),g(u)>) Un range(%v. <v,v>))	\
   711 \    |]	==> f(l) = g(l :: 'a llist) :: 'b llist";
   719 \    |]	==> f(l) = g(l :: 'a llist) :: 'b llist";
   712 by (res_inst_tac [("r", "range(%u. <f(u),g(u)>)")] llist_equalityI2 1);
   720 by (res_inst_tac [("r", "range(%u. <f(u),g(u)>)")] llist_equalityI 1);
   713 by (rtac rangeI 1);
   721 by (rtac rangeI 1);
   714 by (rtac subsetI 1);
   722 by (rtac subsetI 1);
   715 by (etac rangeE 1);
   723 by (etac rangeE 1);
   716 by (etac ssubst 1);
   724 by (etac ssubst 1);
   717 by (res_inst_tac [("l", "u")] llistE 1);
   725 by (res_inst_tac [("l", "u")] llistE 1);
   718 by (etac ssubst 1);
   726 by (etac ssubst 1);
   719 by (stac prem1 1);
   727 by (stac prem1 1);
   720 by (fast_tac set_cs 1);
   728 by (rtac llistD_Fun_range_I 1);
   721 by (etac ssubst 1);
   729 by (etac ssubst 1);
   722 by (rtac prem2 1);
   730 by (rtac prem2 1);
   723 val llist_fun_equalityI = result();
   731 val llist_fun_equalityI = result();
   724 
   732 
   725 (*simpset for llist bisimulations*)
   733 (*simpset for llist bisimulations*)
   730 
   738 
   731 (*** The functional "lmap" ***)
   739 (*** The functional "lmap" ***)
   732 
   740 
   733 goal LList.thy "lmap(f,LNil) = LNil";
   741 goal LList.thy "lmap(f,LNil) = LNil";
   734 by (rtac (lmap_def RS def_llist_corec RS trans) 1);
   742 by (rtac (lmap_def RS def_llist_corec RS trans) 1);
   735 by(simp_tac llistD_ss 1);
   743 by (simp_tac llistD_ss 1);
   736 val lmap_LNil = result();
   744 val lmap_LNil = result();
   737 
   745 
   738 goal LList.thy "lmap(f, LCons(M,N)) = LCons(f(M), lmap(f,N))";
   746 goal LList.thy "lmap(f, LCons(M,N)) = LCons(f(M), lmap(f,N))";
   739 by (rtac (lmap_def RS def_llist_corec RS trans) 1);
   747 by (rtac (lmap_def RS def_llist_corec RS trans) 1);
   740 by(simp_tac llistD_ss 1);
   748 by (simp_tac llistD_ss 1);
   741 val lmap_LCons = result();
   749 val lmap_LCons = result();
   742 
   750 
   743 
   751 
   744 (** Two easy results about lmap **)
   752 (** Two easy results about lmap **)
   745 
   753 
   756 
   764 
   757 (*** iterates -- llist_fun_equalityI cannot be used! ***)
   765 (*** iterates -- llist_fun_equalityI cannot be used! ***)
   758 
   766 
   759 goal LList.thy "iterates(f,x) = LCons(x, iterates(f,f(x)))";
   767 goal LList.thy "iterates(f,x) = LCons(x, iterates(f,f(x)))";
   760 by (rtac (iterates_def RS def_llist_corec RS trans) 1);
   768 by (rtac (iterates_def RS def_llist_corec RS trans) 1);
   761 by(simp_tac sum_ss 1);
   769 by (simp_tac sum_ss 1);
   762 val iterates = result();
   770 val iterates = result();
   763 
   771 
   764 goal LList.thy "lmap(f, iterates(f,x)) = iterates(f,f(x))";
   772 goal LList.thy "lmap(f, iterates(f,x)) = iterates(f,f(x))";
   765 by (res_inst_tac [("r", "range(%u.<lmap(f,iterates(f,u)),iterates(f,f(u))>)")] 
   773 by (res_inst_tac [("r", "range(%u.<lmap(f,iterates(f,u)),iterates(f,f(u))>)")] 
   766     llist_equalityI 1);
   774     llist_equalityI 1);
   782 
   790 
   783 goal LList.thy
   791 goal LList.thy
   784     "nat_rec(n, LCons(b, l), %m. lmap(f)) =	\
   792     "nat_rec(n, LCons(b, l), %m. lmap(f)) =	\
   785 \    LCons(nat_rec(n, b, %m. f), nat_rec(n, l, %m. lmap(f)))";
   793 \    LCons(nat_rec(n, b, %m. f), nat_rec(n, l, %m. lmap(f)))";
   786 by (nat_ind_tac "n" 1);
   794 by (nat_ind_tac "n" 1);
   787 by(ALLGOALS (asm_simp_tac (nat_ss addsimps [lmap_LCons])));
   795 by (ALLGOALS (asm_simp_tac (nat_ss addsimps [lmap_LCons])));
   788 val fun_power_lmap = result();
   796 val fun_power_lmap = result();
   789 
   797 
   790 goal Nat.thy "nat_rec(n, g(x), %m. g) = nat_rec(Suc(n), x, %m. g)";
   798 goal Nat.thy "nat_rec(n, g(x), %m. g) = nat_rec(Suc(n), x, %m. g)";
   791 by (nat_ind_tac "n" 1);
   799 by (nat_ind_tac "n" 1);
   792 by(ALLGOALS (asm_simp_tac nat_ss));
   800 by (ALLGOALS (asm_simp_tac nat_ss));
   793 val fun_power_Suc = result();
   801 val fun_power_Suc = result();
   794 
   802 
   795 val Pair_cong = read_instantiate_sg (sign_of Prod.thy)
   803 val Pair_cong = read_instantiate_sg (sign_of Prod.thy)
   796  [("f","Pair")] (standard(refl RS cong RS cong));
   804  [("f","Pair")] (standard(refl RS cong RS cong));
   797 
   805 
   812 by (stac fun_power_lmap 1);
   820 by (stac fun_power_lmap 1);
   813 br llistD_Fun_LCons_I 1;
   821 br llistD_Fun_LCons_I 1;
   814 by (rtac (lmap_iterates RS subst) 1);
   822 by (rtac (lmap_iterates RS subst) 1);
   815 by (stac fun_power_Suc 1);
   823 by (stac fun_power_Suc 1);
   816 by (stac fun_power_Suc 1);
   824 by (stac fun_power_Suc 1);
   817 br UN1_I 1;
   825 br (UN1_I RS UnI1) 1;
   818 br rangeI 1;
   826 br rangeI 1;
   819 val iterates_equality = result();
   827 val iterates_equality = result();
   820 
   828 
   821 
   829 
   822 (*** lappend -- its two arguments cause some complications! ***)
   830 (*** lappend -- its two arguments cause some complications! ***)
   823 
   831 
   824 goalw LList.thy [lappend_def] "lappend(LNil,LNil) = LNil";
   832 goalw LList.thy [lappend_def] "lappend(LNil,LNil) = LNil";
   825 by (rtac (llist_corec RS trans) 1);
   833 by (rtac (llist_corec RS trans) 1);
   826 by(simp_tac llistD_ss 1);
   834 by (simp_tac llistD_ss 1);
   827 val lappend_LNil_LNil = result();
   835 val lappend_LNil_LNil = result();
   828 
   836 
   829 goalw LList.thy [lappend_def]
   837 goalw LList.thy [lappend_def]
   830     "lappend(LNil,LCons(l,l')) = LCons(l, lappend(LNil,l'))";
   838     "lappend(LNil,LCons(l,l')) = LCons(l, lappend(LNil,l'))";
   831 by (rtac (llist_corec RS trans) 1);
   839 by (rtac (llist_corec RS trans) 1);
   832 by(simp_tac llistD_ss 1);
   840 by (simp_tac llistD_ss 1);
   833 (* 3.3(5.7) vs 1.3 !by (SIMP_TAC llistD_ss 1);*)
   841 (* 3.3(5.7) vs 1.3 !by (SIMP_TAC llistD_ss 1);*)
   834 val lappend_LNil_LCons = result();
   842 val lappend_LNil_LCons = result();
   835 
   843 
   836 goalw LList.thy [lappend_def]
   844 goalw LList.thy [lappend_def]
   837     "lappend(LCons(l,l'), N) = LCons(l, lappend(l',N))";
   845     "lappend(LCons(l,l'), N) = LCons(l, lappend(l',N))";
   838 by (rtac (llist_corec RS trans) 1);
   846 by (rtac (llist_corec RS trans) 1);
   839 by(simp_tac llistD_ss 1);
   847 by (simp_tac llistD_ss 1);
   840 (* 5(5.5) vs 1.3 !by (SIMP_TAC llistD_ss 1);*)
   848 (* 5(5.5) vs 1.3 !by (SIMP_TAC llistD_ss 1);*)
   841 val lappend_LCons = result();
   849 val lappend_LCons = result();
   842 
   850 
   843 goal LList.thy "lappend(LNil,l) = l";
   851 goal LList.thy "lappend(LNil,l) = l";
   844 by (res_inst_tac [("l","l")] llist_fun_equalityI 1);
   852 by (res_inst_tac [("l","l")] llist_fun_equalityI 1);
   875 by (res_inst_tac [("l", "l")] llistE 1);
   883 by (res_inst_tac [("l", "l")] llistE 1);
   876 by (res_inst_tac [("l", "n")] llistE 1);
   884 by (res_inst_tac [("l", "n")] llistE 1);
   877 by (ALLGOALS (asm_simp_tac (llistD_ss addsimps
   885 by (ALLGOALS (asm_simp_tac (llistD_ss addsimps
   878       [lappend_LNil_LNil,lappend_LCons,lappend_LNil_LCons,
   886       [lappend_LNil_LNil,lappend_LCons,lappend_LNil_LCons,
   879        lmap_LNil,lmap_LCons])));
   887        lmap_LNil,lmap_LCons])));
   880 by (REPEAT_SOME (ares_tac [llistD_Fun_LCons_I, UN1_I, rangeI]));
   888 by (REPEAT_SOME (ares_tac [llistD_Fun_LCons_I, UN1_I RS UnI1, rangeI]));
   881 by (rtac range_eqI 1);
   889 by (rtac range_eqI 1);
   882 by (rtac (refl RS Pair_cong) 1);
   890 by (rtac (refl RS Pair_cong) 1);
   883 by (stac lmap_LNil 1);
   891 by (stac lmap_LNil 1);
   884 by (rtac refl 1);
   892 by (rtac refl 1);
   885 val lmap_lappend_distrib = result();
   893 val lmap_lappend_distrib = result();
   886 
   894 
   887 (*Shorter proof of the theorem above using llist_equalityI2*)
   895 (*Shorter proof of theorem above using llist_equalityI as strong coinduction*)
   888 goal LList.thy "lmap(f, lappend(l,n)) = lappend(lmap(f,l), lmap(f,n))";
   896 goal LList.thy "lmap(f, lappend(l,n)) = lappend(lmap(f,l), lmap(f,n))";
   889 by (res_inst_tac [("l","l")] llist_fun_equalityI 1);
   897 by (res_inst_tac [("l","l")] llist_fun_equalityI 1);
   890 by (simp_tac (llistD_ss addsimps [lappend_LNil, lmap_LNil])1);
   898 by (simp_tac (llistD_ss addsimps [lappend_LNil, lmap_LNil])1);
   891 by (simp_tac (llistD_ss addsimps [lappend_LCons, lmap_LCons]) 1);
   899 by (simp_tac (llistD_ss addsimps [lappend_LCons, lmap_LCons]) 1);
   892 val lmap_lappend_distrib = result();
   900 val lmap_lappend_distrib = result();
   893 
   901 
   894 (*Without llist_equalityI2, three case analyses might be needed*)
   902 (*Without strong coinduction, three case analyses might be needed*)
   895 goal LList.thy "lappend(lappend(l1,l2) ,l3) = lappend(l1, lappend(l2,l3))";
   903 goal LList.thy "lappend(lappend(l1,l2) ,l3) = lappend(l1, lappend(l2,l3))";
   896 by (res_inst_tac [("l","l1")] llist_fun_equalityI 1);
   904 by (res_inst_tac [("l","l1")] llist_fun_equalityI 1);
   897 by (simp_tac (llistD_ss addsimps [lappend_LNil])1);
   905 by (simp_tac (llistD_ss addsimps [lappend_LNil])1);
   898 by (simp_tac (llistD_ss addsimps [lappend_LCons]) 1);
   906 by (simp_tac (llistD_ss addsimps [lappend_LCons]) 1);
   899 val lappend_assoc = result();
   907 val lappend_assoc = result();