src/Pure/Syntax/syntax.ML
changeset 46989 88b0a8052c75
parent 46849 36f392239b6a
child 49674 dbadb4d03cbc
--- a/src/Pure/Syntax/syntax.ML	Sat Mar 17 12:52:40 2012 +0100
+++ b/src/Pure/Syntax/syntax.ML	Sat Mar 17 13:06:23 2012 +0100
@@ -43,6 +43,7 @@
   val read_typ: Proof.context -> string -> typ
   val read_term: Proof.context -> string -> term
   val read_prop: Proof.context -> string -> term
+  val read_typs: Proof.context -> string list -> typ list
   val read_terms: Proof.context -> string list -> term list
   val read_props: Proof.context -> string list -> term list
   val read_sort_global: theory -> string -> sort
@@ -265,11 +266,17 @@
 (* read = parse + check *)
 
 fun read_sort ctxt = parse_sort ctxt #> check_sort ctxt;
-fun read_typ ctxt = parse_typ ctxt #> singleton (check_typs ctxt);
+
+fun read_typs ctxt =
+  grouped 10 (Par_List.map_name "Syntax.read_typs") (parse_typ ctxt) #> check_typs ctxt;
 
-fun read_terms ctxt = Par_List.map_name "Syntax.read_terms" (parse_term ctxt) #> check_terms ctxt;
-fun read_props ctxt = Par_List.map_name "Syntax.read_props" (parse_prop ctxt) #> check_props ctxt;
+fun read_terms ctxt =
+  grouped 10 (Par_List.map_name "Syntax.read_terms") (parse_term ctxt) #> check_terms ctxt;
 
+fun read_props ctxt =
+  grouped 10 (Par_List.map_name "Syntax.read_props") (parse_prop ctxt) #> check_props ctxt;
+
+val read_typ = singleton o read_typs;
 val read_term = singleton o read_terms;
 val read_prop = singleton o read_props;