src/Pure/Syntax/ROOT.ML
changeset 18 c9ec452ff08f
parent 0 a5a9c433f639
child 42 d981488bda7b
--- a/src/Pure/Syntax/ROOT.ML	Fri Oct 01 13:26:22 1993 +0100
+++ b/src/Pure/Syntax/ROOT.ML	Mon Oct 04 15:30:49 1993 +0100
@@ -1,40 +1,48 @@
-(*  Title:      Pure/Syntax/ROOT
+(*  Title:      Pure/Syntax/ROOT.ML
     ID:         $Id$
     Author:     Tobias Nipkow and Markus Wenzel, TU Muenchen
 
-This file builds the syntax module. It assumes the standard Isabelle
-environment.
+This file builds the syntax module.
 *)
 
+use "lib.ML";   (* FIXME *)
+
+use "pretty.ML";
+
 use "ast.ML";
 use "xgram.ML";
+use "lexicon.ML";
 use "extension.ML";
-use "lexicon.ML";
 use "parse_tree.ML";
-use "earley0A.ML";
+use "parser.ML";
+use "earley0A.ML";        (* FIXME *)
 use "type_ext.ML";
 use "sextension.ML";
-use "pretty.ML";
 use "printer.ML";
 use "syntax.ML";
 
-structure Ast = AstFun();
+structure Pretty = PrettyFun();
+
+structure Ast = AstFun(Pretty);
 structure XGram = XGramFun(Ast);
-structure Extension = ExtensionFun(XGram);
-structure Lexicon = LexiconFun(Extension);
-structure ParseTree = ParseTreeFun(structure Lexicon = Lexicon and Ast = Ast);
-structure Earley = EarleyFun(structure XGram = XGram and ParseTree = ParseTree);
+structure Lexicon = LexiconFun();
+structure Extension = ExtensionFun(structure XGram = XGram and Lexicon = Lexicon);
+structure ParseTree = ParseTreeFun(structure Ast = Ast and Lexicon = Lexicon);
+structure Parser = ParserFun(structure Symtab = Symtab and XGram = XGram
+  and ParseTree = ParseTree);
+structure Earley = EarleyFun(structure Symtab = Symtab and XGram = XGram
+  and ParseTree = ParseTree);     (* FIXME *)
 structure TypeExt = TypeExtFun(structure Extension = Extension
   and Lexicon = Lexicon);
 structure SExtension = SExtensionFun(structure TypeExt = TypeExt
   and Lexicon = Lexicon);
-structure Pretty = PrettyFun();
 structure Printer = PrinterFun(structure Symtab = Symtab and Lexicon = Lexicon
   and TypeExt = TypeExt and SExtension = SExtension and Pretty = Pretty);
 structure Syntax = SyntaxFun(structure Symtab = Symtab and TypeExt = TypeExt
   and Parser = Earley and SExtension = SExtension and Printer = Printer);
+     (* FIXME  ^^^^^^ *)
 
-(*Basic_Syntax has the most important primitives, which are made pervasive*)
+(*BasicSyntax has the most important primitives, which are made pervasive*)
 signature BASIC_SYNTAX = sig include SEXTENSION0 include PRINTER0 end;
-structure Basic_Syntax: BASIC_SYNTAX = Syntax;
+structure BasicSyntax: BASIC_SYNTAX = Syntax;