src/HOL/MicroJava/J/Eval.thy
changeset 8011 d14c4e9e9c8e
child 8034 6fc37b5c5e98
     1.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
     1.2 +++ b/src/HOL/MicroJava/J/Eval.thy	Thu Nov 11 12:23:45 1999 +0100
     1.3 @@ -0,0 +1,121 @@
     1.4 +(*  Title:      HOL/MicroJava/J/Eval.thy
     1.5 +    ID:         $Id$
     1.6 +    Author:     David von Oheimb
     1.7 +    Copyright   1999 Technische Universitaet Muenchen
     1.8 +
     1.9 +Operational evaluation (big-step) semantics of the 
    1.10 +execution of Java expressions and statements
    1.11 +*)
    1.12 +
    1.13 +Eval = State + 
    1.14 +
    1.15 +consts
    1.16 +  eval  :: "javam prog \\<Rightarrow> (xstate \\<times> expr      \\<times> val      \\<times> xstate) set"
    1.17 +  evals :: "javam prog \\<Rightarrow> (xstate \\<times> expr list \\<times> val list \\<times> xstate) set"
    1.18 +  exec  :: "javam prog \\<Rightarrow> (xstate \\<times> stmt                 \\<times> xstate) set"
    1.19 +
    1.20 +syntax
    1.21 +  eval :: "[javam prog,xstate,expr,val,xstate] \\<Rightarrow> bool "("_\\<turnstile>_ -_\\<succ>_\\<rightarrow> _"[51,82,82,82,82]81)
    1.22 +  evals:: "[javam prog,xstate,expr list,
    1.23 +	                      val list,xstate] \\<Rightarrow> bool "("_\\<turnstile>_ -_[\\<succ>]_\\<rightarrow> _"[51,82,51,51,82]81)
    1.24 +  exec :: "[javam prog,xstate,stmt,    xstate] \\<Rightarrow> bool "("_\\<turnstile>_ -_\\<rightarrow> _"  [51,82,82,   82]81)
    1.25 +
    1.26 +translations
    1.27 +  "G\\<turnstile>s -e \\<succ> v\\<rightarrow> (x,s')" == "(s, e, v, _args x s') \\<in> eval  G"
    1.28 +  "G\\<turnstile>s -e \\<succ> v\\<rightarrow>    s' " == "(s, e, v,         s') \\<in> eval  G"
    1.29 +  "G\\<turnstile>s -e[\\<succ>]v\\<rightarrow> (x,s')" == "(s, e, v, _args x s') \\<in> evals G"
    1.30 +  "G\\<turnstile>s -e[\\<succ>]v\\<rightarrow>    s' " == "(s, e, v,         s') \\<in> evals G"
    1.31 +  "G\\<turnstile>s -c     \\<rightarrow> (x,s')" == "(s, c,    _args x s') \\<in> exec  G"
    1.32 +  "G\\<turnstile>s -c     \\<rightarrow>    s' " == "(s, c,            s') \\<in> exec  G"
    1.33 +
    1.34 +inductive "eval G" "evals G" "exec G" intrs
    1.35 +
    1.36 +(* evaluation of expressions *)
    1.37 +
    1.38 +  (* cf. 15.5 *)
    1.39 +  XcptE				  "G\\<turnstile>(Some xc,s) -e\\<succ>arbitrary\\<rightarrow> (Some xc,s)"
    1.40 +
    1.41 +  (* cf. 15.8.1 *)
    1.42 +  NewC	"\\<lbrakk>h = heap s; (a,x) = new_Addr h;
    1.43 +	  h'= h(a\\<mapsto>(C,init_vars (fields (G,C))))\\<rbrakk> \\<Longrightarrow>
    1.44 +				   G\\<turnstile>Norm s -NewC C\\<succ>Addr a\\<rightarrow> c_hupd h' (x,s)"
    1.45 +
    1.46 +  (* cf. 15.15 *)
    1.47 +  Cast	"\\<lbrakk>G\\<turnstile>Norm s0 -e\\<succ>v\\<rightarrow> (x1,s1);
    1.48 +	  x2=raise_if (\\<not> cast_ok G T (heap s1) v) ClassCast x1\\<rbrakk> \\<Longrightarrow>
    1.49 +			        G\\<turnstile>Norm s0 -Cast T e\\<succ>v\\<rightarrow> (x2,s1)"
    1.50 +
    1.51 +  (* cf. 15.7.1 *)
    1.52 +  Lit				   "G\\<turnstile>Norm s -Lit v\\<succ>v\\<rightarrow> Norm s"
    1.53 +
    1.54 +  (* cf. 15.13.1, 15.2 *)
    1.55 +  LAcc				  "G\\<turnstile>Norm s -LAcc v\\<succ>the (locals s v)\\<rightarrow> Norm s"
    1.56 +
    1.57 +  (* cf. 15.25.1 *)
    1.58 +  LAss  "\\<lbrakk>G\\<turnstile>Norm s -e\\<succ>v\\<rightarrow>  (x,(h,l));
    1.59 +	  l' = (if x = None then l(va\\<mapsto>v) else l)\\<rbrakk> \\<Longrightarrow>
    1.60 +				   G\\<turnstile>Norm s -va\\<Colon>=e\\<succ>v\\<rightarrow> (x,(h,l'))"
    1.61 +
    1.62 +
    1.63 +  (* cf. 15.10.1, 15.2 *)
    1.64 +  FAcc	"\\<lbrakk>G\\<turnstile>Norm s0 -e\\<succ>a'\\<rightarrow> (x1,s1); 
    1.65 +	  v = the (snd (the (heap s1 (the_Addr a'))) (fn,T))\\<rbrakk> \\<Longrightarrow>
    1.66 +				 G\\<turnstile>Norm s0 -{T}e..fn\\<succ>v\\<rightarrow> (np a' x1,s1)"
    1.67 +
    1.68 +  (* cf. 15.25.1 *)
    1.69 +  FAss  "\\<lbrakk>G\\<turnstile>     Norm s0  -e1\\<succ>a'\\<rightarrow> (x1,s1); a = the_Addr a';
    1.70 +	  G\\<turnstile>(np a' x1,s1) -e2\\<succ>v \\<rightarrow> (x2,s2);
    1.71 +	  h = heap s2; (c,fs) = the (h a);
    1.72 +	  h' = h(a\\<mapsto>(c,(fs((fn,T)\\<mapsto>v))))\\<rbrakk> \\<Longrightarrow>
    1.73 +			  G\\<turnstile>Norm s0 -{T}e1..fn\\<in>=e2\\<succ>v\\<rightarrow> c_hupd h' (x2,s2)"
    1.74 +
    1.75 +  (* cf. 15.11.4.1, 15.11.4.2, 15.11.4.4, 15.11.4.5, 14.15 *)
    1.76 +  Call	"\\<lbrakk>G\\<turnstile>Norm s0 -e\\<succ>a'\\<rightarrow> s1; a = the_Addr a';
    1.77 +	   G\\<turnstile>s1 -ps[\\<succ>]pvs\\<rightarrow> (x,(h,l)); dynT = fst (the (h a));
    1.78 +	   (md,rT,pns,lvars,blk,res) = the (cmethd (G,dynT) (mn,pTs));
    1.79 +	   G\\<turnstile>(np a' x,(h,(init_vars lvars)(pns[\\<mapsto>]pvs)(This\\<mapsto>a'))) -blk\\<rightarrow> s3;
    1.80 +	   G\\<turnstile>     s3 -res\\<succ>v \\<rightarrow> (x4,s4)\\<rbrakk> \\<Longrightarrow>
    1.81 +			    G\\<turnstile>Norm s0 -e..mn({pTs}ps)\\<succ>v\\<rightarrow> (x4,(heap s4,l))"
    1.82 +
    1.83 +
    1.84 +(* evaluation of expression lists *)
    1.85 +
    1.86 +  (* cf. 15.5 *)
    1.87 +  XcptEs			  "G\\<turnstile>(Some xc,s) -e[\\<succ>]arbitrary\\<rightarrow> (Some xc,s)"
    1.88 +
    1.89 +  (* cf. 15.11.??? *)
    1.90 +  Nil
    1.91 +				    "G\\<turnstile>Norm s0 -[][\\<succ>][]\\<rightarrow> Norm s0"
    1.92 +
    1.93 +  (* cf. 15.6.4 *)
    1.94 +  Cons	"\\<lbrakk>G\\<turnstile>Norm s0 -e  \\<succ> v \\<rightarrow> s1;
    1.95 +           G\\<turnstile>     s1 -es[\\<succ>]vs\\<rightarrow> s2\\<rbrakk> \\<Longrightarrow>
    1.96 +				   G\\<turnstile>Norm s0 -e#es[\\<succ>]v#vs\\<rightarrow> s2"
    1.97 +
    1.98 +(* execution of statements *)
    1.99 +
   1.100 +  (* cf. 14.1 *)
   1.101 +  XcptS				 "G\\<turnstile>(Some xc,s) -s0\\<rightarrow> (Some xc,s)"
   1.102 +
   1.103 +  (* cf. 14.5 *)
   1.104 +  Skip	 			    "G\\<turnstile>Norm s -Skip\\<rightarrow> Norm s"
   1.105 +
   1.106 +  (* cf. 14.7 *)
   1.107 +  Expr	"\\<lbrakk>G\\<turnstile>Norm s0 -e\\<succ>v\\<rightarrow> s1\\<rbrakk> \\<Longrightarrow>
   1.108 +				  G\\<turnstile>Norm s0 -Expr e\\<rightarrow> s1"
   1.109 +
   1.110 +  (* cf. 14.2 *)
   1.111 +  Comp	"\\<lbrakk>G\\<turnstile>Norm s0 -s \\<rightarrow> s1;
   1.112 +	  G\\<turnstile>     s1 -t \\<rightarrow> s2\\<rbrakk> \\<Longrightarrow>
   1.113 +				 G\\<turnstile>Norm s0 -(s;; t)\\<rightarrow> s2"
   1.114 +
   1.115 +  (* cf. 14.8.2 *)
   1.116 +  Cond	"\\<lbrakk>G\\<turnstile>Norm s0  -e \\<succ>v\\<rightarrow> s1;
   1.117 +	  G\\<turnstile>     s1 -(if  the_Bool v then s else t)\\<rightarrow> s2\\<rbrakk> \\<Longrightarrow>
   1.118 +		        G\\<turnstile>Norm s0 -(If(e) s Else t)\\<rightarrow> s2"
   1.119 +
   1.120 +  (* cf. 14.10, 14.10.1 *)
   1.121 +  Loop	"\\<lbrakk>G\\<turnstile>Norm s0 -(If(e) (s;; While(e) s) Else Skip)\\<rightarrow> s1\\<rbrakk> \\<Longrightarrow>
   1.122 +			    G\\<turnstile>Norm s0 -(While(e) s)\\<rightarrow> s1"
   1.123 +
   1.124 +end