src/HOL/MicroJava/J/Eval.thy
changeset 10056 9f84ffa4a8d0
parent 10042 7164dc0d24d8
child 10061 fe82134773dc
equal deleted inserted replaced
10055:2264bdd8becc 10056:9f84ffa4a8d0
    13   eval  :: "java_mb prog => (xstate \\<times> expr      \\<times> val      \\<times> xstate) set"
    13   eval  :: "java_mb prog => (xstate \\<times> expr      \\<times> val      \\<times> xstate) set"
    14   evals :: "java_mb prog => (xstate \\<times> expr list \\<times> val list \\<times> xstate) set"
    14   evals :: "java_mb prog => (xstate \\<times> expr list \\<times> val list \\<times> xstate) set"
    15   exec  :: "java_mb prog => (xstate \\<times> stmt                 \\<times> xstate) set"
    15   exec  :: "java_mb prog => (xstate \\<times> stmt                 \\<times> xstate) set"
    16 
    16 
    17 syntax
    17 syntax
    18   eval :: "[java_mb prog,xstate,expr,val,xstate] => bool "("_\\<turnstile>_ -_\\<succ>_-> _"[51,82,82,82,82]81)
    18   eval :: "[java_mb prog,xstate,expr,val,xstate] => bool "
       
    19           ("_\\<turnstile>_ -_\\<succ>_-> _"[51,82,82,82,82]81)
    19   evals:: "[java_mb prog,xstate,expr list,
    20   evals:: "[java_mb prog,xstate,expr list,
    20 	                      val list,xstate] => bool "("_\\<turnstile>_ -_[\\<succ>]_-> _"[51,82,51,51,82]81)
    21                         val list,xstate] => bool "
    21   exec :: "[java_mb prog,xstate,stmt,    xstate] => bool "("_\\<turnstile>_ -_-> _"  [51,82,82,   82]81)
    22           ("_\\<turnstile>_ -_[\\<succ>]_-> _"[51,82,51,51,82]81)
       
    23   exec :: "[java_mb prog,xstate,stmt,    xstate] => bool "
       
    24           ("_\\<turnstile>_ -_-> _"  [51,82,82,   82]81)
    22 
    25 
    23 translations
    26 translations
    24   "G\\<turnstile>s -e \\<succ> v-> (x,s')" <= "(s, e, v, x, s') \\<in> eval  G"
    27   "G\\<turnstile>s -e \\<succ> v-> (x,s')" <= "(s, e, v, x, s') \\<in> eval  G"
    25   "G\\<turnstile>s -e \\<succ> v->    s' " == "(s, e, v,    s' ) \\<in> eval  G"
    28   "G\\<turnstile>s -e \\<succ> v->    s' " == "(s, e, v,    s') \\<in> eval  G"
    26   "G\\<turnstile>s -e[\\<succ>]v-> (x,s')" <= "(s, e, v, x, s') \\<in> evals G"
    29   "G\\<turnstile>s -e[\\<succ>]v-> (x,s')" <= "(s, e, v, x, s') \\<in> evals G"
    27   "G\\<turnstile>s -e[\\<succ>]v->    s' " == "(s, e, v,    s' ) \\<in> evals G"
    30   "G\\<turnstile>s -e[\\<succ>]v->    s' " == "(s, e, v,    s') \\<in> evals G"
    28   "G\\<turnstile>s -c    -> (x,s')" <= "(s, c, x, s') \\<in> exec  G"
    31   "G\\<turnstile>s -c    -> (x,s')" <= "(s, c, x, s') \\<in> exec G"
    29   "G\\<turnstile>s -c    ->    s' " == "(s, c,    s') \\<in> exec  G"
    32   "G\\<turnstile>s -c    ->    s' " == "(s, c,    s') \\<in> exec G"
    30 
    33 
    31 inductive "eval G" "evals G" "exec G" intrs
    34 inductive "eval G" "evals G" "exec G" intrs
    32 
    35 
    33 (* evaluation of expressions *)
    36 (* evaluation of expressions *)
    34 
    37 
    35   (* cf. 15.5 *)
    38   (* cf. 15.5 *)
    36   XcptE				  "G\\<turnstile>(Some xc,s) -e\\<succ>arbitrary-> (Some xc,s)"
    39   XcptE "G\\<turnstile>(Some xc,s) -e\\<succ>arbitrary-> (Some xc,s)"
    37 
    40 
    38   (* cf. 15.8.1 *)
    41   (* cf. 15.8.1 *)
    39   NewC	"[|h = heap s; (a,x) = new_Addr h;
    42   NewC  "[| h = heap s; (a,x) = new_Addr h;
    40 	  h'= h(a\\<mapsto>(C,init_vars (fields (G,C))))|] ==>
    43             h'= h(a\\<mapsto>(C,init_vars (fields (G,C)))) |] ==>
    41 				   G\\<turnstile>Norm s -NewC C\\<succ>Addr a-> c_hupd h' (x,s)"
    44          G\\<turnstile>Norm s -NewC C\\<succ>Addr a-> c_hupd h' (x,s)"
    42 
    45 
    43   (* cf. 15.15 *)
    46   (* cf. 15.15 *)
    44   Cast	"[|G\\<turnstile>Norm s0 -e\\<succ>v-> (x1,s1);
    47   Cast  "[| G\\<turnstile>Norm s0 -e\\<succ>v-> (x1,s1);
    45 	  x2=raise_if (\\<not> cast_ok G C (heap s1) v) ClassCast x1|] ==>
    48             x2 = raise_if (\\<not> cast_ok G C (heap s1) v) ClassCast x1 |] ==>
    46 			        G\\<turnstile>Norm s0 -Cast C e\\<succ>v-> (x2,s1)"
    49          G\\<turnstile>Norm s0 -Cast C e\\<succ>v-> (x2,s1)"
    47 
    50 
    48   (* cf. 15.7.1 *)
    51   (* cf. 15.7.1 *)
    49   Lit				   "G\\<turnstile>Norm s -Lit v\\<succ>v-> Norm s"
    52   Lit   "G\\<turnstile>Norm s -Lit v\\<succ>v-> Norm s"
    50 
    53 
    51   BinOp "[|G\\<turnstile>Norm s -e1\\<succ>v1-> s1;
    54   BinOp "[| G\\<turnstile>Norm s -e1\\<succ>v1-> s1;
    52 	  G\\<turnstile>s1     -e2\\<succ>v2-> s2;
    55             G\\<turnstile>s1     -e2\\<succ>v2-> s2;
    53 	  v = (case bop of Eq  => Bool (v1 = v2)
    56             v = (case bop of Eq  => Bool (v1 = v2)
    54 	                 | Add => Intg (the_Intg v1 + the_Intg v2))|] ==>
    57                            | Add => Intg (the_Intg v1 + the_Intg v2)) |] ==>
    55 				   G\\<turnstile>Norm s -BinOp bop e1 e2\\<succ>v-> s2"
    58          G\\<turnstile>Norm s -BinOp bop e1 e2\\<succ>v-> s2"
    56 
    59 
    57   (* cf. 15.13.1, 15.2 *)
    60   (* cf. 15.13.1, 15.2 *)
    58   LAcc				  "G\\<turnstile>Norm s -LAcc v\\<succ>the (locals s v)-> Norm s"
    61   LAcc  "G\\<turnstile>Norm s -LAcc v\\<succ>the (locals s v)-> Norm s"
    59 
    62 
    60   (* cf. 15.25.1 *)
    63   (* cf. 15.25.1 *)
    61   LAss  "[|G\\<turnstile>Norm s -e\\<succ>v->  (x,(h,l));
    64   LAss  "[| G\\<turnstile>Norm s -e\\<succ>v-> (x,(h,l));
    62 	  l' = (if x = None then l(va\\<mapsto>v) else l)|] ==>
    65             l' = (if x = None then l(va\\<mapsto>v) else l) |] ==>
    63 				   G\\<turnstile>Norm s -va::=e\\<succ>v-> (x,(h,l'))"
    66          G\\<turnstile>Norm s -va::=e\\<succ>v-> (x,(h,l'))"
    64 
    67 
    65 
    68 
    66   (* cf. 15.10.1, 15.2 *)
    69   (* cf. 15.10.1, 15.2 *)
    67   FAcc	"[|G\\<turnstile>Norm s0 -e\\<succ>a'-> (x1,s1); 
    70   FAcc  "[| G\\<turnstile>Norm s0 -e\\<succ>a'-> (x1,s1); 
    68 	  v = the (snd (the (heap s1 (the_Addr a'))) (fn,T))|] ==>
    71             v = the (snd (the (heap s1 (the_Addr a'))) (fn,T)) |] ==>
    69 				 G\\<turnstile>Norm s0 -{T}e..fn\\<succ>v-> (np a' x1,s1)"
    72          G\\<turnstile>Norm s0 -{T}e..fn\\<succ>v-> (np a' x1,s1)"
    70 
    73 
    71   (* cf. 15.25.1 *)
    74   (* cf. 15.25.1 *)
    72   FAss  "[|G\\<turnstile>     Norm s0  -e1\\<succ>a'-> (x1,s1); a = the_Addr a';
    75   FAss  "[| G\\<turnstile>     Norm s0  -e1\\<succ>a'-> (x1,s1); a = the_Addr a';
    73 	  G\\<turnstile>(np a' x1,s1) -e2\\<succ>v -> (x2,s2);
    76             G\\<turnstile>(np a' x1,s1) -e2\\<succ>v -> (x2,s2);
    74 	  h = heap s2; (c,fs) = the (h a);
    77             h  = heap s2; (c,fs) = the (h a);
    75 	  h' = h(a\\<mapsto>(c,(fs((fn,T)\\<mapsto>v))))|] ==>
    78             h' = h(a\\<mapsto>(c,(fs((fn,T)\\<mapsto>v)))) |] ==>
    76 			  G\\<turnstile>Norm s0 -{T}e1..fn:=e2\\<succ>v-> c_hupd h' (x2,s2)"
    79          G\\<turnstile>Norm s0 -{T}e1..fn:=e2\\<succ>v-> c_hupd h' (x2,s2)"
    77 
    80 
    78   (* cf. 15.11.4.1, 15.11.4.2, 15.11.4.4, 15.11.4.5, 14.15 *)
    81   (* cf. 15.11.4.1, 15.11.4.2, 15.11.4.4, 15.11.4.5, 14.15 *)
    79   Call	"[|G\\<turnstile>Norm s0 -e\\<succ>a'-> s1; a = the_Addr a';
    82   Call  "[| G\\<turnstile>Norm s0 -e\\<succ>a'-> s1; a = the_Addr a';
    80 	   G\\<turnstile>s1 -ps[\\<succ>]pvs-> (x,(h,l)); dynT = fst (the (h a));
    83             G\\<turnstile>s1 -ps[\\<succ>]pvs-> (x,(h,l)); dynT = fst (the (h a));
    81 	   (md,rT,pns,lvars,blk,res) = the (method (G,dynT) (mn,pTs));
    84             (md,rT,pns,lvars,blk,res) = the (method (G,dynT) (mn,pTs));
    82 	   G\\<turnstile>(np a' x,(h,(init_vars lvars)(pns[\\<mapsto>]pvs)(This\\<mapsto>a'))) -blk-> s3;
    85             G\\<turnstile>(np a' x,(h,(init_vars lvars)(pns[\\<mapsto>]pvs)(This\\<mapsto>a'))) -blk-> s3;
    83 	   G\\<turnstile>     s3 -res\\<succ>v -> (x4,s4)|] ==>
    86             G\\<turnstile> s3 -res\\<succ>v -> (x4,s4) |] ==>
    84 			    G\\<turnstile>Norm s0 -e..mn({pTs}ps)\\<succ>v-> (x4,(heap s4,l))"
    87          G\\<turnstile>Norm s0 -e..mn({pTs}ps)\\<succ>v-> (x4,(heap s4,l))"
    85 
    88 
    86 
    89 
    87 (* evaluation of expression lists *)
    90 (* evaluation of expression lists *)
    88 
    91 
    89   (* cf. 15.5 *)
    92   (* cf. 15.5 *)
    90   XcptEs			  "G\\<turnstile>(Some xc,s) -e[\\<succ>]arbitrary-> (Some xc,s)"
    93   XcptEs "G\\<turnstile>(Some xc,s) -e[\\<succ>]arbitrary-> (Some xc,s)"
    91 
    94 
    92   (* cf. 15.11.??? *)
    95   (* cf. 15.11.??? *)
    93   Nil
    96   Nil   "G\\<turnstile>Norm s0 -[][\\<succ>][]-> Norm s0"
    94 				    "G\\<turnstile>Norm s0 -[][\\<succ>][]-> Norm s0"
       
    95 
    97 
    96   (* cf. 15.6.4 *)
    98   (* cf. 15.6.4 *)
    97   Cons	"[|G\\<turnstile>Norm s0 -e  \\<succ> v -> s1;
    99   Cons  "[| G\\<turnstile>Norm s0 -e  \\<succ> v -> s1;
    98            G\\<turnstile>     s1 -es[\\<succ>]vs-> s2|] ==>
   100             G\\<turnstile>     s1 -es[\\<succ>]vs-> s2 |] ==>
    99 				   G\\<turnstile>Norm s0 -e#es[\\<succ>]v#vs-> s2"
   101          G\\<turnstile>Norm s0 -e#es[\\<succ>]v#vs-> s2"
   100 
   102 
   101 (* execution of statements *)
   103 (* execution of statements *)
   102 
   104 
   103   (* cf. 14.1 *)
   105   (* cf. 14.1 *)
   104   XcptS				 "G\\<turnstile>(Some xc,s) -s0-> (Some xc,s)"
   106   XcptS "G\\<turnstile>(Some xc,s) -s0-> (Some xc,s)"
   105 
   107 
   106   (* cf. 14.5 *)
   108   (* cf. 14.5 *)
   107   Skip	 			    "G\\<turnstile>Norm s -Skip-> Norm s"
   109   Skip  "G\\<turnstile>Norm s -Skip-> Norm s"
   108 
   110 
   109   (* cf. 14.7 *)
   111   (* cf. 14.7 *)
   110   Expr	"[|G\\<turnstile>Norm s0 -e\\<succ>v-> s1|] ==>
   112   Expr  "[| G\\<turnstile>Norm s0 -e\\<succ>v-> s1 |] ==>
   111 				  G\\<turnstile>Norm s0 -Expr e-> s1"
   113          G\\<turnstile>Norm s0 -Expr e-> s1"
   112 
   114 
   113   (* cf. 14.2 *)
   115   (* cf. 14.2 *)
   114   Comp	"[|G\\<turnstile>Norm s0 -s -> s1;
   116   Comp  "[| G\\<turnstile>Norm s0 -s -> s1;
   115 	  G\\<turnstile>     s1 -t -> s2|] ==>
   117             G\\<turnstile> s1 -t -> s2|] ==>
   116 				 G\\<turnstile>Norm s0 -(s;; t)-> s2"
   118          G\\<turnstile>Norm s0 -(s;; t)-> s2"
   117 
   119 
   118   (* cf. 14.8.2 *)
   120   (* cf. 14.8.2 *)
   119   Cond	"[|G\\<turnstile>Norm s0  -e \\<succ>v-> s1;
   121   Cond  "[| G\\<turnstile>Norm s0  -e \\<succ>v-> s1;
   120 	  G\\<turnstile>     s1 -(if  the_Bool v then s else t)-> s2|] ==>
   122             G\\<turnstile> s1 -(if  the_Bool v then s else t)-> s2|] ==>
   121 		        G\\<turnstile>Norm s0 -(If(e) s Else t)-> s2"
   123          G\\<turnstile>Norm s0 -(If(e) s Else t)-> s2"
   122 
   124 
   123   (* cf. 14.10, 14.10.1 *)
   125   (* cf. 14.10, 14.10.1 *)
   124   Loop	"[|G\\<turnstile>Norm s0 -(If(e) (s;; While(e) s) Else Skip)-> s1|] ==>
   126   Loop  "[| G\\<turnstile>Norm s0 -(If(e) (s;; While(e) s) Else Skip)-> s1 |] ==>
   125 			    G\\<turnstile>Norm s0 -(While(e) s)-> s1"
   127          G\\<turnstile>Norm s0 -(While(e) s)-> s1"
   126 
   128 
   127 end
   129 end