src/HOL/MicroJava/J/Decl.thy
author oheimb
Wed, 06 Dec 2000 19:10:36 +0100
changeset 10613 78b1d6c3ee9c
parent 10042 7164dc0d24d8
child 10925 5ffe7ed8899a
permissions -rw-r--r--
improved superclass entry for classes and definition status of is_class, class corrected recursive definitions of "method" and "fields" cleanup of many proofs, removed superfluous tactics and theorems

(*  Title:      HOL/MicroJava/J/Decl.thy
    ID:         $Id$
    Author:     David von Oheimb
    Copyright   1997 Technische Universitaet Muenchen

Class declarations and programs
*)

Decl = Type +

types	fdecl		(* field declaration, cf. 8.3 (, 9.3) *)
	= "vname \\<times> ty"


types	sig		(* signature of a method, cf. 8.4.2 *)
	= "mname \\<times> ty list"

	'c mdecl		(* method declaration in a class *)
	= "sig \\<times> ty \\<times> 'c"

types	'c class		(* class *)
	= "cname \\<times> fdecl list \\<times> 'c mdecl list"
	(* superclass, fields, methods*)

	'c cdecl		(* class declaration, cf. 8.1 *)
	= "cname \\<times> 'c class"

consts

  Object  :: cname	(* name of root class *)
  ObjectC :: 'c cdecl	(* decl of root class *)

defs 

 ObjectC_def "ObjectC == (Object, (arbitrary, [], []))"


types 'c prog = "'c cdecl list"


translations
  "fdecl"   <= (type) "vname \\<times> ty"
  "sig"     <= (type) "mname \\<times> ty list"
  "mdecl c" <= (type) "sig \\<times> ty \\<times> c"
  "class c" <= (type) "cname \\<times> fdecl list \\<times> (c mdecl) list"
  "cdecl c" <= (type) "cname \\<times> (c class)"
  "prog  c" <= (type) "(c cdecl) list"

consts

  class		:: "'c prog => (cname \\<leadsto> 'c class)"
  is_class	:: "'c prog => cname => bool"

translations

  "class"        => "map_of"
  "is_class G C" == "class G C \\<noteq> None"

consts

  is_type	:: "'c prog => ty    => bool"

primrec
"is_type G (PrimT pt) = True"
"is_type G (RefT t) = (case t of NullT => True | ClassT C => is_class G C)"

end