diff -r 52a87574bca9 -r 6c6ccf573479 src/HOL/MicroJava/J/Eval.thy --- a/src/HOL/MicroJava/J/Eval.thy Sat Jan 02 18:46:36 2016 +0100 +++ b/src/HOL/MicroJava/J/Eval.thy Sat Jan 02 18:48:45 2016 +0100 @@ -8,7 +8,7 @@ theory Eval imports State WellType begin - -- "Auxiliary notions" + \ "Auxiliary notions" definition fits :: "java_mb prog \ state \ val \ ty \ bool" ("_,_\_ fits _"[61,61,61,61]60) where "G,s\a' fits T \ case T of PrimT T' \ False | RefT T' \ a'=Null \ G\obj_ty(lookup_obj s a')\T" @@ -23,7 +23,7 @@ "new_xcpt_var vn \ \(x,s). Norm (lupd(vn\the x) s)" - -- "Evaluation relations" + \ "Evaluation relations" inductive eval :: "[java_mb prog,xstate,expr,val,xstate] => bool " @@ -36,21 +36,21 @@ for G :: "java_mb prog" where - -- "evaluation of expressions" + \ "evaluation of expressions" - XcptE:"G\(Some xc,s) -e\undefined-> (Some xc,s)" -- "cf. 15.5" + XcptE:"G\(Some xc,s) -e\undefined-> (Some xc,s)" \ "cf. 15.5" - -- "cf. 15.8.1" + \ "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" + \ "cf. 15.15" | Cast: "[| G\Norm s0 -e\v-> (x1,s1); x2 = raise_if (\ cast_ok G C (heap s1) v) ClassCast x1 |] ==> G\Norm s0 -Cast C e\v-> (x2,s1)" - -- "cf. 15.7.1" + \ "cf. 15.7.1" | Lit: "G\Norm s -Lit v\v-> Norm s" | BinOp:"[| G\Norm s -e1\v1-> s1; @@ -59,27 +59,27 @@ | Add => Intg (the_Intg v1 + the_Intg v2)) |] ==> G\Norm s -BinOp bop e1 e2\v-> s2" - -- "cf. 15.13.1, 15.2" + \ "cf. 15.13.1, 15.2" | LAcc: "G\Norm s -LAcc v\the (locals s v)-> Norm s" - -- "cf. 15.25.1" + \ "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" + \ "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" + \ "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" + \ "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 (method (G,dynT) (mn,pTs)); @@ -88,43 +88,43 @@ G\Norm s0 -{C}e..mn({pTs}ps)\v-> (x4,(heap s4,l))" - -- "evaluation of expression lists" + \ "evaluation of expression lists" - -- "cf. 15.5" + \ "cf. 15.5" | XcptEs:"G\(Some xc,s) -e[\]undefined-> (Some xc,s)" - -- "cf. 15.11.???" + \ "cf. 15.11.???" | Nil: "G\Norm s0 -[][\][]-> Norm s0" - -- "cf. 15.6.4" + \ "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" + \ "execution of statements" - -- "cf. 14.1" + \ "cf. 14.1" | XcptS:"G\(Some xc,s) -c-> (Some xc,s)" - -- "cf. 14.5" + \ "cf. 14.5" | Skip: "G\Norm s -Skip-> Norm s" - -- "cf. 14.7" + \ "cf. 14.7" | Expr: "[| G\Norm s0 -e\v-> s1 |] ==> G\Norm s0 -Expr e-> s1" - -- "cf. 14.2" + \ "cf. 14.2" | Comp: "[| G\Norm s0 -c1-> s1; G\ s1 -c2-> s2|] ==> G\Norm s0 -c1;; c2-> s2" - -- "cf. 14.8.2" + \ "cf. 14.8.2" | Cond: "[| G\Norm s0 -e\v-> s1; 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" + \ "cf. 14.10, 14.10.1" | 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;