author | blanchet |
Wed, 15 Dec 2010 18:10:32 +0100 | |
changeset 41171 | 043f8dc3b51f |
parent 35431 | 8758fe1fc9f8 |
child 42463 | f270e3e18be5 |
permissions | -rw-r--r-- |
8011 | 1 |
(* Title: HOL/MicroJava/J/Decl.thy |
2 |
Author: David von Oheimb |
|
11372 | 3 |
Copyright 1999 Technische Universitaet Muenchen |
11070 | 4 |
*) |
8011 | 5 |
|
12911 | 6 |
header {* \isaheader{Class Declarations and Programs} *} |
8011 | 7 |
|
16417 | 8 |
theory Decl imports Type begin |
8011 | 9 |
|
12517 | 10 |
types |
11 |
fdecl = "vname \<times> ty" -- "field declaration, cf. 8.3 (, 9.3)" |
|
12 |
||
13 |
sig = "mname \<times> ty list" -- "signature of a method, cf. 8.4.2" |
|
14 |
||
15 |
'c mdecl = "sig \<times> ty \<times> 'c" -- "method declaration in a class" |
|
16 |
||
18551 | 17 |
'c "class" = "cname \<times> fdecl list \<times> 'c mdecl list" |
12517 | 18 |
-- "class = superclass, fields, methods" |
19 |
||
20 |
'c cdecl = "cname \<times> 'c class" -- "class declaration, cf. 8.1" |
|
21 |
||
22 |
'c prog = "'c cdecl list" -- "program" |
|
8011 | 23 |
|
24 |
||
10613
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
25 |
translations |
35431 | 26 |
(type) "fdecl" <= (type) "vname \<times> ty" |
27 |
(type) "sig" <= (type) "mname \<times> ty list" |
|
28 |
(type) "'c mdecl" <= (type) "sig \<times> ty \<times> 'c" |
|
29 |
(type) "'c class" <= (type) "cname \<times> fdecl list \<times> ('c mdecl) list" |
|
30 |
(type) "'c cdecl" <= (type) "cname \<times> ('c class)" |
|
31 |
(type) "'c prog" <= (type) "('c cdecl) list" |
|
10613
78b1d6c3ee9c
improved superclass entry for classes and definition status of is_class, class
oheimb
parents:
10042
diff
changeset
|
32 |
|
9346
297dcbf64526
re-structuring MicroJava; added Example; corrected := syntax; simplfied cast
oheimb
parents:
8011
diff
changeset
|
33 |
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
18551
diff
changeset
|
34 |
definition "class" :: "'c prog => (cname \<rightharpoonup> 'c class)" where |
12517 | 35 |
"class \<equiv> map_of" |
36 |
||
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
18551
diff
changeset
|
37 |
definition is_class :: "'c prog => cname => bool" where |
12517 | 38 |
"is_class G C \<equiv> class G C \<noteq> None" |
39 |
||
11026
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10925
diff
changeset
|
40 |
|
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10925
diff
changeset
|
41 |
lemma finite_is_class: "finite {C. is_class G C}" |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10925
diff
changeset
|
42 |
apply (unfold is_class_def class_def) |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10925
diff
changeset
|
43 |
apply (fold dom_def) |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10925
diff
changeset
|
44 |
apply (rule finite_dom_map_of) |
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
oheimb
parents:
10925
diff
changeset
|
45 |
done |
9346
297dcbf64526
re-structuring MicroJava; added Example; corrected := syntax; simplfied cast
oheimb
parents:
8011
diff
changeset
|
46 |
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
18551
diff
changeset
|
47 |
primrec is_type :: "'c prog => ty => bool" where |
12517 | 48 |
"is_type G (PrimT pt) = True" |
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
18551
diff
changeset
|
49 |
| "is_type G (RefT t) = (case t of NullT => True | ClassT C => is_class G C)" |
9346
297dcbf64526
re-structuring MicroJava; added Example; corrected := syntax; simplfied cast
oheimb
parents:
8011
diff
changeset
|
50 |
|
8011 | 51 |
end |