src/HOL/MicroJava/J/Prog.thy
changeset 8011 d14c4e9e9c8e
     1.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
     1.2 +++ b/src/HOL/MicroJava/J/Prog.thy	Thu Nov 11 12:23:45 1999 +0100
     1.3 @@ -0,0 +1,30 @@
     1.4 +(*  Title:      HOL/MicroJava/J/Prog.thy
     1.5 +    ID:         $Id$
     1.6 +    Author:     David von Oheimb
     1.7 +    Copyright   1999 Technische Universitaet Muenchen
     1.8 +
     1.9 +Java program = list of class declarations
    1.10 +*)
    1.11 +
    1.12 +Prog = Decl +
    1.13 +
    1.14 +types 'c prog = "'c cdecl list"
    1.15 +
    1.16 +consts
    1.17 +
    1.18 +  class		:: "'c prog \\<Rightarrow> (cname \\<leadsto> 'c class)"
    1.19 +
    1.20 +  is_class	:: "'c prog \\<Rightarrow> cname \\<Rightarrow> bool"
    1.21 +  is_type	:: "'c prog \\<Rightarrow> ty    \\<Rightarrow> bool"
    1.22 +
    1.23 +defs
    1.24 +
    1.25 +  class_def	"class        \\<equiv> map_of"
    1.26 +
    1.27 +  is_class_def	"is_class G C \\<equiv> class G C \\<noteq> None"
    1.28 +
    1.29 +primrec
    1.30 +"is_type G (PrimT pt) = True"
    1.31 +"is_type G (RefT t) = (case t of NullT \\<Rightarrow> True | ClassT C \\<Rightarrow> is_class G C)"
    1.32 +
    1.33 +end