8011
|
1 |
(* Title: HOL/MicroJava/J/Term.thy
|
|
2 |
ID: $Id$
|
|
3 |
Author: David von Oheimb
|
|
4 |
Copyright 1999 Technische Universitaet Muenchen
|
|
5 |
|
|
6 |
Java expressions and statements
|
|
7 |
*)
|
|
8 |
|
|
9 |
Term = Type +
|
|
10 |
|
|
11 |
types loc (* locations, i.e. abstract references on objects *)
|
|
12 |
arities loc :: term
|
|
13 |
|
|
14 |
datatype val_ (* name non 'val' because of clash with ML token *)
|
|
15 |
= Unit (* dummy result value of void methods *)
|
|
16 |
| Null (* null reference *)
|
|
17 |
| Bool bool (* Boolean value *)
|
|
18 |
| Intg int (* integer value *) (* Name Intg instead of Int because
|
|
19 |
of clash with HOL/Set.thy *)
|
|
20 |
| Addr loc (* addresses, i.e. locations of objects *)
|
|
21 |
types val = val_
|
|
22 |
translations "val" <= (type) "val_"
|
|
23 |
|
9240
|
24 |
datatype binop = Eq | Add (* function codes for binary operation *)
|
|
25 |
|
8011
|
26 |
datatype expr
|
|
27 |
= NewC cname (* class instance creation *)
|
|
28 |
| Cast ty expr (* type cast *)
|
|
29 |
| Lit val (* literal value, also references *)
|
9240
|
30 |
| BinOp binop expr expr (* binary operation *)
|
8011
|
31 |
| LAcc vname (* local (incl. parameter) access *)
|
|
32 |
| LAss vname expr (* local assign *) ("_\\<Colon>=_" [ 90,90]90)
|
|
33 |
| FAcc cname expr vname (* field access *) ("{_}_.._"[10,90,99 ]90)
|
|
34 |
| FAss cname expr vname
|
|
35 |
expr (* field ass. *)("{_}_.._\\<in>=_"[10,90,99,90]90)
|
|
36 |
| Call expr mname (ty list) (expr list)(* method call*)
|
|
37 |
("_.._'({_}_')" [90,99,10,10] 90)
|
|
38 |
and stmt
|
|
39 |
= Skip (* empty statement *)
|
|
40 |
| Expr expr (* expression statement *)
|
|
41 |
| Comp stmt stmt ("_;; _" [ 61,60]60)
|
|
42 |
| Cond expr stmt stmt ("If'(_') _ Else _" [ 80,79,79]70)
|
|
43 |
| Loop expr stmt ("While'(_') _" [ 80,79]70)
|
|
44 |
end
|