src/HOLCF/ROOT.ML
author berghofe
Fri, 02 May 1997 16:41:35 +0200
changeset 3098 a31170b67367
parent 2394 91d8abf108be
child 3511 da4dd8b7ced4
permissions -rw-r--r--
Updated to LaTeX 2e

(*  Title:      HOLCF/ROOT
    ID:         $Id$
    Author:     Franz Regensburger
    Copyright   1993 Technische Universitaet Muenchen

ROOT file for the conservative extension of HOL by the LCF logic.
Should be executed in subdirectory HOLCF.
*)

val banner = "HOLCF with sections axioms,ops,domain,generated";
writeln banner;

print_depth 1;

use_thy "HOLCF";

(* install sections: axioms, ops *)

use "ax_ops/holcflogic.ML";
use "ax_ops/thy_axioms.ML";
use "ax_ops/thy_ops.ML";
use "ax_ops/thy_syntax.ML";


(* install sections: domain, generated *)

use "domain/library.ML";
use "domain/syntax.ML";
use "domain/axioms.ML";
use "domain/theorems.ML";
use "domain/extender.ML";
use "domain/interface.ML";

init_thy_reader();
init_pps ();

fun qed_spec_mp name =
  let val thm = normalize_thm [RSspec,RSmp] (result())
  in bind_thm(name, thm) end;

print_depth 10;  

val HOLCF_build_completed = (); (*indicate successful build*)