author | haftmann |
Mon, 01 Mar 2010 13:40:23 +0100 | |
changeset 35416 | d8d7d1b785af |
parent 28524 | 644b62cf678f |
child 44035 | 322d1657c40c |
permissions | -rw-r--r-- |
12951 | 1 |
(* Title: HOL/MicroJava/J/SystemClasses.thy |
2 |
Author: Gerwin Klein |
|
3 |
Copyright 2002 Technische Universitaet Muenchen |
|
4 |
*) |
|
5 |
||
6 |
header {* \isaheader{System Classes} *} |
|
7 |
||
16417 | 8 |
theory SystemClasses imports Decl begin |
12951 | 9 |
|
10 |
text {* |
|
11 |
This theory provides definitions for the @{text Object} class, |
|
12 |
and the system exceptions. |
|
13 |
*} |
|
14 |
||
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
28524
diff
changeset
|
15 |
definition ObjectC :: "'c cdecl" where |
28524 | 16 |
"ObjectC \<equiv> (Object, (undefined,[],[]))" |
12951 | 17 |
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
28524
diff
changeset
|
18 |
definition NullPointerC :: "'c cdecl" where |
12951 | 19 |
"NullPointerC \<equiv> (Xcpt NullPointer, (Object,[],[]))" |
20 |
||
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
28524
diff
changeset
|
21 |
definition ClassCastC :: "'c cdecl" where |
12951 | 22 |
"ClassCastC \<equiv> (Xcpt ClassCast, (Object,[],[]))" |
23 |
||
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
28524
diff
changeset
|
24 |
definition OutOfMemoryC :: "'c cdecl" where |
12951 | 25 |
"OutOfMemoryC \<equiv> (Xcpt OutOfMemory, (Object,[],[]))" |
26 |
||
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
28524
diff
changeset
|
27 |
definition SystemClasses :: "'c cdecl list" where |
12951 | 28 |
"SystemClasses \<equiv> [ObjectC, NullPointerC, ClassCastC, OutOfMemoryC]" |
29 |
||
30 |
end |