src/Pure/ROOT.ML
author kleing
Sat, 30 Apr 2005 14:18:36 +0200
changeset 15900 d6156cb8dc2e
parent 15825 1576f9d3ffae
child 16108 cf468b93a02e
permissions -rw-r--r--
fixed typo

(*  Title:      Pure/ROOT.ML
    ID:         $Id$
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
    Copyright   1993  University of Cambridge

Root file for Pure Isabelle.
*)

val banner = "Pure Isabelle";
val version = "Isabelle repository version";    (*filled in automatically!*)


print_depth 10;

(*fake hiding of private structures*)
structure Hidden = struct end;

(*basic tools*)
use "library.ML";
cd "General"; use "ROOT.ML"; cd "..";

(*fundamental structures*)
use "term.ML";
use "General/pretty.ML";
use "sorts.ML";
use "type.ML";

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

(*core of tactical proof system*)
use "type_infer.ML";
use "sign.ML";
use "envir.ML";
use "pattern.ML";
use "unify.ML";
use "net.ML";
use "logic.ML";
use "theory.ML";
use "theory_data.ML";
use "context.ML";
use "proofterm.ML";
use "thm.ML";
use "display.ML";
use "fact_index.ML";
use "pure_thy.ML";
use "drule.ML";
use "tctical.ML";
use "search.ML";
use "meta_simplifier.ML";
use "tactic.ML";

(*proof term operations*)
use "Proof/reconstruct.ML";
use "Proof/proof_syntax.ML";
use "Proof/proof_rewrite_rules.ML";
use "Proof/proofchecker.ML";

(*theory auto loader database*)
use "Thy/thy_load.ML";
use "Thy/thy_info.ML";

(*theory syntax -- old format*)
use "Thy/thy_scan.ML";
use "Thy/thy_parse.ML";
use "Thy/thy_syn.ML";

(*theory syntax -- new format*)
use "Isar/outer_lex.ML";

(*theory presentation*)
use "Thy/html.ML";
use "Thy/latex.ML";
use "Thy/present.ML";
use "Thy/thm_deps.ML";

(*theorem database ML interface*)
use "Thy/thm_database.ML";

(*the Isar subsystem*)
cd "Isar"; use "ROOT.ML"; cd "..";

use "axclass.ML";
use "codegen.ML";
use "Proof/extraction.ML";

(*old goal package -- obsolete*)
use "goals.ML";

(*the IsaPlanner subsystem*)
cd "IsaPlanner"; use "ROOT.ML"; cd "..";

(*configuration for Proof General*)
use "proof_general.ML";

(*the Pure theories*)
use_thy "Pure"; structure Pure = struct val thy = theory "Pure" end;
use_thy "CPure"; structure CPure = struct val thy = theory "CPure" end;


(*several object-logics declare theories that hide basis library structures*)
structure BasisLibrary =
struct
  structure List = List;
  structure Option = Option;
  structure Bool = Bool;
  structure String = String;
  structure Int = Int;
  structure Real = Real;
end;

use "install_pp.ML";

val use = ThyInfo.use;
val cd = File.cd o Path.unpack;

ml_prompts "ML> " "ML# ";

proofs := 0;