src/HOLCF/HOLCF.thy
author huffman
Tue, 26 Jul 2005 18:29:59 +0200
changeset 16922 2128ac2aa5db
parent 16841 228d663cc9b3
child 17612 5b37787d2d9e
permissions -rw-r--r--
brought ML files up to date with new lemmas

(*  Title:      HOLCF/HOLCF.thy
    ID:         $Id$
    Author:     Franz Regensburger

HOLCF -- a semantic extension of HOL by the LCF logic.
*)

theory HOLCF
imports Sprod Ssum Up Lift Discrete One Tr Domain
uses
  "holcf_logic.ML"
  "cont_consts.ML"
  "domain/library.ML"
  "domain/syntax.ML"
  "domain/axioms.ML"
  "domain/theorems.ML"
  "domain/extender.ML"
  "domain/interface.ML"
  "adm_tac.ML"

begin

ML_setup {*
  simpset_ref() := simpset() addSolver
     (mk_solver "adm_tac" (fn thms => (adm_tac (cut_facts_tac thms THEN' cont_tacRs))));
*}

end