src/Pure/ROOT.ML
changeset 19 929ad32d63fc
parent 2 c67f44be880f
child 73 075db6ac7f2f
--- a/src/Pure/ROOT.ML	Mon Oct 04 15:30:49 1993 +0100
+++ b/src/Pure/ROOT.ML	Mon Oct 04 15:36:31 1993 +0100
@@ -1,16 +1,14 @@
-(*  Title: 	ROOT
+(*  Title:      Pure/ROOT.ML
     ID:         $Id$
-    Author: 	Lawrence C Paulson, Cambridge University Computer Laboratory
+    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1993  University of Cambridge
 
-Root file for pure Isabelle: all modules in proper order for loading
-Loads pure Isabelle into an empty database.
+Root file for Pure Isabelle: all modules in proper order for loading.
+Loads pure Isabelle into an empty database (see also Makefile).
 
-To build system, use Makefile (Poly/ML) or Makefile.NJ (New Jersey ML)
-
-TO DO: 
+TO DO:
 instantiation of theorems can lead to inconsistent sorting of type vars if
-'a::S is already present and 'a::T is introduced. 
+'a::S is already present and 'a::T is introduced.
 *)
 
 val banner = "Pure Isabelle";
@@ -22,13 +20,18 @@
 use "term.ML";
 use "symtab.ML";
 
-(*Used for building the parser generator*)
 structure Symtab = SymtabFun();
+
+(*Syntax module*)
 cd "Syntax";
 use "ROOT.ML";
 cd "..";
 
-print_depth 1;
+(*Theory parser*)
+cd "Thy";
+use "ROOT.ML";
+cd "..";
+
 use "type.ML";
 use "sign.ML";
 use "sequence.ML";
@@ -50,23 +53,20 @@
 structure Envir = EnvirFun();
 structure Pattern = PatternFun(structure Sign=Sign and Envir=Envir);
 structure Unify = UnifyFun(structure Sign=Sign and Envir=Envir
-			   and Sequence=Sequence and Pattern=Pattern);
+                           and Sequence=Sequence and Pattern=Pattern);
 structure Net = NetFun();
 structure Logic = LogicFun(structure Unify=Unify and Net=Net);
 structure Thm = ThmFun(structure Logic=Logic and Unify=Unify and Net=Net
-                             and Pattern=Pattern); 
+                             and Pattern=Pattern);
 structure Drule = DruleFun(structure Logic=Logic and Thm=Thm);
 structure Tactical = TacticalFun(structure Logic=Logic and Drule=Drule);
-structure Tactic = TacticFun(structure Logic=Logic and Drule=Drule 
-			     and Tactical=Tactical and Net=Net);
+structure Tactic = TacticFun(structure Logic=Logic and Drule=Drule
+                             and Tactical=Tactical and Net=Net);
 structure Goals = GoalsFun(structure Logic=Logic and Drule=Drule
-			   and Tactic=Tactic and Pattern=Pattern);
-open Basic_Syntax Thm Drule Tactical Tactic Goals;
+                           and Tactic=Tactic and Pattern=Pattern);
+open BasicSyntax Thm Drule Tactical Tactic Goals;
 
 structure Pure = struct val thy = pure_thy end;
 
-(* Theory parser *)
-cd "Thy";
-use "ROOT.ML";
-cd "..";
+use "install_pp.ML";