src/Pure/ROOT.ML
author wenzelm
Thu, 28 May 1998 12:22:05 +0200
changeset 4978 f14ec8ec1db1
parent 4962 e9217cb15b42
child 4986 d4f257d3445a
permissions -rw-r--r--
version under control of Admin/makedist; ml_prompts;

(*  Title:      Pure/ROOT.ML
    ID:         $Id$
    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 ML database (see also README).
*)

val banner = "Pure Isabelle";
val version = "Internal working version of Isabelle";

print_depth 1;
ml_prompts "> " "# ";


(*basic utils*)
use "library.ML";
use "table.ML";
use "seq.ML";
use "name_space.ML";
use "term.ML";

(*inner syntax module*)
cd "Syntax";
use "ROOT.ML";
cd "..";

(*main system*)
use "sorts.ML";
use "type_infer.ML";
use "type.ML";
use "sign.ML";
use "envir.ML";
use "pattern.ML";
use "unify.ML";
use "net.ML";
use "logic.ML";
use "theory.ML";
use "thm.ML";
use "display.ML";
use "attribute.ML";
use "pure_thy.ML";
use "deriv.ML";
use "drule.ML";
use "tctical.ML";
use "search.ML";
use "tactic.ML";
use "goals.ML";
use "axclass.ML";

(*theory parser and loader*)
cd "Thy";
use "ROOT.ML";
cd "..";

use "install_pp.ML";

(*several object-logics declare theories named List or Option, hiding
  the eponymous basis library structures*)
structure BasisLibrary =
struct
  structure List = List and Option = Option;
end;

open Use;

print_depth 8;