src/Pure/ROOT.ML
author aspinall
Thu, 21 Dec 2006 08:42:53 +0100
changeset 21890 3fb9762ba701
parent 21889 682dbe947862
child 21930 918fb0fb5c72
permissions -rw-r--r--
Disable new Proof General code until SML/NJ compile fixed.

(*  Title:      Pure/ROOT.ML
    ID:         $Id$

Pure Isabelle.
*)

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

(*if true then some tools will OMIT some proofs*)
val quick_and_dirty = ref false;

print_depth 10;

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

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

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

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

(*core of tactical proof system*)
use "envir.ML";
use "logic.ML";
use "type_infer.ML";
use "consts.ML";
use "sign.ML";
use "pattern.ML";
use "unify.ML";
use "net.ML";
use "defs.ML";
use "theory.ML";
use "proofterm.ML";
use "thm.ML";
use "morphism.ML";
use "fact_index.ML";
use "pure_thy.ML";
use "display.ML";
use "drule.ML";
use "variable.ML";
use "tctical.ML";
use "search.ML";
use "tactic.ML";
use "meta_simplifier.ML";
use "conjunction.ML";
use "assumption.ML";
use "goal.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*)
use "Isar/outer_lex.ML";

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

use "Proof/extraction.ML";

cd "Tools"; use "ROOT.ML"; cd "..";

(*configuration for Proof General*)
(* Next line is OLD CODE: in case you have problems, uncomment next line and 
   comment out line after. Please report any problems to da@inf.ed.ac.uk.
   Plan is to remove old code very soon. *)
(use |> setmp proofs 1 |> setmp quick_and_dirty true) "proof_general.ML"; 
(* Next line is NEW CODE.  Currently broken on SMLNJ. *)
(* cd "ProofGeneral"; use "ROOT.ML"; cd "..";  new code *)

use_thy "Pure";
structure Pure = struct val thy = theory "Pure" end;

Context.add_setup
 (Theory.del_modesyntax Syntax.default_mode Syntax.appl_syntax #>
  Theory.add_syntax Syntax.applC_syntax);
use_thy "CPure";
structure CPure = struct val thy = theory "CPure" end;

(*final ML setup*)
use "install_pp.ML";
val use = ThyInfo.use;
val cd = File.cd o Path.explode;
ml_prompts "ML> " "ML# ";

proofs := 0;