| author | haftmann | 
| Tue, 28 Jun 2005 10:24:53 +0200 | |
| changeset 16573 | cc86fd4eeee4 | 
| parent 16417 | 9bc16273c2d4 | 
| child 18551 | be0705186ff5 | 
| permissions | -rw-r--r-- | 
| 11376 | 1 | (* Title: HOL/NanoJava/Decl.thy | 
| 2 | ID: $Id$ | |
| 3 | Author: David von Oheimb | |
| 4 | Copyright 2001 Technische Universitaet Muenchen | |
| 5 | *) | |
| 6 | ||
| 7 | header "Types, class Declarations, and whole programs" | |
| 8 | ||
| 16417 | 9 | theory Decl imports Term begin | 
| 11376 | 10 | |
| 11 | datatype ty | |
| 11558 
6539627881e8
simplified vnam/vname, introduced fname, improved comments
 oheimb parents: 
11486diff
changeset | 12 |   = NT           --{* null type  *}
 | 
| 
6539627881e8
simplified vnam/vname, introduced fname, improved comments
 oheimb parents: 
11486diff
changeset | 13 |   | Class cname  --{* class type *}
 | 
| 11376 | 14 | |
| 11565 | 15 | text{* Field declaration *}
 | 
| 11558 
6539627881e8
simplified vnam/vname, introduced fname, improved comments
 oheimb parents: 
11486diff
changeset | 16 | types fdecl | 
| 
6539627881e8
simplified vnam/vname, introduced fname, improved comments
 oheimb parents: 
11486diff
changeset | 17 | = "fname \<times> ty" | 
| 11376 | 18 | |
| 11558 
6539627881e8
simplified vnam/vname, introduced fname, improved comments
 oheimb parents: 
11486diff
changeset | 19 | record methd | 
| 11376 | 20 | = par :: ty | 
| 21 | res :: ty | |
| 22 | lcl ::"(vname \<times> ty) list" | |
| 23 | bdy :: stmt | |
| 24 | ||
| 11565 | 25 | text{* Method declaration *}
 | 
| 11558 
6539627881e8
simplified vnam/vname, introduced fname, improved comments
 oheimb parents: 
11486diff
changeset | 26 | types mdecl | 
| 11376 | 27 | = "mname \<times> methd" | 
| 28 | ||
| 11558 
6539627881e8
simplified vnam/vname, introduced fname, improved comments
 oheimb parents: 
11486diff
changeset | 29 | record class | 
| 11376 | 30 | = super :: cname | 
| 12264 
9c356e2da72f
renamed "fields" to "flds" (avoid clash with new "fields" operation);
 wenzelm parents: 
11565diff
changeset | 31 | flds ::"fdecl list" | 
| 11376 | 32 | methods ::"mdecl list" | 
| 33 | ||
| 11565 | 34 | text{* Class declaration *}
 | 
| 11558 
6539627881e8
simplified vnam/vname, introduced fname, improved comments
 oheimb parents: 
11486diff
changeset | 35 | types cdecl | 
| 11376 | 36 | = "cname \<times> class" | 
| 37 | ||
| 38 | types prog | |
| 39 | = "cdecl list" | |
| 40 | ||
| 41 | translations | |
| 11558 
6539627881e8
simplified vnam/vname, introduced fname, improved comments
 oheimb parents: 
11486diff
changeset | 42 | "fdecl" \<leftharpoondown> (type)"fname \<times> ty" | 
| 11376 | 43 | "mdecl" \<leftharpoondown> (type)"mname \<times> ty \<times> ty \<times> stmt" | 
| 44 | "class" \<leftharpoondown> (type)"cname \<times> fdecl list \<times> mdecl list" | |
| 45 | "cdecl" \<leftharpoondown> (type)"cname \<times> class" | |
| 46 | "prog " \<leftharpoondown> (type)"cdecl list" | |
| 47 | ||
| 48 | consts | |
| 49 | ||
| 11558 
6539627881e8
simplified vnam/vname, introduced fname, improved comments
 oheimb parents: 
11486diff
changeset | 50 |   Prog    :: prog	--{* program as a global value *}
 | 
| 
6539627881e8
simplified vnam/vname, introduced fname, improved comments
 oheimb parents: 
11486diff
changeset | 51 |   Object  :: cname	--{* name of root class *}
 | 
| 11376 | 52 | |
| 53 | ||
| 54 | constdefs | |
| 14134 | 55 | class :: "cname \<rightharpoonup> class" | 
| 11376 | 56 | "class \<equiv> map_of Prog" | 
| 57 | ||
| 58 | is_class :: "cname => bool" | |
| 59 | "is_class C \<equiv> class C \<noteq> None" | |
| 60 | ||
| 61 | lemma finite_is_class: "finite {C. is_class C}"
 | |
| 62 | apply (unfold is_class_def class_def) | |
| 63 | apply (fold dom_def) | |
| 64 | apply (rule finite_dom_map_of) | |
| 65 | done | |
| 66 | ||
| 67 | end |