diff -r 27793282952c -r 33e290a70445 src/HOL/MicroJava/J/Eval.thy --- a/src/HOL/MicroJava/J/Eval.thy Thu Jan 18 18:39:43 2001 +0100 +++ b/src/HOL/MicroJava/J/Eval.thy Thu Jan 18 20:23:51 2001 +0100 @@ -112,7 +112,7 @@ (* execution of statements *) (* cf. 14.1 *) - XcptS "G\\(Some xc,s) -s0-> (Some xc,s)" + XcptS "G\\(Some xc,s) -c-> (Some xc,s)" (* cf. 14.5 *) Skip "G\\Norm s -Skip-> Norm s" @@ -122,20 +122,21 @@ 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" + Comp "[| G\\Norm s0 -c1-> s1; + G\\ s1 -c2-> s2|] ==> + G\\Norm s0 -c1;; c2-> 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" + G\\ s1 -(if the_Bool v then c1 else c2)-> s2|] ==> + G\\Norm s0 -If(e) c1 Else c2-> s2" (* cf. 14.10, 14.10.1 *) - Loop "[| G\\Norm s0 -e\\v-> s1; - if the_Bool v then G\\s1 -s-> s2 \\ G\\s2 -While(e) s-> s3 - else s3 = s1 |] ==> - G\\Norm s0 -While(e) s-> s3" + LoopF "[| G\\Norm s0 -e\\v-> s1; \\the_Bool v |] ==> + G\\Norm s0 -While(e) c-> s1" + LoopT "[| G\\Norm s0 -e\\v-> s1; the_Bool v; + G\\s1 -c-> s2; G\\s2 -While(e) c-> s3 |] ==> + G\\Norm s0 -While(e) c-> s3" monos if_def2