| 8011 |      1 | (*  Title:      HOL/MicroJava/J/Prog.thy
 | 
|  |      2 |     ID:         $Id$
 | 
|  |      3 |     Author:     David von Oheimb
 | 
|  |      4 |     Copyright   1999 Technische Universitaet Muenchen
 | 
|  |      5 | 
 | 
|  |      6 | Java program = list of class declarations
 | 
|  |      7 | *)
 | 
|  |      8 | 
 | 
|  |      9 | Prog = Decl +
 | 
|  |     10 | 
 | 
|  |     11 | types 'c prog = "'c cdecl list"
 | 
|  |     12 | 
 | 
|  |     13 | consts
 | 
|  |     14 | 
 | 
|  |     15 |   class		:: "'c prog \\<Rightarrow> (cname \\<leadsto> 'c class)"
 | 
|  |     16 | 
 | 
|  |     17 |   is_class	:: "'c prog \\<Rightarrow> cname \\<Rightarrow> bool"
 | 
|  |     18 |   is_type	:: "'c prog \\<Rightarrow> ty    \\<Rightarrow> bool"
 | 
|  |     19 | 
 | 
|  |     20 | defs
 | 
|  |     21 | 
 | 
|  |     22 |   class_def	"class        \\<equiv> map_of"
 | 
|  |     23 | 
 | 
|  |     24 |   is_class_def	"is_class G C \\<equiv> class G C \\<noteq> None"
 | 
|  |     25 | 
 | 
|  |     26 | primrec
 | 
|  |     27 | "is_type G (PrimT pt) = True"
 | 
|  |     28 | "is_type G (RefT t) = (case t of NullT \\<Rightarrow> True | ClassT C \\<Rightarrow> is_class G C)"
 | 
|  |     29 | 
 | 
|  |     30 | end
 |