author | oheimb |
Thu, 18 Jan 2001 17:38:56 +0100 | |
changeset 10925 | 5ffe7ed8899a |
parent 10897 | 697fab84709e |
child 12519 | a955fe2879ba |
permissions | -rw-r--r-- |
9376 | 1 |
(* Title: HOL/MicroJava/JVM/JVMInstructions.thy |
2 |
ID: $Id$ |
|
3 |
Author: Gerwin Klein |
|
4 |
Copyright 2000 Technische Universitaet Muenchen |
|
5 |
*) |
|
6 |
||
10057 | 7 |
header {* Instructions of the JVM *} |
8 |
||
9 |
||
10 |
theory JVMInstructions = JVMState: |
|
9376 | 11 |
|
9550
19a6d1289f5e
Invoke instruction gets fully qualified method name (class+name+sig) as
kleing
parents:
9376
diff
changeset
|
12 |
|
9376 | 13 |
datatype |
14 |
instr = Load nat (* load from local variable *) |
|
15 |
| Store nat (* store into local variable *) |
|
10897
697fab84709e
removed instructions Aconst_null+Bipush, introduced LitPush
kleing
parents:
10591
diff
changeset
|
16 |
| LitPush val (* push a literal (constant) *) |
9376 | 17 |
| New cname (* create object *) |
18 |
| Getfield vname cname (* Fetch field from object *) |
|
19 |
| Putfield vname cname (* Set field in object *) |
|
20 |
| Checkcast cname (* Check whether object is of given type *) |
|
9550
19a6d1289f5e
Invoke instruction gets fully qualified method name (class+name+sig) as
kleing
parents:
9376
diff
changeset
|
21 |
| Invoke cname mname "(ty list)" (* inv. instance meth of an object *) |
10057 | 22 |
| Return (* return from method *) |
23 |
| Pop (* pop top element from opstack *) |
|
24 |
| Dup (* duplicate top element of opstack *) |
|
25 |
| Dup_x1 (* duplicate next to top element *) |
|
26 |
| Dup_x2 (* duplicate 3rd element *) |
|
27 |
| Swap (* swap top and next to top element *) |
|
28 |
| IAdd (* integer addition *) |
|
29 |
| Goto int (* goto relative address *) |
|
9376 | 30 |
| Ifcmpeq int (* Branch if int/ref comparison succeeds *) |
31 |
||
32 |
||
33 |
types |
|
34 |
bytecode = "instr list" |
|
10591 | 35 |
jvm_prog = "(nat \<times> nat \<times> bytecode) prog" |
9376 | 36 |
|
37 |
end |