src/HOL/MicroJava/J/WellForm.thy
author nipkow
Fri, 26 Nov 1999 08:46:59 +0100
changeset 8034 6fc37b5c5e98
parent 8011 d14c4e9e9c8e
child 8082 381716a86fcb
permissions -rw-r--r--
Various little changes like cmethd -> method and cfield -> field.
nipkow@8011
     1
(*  Title:      HOL/MicroJava/J/WellForm.thy
nipkow@8011
     2
    ID:         $Id$
nipkow@8011
     3
    Author:     David von Oheimb
nipkow@8011
     4
    Copyright   1999 Technische Universitaet Muenchen
nipkow@8011
     5
nipkow@8011
     6
Well-formedness of Java programs
nipkow@8011
     7
for static checks on expressions and statements, see WellType.thy
nipkow@8011
     8
nipkow@8011
     9
improvements over Java Specification 1.0 (cf. 8.4.6.3, 8.4.6.4, 9.4.1):
nipkow@8011
    10
* a method implementing or overwriting another method may have a result type 
nipkow@8011
    11
  that widens to the result type of the other method (instead of identical type)
nipkow@8011
    12
nipkow@8011
    13
simplifications:
nipkow@8011
    14
* for uniformity, Object is assumed to be declared like any other class
nipkow@8011
    15
*)
nipkow@8011
    16
nipkow@8011
    17
WellForm = TypeRel +
nipkow@8011
    18
nipkow@8011
    19
types 'c wtm = 'c prog => cname => 'c mdecl => bool
nipkow@8011
    20
nipkow@8011
    21
constdefs
nipkow@8011
    22
nipkow@8011
    23
 wf_fdecl	:: "'c prog \\<Rightarrow>          fdecl \\<Rightarrow> bool"
nipkow@8011
    24
"wf_fdecl G \\<equiv> \\<lambda>(fn,ft). is_type G ft"
nipkow@8011
    25
nipkow@8011
    26
 wf_mhead	:: "'c prog \\<Rightarrow> sig   \\<Rightarrow> ty \\<Rightarrow> bool"
nipkow@8011
    27
"wf_mhead G \\<equiv> \\<lambda>(mn,pTs) rT. (\\<forall>T\\<in>set pTs. is_type G T) \\<and> is_type G rT"
nipkow@8011
    28
nipkow@8011
    29
 wf_mdecl	:: "'c wtm \\<Rightarrow> 'c wtm"
nipkow@8011
    30
"wf_mdecl wtm G C \\<equiv> \\<lambda>(sig,rT,mb). wf_mhead G sig rT \\<and> wtm G C (sig,rT,mb)"
nipkow@8011
    31
nipkow@8011
    32
  wf_cdecl	:: "'c wtm \\<Rightarrow> 'c prog \\<Rightarrow> 'c cdecl \\<Rightarrow> bool"
nipkow@8011
    33
"wf_cdecl wtm G \\<equiv>
nipkow@8011
    34
   \\<lambda>(C,(sc,fs,ms)).
nipkow@8011
    35
	(\\<forall>f\\<in>set fs. wf_fdecl G   f    ) \\<and>  unique fs \\<and>
nipkow@8011
    36
	(\\<forall>m\\<in>set ms. wf_mdecl wtm G C m) \\<and>  unique ms \\<and>
nipkow@8011
    37
	(case sc of None \\<Rightarrow> C = Object
nipkow@8011
    38
         | Some D \\<Rightarrow>
nipkow@8011
    39
             is_class G D \\<and>  \\<not>  G\\<turnstile>D\\<prec>C C \\<and>
nipkow@8011
    40
             (\\<forall>(sig,rT,b)\\<in>set ms. \\<forall>D' rT' b'.
nipkow@8034
    41
                 method(G,D) sig = Some(D',rT',b') \\<longrightarrow> G\\<turnstile>rT\\<preceq>rT'))"
nipkow@8011
    42
nipkow@8011
    43
 wf_prog	:: "'c wtm \\<Rightarrow> 'c prog \\<Rightarrow> bool"
nipkow@8011
    44
"wf_prog wtm G \\<equiv>
nipkow@8011
    45
   let cs = set G in ObjectC \\<in> cs \\<and> (\\<forall>c\\<in>cs. wf_cdecl wtm G c) \\<and> unique G"
nipkow@8011
    46
nipkow@8011
    47
end