author | Cezary Kaliszyk <kaliszyk@in.tum.de> |
Mon, 26 Apr 2010 15:14:14 +0200 | |
changeset 36352 | f71978e47cd5 |
parent 35431 | 8758fe1fc9f8 |
child 42463 | f270e3e18be5 |
permissions | -rw-r--r-- |
11376 | 1 |
(* Title: HOL/NanoJava/Decl.thy |
2 |
Author: David von Oheimb |
|
3 |
Copyright 2001 Technische Universitaet Muenchen |
|
4 |
*) |
|
5 |
||
6 |
header "Types, class Declarations, and whole programs" |
|
7 |
||
16417 | 8 |
theory Decl imports Term begin |
11376 | 9 |
|
10 |
datatype ty |
|
11558
6539627881e8
simplified vnam/vname, introduced fname, improved comments
oheimb
parents:
11486
diff
changeset
|
11 |
= NT --{* null type *} |
6539627881e8
simplified vnam/vname, introduced fname, improved comments
oheimb
parents:
11486
diff
changeset
|
12 |
| Class cname --{* class type *} |
11376 | 13 |
|
11565 | 14 |
text{* Field declaration *} |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
18551
diff
changeset
|
15 |
types fdecl |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
18551
diff
changeset
|
16 |
= "fname \<times> ty" |
11376 | 17 |
|
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
18551
diff
changeset
|
18 |
record methd |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
18551
diff
changeset
|
19 |
= par :: ty |
11376 | 20 |
res :: ty |
21 |
lcl ::"(vname \<times> ty) list" |
|
22 |
bdy :: stmt |
|
23 |
||
11565 | 24 |
text{* Method declaration *} |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
18551
diff
changeset
|
25 |
types mdecl |
11376 | 26 |
= "mname \<times> methd" |
27 |
||
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
18551
diff
changeset
|
28 |
record "class" |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
18551
diff
changeset
|
29 |
= super :: cname |
12264
9c356e2da72f
renamed "fields" to "flds" (avoid clash with new "fields" operation);
wenzelm
parents:
11565
diff
changeset
|
30 |
flds ::"fdecl list" |
11376 | 31 |
methods ::"mdecl list" |
32 |
||
11565 | 33 |
text{* Class declaration *} |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
18551
diff
changeset
|
34 |
types cdecl |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
18551
diff
changeset
|
35 |
= "cname \<times> class" |
11376 | 36 |
|
37 |
types prog |
|
38 |
= "cdecl list" |
|
39 |
||
40 |
translations |
|
35431 | 41 |
(type) "fdecl" \<leftharpoondown> (type) "fname \<times> ty" |
42 |
(type) "mdecl" \<leftharpoondown> (type) "mname \<times> ty \<times> ty \<times> stmt" |
|
43 |
(type) "class" \<leftharpoondown> (type) "cname \<times> fdecl list \<times> mdecl list" |
|
44 |
(type) "cdecl" \<leftharpoondown> (type) "cname \<times> class" |
|
45 |
(type) "prog " \<leftharpoondown> (type) "cdecl list" |
|
11376 | 46 |
|
47 |
consts |
|
48 |
||
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
18551
diff
changeset
|
49 |
Prog :: prog --{* program as a global value *} |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
18551
diff
changeset
|
50 |
Object :: cname --{* name of root class *} |
11376 | 51 |
|
52 |
||
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
32960
diff
changeset
|
53 |
definition "class" :: "cname \<rightharpoonup> class" where |
11376 | 54 |
"class \<equiv> map_of Prog" |
55 |
||
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
32960
diff
changeset
|
56 |
definition is_class :: "cname => bool" where |
11376 | 57 |
"is_class C \<equiv> class C \<noteq> None" |
58 |
||
59 |
lemma finite_is_class: "finite {C. is_class C}" |
|
60 |
apply (unfold is_class_def class_def) |
|
61 |
apply (fold dom_def) |
|
62 |
apply (rule finite_dom_map_of) |
|
63 |
done |
|
64 |
||
65 |
end |