src/HOL/MicroJava/JVM/JVMInstructions.thy
author kleing
Sun, 07 Jan 2001 18:43:13 +0100
changeset 10812 ead84e90bfeb
parent 10591 6d6cb845e841
child 10897 697fab84709e
permissions -rw-r--r--
merged semilattice orders with <=' from Convert.thy (now defined in JVMType.thy)

(*  Title:      HOL/MicroJava/JVM/JVMInstructions.thy
    ID:         $Id$
    Author:     Gerwin Klein
    Copyright   2000 Technische Universitaet Muenchen
*)

header {* Instructions of the JVM *}


theory JVMInstructions = JVMState:


datatype 
  instr = Load nat                  (* load from local variable *)
        | Store nat                 (* store into local variable *)
        | Bipush int                (* push int *)
        | Aconst_null               (* push null *)
        | New cname                 (* create object *)
        | Getfield vname cname	  	(* Fetch field from object *)
        | Putfield vname cname		  (* Set field in object     *)
        | Checkcast cname	          (* Check whether object is of given type *)
        | Invoke cname mname "(ty list)"  (* inv. instance meth of an object *)
        | Return                    (* return from method *)
        | Pop                       (* pop top element from opstack *)
        | Dup                       (* duplicate top element of opstack *)
        | Dup_x1                    (* duplicate next to top element *)
        | Dup_x2                    (* duplicate 3rd element *)
        | Swap                      (* swap top and next to top element *)
        | IAdd                      (* integer addition *)
        | Goto int                  (* goto relative address *)
        | Ifcmpeq int               (* Branch if int/ref comparison succeeds *)


types
  bytecode = "instr list"
  jvm_prog = "(nat \<times> nat \<times> bytecode) prog" 

end