diff -r 69032a618aa9 -r d14c4e9e9c8e src/HOL/MicroJava/J/Eval.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/MicroJava/J/Eval.thy Thu Nov 11 12:23:45 1999 +0100 @@ -0,0 +1,121 @@ +(* Title: HOL/MicroJava/J/Eval.thy + ID: $Id$ + Author: David von Oheimb + Copyright 1999 Technische Universitaet Muenchen + +Operational evaluation (big-step) semantics of the +execution of Java expressions and statements +*) + +Eval = State + + +consts + eval :: "javam prog \\ (xstate \\ expr \\ val \\ xstate) set" + evals :: "javam prog \\ (xstate \\ expr list \\ val list \\ xstate) set" + exec :: "javam prog \\ (xstate \\ stmt \\ xstate) set" + +syntax + eval :: "[javam prog,xstate,expr,val,xstate] \\ bool "("_\\_ -_\\_\\ _"[51,82,82,82,82]81) + evals:: "[javam prog,xstate,expr list, + val list,xstate] \\ bool "("_\\_ -_[\\]_\\ _"[51,82,51,51,82]81) + exec :: "[javam prog,xstate,stmt, xstate] \\ bool "("_\\_ -_\\ _" [51,82,82, 82]81) + +translations + "G\\s -e \\ v\\ (x,s')" == "(s, e, v, _args x s') \\ eval G" + "G\\s -e \\ v\\ s' " == "(s, e, v, s') \\ eval G" + "G\\s -e[\\]v\\ (x,s')" == "(s, e, v, _args x s') \\ evals G" + "G\\s -e[\\]v\\ s' " == "(s, e, v, s') \\ evals G" + "G\\s -c \\ (x,s')" == "(s, c, _args x s') \\ exec G" + "G\\s -c \\ s' " == "(s, c, s') \\ exec G" + +inductive "eval G" "evals G" "exec G" intrs + +(* evaluation of expressions *) + + (* cf. 15.5 *) + XcptE "G\\(Some xc,s) -e\\arbitrary\\ (Some xc,s)" + + (* cf. 15.8.1 *) + NewC "\\h = heap s; (a,x) = new_Addr h; + h'= h(a\\(C,init_vars (fields (G,C))))\\ \\ + G\\Norm s -NewC C\\Addr a\\ c_hupd h' (x,s)" + + (* cf. 15.15 *) + Cast "\\G\\Norm s0 -e\\v\\ (x1,s1); + x2=raise_if (\\ cast_ok G T (heap s1) v) ClassCast x1\\ \\ + G\\Norm s0 -Cast T e\\v\\ (x2,s1)" + + (* cf. 15.7.1 *) + Lit "G\\Norm s -Lit v\\v\\ Norm s" + + (* cf. 15.13.1, 15.2 *) + LAcc "G\\Norm s -LAcc v\\the (locals s v)\\ Norm s" + + (* cf. 15.25.1 *) + LAss "\\G\\Norm s -e\\v\\ (x,(h,l)); + l' = (if x = None then l(va\\v) else l)\\ \\ + G\\Norm s -va\\=e\\v\\ (x,(h,l'))" + + + (* cf. 15.10.1, 15.2 *) + FAcc "\\G\\Norm s0 -e\\a'\\ (x1,s1); + v = the (snd (the (heap s1 (the_Addr a'))) (fn,T))\\ \\ + G\\Norm s0 -{T}e..fn\\v\\ (np a' x1,s1)" + + (* cf. 15.25.1 *) + FAss "\\G\\ Norm s0 -e1\\a'\\ (x1,s1); a = the_Addr a'; + G\\(np a' x1,s1) -e2\\v \\ (x2,s2); + h = heap s2; (c,fs) = the (h a); + h' = h(a\\(c,(fs((fn,T)\\v))))\\ \\ + G\\Norm s0 -{T}e1..fn\\=e2\\v\\ c_hupd h' (x2,s2)" + + (* cf. 15.11.4.1, 15.11.4.2, 15.11.4.4, 15.11.4.5, 14.15 *) + Call "\\G\\Norm s0 -e\\a'\\ s1; a = the_Addr a'; + G\\s1 -ps[\\]pvs\\ (x,(h,l)); dynT = fst (the (h a)); + (md,rT,pns,lvars,blk,res) = the (cmethd (G,dynT) (mn,pTs)); + G\\(np a' x,(h,(init_vars lvars)(pns[\\]pvs)(This\\a'))) -blk\\ s3; + G\\ s3 -res\\v \\ (x4,s4)\\ \\ + G\\Norm s0 -e..mn({pTs}ps)\\v\\ (x4,(heap s4,l))" + + +(* evaluation of expression lists *) + + (* cf. 15.5 *) + XcptEs "G\\(Some xc,s) -e[\\]arbitrary\\ (Some xc,s)" + + (* cf. 15.11.??? *) + Nil + "G\\Norm s0 -[][\\][]\\ Norm s0" + + (* cf. 15.6.4 *) + Cons "\\G\\Norm s0 -e \\ v \\ s1; + G\\ s1 -es[\\]vs\\ s2\\ \\ + G\\Norm s0 -e#es[\\]v#vs\\ s2" + +(* execution of statements *) + + (* cf. 14.1 *) + XcptS "G\\(Some xc,s) -s0\\ (Some xc,s)" + + (* cf. 14.5 *) + Skip "G\\Norm s -Skip\\ Norm s" + + (* cf. 14.7 *) + Expr "\\G\\Norm s0 -e\\v\\ s1\\ \\ + G\\Norm s0 -Expr e\\ s1" + + (* cf. 14.2 *) + Comp "\\G\\Norm s0 -s \\ s1; + G\\ s1 -t \\ s2\\ \\ + G\\Norm s0 -(s;; t)\\ s2" + + (* cf. 14.8.2 *) + Cond "\\G\\Norm s0 -e \\v\\ s1; + G\\ s1 -(if the_Bool v then s else t)\\ s2\\ \\ + G\\Norm s0 -(If(e) s Else t)\\ s2" + + (* cf. 14.10, 14.10.1 *) + Loop "\\G\\Norm s0 -(If(e) (s;; While(e) s) Else Skip)\\ s1\\ \\ + G\\Norm s0 -(While(e) s)\\ s1" + +end