src/Pure/Isar/ROOT.ML
author ballarin
Fri, 15 Apr 2005 12:00:00 +0200
changeset 15735 953f188e16c6
parent 15709 f04c3d668c65
child 15827 5fdf2d8dab9c
permissions -rw-r--r--
Removed most of the atp interface from Pure.

(*  Title:      Pure/Isar/ROOT.ML
    ID:         $Id$
    Author:     Markus Wenzel, TU Muenchen

Isar -- Intelligible Semi-Automated Reasoning for Isabelle.
*)

(*basic proof engine*)
use "object_logic.ML";
use "auto_bind.ML";
use "rule_cases.ML";
use "proof_context.ML";
use "proof_data.ML";
use "context_rules.ML";
use "args.ML";
use "attrib.ML";
use "locale.ML";
use "proof.ML";
use "proof_history.ML";
use "net_rules.ML";
use "induct_attrib.ML";
use "method.ML";

(*derived proof elements*)
use "calculation.ML";
use "obtain.ML";
use "skip_proof.ML";

(*outer syntax*)
(*use "outer_lex.ML";*)   (*see ../Thy/ROOT.ML*)
use "antiquote.ML";
use "outer_parse.ML";

(*toplevel environment*)
use "toplevel.ML";
use "isar_output.ML";
use "session.ML";

(*theory syntax*)
use "thy_header.ML";
use "outer_syntax.ML";

(*theory and proof operations*)
use "isar_thy.ML";
use "constdefs.ML";
use "isar_cmd.ML";
use "isar_syn.ML";

(*main ML interfaces*)
use "isar.ML";

structure PureIsar =
struct
  structure ObjectLogic = ObjectLogic;
  structure AutoBind = AutoBind;
  structure RuleCases = RuleCases;
  structure ProofContext = ProofContext;
  structure ContextRules = ContextRules;
  structure Args = Args;
  structure Attrib = Attrib;
  structure Locale = Locale;
  structure Proof = Proof;
  structure ProofHistory = ProofHistory;
  structure NetRules = NetRules;
  structure InductAttrib = InductAttrib;
  structure Method = Method;
  structure Calculation = Calculation;
  structure Obtain = Obtain;
  structure SkipProof = SkipProof;
  structure OuterLex = OuterLex;
  structure Antiquote = Antiquote;
  structure OuterParse = OuterParse;
  structure Toplevel = Toplevel;
  structure Session = Session;
  structure IsarOutput = IsarOutput;
  structure ThyHeader = ThyHeader;
  structure OuterSyntax = OuterSyntax;
  structure IsarThy = IsarThy;
  structure Constdefs = Constdefs;
  structure IsarCmd = IsarCmd;
  structure IsarSyn = IsarSyn;
  structure Isar = Isar;
end;