Moved the old List to ex and replaced it by one defined via
authornipkow
Fri, 02 Dec 1994 16:13:34 +0100
changeset 196 61620d959717
parent 195 df6b3bd14dcb
child 197 2757544bbe6d
Moved the old List to ex and replaced it by one defined via datatype and primrec. List does not depend on Sexp any more. Had to modify the datatype-grammar to allow "[]" ("[]"). Does not allow strings but only ids as type constructors. This is consistent with the documentation and merely undoes an extension of Markus's.
List.ML
List.thy
ROOT.ML
thy_syntax.ML
--- a/List.ML	Fri Dec 02 16:09:49 1994 +0100
+++ b/List.ML	Fri Dec 02 16:13:34 1994 +0100
@@ -1,296 +1,38 @@
 (*  Title: 	HOL/List
     ID:         $Id$
-    Author: 	Lawrence C Paulson, Cambridge University Computer Laboratory
-    Copyright   1993  University of Cambridge
+    Author: 	Tobias Nipkow
+    Copyright   1994 TU Muenchen
 
-Definition of type 'a list by a least fixed point
+List lemmas
 *)
 
 open List;
 
-val list_con_defs = [NIL_def, CONS_def];
-
-goal List.thy "list(A) = {Numb(0)} <+> (A <*> list(A))";
-let val rew = rewrite_rule list_con_defs in  
-by (fast_tac (univ_cs addSIs (equalityI :: map rew list.intrs)
-                      addEs [rew list.elim]) 1)
-end;
-qed "list_unfold";
-
-(*This justifies using list in other recursive type definitions*)
-goalw List.thy list.defs "!!A B. A<=B ==> list(A) <= list(B)";
-by (rtac lfp_mono 1);
-by (REPEAT (ares_tac basic_monos 1));
-qed "list_mono";
-
-(*Type checking -- list creates well-founded sets*)
-goalw List.thy (list_con_defs @ list.defs) "list(sexp) <= sexp";
-by (rtac lfp_lowerbound 1);
-by (fast_tac (univ_cs addIs sexp.intrs@[sexp_In0I,sexp_In1I]) 1);
-qed "list_sexp";
-
-(* A <= sexp ==> list(A) <= sexp *)
-val list_subset_sexp = standard ([list_mono, list_sexp] MRS subset_trans);
+val [Nil_not_Cons,Cons_not_Nil] = list.distinct;
 
-(*Induction for the type 'a list *)
-val prems = goalw List.thy [Nil_def,Cons_def]
-    "[| P(Nil);   \
-\       !!x xs. P(xs) ==> P(x # xs) |]  ==> P(l)";
-by (rtac (Rep_list_inverse RS subst) 1);   (*types force good instantiation*)
-by (rtac (Rep_list RS list.induct) 1);
-by (REPEAT (ares_tac prems 1
-     ORELSE eresolve_tac [rangeE, ssubst, Abs_list_inverse RS subst] 1));
-qed "list_induct";
-
-(*Perform induction on xs. *)
-fun list_ind_tac a M = 
-    EVERY [res_inst_tac [("l",a)] list_induct M,
-	   rename_last_tac a ["1"] (M+1)];
-
-(*** Isomorphisms ***)
-
-goal List.thy "inj(Rep_list)";
-by (rtac inj_inverseI 1);
-by (rtac Rep_list_inverse 1);
-qed "inj_Rep_list";
-
-goal List.thy "inj_onto(Abs_list,list(range(Leaf)))";
-by (rtac inj_onto_inverseI 1);
-by (etac Abs_list_inverse 1);
-qed "inj_onto_Abs_list";
-
-(** Distinctness of constructors **)
-
-goalw List.thy list_con_defs "CONS(M,N) ~= NIL";
-by (rtac In1_not_In0 1);
-qed "CONS_not_NIL";
-val NIL_not_CONS = standard (CONS_not_NIL RS not_sym);
+val Cons_neq_Nil = store_thm("Cons_neq_Nil", standard (Cons_not_Nil RS notE));
+val Nil_neq_Cons = store_thm("Nil_neq_Cons", sym RS Cons_neq_Nil);
 
-val CONS_neq_NIL = standard (CONS_not_NIL RS notE);
-val NIL_neq_CONS = sym RS CONS_neq_NIL;
-
-goalw List.thy [Nil_def,Cons_def] "x # xs ~= Nil";
-by (rtac (CONS_not_NIL RS (inj_onto_Abs_list RS inj_onto_contraD)) 1);
-by (REPEAT (resolve_tac (list.intrs @ [rangeI, Rep_list]) 1));
-qed "Cons_not_Nil";
-
-val Nil_not_Cons = standard (Cons_not_Nil RS not_sym);
-
-val Cons_neq_Nil = standard (Cons_not_Nil RS notE);
-val Nil_neq_Cons = sym RS Cons_neq_Nil;
-
-(** Injectiveness of CONS and Cons **)
-
-goalw List.thy [CONS_def] "(CONS(K,M)=CONS(L,N)) = (K=L & M=N)";
-by (fast_tac (HOL_cs addSEs [Scons_inject, make_elim In1_inject]) 1);
-qed "CONS_CONS_eq";
-
-val CONS_inject = standard (CONS_CONS_eq RS iffD1 RS conjE);
-
-(*For reasoning about abstract list constructors*)
-val list_cs = set_cs addIs [Rep_list] @ list.intrs
-	             addSEs [CONS_neq_NIL,NIL_neq_CONS,CONS_inject]
-		     addSDs [inj_onto_Abs_list RS inj_ontoD,
-			     inj_Rep_list RS injD, Leaf_inject];
+val Cons_inject =
+  store_thm("Cons_inject", standard ((hd list.inject) RS iffD1 RS conjE));
 
-goalw List.thy [Cons_def] "(x#xs=y#ys) = (x=y & xs=ys)";
-by (fast_tac list_cs 1);
-qed "Cons_Cons_eq";
-val Cons_inject = standard (Cons_Cons_eq RS iffD1 RS conjE);
-
-val [major] = goal List.thy "CONS(M,N): list(A) ==> M: A & N: list(A)";
-by (rtac (major RS setup_induction) 1);
-by (etac list.induct 1);
-by (ALLGOALS (fast_tac list_cs));
-qed "CONS_D";
+val list_ss = HOL_ss addsimps list.simps;
 
-val prems = goalw List.thy [CONS_def,In1_def]
-    "CONS(M,N): sexp ==> M: sexp & N: sexp";
-by (cut_facts_tac prems 1);
-by (fast_tac (set_cs addSDs [Scons_D]) 1);
-qed "sexp_CONS_D";
-
-
-(*Basic ss with constructors and their freeness*)
-val list_free_simps = [Cons_not_Nil, Nil_not_Cons, Cons_Cons_eq,
-		       CONS_not_NIL, NIL_not_CONS, CONS_CONS_eq]
-                      @ list.intrs;
-val list_free_ss = HOL_ss  addsimps  list_free_simps;
-
-goal List.thy "!!N. N: list(A) ==> !M. N ~= CONS(M,N)";
-by (etac list.induct 1);
-by (ALLGOALS (asm_simp_tac list_free_ss));
-qed "not_CONS_self";
-
-goal List.thy "!x. l ~= x#l";
-by (list_ind_tac "l" 1);
-by (ALLGOALS (asm_simp_tac list_free_ss));
+goal List.thy "!x. xs ~= x#xs";
+by (list.induct_tac "xs" 1);
+by (ALLGOALS (asm_simp_tac list_ss));
 qed "not_Cons_self";
 
-
 goal List.thy "(xs ~= []) = (? y ys. xs = y#ys)";
-by(list_ind_tac "xs" 1);
-by(simp_tac list_free_ss 1);
-by(asm_simp_tac list_free_ss 1);
+by (list.induct_tac "xs" 1);
+by(simp_tac list_ss 1);
+by(asm_simp_tac list_ss 1);
 by(REPEAT(resolve_tac [exI,refl,conjI] 1));
 qed "neq_Nil_conv";
 
-(** Conversion rules for List_case: case analysis operator **)
-
-goalw List.thy [List_case_def,NIL_def] "List_case(c, h, NIL) = c";
-by (rtac Case_In0 1);
-qed "List_case_NIL";
-
-goalw List.thy [List_case_def,CONS_def]  "List_case(c, h, CONS(M,N)) = h(M,N)";
-by (simp_tac (HOL_ss addsimps [Split,Case_In1]) 1);
-qed "List_case_CONS";
-
-(*** List_rec -- by wf recursion on pred_sexp ***)
-
-(* The trancl(pred_sexp) is essential because pred_sexp_CONS_I1,2 would not
-   hold if pred_sexp^+ were changed to pred_sexp. *)
-
-val List_rec_unfold = [List_rec_def, wf_pred_sexp RS wf_trancl] MRS def_wfrec
-                      |> standard;
-
-(** pred_sexp lemmas **)
-
-goalw List.thy [CONS_def,In1_def]
-    "!!M. [| M: sexp;  N: sexp |] ==> <M, CONS(M,N)> : pred_sexp^+";
-by (asm_simp_tac pred_sexp_ss 1);
-qed "pred_sexp_CONS_I1";
-
-goalw List.thy [CONS_def,In1_def]
-    "!!M. [| M: sexp;  N: sexp |] ==> <N, CONS(M,N)> : pred_sexp^+";
-by (asm_simp_tac pred_sexp_ss 1);
-qed "pred_sexp_CONS_I2";
-
-val [prem] = goal List.thy
-    "<CONS(M1,M2), N> : pred_sexp^+ ==> \
-\    <M1,N> : pred_sexp^+ & <M2,N> : pred_sexp^+";
-by (rtac (prem RS (pred_sexp_subset_Sigma RS trancl_subset_Sigma RS 
-		   subsetD RS SigmaE2)) 1);
-by (etac (sexp_CONS_D RS conjE) 1);
-by (REPEAT (ares_tac [conjI, pred_sexp_CONS_I1, pred_sexp_CONS_I2,
-		      prem RSN (2, trans_trancl RS transD)] 1));
-qed "pred_sexp_CONS_D";
-
-(** Conversion rules for List_rec **)
-
-goal List.thy "List_rec(NIL,c,h) = c";
-by (rtac (List_rec_unfold RS trans) 1);
-by (simp_tac (HOL_ss addsimps [List_case_NIL]) 1);
-qed "List_rec_NIL";
-
-goal List.thy "!!M. [| M: sexp;  N: sexp |] ==> \
-\    List_rec(CONS(M,N), c, h) = h(M, N, List_rec(N,c,h))";
-by (rtac (List_rec_unfold RS trans) 1);
-by (asm_simp_tac
-    (HOL_ss addsimps [List_case_CONS, list.CONS_I, pred_sexp_CONS_I2, 
-		      cut_apply])1);
-qed "List_rec_CONS";
-
-(*** list_rec -- by List_rec ***)
-
-val Rep_list_in_sexp =
-    [range_Leaf_subset_sexp RS list_subset_sexp, Rep_list] MRS subsetD;
-
-local
-  val list_rec_simps = list_free_simps @
-	          [List_rec_NIL, List_rec_CONS, 
-		   Abs_list_inverse, Rep_list_inverse,
-		   Rep_list, rangeI, inj_Leaf, Inv_f_f,
-		   sexp.LeafI, Rep_list_in_sexp]
-in
-  val list_rec_Nil = prove_goalw List.thy [list_rec_def, Nil_def]
-      "list_rec(Nil,c,h) = c"
-   (fn _=> [simp_tac (HOL_ss addsimps list_rec_simps) 1]);
-
-  val list_rec_Cons = prove_goalw List.thy [list_rec_def, Cons_def]
-      "list_rec(a#l, c, h) = h(a, l, list_rec(l,c,h))"
-   (fn _=> [simp_tac (HOL_ss addsimps list_rec_simps) 1]);
-end;
-
-val list_simps = [List_rec_NIL, List_rec_CONS,
-		  list_rec_Nil, list_rec_Cons];
-val list_ss = list_free_ss addsimps list_simps;
-
-
-(*Type checking.  Useful?*)
-val major::A_subset_sexp::prems = goal List.thy
-    "[| M: list(A);    	\
-\       A<=sexp;      	\
-\       c: C(NIL);      \
-\       !!x y r. [| x: A;  y: list(A);  r: C(y) |] ==> h(x,y,r): C(CONS(x,y)) \
-\    |] ==> List_rec(M,c,h) : C(M :: 'a item)";
-val sexp_ListA_I = A_subset_sexp RS list_subset_sexp RS subsetD;
-val sexp_A_I = A_subset_sexp RS subsetD;
-by (rtac (major RS list.induct) 1);
-by (ALLGOALS(asm_simp_tac (list_ss addsimps ([sexp_A_I,sexp_ListA_I]@prems))));
-qed "List_rec_type";
-
-(** Generalized map functionals **)
-
-goalw List.thy [Rep_map_def] "Rep_map(f,Nil) = NIL";
-by (rtac list_rec_Nil 1);
-qed "Rep_map_Nil";
-
-goalw List.thy [Rep_map_def]
-    "Rep_map(f, x#xs) = CONS(f(x), Rep_map(f,xs))";
-by (rtac list_rec_Cons 1);
-qed "Rep_map_Cons";
-
-goalw List.thy [Rep_map_def] "!!f. (!!x. f(x): A) ==> Rep_map(f,xs): list(A)";
-by (rtac list_induct 1);
-by(ALLGOALS(asm_simp_tac list_ss));
-qed "Rep_map_type";
-
-goalw List.thy [Abs_map_def] "Abs_map(g,NIL) = Nil";
-by (rtac List_rec_NIL 1);
-qed "Abs_map_NIL";
-
-val prems = goalw List.thy [Abs_map_def]
-    "[| M: sexp;  N: sexp |] ==> \
-\    Abs_map(g, CONS(M,N)) = g(M) # Abs_map(g,N)";
-by (REPEAT (resolve_tac (List_rec_CONS::prems) 1));
-qed "Abs_map_CONS";
-
-(*These 2 rules ease the use of primitive recursion.  NOTE USE OF == *)
-val [rew] = goal List.thy
-    "[| !!xs. f(xs) == list_rec(xs,c,h) |] ==> f([]) = c";
-by (rewtac rew);
-by (rtac list_rec_Nil 1);
-qed "def_list_rec_Nil";
-
-val [rew] = goal List.thy
-    "[| !!xs. f(xs) == list_rec(xs,c,h) |] ==> f(x#xs) = h(x,xs,f(xs))";
-by (rewtac rew);
-by (rtac list_rec_Cons 1);
-qed "def_list_rec_Cons";
-
-fun list_recs def =
-      [standard (def RS def_list_rec_Nil),
-       standard (def RS def_list_rec_Cons)];
-
-(*** Unfolding the basic combinators ***)
-
-val [null_Nil,null_Cons] = list_recs null_def;
-val [_,hd_Cons] = list_recs hd_def;
-val [_,tl_Cons] = list_recs tl_def;
-val [ttl_Nil,ttl_Cons] = list_recs ttl_def;
-val [append_Nil,append_Cons] = list_recs append_def;
-val [mem_Nil, mem_Cons] = list_recs mem_def;
-val [map_Nil,map_Cons] = list_recs map_def;
-val [list_case_Nil,list_case_Cons] = list_recs list_case_def;
-val [filter_Nil,filter_Cons] = list_recs filter_def;
-val [list_all_Nil,list_all_Cons] = list_recs list_all_def;
-
-val list_ss = arith_ss addsimps
-  [Cons_not_Nil, Nil_not_Cons, Cons_Cons_eq,
-   list_rec_Nil, list_rec_Cons,
-   null_Nil, null_Cons, hd_Cons, tl_Cons, ttl_Nil, ttl_Cons,
+val list_ss = arith_ss addsimps list.simps @
+  [null_Nil, null_Cons, hd_Cons, tl_Cons, ttl_Nil, ttl_Cons,
    mem_Nil, mem_Cons,
-   list_case_Nil, list_case_Cons,
    append_Nil, append_Cons,
    map_Nil, map_Cons,
    list_all_Nil, list_all_Cons,
@@ -300,68 +42,52 @@
 (** @ - append **)
 
 goal List.thy "(xs@ys)@zs = xs@(ys@zs)";
-by(list_ind_tac "xs" 1);
+by (list.induct_tac "xs" 1);
 by(ALLGOALS(asm_simp_tac list_ss));
 qed "append_assoc";
 
 goal List.thy "xs @ [] = xs";
-by(list_ind_tac "xs" 1);
+by (list.induct_tac "xs" 1);
 by(ALLGOALS(asm_simp_tac list_ss));
 qed "append_Nil2";
 
 (** mem **)
 
 goal List.thy "x mem (xs@ys) = (x mem xs | x mem ys)";
-by(list_ind_tac "xs" 1);
+by (list.induct_tac "xs" 1);
 by(ALLGOALS(asm_simp_tac (list_ss setloop (split_tac [expand_if]))));
 qed "mem_append";
 
 goal List.thy "x mem [x:xs.P(x)] = (x mem xs & P(x))";
-by(list_ind_tac "xs" 1);
+by (list.induct_tac "xs" 1);
 by(ALLGOALS(asm_simp_tac (list_ss setloop (split_tac [expand_if]))));
 qed "mem_filter";
 
 (** list_all **)
 
 goal List.thy "(Alls x:xs.True) = True";
-by(list_ind_tac "xs" 1);
+by (list.induct_tac "xs" 1);
 by(ALLGOALS(asm_simp_tac list_ss));
 qed "list_all_True";
 
 goal List.thy "list_all(p,xs@ys) = (list_all(p,xs) & list_all(p,ys))";
-by(list_ind_tac "xs" 1);
+by (list.induct_tac "xs" 1);
 by(ALLGOALS(asm_simp_tac list_ss));
 qed "list_all_conj";
 
 goal List.thy "(Alls x:xs.P(x)) = (!x. x mem xs --> P(x))";
-by(list_ind_tac "xs" 1);
+by (list.induct_tac "xs" 1);
 by(ALLGOALS(asm_simp_tac (list_ss setloop (split_tac [expand_if]))));
 by(fast_tac HOL_cs 1);
 qed "list_all_mem_conv";
 
 
-(** The functional "map" **)
-
-val map_simps = [Abs_map_NIL, Abs_map_CONS, 
-		 Rep_map_Nil, Rep_map_Cons, 
-		 map_Nil, map_Cons];
-val map_ss = list_free_ss addsimps map_simps;
-
-val [major,A_subset_sexp,minor] = goal List.thy 
-    "[| M: list(A);  A<=sexp;  !!z. z: A ==> f(g(z)) = z |] \
-\    ==> Rep_map(f, Abs_map(g,M)) = M";
-by (rtac (major RS list.induct) 1);
-by (ALLGOALS (asm_simp_tac(map_ss addsimps [sexp_A_I,sexp_ListA_I,minor])));
-qed "Abs_map_inverse";
-
-(*Rep_map_inverse is obtained via Abs_Rep_map and map_ident*)
-
 (** list_case **)
 
 goal List.thy
  "P(list_case(a,f,xs)) = ((xs=[] --> P(a)) & \
 \                         (!y ys. xs=y#ys --> P(f(y,ys))))";
-by(list_ind_tac "xs" 1);
+by (list.induct_tac "xs" 1);
 by(ALLGOALS(asm_simp_tac list_ss));
 by(fast_tac HOL_cs 1);
 qed "expand_list_case";
@@ -370,28 +96,22 @@
 (** Additional mapping lemmas **)
 
 goal List.thy "map(%x.x, xs) = xs";
-by (list_ind_tac "xs" 1);
-by (ALLGOALS (asm_simp_tac map_ss));
+by (list.induct_tac "xs" 1);
+by (ALLGOALS (asm_simp_tac list_ss));
 qed "map_ident";
 
 goal List.thy "map(f, xs@ys) = map(f,xs) @ map(f,ys)";
-by (list_ind_tac "xs" 1);
-by (ALLGOALS (asm_simp_tac (map_ss addsimps [append_Nil,append_Cons])));
+by (list.induct_tac "xs" 1);
+by (ALLGOALS (asm_simp_tac list_ss));
 qed "map_append";
 
 goalw List.thy [o_def] "map(f o g, xs) = map(f, map(g, xs))";
-by (list_ind_tac "xs" 1);
-by (ALLGOALS (asm_simp_tac map_ss));
+by (list.induct_tac "xs" 1);
+by (ALLGOALS (asm_simp_tac list_ss));
 qed "map_compose";
 
-goal List.thy "!!f. (!!x. f(x): sexp) ==> \
-\	Abs_map(g, Rep_map(f,xs)) = map(%t. g(f(t)), xs)";
-by (list_ind_tac "xs" 1);
-by(ALLGOALS(asm_simp_tac(map_ss addsimps
-       [Rep_map_type,list_sexp RS subsetD])));
-qed "Abs_Rep_map";
-
 val list_ss = list_ss addsimps
-  [mem_append, mem_filter, append_assoc, append_Nil2, map_ident,
+  [not_Cons_self, append_assoc, append_Nil2, mem_append, mem_filter,
+   map_ident, map_append, map_compose,
    list_all_True, list_all_conj];
 
--- a/List.thy	Fri Dec 02 16:09:49 1994 +0100
+++ b/List.thy	Fri Dec 02 16:13:34 1994 +0100
@@ -1,52 +1,32 @@
-(*  Title:      HOL/list
+(*  Title:      HOL/List.thy
     ID:         $Id$
-    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
-    Copyright   1993  University of Cambridge
+    Author:     Tobias Nipkow
+    Copyright   1994 TU Muenchen
 
-Definition of type 'a list by a least fixed point
+Definition of type 'a list as a datatype. This allows primrec to work.
 
-We use          list(A) == lfp(%Z. {NUMB(0)} <+> A <*> Z)
-and not         list    == lfp(%Z. {NUMB(0)} <+> range(Leaf) <*> Z)
-so that list can serve as a "functor" for defining other recursive types
 *)
 
-List = Sexp +
+List = Arith +
 
-types
-  'a list
-
-arities
-  list :: (term) term
-
+datatype 'a list = "[]" ("[]") | "#"('a,'a list) (infixr 65)
 
 consts
 
-  list      :: "'a item set => 'a item set"
-  Rep_list  :: "'a list => 'a item"
-  Abs_list  :: "'a item => 'a list"
-  NIL       :: "'a item"
-  CONS      :: "['a item, 'a item] => 'a item"
-  Nil       :: "'a list"
-  "#"       :: "['a, 'a list] => 'a list"                   	(infixr 65)
-  List_case :: "['b, ['a item, 'a item]=>'b, 'a item] => 'b"
-  List_rec  :: "['a item, 'b, ['a item, 'a item, 'b]=>'b] => 'b"
-  list_case :: "['b, ['a, 'a list]=>'b, 'a list] => 'b"
-  list_rec  :: "['a list, 'b, ['a, 'a list, 'b]=>'b] => 'b"
-  Rep_map   :: "('b => 'a item) => ('b list => 'a item)"
-  Abs_map   :: "('a item => 'b) => 'a item => 'b list"
   null      :: "'a list => bool"
   hd        :: "'a list => 'a"
   tl,ttl    :: "'a list => 'a list"
-  mem		:: "['a, 'a list] => bool"			(infixl 55)
+  mem       :: "['a, 'a list] => bool"			(infixl 55)
   list_all  :: "('a => bool) => ('a list => bool)"
   map       :: "('a=>'b) => ('a list => 'b list)"
-  "@"	    :: "['a list, 'a list] => 'a list"			(infixr 65)
+  "@"	    :: "['a list, 'a list] => 'a list"		(infixr 65)
   filter    :: "['a => bool, 'a list] => 'a list"
+  foldl     :: "[['b,'a] => 'b, 'b, 'a list] => 'b"
+  length    :: "'a list => nat"
 
+syntax
   (* list Enumeration *)
-
-  "[]"      :: "'a list"                            ("[]")
-  "@list"   :: "args => 'a list"                    ("[(_)]")
+  "@list"   :: "args => 'a list"                        ("[(_)]")
 
   (* Special syntax for list_all and filter *)
   "@Alls"	:: "[idt, 'a list, bool] => bool"	("(2Alls _:_./ _)" 10)
@@ -55,66 +35,42 @@
 translations
   "[x, xs]"     == "x#[xs]"
   "[x]"         == "x#[]"
-  "[]"          == "Nil"
-
-  "case xs of Nil => a | y#ys => b" == "list_case(a, %y ys.b, xs)"
 
   "[x:xs . P]"	== "filter(%x.P,xs)"
   "Alls x:xs.P"	== "list_all(%x.P,xs)"
 
-defs
-  (* Defining the Concrete Constructors *)
-  NIL_def       "NIL == In0(Numb(0))"
-  CONS_def      "CONS(M, N) == In1(M $ N)"
-
-inductive "list(A)"
-  intrs
-    NIL_I  "NIL: list(A)"
-    CONS_I "[| a: A;  M: list(A) |] ==> CONS(a,M) : list(A)"
-
-rules
-  (* Faking a Type Definition ... *)
-  Rep_list          "Rep_list(xs): list(range(Leaf))"
-  Rep_list_inverse  "Abs_list(Rep_list(xs)) = xs"
-  Abs_list_inverse  "M: list(range(Leaf)) ==> Rep_list(Abs_list(M)) = M"
-
-
-defs
-  (* Defining the Abstract Constructors *)
-  Nil_def       "Nil == Abs_list(NIL)"
-  Cons_def      "x#xs == Abs_list(CONS(Leaf(x), Rep_list(xs)))"
-
-  List_case_def "List_case(c, d) == Case(%x.c, Split(d))"
-
-  (* list Recursion -- the trancl is Essential; see list.ML *)
-
-  List_rec_def
-   "List_rec(M, c, d) == wfrec(trancl(pred_sexp), M, \
-\                         List_case(%g.c, %x y g. d(x, y, g(y))))"
-
-  list_rec_def
-   "list_rec(l, c, d) == \
-\   List_rec(Rep_list(l), c, %x y r. d(Inv(Leaf, x), Abs_list(y), r))"
-
-  (* Generalized Map Functionals *)
-
-  Rep_map_def "Rep_map(f, xs) == list_rec(xs, NIL, %x l r. CONS(f(x), r))"
-  Abs_map_def "Abs_map(g, M) == List_rec(M, Nil, %N L r. g(N)#r)"
-
-  null_def      "null(xs)            == list_rec(xs, True, %x xs r.False)"
-  hd_def        "hd(xs)              == list_rec(xs, @x.True, %x xs r.x)"
-  tl_def        "tl(xs)              == list_rec(xs, @xs.True, %x xs r.xs)"
-  (* a total version of tl: *)
-  ttl_def	"ttl(xs)             == list_rec(xs, [], %x xs r.xs)"
-
-  mem_def	"x mem xs            == \
-\		   list_rec(xs, False, %y ys r. if(y=x, True, r))"
-  list_all_def  "list_all(P, xs)     == list_rec(xs, True, %x l r. P(x) & r)"
-  map_def       "map(f, xs)          == list_rec(xs, [], %x l r. f(x)#r)"
-  append_def	"xs@ys               == list_rec(xs, ys, %x l r. x#r)"
-  filter_def	"filter(P,xs)        == \
-\                  list_rec(xs, [], %x xs r. if(P(x), x#r, r))"
-
-  list_case_def "list_case(a, f, xs) == list_rec(xs, a, %x xs r.f(x, xs))"
-
+primrec null list
+  null_Nil "null([]) = True"
+  null_Cons "null(x#xs) = False"
+primrec hd list
+  hd_Nil  "hd([]) = (@x.False)"
+  hd_Cons "hd(x#xs) = x"
+primrec tl list
+  tl_Nil  "tl([]) = (@x.False)"
+  tl_Cons "tl(x#xs) = xs"
+primrec ttl list
+  (* a "total" version of tl: *)
+  ttl_Nil  "ttl([]) = []"
+  ttl_Cons "ttl(x#xs) = xs"
+primrec "op mem" list
+  mem_Nil  "x mem [] = False"
+  mem_Cons "x mem (y#ys) = if(y=x, True, x mem ys)"
+primrec list_all list
+  list_all_Nil  "list_all(P,[]) = True"
+  list_all_Cons "list_all(P,x#xs) = (P(x) & list_all(P,xs))"
+primrec map list
+  map_Nil  "map(f,[]) = []"
+  map_Cons "map(f,x#xs) = f(x)#map(f,xs)"
+primrec "op @" list
+  append_Nil  "[] @ ys = ys"
+  append_Cons "(x#xs)@ys = x#(xs@ys)"
+primrec filter list
+  filter_Nil  "filter(P,[]) = []"
+  filter_Cons "filter(P,x#xs) = if(P(x), x#filter(P,xs), filter(P,xs))"
+primrec foldl list
+  foldl_Nil  "foldl(f,a,[]) = a"
+  foldl_Cons "foldl(f,a,x#xs) = foldl(f, f(a,x), xs)"
+primrec length list
+  length_Nil  "length([]) = 0"
+  length_Cons "length(x#xs) = Suc(length(xs))"
 end
--- a/ROOT.ML	Fri Dec 02 16:09:49 1994 +0100
+++ b/ROOT.ML	Fri Dec 02 16:13:34 1994 +0100
@@ -78,6 +78,7 @@
 use_thy "Inductive";
 
 use_thy "Finite";
+use_thy "Sexp";
 use_thy "List";
 
 init_pps ();
--- a/thy_syntax.ML	Fri Dec 02 16:09:49 1994 +0100
+++ b/thy_syntax.ML	Fri Dec 02 16:13:34 1994 +0100
@@ -139,7 +139,7 @@
   (*parsers*)
   val tvars = type_args >> map (cat "dtVar");
   val typ =
-    tvars -- name >> (cat "dtTyp" o mk_pair o apfst mk_list) ||
+    tvars -- (ident>>quote) >> (cat "dtTyp" o mk_pair o apfst mk_list) ||
     type_var >> cat "dtVar";
   val opt_typs = optional ("(" $$-- list1 typ --$$ ")") [];
   val constructor = name -- opt_typs -- opt_mixfix;