src/ZF/IMP/Assign.thy
author clasohm
Thu, 19 Oct 1995 13:25:03 +0100
changeset 1287 84f44b84d584
parent 482 3a4e092ba69c
permissions -rw-r--r--
corrected spelling of title (to test new CVS loginfo)

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

Assign = Aexp +
consts 
	"assign" :: "[i,i,i] => i"	("_[_'/_]" [900,0,0] 900)

rules 
	assign_def	"sigma[m/x] == lam y:loc. if(y=x,m,sigma`y)"
end