src/HOL/MicroJava/JVM/Store.thy
author oheimb
Fri, 14 Jul 2000 16:32:51 +0200
changeset 9346 297dcbf64526
parent 8038 a13c3b80d3d4
child 10042 7164dc0d24d8
permissions -rw-r--r--
re-structuring MicroJava; added Example; corrected := syntax; simplfied cast

(*  Title:      HOL/MicroJava/JVM/Store.thy
    ID:         $Id$
    Author:     Cornelia Pusch
    Copyright   1999 Technische Universitaet Muenchen

The store.

The JVM builds on many notions already defined in Java.
Conform provides notions for the type safety proof of the Bytecode Verifier.
*)

Store = Conform +  

constdefs
 newref :: "('a \\<leadsto> 'b) \\<Rightarrow> 'a"
 "newref s \\<equiv> \\<epsilon>v. s v = None"

end