src/ZF/IMP/Evalc0.thy
author paulson
Wed, 02 Apr 1997 15:39:44 +0200
changeset 2877 6476784dba1c
parent 482 3a4e092ba69c
permissions -rw-r--r--
Converted to call blast_tac. Proves theorems in ZF.thy to make that theory usable again

(*  Title: 	ZF/IMP/Evalc0.thy
    ID:         $Id$
    Author: 	Heiko Loetzbeyer & Robert Sandner, TUM
    Copyright   1994 TUM
*)

Evalc0 = Evalb + Com + Assign +

consts
       evalc    :: "i"
       "@evalc" :: "[i,i,i] => o"   ("<_,_>/ -c-> _")

translations
       "<ce,sig> -c-> s" == "<ce,sig,s> : evalc"
end