src/HOL/MicroJava/J/WellType.thy
changeset 8011 d14c4e9e9c8e
child 8034 6fc37b5c5e98
equal deleted inserted replaced
8010:69032a618aa9 8011:d14c4e9e9c8e
       
     1 (*  Title:      HOL/MicroJava/J/WellType.thy
       
     2     ID:         $Id$
       
     3     Author:     David von Oheimb
       
     4     Copyright   1999 Technische Universitaet Muenchen
       
     5 
       
     6 Well-typedness of Java programs
       
     7 
       
     8 the formulation of well-typedness of method calls given below (as well as
       
     9 the Java Specification 1.0) is a little too restrictive: Is does not allow
       
    10 methods of class Object to be called upon references of interface type.
       
    11 
       
    12 simplifications:
       
    13 * the type rules include all static checks on expressions and statements, e.g.
       
    14   definedness of names (of parameters, locals, fields, methods)
       
    15 
       
    16 *)
       
    17 
       
    18 WellType = Term + WellForm +
       
    19 
       
    20 types	lenv (* local variables, including method parameters and This *)
       
    21 	= "vname \\<leadsto> ty"
       
    22         'c env
       
    23 	= "'c prog \\<times> lenv"
       
    24 
       
    25 syntax
       
    26 
       
    27   prg		:: "'c env \\<Rightarrow> 'c prog"
       
    28   localT	:: "'c env \\<Rightarrow> (vname \\<leadsto> ty)"
       
    29 
       
    30 translations	
       
    31 
       
    32   "prg"		=> "fst"
       
    33   "localT"	=> "snd"
       
    34 
       
    35 consts
       
    36 
       
    37   more_spec	:: "'c prog \\<Rightarrow> (ty \\<times> 'x) \\<times> ty list \\<Rightarrow>
       
    38 		               (ty \\<times> 'x) \\<times> ty list \\<Rightarrow> bool"
       
    39   m_head	:: "'c prog \\<Rightarrow>  ref_ty \\<Rightarrow> sig \\<Rightarrow>  (ty \\<times> ty) option"
       
    40   appl_methds	:: "'c prog \\<Rightarrow>  ref_ty \\<Rightarrow> sig \\<Rightarrow> ((ty \\<times> ty) \\<times> ty list) set"
       
    41   max_spec	:: "'c prog \\<Rightarrow>  ref_ty \\<Rightarrow> sig \\<Rightarrow> ((ty \\<times> ty) \\<times> ty list) set"
       
    42 
       
    43 defs
       
    44 
       
    45   m_head_def  "m_head G t sig \\<equiv> case t of NullT \\<Rightarrow> None | ClassT C \\<Rightarrow>
       
    46 		 option_map (\\<lambda>(md,(rT,mb)). (Class md,rT)) (cmethd (G,C) sig)"
       
    47                                                                
       
    48   more_spec_def	  "more_spec G \\<equiv> \\<lambda>((d,h),pTs). \\<lambda>((d',h'),pTs'). G\\<turnstile>d\\<preceq>d' \\<and>
       
    49 		                  list_all2 (\\<lambda>T T'. G\\<turnstile>T\\<preceq>T') pTs pTs'"
       
    50   
       
    51   (* applicable methods, cf. 15.11.2.1 *)
       
    52   appl_methds_def "appl_methds G T \\<equiv> \\<lambda>(mn, pTs). {(mh,pTs') |mh pTs'.
       
    53 		                  m_head G T (mn, pTs') = Some mh \\<and>
       
    54 		                  list_all2 (\\<lambda>T T'. G\\<turnstile>T\\<preceq>T') pTs pTs'}"
       
    55 
       
    56   (* maximally specific methods, cf. 15.11.2.2 *)
       
    57    max_spec_def	  "max_spec G rT sig \\<equiv> {m. m \\<in>appl_methds G rT sig \\<and> 
       
    58 				          (\\<forall>m'\\<in>appl_methds G rT sig.
       
    59 				                   more_spec G m' m \\<longrightarrow> m' = m)}"
       
    60 consts
       
    61 
       
    62   typeof :: "(loc \\<Rightarrow> ty option) \\<Rightarrow> val \\<Rightarrow> ty option"
       
    63 
       
    64 primrec
       
    65 	"typeof dt  Unit    = Some (PrimT Void)"
       
    66 	"typeof dt  Null    = Some NT"
       
    67 	"typeof dt (Bool b) = Some (PrimT Boolean)"
       
    68 	"typeof dt (Intg i) = Some (PrimT Integer)"
       
    69 	"typeof dt (Addr a) = dt a"
       
    70 
       
    71 types
       
    72 	javam = "vname list \\<times> (vname \\<times> ty) list \\<times> stmt \\<times> expr"
       
    73 	(* method body with parameter names, local variables, block, result expression *)
       
    74 
       
    75 consts
       
    76 
       
    77   ty_expr :: "javam env \\<Rightarrow> (expr      \\<times> ty     ) set"
       
    78   ty_exprs:: "javam env \\<Rightarrow> (expr list \\<times> ty list) set"
       
    79   wt_stmt :: "javam env \\<Rightarrow>  stmt                 set"
       
    80 
       
    81 syntax
       
    82 
       
    83 ty_expr :: "javam env \\<Rightarrow> [expr     , ty     ] \\<Rightarrow> bool" ("_\\<turnstile>_\\<Colon>_"  [51,51,51]50)
       
    84 ty_exprs:: "javam env \\<Rightarrow> [expr list, ty list] \\<Rightarrow> bool" ("_\\<turnstile>_[\\<Colon>]_"[51,51,51]50)
       
    85 wt_stmt :: "javam env \\<Rightarrow>  stmt                \\<Rightarrow> bool" ("_\\<turnstile>_ \\<surd>" [51,51   ]50)
       
    86 
       
    87 translations
       
    88 	"E\\<turnstile>e \\<Colon> T" == "(e,T) \\<in> ty_expr  E"
       
    89 	"E\\<turnstile>e[\\<Colon>]T" == "(e,T) \\<in> ty_exprs E"
       
    90 	"E\\<turnstile>c \\<surd>"    == "c     \\<in> wt_stmt  E"
       
    91   
       
    92 inductive "ty_expr E" "ty_exprs E" "wt_stmt E" intrs
       
    93 
       
    94 (* well-typed expressions *)
       
    95 
       
    96   (* cf. 15.8 *)
       
    97   NewC	"\\<lbrakk>is_class (prg E) C\\<rbrakk> \\<Longrightarrow>
       
    98 						 E\\<turnstile>NewC C\\<Colon>Class C"
       
    99 
       
   100   (* cf. 15.15 *)
       
   101   Cast	"\\<lbrakk>E\\<turnstile>e\\<Colon>T;
       
   102 	  prg E\\<turnstile>T\\<Rightarrow>? T'\\<rbrakk> \\<Longrightarrow>
       
   103 						 E\\<turnstile>Cast T' e\\<Colon>T'"
       
   104 
       
   105   (* cf. 15.7.1 *)
       
   106   Lit	"\\<lbrakk>typeof (\\<lambda>v. None) x = Some T\\<rbrakk> \\<Longrightarrow>
       
   107 						 E\\<turnstile>Lit x\\<Colon>T"
       
   108 
       
   109   (* cf. 15.13.1 *)
       
   110   LAcc	"\\<lbrakk>localT E v = Some T; is_type (prg E) T\\<rbrakk> \\<Longrightarrow>
       
   111 						 E\\<turnstile>LAcc v\\<Colon>T"
       
   112   
       
   113   (* cf. 15.25, 15.25.1 *)
       
   114   LAss  "\\<lbrakk>E\\<turnstile>LAcc v\\<Colon>T;
       
   115 	   E\\<turnstile>e\\<Colon>T';
       
   116 	   prg E\\<turnstile>T'\\<preceq>T\\<rbrakk> \\<Longrightarrow>
       
   117 						 E\\<turnstile>v\\<Colon>=e\\<Colon>T'"
       
   118 
       
   119   (* cf. 15.10.1 *)
       
   120   FAcc	"\\<lbrakk>E\\<turnstile>a\\<Colon>Class C; 
       
   121 	  cfield (prg E,C) fn = Some (fd,fT)\\<rbrakk> \\<Longrightarrow>
       
   122 						 E\\<turnstile>{fd}a..fn\\<Colon>fT"
       
   123 
       
   124   (* cf. 15.25, 15.25.1 *)
       
   125   FAss  "\\<lbrakk>E\\<turnstile>{fd}a..fn\\<Colon>T;
       
   126 	  E\\<turnstile>v       \\<Colon>T';
       
   127 	  prg E\\<turnstile>T'\\<preceq>T\\<rbrakk> \\<Longrightarrow>
       
   128 					 	 E\\<turnstile>{fd}a..fn\\<in>=v\\<Colon>T'"
       
   129 
       
   130 
       
   131   (* cf. 15.11.1, 15.11.2, 15.11.3 *)
       
   132   Call	"\\<lbrakk>E\\<turnstile>a\\<Colon>RefT t;
       
   133 	  E\\<turnstile>ps[\\<Colon>]pTs;
       
   134 	  max_spec (prg E) t (mn, pTs) = {((md,rT),pTs')}\\<rbrakk> \\<Longrightarrow>
       
   135 						 E\\<turnstile>a..mn({pTs'}ps)\\<Colon>rT"
       
   136 
       
   137 (* well-typed expression lists *)
       
   138 
       
   139   (* cf. 15.11.??? *)
       
   140   Nil						"E\\<turnstile>[][\\<Colon>][]"
       
   141 
       
   142   (* cf. 15.11.??? *)
       
   143   Cons	"\\<lbrakk>E\\<turnstile>e\\<Colon>T;
       
   144 	   E\\<turnstile>es[\\<Colon>]Ts\\<rbrakk> \\<Longrightarrow>
       
   145 						 E\\<turnstile>e#es[\\<Colon>]T#Ts"
       
   146 
       
   147 (* well-typed statements *)
       
   148 
       
   149   Skip					"E\\<turnstile>Skip\\<surd>"
       
   150 
       
   151   Expr	"\\<lbrakk>E\\<turnstile>e\\<Colon>T\\<rbrakk> \\<Longrightarrow>
       
   152 					 E\\<turnstile>Expr e\\<surd>"
       
   153 
       
   154  Comp	"\\<lbrakk>E\\<turnstile>s1\\<surd>; 
       
   155 	  E\\<turnstile>s2\\<surd>\\<rbrakk> \\<Longrightarrow>
       
   156 					 E\\<turnstile>s1;; s2\\<surd>"
       
   157 
       
   158   (* cf. 14.8 *)
       
   159   Cond	"\\<lbrakk>E\\<turnstile>e\\<Colon>PrimT Boolean;
       
   160 	  E\\<turnstile>s1\\<surd>;
       
   161 	  E\\<turnstile>s2\\<surd>\\<rbrakk> \\<Longrightarrow>
       
   162 					 E\\<turnstile>If(e) s1 Else s2\\<surd>"
       
   163 
       
   164   (* cf. 14.10 *)
       
   165   Loop "\\<lbrakk>E\\<turnstile>e\\<Colon>PrimT Boolean;
       
   166 	 E\\<turnstile>s\\<surd>\\<rbrakk> \\<Longrightarrow>
       
   167 					 E\\<turnstile>While(e) s\\<surd>"
       
   168 
       
   169 constdefs
       
   170 
       
   171  wt_java_mdecl :: javam prog => cname => javam mdecl => bool
       
   172 "wt_java_mdecl G C \\<equiv> \\<lambda>((mn,pTs),rT,(pns,lvars,blk,res)).
       
   173 	length pTs = length pns \\<and>
       
   174 	nodups pns \\<and>
       
   175 	unique lvars \\<and>
       
   176 	(\\<forall>pn\\<in>set pns. map_of lvars pn = None) \\<and>
       
   177 	(\\<forall>(vn,T)\\<in>set lvars. is_type G T) &
       
   178 	(let E = (G,map_of lvars(pns[\\<mapsto>]pTs)(This\\<mapsto>Class C)) in
       
   179 	 E\\<turnstile>blk\\<surd> \\<and> (\\<exists>T. E\\<turnstile>res\\<Colon>T \\<and> G\\<turnstile>T\\<preceq>rT))"
       
   180 
       
   181  wf_java_prog :: javam prog => bool
       
   182 "wf_java_prog G \\<equiv> wf_prog wt_java_mdecl G"
       
   183 
       
   184 end