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