| author | paulson | 
| Tue, 30 Nov 1999 17:53:34 +0100 | |
| changeset 8042 | ecdedff41e67 | 
| parent 7256 | 0a69baf28961 | 
| child 9747 | 043098ba5098 | 
| permissions | -rw-r--r-- | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 1 | (* Title: HOL/ex/SList.ML | 
| 
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 | 
| 4521 | 4 | Copyright 1998 University of Cambridge | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 5 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 6 | Definition of type 'a list by a least fixed point | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 7 | *) | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 8 | |
| 5977 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 paulson parents: 
5535diff
changeset | 9 | |
| 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 paulson parents: 
5535diff
changeset | 10 | Goalw [List_def] "x : list (range Leaf) ==> x : List"; | 
| 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 paulson parents: 
5535diff
changeset | 11 | by (Asm_simp_tac 1); | 
| 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 paulson parents: 
5535diff
changeset | 12 | qed "ListI"; | 
| 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 paulson parents: 
5535diff
changeset | 13 | |
| 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 paulson parents: 
5535diff
changeset | 14 | Goalw [List_def] "x : List ==> x : list (range Leaf)"; | 
| 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 paulson parents: 
5535diff
changeset | 15 | by (Asm_simp_tac 1); | 
| 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 paulson parents: 
5535diff
changeset | 16 | qed "ListD"; | 
| 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 paulson parents: 
5535diff
changeset | 17 | |
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 18 | val list_con_defs = [NIL_def, CONS_def]; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 19 | |
| 7256 | 20 | Goal "list(A) = usum {Numb(0)} (uprod A (list A))";
 | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 21 | let val rew = rewrite_rule list_con_defs in | 
| 4089 | 22 | by (fast_tac (claset() addSIs (equalityI :: map rew list.intrs) | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 23 | addEs [rew list.elim]) 1) | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 24 | end; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 25 | qed "list_unfold"; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 26 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 27 | (*This justifies using list in other recursive type definitions*) | 
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 28 | Goalw list.defs "A<=B ==> list(A) <= list(B)"; | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 29 | by (rtac lfp_mono 1); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 30 | by (REPEAT (ares_tac basic_monos 1)); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 31 | qed "list_mono"; | 
| 
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 | (*Type checking -- list creates well-founded sets*) | 
| 5069 | 34 | Goalw (list_con_defs @ list.defs) "list(sexp) <= sexp"; | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 35 | by (rtac lfp_lowerbound 1); | 
| 4089 | 36 | by (fast_tac (claset() addIs sexp.intrs@[sexp_In0I,sexp_In1I]) 1); | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 37 | qed "list_sexp"; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 38 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 39 | (* A <= sexp ==> list(A) <= sexp *) | 
| 5977 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 paulson parents: 
5535diff
changeset | 40 | bind_thm ("list_subset_sexp", [list_mono, list_sexp] MRS subset_trans);
 | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 41 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 42 | (*Induction for the type 'a list *) | 
| 5977 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 paulson parents: 
5535diff
changeset | 43 | val prems = Goalw [Nil_def,Cons_def] | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 44 | "[| P(Nil); \ | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 45 | \ !!x xs. P(xs) ==> P(x # xs) |] ==> P(l)"; | 
| 5977 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 paulson parents: 
5535diff
changeset | 46 | by (rtac (Rep_List_inverse RS subst) 1); (*types force good instantiation*) | 
| 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 paulson parents: 
5535diff
changeset | 47 | by (rtac (Rep_List RS ListD RS list.induct) 1); | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 48 | by (REPEAT (ares_tac prems 1 | 
| 5977 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 paulson parents: 
5535diff
changeset | 49 | ORELSE eresolve_tac [rangeE, ssubst, | 
| 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 paulson parents: 
5535diff
changeset | 50 | ListI RS Abs_List_inverse RS subst] 1)); | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 51 | qed "list_induct2"; | 
| 
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 | (*Perform induction on xs. *) | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 54 | fun list_ind_tac a M = | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 55 |     EVERY [res_inst_tac [("l",a)] list_induct2 M,
 | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 56 | rename_last_tac a ["1"] (M+1)]; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 57 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 58 | (*** Isomorphisms ***) | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 59 | |
| 5977 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 paulson parents: 
5535diff
changeset | 60 | Goal "inj Rep_List"; | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 61 | by (rtac inj_inverseI 1); | 
| 5977 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 paulson parents: 
5535diff
changeset | 62 | by (rtac Rep_List_inverse 1); | 
| 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 paulson parents: 
5535diff
changeset | 63 | qed "inj_Rep_List"; | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 64 | |
| 5977 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 paulson parents: 
5535diff
changeset | 65 | Goal "inj_on Abs_List List"; | 
| 4831 | 66 | by (rtac inj_on_inverseI 1); | 
| 5977 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 paulson parents: 
5535diff
changeset | 67 | by (etac Abs_List_inverse 1); | 
| 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 paulson parents: 
5535diff
changeset | 68 | qed "inj_on_Abs_List"; | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 69 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 70 | (** Distinctness of constructors **) | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 71 | |
| 5069 | 72 | Goalw list_con_defs "CONS M N ~= NIL"; | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 73 | by (rtac In1_not_In0 1); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 74 | qed "CONS_not_NIL"; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 75 | bind_thm ("NIL_not_CONS", (CONS_not_NIL RS not_sym));
 | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 76 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 77 | bind_thm ("CONS_neq_NIL", (CONS_not_NIL RS notE));
 | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 78 | val NIL_neq_CONS = sym RS CONS_neq_NIL; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 79 | |
| 5069 | 80 | Goalw [Nil_def,Cons_def] "x # xs ~= Nil"; | 
| 5977 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 paulson parents: 
5535diff
changeset | 81 | by (rtac (CONS_not_NIL RS (inj_on_Abs_List RS inj_on_contraD)) 1); | 
| 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 paulson parents: 
5535diff
changeset | 82 | by (REPEAT (resolve_tac (list.intrs @ [rangeI, Rep_List RS ListD, ListI]) 1)); | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 83 | qed "Cons_not_Nil"; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 84 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 85 | bind_thm ("Nil_not_Cons", Cons_not_Nil RS not_sym);
 | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 86 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 87 | (** Injectiveness of CONS and Cons **) | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 88 | |
| 5069 | 89 | Goalw [CONS_def] "(CONS K M=CONS L N) = (K=L & M=N)"; | 
| 4089 | 90 | by (fast_tac (claset() addSEs [Scons_inject, make_elim In1_inject]) 1); | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 91 | qed "CONS_CONS_eq"; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 92 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 93 | (*For reasoning about abstract list constructors*) | 
| 5977 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 paulson parents: 
5535diff
changeset | 94 | AddIs [Rep_List RS ListD, ListI]; | 
| 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 paulson parents: 
5535diff
changeset | 95 | AddIs list.intrs; | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 96 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 97 | AddIffs [CONS_not_NIL, NIL_not_CONS, CONS_CONS_eq]; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 98 | |
| 5977 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 paulson parents: 
5535diff
changeset | 99 | AddSDs [inj_on_Abs_List RS inj_onD, | 
| 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 paulson parents: 
5535diff
changeset | 100 | inj_Rep_List RS injD, Leaf_inject]; | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 101 | |
| 5069 | 102 | Goalw [Cons_def] "(x#xs=y#ys) = (x=y & xs=ys)"; | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 103 | by (Fast_tac 1); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 104 | qed "Cons_Cons_eq"; | 
| 5977 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 paulson parents: 
5535diff
changeset | 105 | bind_thm ("Cons_inject2", Cons_Cons_eq RS iffD1 RS conjE);
 | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 106 | |
| 5977 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 paulson parents: 
5535diff
changeset | 107 | Goal "CONS M N: list(A) ==> M: A & N: list(A)"; | 
| 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 paulson parents: 
5535diff
changeset | 108 | by (etac setup_induction 1); | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 109 | by (etac list.induct 1); | 
| 5977 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 paulson parents: 
5535diff
changeset | 110 | by (ALLGOALS Fast_tac); | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 111 | qed "CONS_D"; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 112 | |
| 5977 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 paulson parents: 
5535diff
changeset | 113 | Goalw [CONS_def,In1_def] "CONS M N: sexp ==> M: sexp & N: sexp"; | 
| 4089 | 114 | by (fast_tac (claset() addSDs [Scons_D]) 1); | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 115 | qed "sexp_CONS_D"; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 116 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 117 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 118 | (*Reasoning about constructors and their freeness*) | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 119 | Addsimps list.intrs; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 120 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 121 | AddIffs [Cons_not_Nil, Nil_not_Cons, Cons_Cons_eq]; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 122 | |
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 123 | Goal "N: list(A) ==> !M. N ~= CONS M N"; | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 124 | by (etac list.induct 1); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 125 | by (ALLGOALS Asm_simp_tac); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 126 | qed "not_CONS_self"; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 127 | |
| 5069 | 128 | Goal "!x. l ~= x#l"; | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 129 | by (list_ind_tac "l" 1); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 130 | by (ALLGOALS Asm_simp_tac); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 131 | qed "not_Cons_self2"; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 132 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 133 | |
| 5069 | 134 | Goal "(xs ~= []) = (? y ys. xs = y#ys)"; | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 135 | by (list_ind_tac "xs" 1); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 136 | by (Simp_tac 1); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 137 | by (Asm_simp_tac 1); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 138 | by (REPEAT(resolve_tac [exI,refl,conjI] 1)); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 139 | qed "neq_Nil_conv2"; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 140 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 141 | (** Conversion rules for List_case: case analysis operator **) | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 142 | |
| 5069 | 143 | Goalw [List_case_def,NIL_def] "List_case c h NIL = c"; | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 144 | by (rtac Case_In0 1); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 145 | qed "List_case_NIL"; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 146 | |
| 5069 | 147 | Goalw [List_case_def,CONS_def] "List_case c h (CONS M N) = h M N"; | 
| 4521 | 148 | by (Simp_tac 1); | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 149 | qed "List_case_CONS"; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 150 | |
| 4521 | 151 | Addsimps [List_case_NIL, List_case_CONS]; | 
| 152 | ||
| 153 | ||
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 154 | (*** List_rec -- by wf recursion on pred_sexp ***) | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 155 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 156 | (* The trancl(pred_sexp) is essential because pred_sexp_CONS_I1,2 would not | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 157 | hold if pred_sexp^+ were changed to pred_sexp. *) | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 158 | |
| 5278 | 159 | Goal "(%M. List_rec M c d) = wfrec (trancl pred_sexp) \ | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 160 | \ (%g. List_case c (%x y. d x y (g y)))"; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 161 | by (simp_tac (HOL_ss addsimps [List_rec_def]) 1); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 162 | val List_rec_unfold = standard | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 163 | ((wf_pred_sexp RS wf_trancl) RS ((result() RS eq_reflection) RS def_wfrec)); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 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 | * Old: | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 167 | * val List_rec_unfold = [List_rec_def,wf_pred_sexp RS wf_trancl] MRS def_wfrec | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 168 | * |> standard; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 169 | *---------------------------------------------------------------------------*) | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 170 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 171 | (** pred_sexp lemmas **) | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 172 | |
| 5069 | 173 | Goalw [CONS_def,In1_def] | 
| 5148 
74919e8f221c
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5143diff
changeset | 174 | "[| M: sexp; N: sexp |] ==> (M, CONS M N) : pred_sexp^+"; | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 175 | by (Asm_simp_tac 1); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 176 | qed "pred_sexp_CONS_I1"; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 177 | |
| 5069 | 178 | Goalw [CONS_def,In1_def] | 
| 5148 
74919e8f221c
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5143diff
changeset | 179 | "[| M: sexp; N: sexp |] ==> (N, CONS M N) : pred_sexp^+"; | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 180 | by (Asm_simp_tac 1); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 181 | qed "pred_sexp_CONS_I2"; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 182 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 183 | val [prem] = goal SList.thy | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 184 | "(CONS M1 M2, N) : pred_sexp^+ ==> \ | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 185 | \ (M1,N) : pred_sexp^+ & (M2,N) : pred_sexp^+"; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 186 | by (rtac (prem RS (pred_sexp_subset_Sigma RS trancl_subset_Sigma RS | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 187 | subsetD RS SigmaE2)) 1); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 188 | by (etac (sexp_CONS_D RS conjE) 1); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 189 | by (REPEAT (ares_tac [conjI, pred_sexp_CONS_I1, pred_sexp_CONS_I2, | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 190 | prem RSN (2, trans_trancl RS transD)] 1)); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 191 | qed "pred_sexp_CONS_D"; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 192 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 193 | (** Conversion rules for List_rec **) | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 194 | |
| 5069 | 195 | Goal "List_rec NIL c h = c"; | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 196 | by (rtac (List_rec_unfold RS trans) 1); | 
| 4521 | 197 | by (Simp_tac 1); | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 198 | qed "List_rec_NIL"; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 199 | |
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 200 | Goal "[| M: sexp; N: sexp |] ==> \ | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 201 | \ List_rec (CONS M N) c h = h M N (List_rec N c h)"; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 202 | by (rtac (List_rec_unfold RS trans) 1); | 
| 4521 | 203 | by (asm_simp_tac (simpset() addsimps [pred_sexp_CONS_I2]) 1); | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 204 | qed "List_rec_CONS"; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 205 | |
| 4521 | 206 | Addsimps [List_rec_NIL, List_rec_CONS]; | 
| 207 | ||
| 208 | ||
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 209 | (*** list_rec -- by List_rec ***) | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 210 | |
| 5977 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 paulson parents: 
5535diff
changeset | 211 | val Rep_List_in_sexp = | 
| 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 paulson parents: 
5535diff
changeset | 212 | [range_Leaf_subset_sexp RS list_subset_sexp, Rep_List RS ListD] | 
| 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 paulson parents: 
5535diff
changeset | 213 | MRS subsetD; | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 214 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 215 | local | 
| 5977 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 paulson parents: 
5535diff
changeset | 216 | val list_rec_simps = [ListI RS Abs_List_inverse, Rep_List_inverse, | 
| 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 paulson parents: 
5535diff
changeset | 217 | Rep_List RS ListD, rangeI, inj_Leaf, inv_f_f, | 
| 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 paulson parents: 
5535diff
changeset | 218 | sexp.LeafI, Rep_List_in_sexp] | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 219 | in | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 220 | val list_rec_Nil = prove_goalw SList.thy [list_rec_def, Nil_def] | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 221 | "list_rec Nil c h = c" | 
| 4089 | 222 | (fn _=> [simp_tac (simpset() addsimps list_rec_simps) 1]); | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 223 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 224 | val list_rec_Cons = prove_goalw SList.thy [list_rec_def, Cons_def] | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 225 | "list_rec (a#l) c h = h a l (list_rec l c h)" | 
| 4089 | 226 | (fn _=> [simp_tac (simpset() addsimps list_rec_simps) 1]); | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 227 | end; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 228 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 229 | Addsimps [List_rec_NIL, List_rec_CONS, list_rec_Nil, list_rec_Cons]; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 230 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 231 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 232 | (*Type checking. Useful?*) | 
| 5977 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 paulson parents: 
5535diff
changeset | 233 | val major::A_subset_sexp::prems = | 
| 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 paulson parents: 
5535diff
changeset | 234 | Goal "[| M: list(A); \ | 
| 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 paulson parents: 
5535diff
changeset | 235 | \ A<=sexp; \ | 
| 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 paulson parents: 
5535diff
changeset | 236 | \ c: C(NIL); \ | 
| 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 paulson parents: 
5535diff
changeset | 237 | \ !!x y r. [| x: A; y: list(A); r: C(y) |] ==> h x y r: C(CONS x y) \ | 
| 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 paulson parents: 
5535diff
changeset | 238 | \ |] ==> List_rec M c h : C(M :: 'a item)"; | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 239 | val sexp_ListA_I = A_subset_sexp RS list_subset_sexp RS subsetD; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 240 | val sexp_A_I = A_subset_sexp RS subsetD; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 241 | by (rtac (major RS list.induct) 1); | 
| 5535 | 242 | by (ALLGOALS(asm_simp_tac (simpset() addsimps [sexp_A_I,sexp_ListA_I]@prems))); | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 243 | qed "List_rec_type"; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 244 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 245 | (** Generalized map functionals **) | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 246 | |
| 5069 | 247 | Goalw [Rep_map_def] "Rep_map f Nil = NIL"; | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 248 | by (rtac list_rec_Nil 1); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 249 | qed "Rep_map_Nil"; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 250 | |
| 5977 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 paulson parents: 
5535diff
changeset | 251 | Goalw [Rep_map_def] "Rep_map f (x#xs) = CONS (f x) (Rep_map f xs)"; | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 252 | by (rtac list_rec_Cons 1); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 253 | qed "Rep_map_Cons"; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 254 | |
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 255 | val prems = Goalw [Rep_map_def] "(!!x. f(x): A) ==> Rep_map f xs: list(A)"; | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 256 | by (rtac list_induct2 1); | 
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 257 | by (ALLGOALS (asm_simp_tac (simpset() addsimps prems))); | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 258 | qed "Rep_map_type"; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 259 | |
| 5069 | 260 | Goalw [Abs_map_def] "Abs_map g NIL = Nil"; | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 261 | by (rtac List_rec_NIL 1); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 262 | qed "Abs_map_NIL"; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 263 | |
| 5977 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 paulson parents: 
5535diff
changeset | 264 | Goalw [Abs_map_def] | 
| 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 paulson parents: 
5535diff
changeset | 265 | "[| M: sexp; N: sexp |] ==> Abs_map g (CONS M N) = g(M) # Abs_map g N"; | 
| 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 paulson parents: 
5535diff
changeset | 266 | by (REPEAT (ares_tac [List_rec_CONS] 1)); | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 267 | qed "Abs_map_CONS"; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 268 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 269 | (*These 2 rules ease the use of primitive recursion. NOTE USE OF == *) | 
| 5977 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 paulson parents: 
5535diff
changeset | 270 | val [rew] = | 
| 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 paulson parents: 
5535diff
changeset | 271 | Goal "[| !!xs. f(xs) == list_rec xs c h |] ==> f([]) = c"; | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 272 | by (rewtac rew); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 273 | by (rtac list_rec_Nil 1); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 274 | qed "def_list_rec_Nil"; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 275 | |
| 5977 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 paulson parents: 
5535diff
changeset | 276 | val [rew] = | 
| 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 paulson parents: 
5535diff
changeset | 277 | Goal "[| !!xs. f(xs) == list_rec xs c h |] ==> f(x#xs) = h x xs (f xs)"; | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 278 | by (rewtac rew); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 279 | by (rtac list_rec_Cons 1); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 280 | qed "def_list_rec_Cons"; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 281 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 282 | fun list_recs def = | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 283 | [standard (def RS def_list_rec_Nil), | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 284 | standard (def RS def_list_rec_Cons)]; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 285 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 286 | (*** Unfolding the basic combinators ***) | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 287 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 288 | val [null_Nil, null_Cons] = list_recs null_def; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 289 | val [_, hd_Cons] = list_recs hd_def; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 290 | val [_, tl_Cons] = list_recs tl_def; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 291 | val [ttl_Nil, ttl_Cons] = list_recs ttl_def; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 292 | val [append_Nil3, append_Cons] = list_recs append_def; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 293 | val [mem_Nil, mem_Cons] = list_recs mem_def; | 
| 3649 
e530286d4847
Renamed set_of_list to set, and relevant theorems too
 paulson parents: 
3120diff
changeset | 294 | val [set_Nil, set_Cons] = list_recs set_def; | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 295 | val [map_Nil, map_Cons] = list_recs map_def; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 296 | val [list_case_Nil, list_case_Cons] = list_recs list_case_def; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 297 | val [filter_Nil, filter_Cons] = list_recs filter_def; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 298 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 299 | Addsimps | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 300 | [null_Nil, ttl_Nil, | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 301 | mem_Nil, mem_Cons, | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 302 | list_case_Nil, list_case_Cons, | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 303 | append_Nil3, append_Cons, | 
| 3649 
e530286d4847
Renamed set_of_list to set, and relevant theorems too
 paulson parents: 
3120diff
changeset | 304 | set_Nil, set_Cons, | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 305 | map_Nil, map_Cons, | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 306 | filter_Nil, filter_Cons]; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 307 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 308 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 309 | (** @ - append **) | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 310 | |
| 5069 | 311 | Goal "(xs@ys)@zs = xs@(ys@zs)"; | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 312 | by (list_ind_tac "xs" 1); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 313 | by (ALLGOALS Asm_simp_tac); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 314 | qed "append_assoc2"; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 315 | |
| 5069 | 316 | Goal "xs @ [] = xs"; | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 317 | by (list_ind_tac "xs" 1); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 318 | by (ALLGOALS Asm_simp_tac); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 319 | qed "append_Nil4"; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 320 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 321 | (** mem **) | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 322 | |
| 5069 | 323 | Goal "x mem (xs@ys) = (x mem xs | x mem ys)"; | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 324 | by (list_ind_tac "xs" 1); | 
| 4686 | 325 | by (ALLGOALS Asm_simp_tac); | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 326 | qed "mem_append2"; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 327 | |
| 5069 | 328 | Goal "x mem [x:xs. P(x)] = (x mem xs & P(x))"; | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 329 | by (list_ind_tac "xs" 1); | 
| 4686 | 330 | by (ALLGOALS Asm_simp_tac); | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 331 | qed "mem_filter2"; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 332 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 333 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 334 | (** The functional "map" **) | 
| 
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 | Addsimps [Rep_map_Nil, Rep_map_Cons, Abs_map_NIL, Abs_map_CONS]; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 337 | |
| 5977 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 paulson parents: 
5535diff
changeset | 338 | val [major,A_subset_sexp,minor] = | 
| 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 paulson parents: 
5535diff
changeset | 339 | Goal "[| M: list(A); A<=sexp; !!z. z: A ==> f(g(z)) = z |] \ | 
| 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 paulson parents: 
5535diff
changeset | 340 | \ ==> Rep_map f (Abs_map g M) = M"; | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 341 | by (rtac (major RS list.induct) 1); | 
| 5977 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 paulson parents: 
5535diff
changeset | 342 | by (ALLGOALS | 
| 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 paulson parents: 
5535diff
changeset | 343 | (asm_simp_tac (simpset() addsimps [sexp_A_I,sexp_ListA_I,minor]))); | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 344 | qed "Abs_map_inverse"; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 345 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 346 | (*Rep_map_inverse is obtained via Abs_Rep_map and map_ident*) | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 347 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 348 | (** list_case **) | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 349 | |
| 5278 | 350 | Goal "P(list_case a f xs) = ((xs=[] --> P(a)) & \ | 
| 5977 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 paulson parents: 
5535diff
changeset | 351 | \ (!y ys. xs=y#ys --> P(f y ys)))"; | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 352 | by (list_ind_tac "xs" 1); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 353 | by (ALLGOALS Asm_simp_tac); | 
| 4831 | 354 | qed "split_list_case2"; | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 355 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 356 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 357 | (** Additional mapping lemmas **) | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 358 | |
| 5069 | 359 | Goal "map (%x. x) xs = xs"; | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 360 | by (list_ind_tac "xs" 1); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 361 | by (ALLGOALS Asm_simp_tac); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 362 | qed "map_ident2"; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 363 | |
| 5069 | 364 | Goal "map f (xs@ys) = map f xs @ map f ys"; | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 365 | by (list_ind_tac "xs" 1); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 366 | by (ALLGOALS Asm_simp_tac); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 367 | qed "map_append2"; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 368 | |
| 5069 | 369 | Goalw [o_def] "map (f o g) xs = map f (map g xs)"; | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 370 | by (list_ind_tac "xs" 1); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 371 | by (ALLGOALS Asm_simp_tac); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 372 | qed "map_compose2"; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 373 | |
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 374 | val prems = | 
| 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 375 | Goal "(!!x. f(x): sexp) ==> \ | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 376 | \ Abs_map g (Rep_map f xs) = map (%t. g(f(t))) xs"; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 377 | by (list_ind_tac "xs" 1); | 
| 4521 | 378 | by (ALLGOALS (asm_simp_tac(simpset() addsimps | 
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 379 | (prems@[Rep_map_type, list_sexp RS subsetD])))); | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 380 | qed "Abs_Rep_map"; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 381 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 382 | Addsimps [append_Nil4, map_ident2]; |