ROOT.ML
author lcp
Thu, 25 Aug 1994 11:01:45 +0200
changeset 128 89669c58e506
parent 119 93dc86ccee28
child 129 0bba840aa07c
permissions -rw-r--r--
INSTALLATION OF INDUCTIVE DEFINITIONS HOL/Sexp, List, LList: updated for inductive defs; streamlined proofs HOL/List, Subst/UTerm, ex/Simult, ex/Term: updated refs to Sexp intr rules HOL/Univ/diag_eqI: new HOL/intr_elim: now checks that the inductive name does not clash with existing theory names HOL/Sum: now type + is an infixr, to agree with type * HOL/Set: added Pow and the derived rules PowI, PowD, Pow_bottom, Pow_top HOL/Fun/set_cs: now includes Pow rules HOL/mono/Pow_mono: new HOL/Makefile: now has Inductive.thy,.ML and ex/Acc.thy,.ML HOL/Sexp,List,LList,ex/Term: converted as follows node *set -> item Sexp -> sexp LList_corec -> <self> LList_ -> llist_ LList\> -> llist List_case -> <self> List_rec -> <self> List_ -> list_ List\> -> list Term_rec -> <self> Term_ -> term_ Term\> -> term

(*  Title: 	HOL/ROOT
    ID:         $Id$
    Author: 	Tobias Nipkow
    Copyright   1993  University of Cambridge

Adds Classical Higher-order Logic to a database containing pure Isabelle. 
Should be executed in the subdirectory HOL.
*)

val banner = "Higher-Order Logic";
writeln banner;

init_thy_reader();

print_depth 1;  
eta_contract := true;
use_thy "HOL";
use "../Provers/hypsubst.ML";
use "../Provers/classical.ML";
use "../Provers/simplifier.ML";
use "../Provers/splitter.ML";
use "../Provers/ind.ML";

(** Applying HypsubstFun to generate hyp_subst_tac **)

structure Hypsubst_Data =
  struct
  (*Take apart an equality judgement; otherwise raise Match!*)
  fun dest_eq (Const("Trueprop",_) $ (Const("op =",_)  $ t $ u)) = (t,u);

  val imp_intr = impI

  (*etac rev_cut_eq moves an equality to be the last premise. *)
  val rev_cut_eq = prove_goal HOL.thy "[| a=b;  a=b ==> R |] ==> R"
   (fn prems => [ REPEAT(resolve_tac prems 1) ]);

  val rev_mp = rev_mp
  val subst = subst 
  val sym = sym
  end;

structure Hypsubst = HypsubstFun(Hypsubst_Data);
open Hypsubst;

(** Applying ClassicalFun to create a classical prover **)
structure Classical_Data = 
  struct
  val sizef = size_of_thm
  val mp = mp
  val not_elim = notE
  val swap = swap
  val hyp_subst_tacs=[hyp_subst_tac]
  end;

structure Classical = ClassicalFun(Classical_Data);
open Classical;

(*Propositional rules*)
val prop_cs = empty_cs addSIs [refl,TrueI,conjI,disjCI,impI,notI,iffI]
                       addSEs [conjE,disjE,impCE,FalseE,iffE];

(*Quantifier rules*)
val HOL_cs = prop_cs addSIs [allI] addIs [exI,ex1I] 
                     addSEs [exE,ex1E] addEs [allE];

val HOL_dup_cs = prop_cs addSIs [allI] addIs [exCI,ex1I] 
                         addSEs [exE,ex1E] addEs [all_dupE];

structure HOL_Induction = InductionFun(struct val spec=spec end);
open HOL_Induction;

use     "simpdata.ML";
use_thy "Ord";
use_thy "subset";
use_thy "equalities";
use_thy "Prod";
use_thy "Sum";
use_thy "Gfp";
use_thy "Nat";

(* Add user sections *)
use     "Datatype.ML";
use "../Pure/section_utils.ML";
use "ind_syntax.ML";
use "add_ind_def.ML";
use "intr_elim.ML";
use "indrule.ML";
use "Inductive.ML";

structure ThySyn = ThySynFun
 (val user_keywords = ["|", "datatype", "primrec",
		       "inductive", "coinductive", "intrs", 
		       "monos", "con_defs"]
  and user_sections = [("inductive",  inductive_decl ""),
		       ("coinductive",  inductive_decl "Co"),
		       ("datatype",  datatype_decls),
		       ("primrec", primrec_decl)]);
init_thy_reader ();

use_thy "Finite";
use_thy "LList";

use "../Pure/install_pp.ML";
print_depth 8;  

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