src/HOL/MicroJava/J/Eval.thy
changeset 62042 6c6ccf573479
parent 61361 8b5f00202e1a
child 67443 3abf6a722518
equal deleted inserted replaced
62041:52a87574bca9 62042:6c6ccf573479
     6 section \<open>Operational Evaluation (big step) Semantics\<close>
     6 section \<open>Operational Evaluation (big step) Semantics\<close>
     7 
     7 
     8 theory Eval imports State WellType begin
     8 theory Eval imports State WellType begin
     9 
     9 
    10 
    10 
    11   -- "Auxiliary notions"
    11   \<comment> "Auxiliary notions"
    12 
    12 
    13 definition fits :: "java_mb prog \<Rightarrow> state \<Rightarrow> val \<Rightarrow> ty \<Rightarrow> bool" ("_,_\<turnstile>_ fits _"[61,61,61,61]60) where
    13 definition fits :: "java_mb prog \<Rightarrow> state \<Rightarrow> val \<Rightarrow> ty \<Rightarrow> bool" ("_,_\<turnstile>_ fits _"[61,61,61,61]60) where
    14  "G,s\<turnstile>a' fits T  \<equiv> case T of PrimT T' \<Rightarrow> False | RefT T' \<Rightarrow> a'=Null \<or> G\<turnstile>obj_ty(lookup_obj s a')\<preceq>T"
    14  "G,s\<turnstile>a' fits T  \<equiv> case T of PrimT T' \<Rightarrow> False | RefT T' \<Rightarrow> a'=Null \<or> G\<turnstile>obj_ty(lookup_obj s a')\<preceq>T"
    15 
    15 
    16 definition catch :: "java_mb prog \<Rightarrow> xstate \<Rightarrow> cname \<Rightarrow> bool" ("_,_\<turnstile>catch _"[61,61,61]60) where
    16 definition catch :: "java_mb prog \<Rightarrow> xstate \<Rightarrow> cname \<Rightarrow> bool" ("_,_\<turnstile>catch _"[61,61,61]60) where
    21 
    21 
    22 definition new_xcpt_var :: "vname \<Rightarrow> xstate \<Rightarrow> xstate" where
    22 definition new_xcpt_var :: "vname \<Rightarrow> xstate \<Rightarrow> xstate" where
    23  "new_xcpt_var vn \<equiv>  \<lambda>(x,s). Norm (lupd(vn\<mapsto>the x) s)"
    23  "new_xcpt_var vn \<equiv>  \<lambda>(x,s). Norm (lupd(vn\<mapsto>the x) s)"
    24 
    24 
    25 
    25 
    26   -- "Evaluation relations"
    26   \<comment> "Evaluation relations"
    27 
    27 
    28 inductive
    28 inductive
    29   eval :: "[java_mb prog,xstate,expr,val,xstate] => bool "
    29   eval :: "[java_mb prog,xstate,expr,val,xstate] => bool "
    30           ("_ \<turnstile> _ -_\<succ>_-> _" [51,82,60,82,82] 81)
    30           ("_ \<turnstile> _ -_\<succ>_-> _" [51,82,60,82,82] 81)
    31   and evals :: "[java_mb prog,xstate,expr list,
    31   and evals :: "[java_mb prog,xstate,expr list,
    34   and exec :: "[java_mb prog,xstate,stmt,    xstate] => bool "
    34   and exec :: "[java_mb prog,xstate,stmt,    xstate] => bool "
    35           ("_ \<turnstile> _ -_-> _" [51,82,60,82] 81)
    35           ("_ \<turnstile> _ -_-> _" [51,82,60,82] 81)
    36   for G :: "java_mb prog"
    36   for G :: "java_mb prog"
    37 where
    37 where
    38 
    38 
    39   -- "evaluation of expressions"
    39   \<comment> "evaluation of expressions"
    40 
    40 
    41   XcptE:"G\<turnstile>(Some xc,s) -e\<succ>undefined-> (Some xc,s)"  -- "cf. 15.5"
    41   XcptE:"G\<turnstile>(Some xc,s) -e\<succ>undefined-> (Some xc,s)"  \<comment> "cf. 15.5"
    42 
    42 
    43   -- "cf. 15.8.1"
    43   \<comment> "cf. 15.8.1"
    44 | NewC: "[| h = heap s; (a,x) = new_Addr h;
    44 | NewC: "[| h = heap s; (a,x) = new_Addr h;
    45             h'= h(a\<mapsto>(C,init_vars (fields (G,C)))) |] ==>
    45             h'= h(a\<mapsto>(C,init_vars (fields (G,C)))) |] ==>
    46          G\<turnstile>Norm s -NewC C\<succ>Addr a-> c_hupd h' (x,s)"
    46          G\<turnstile>Norm s -NewC C\<succ>Addr a-> c_hupd h' (x,s)"
    47 
    47 
    48   -- "cf. 15.15"
    48   \<comment> "cf. 15.15"
    49 | Cast: "[| G\<turnstile>Norm s0 -e\<succ>v-> (x1,s1);
    49 | Cast: "[| G\<turnstile>Norm s0 -e\<succ>v-> (x1,s1);
    50             x2 = raise_if (\<not> cast_ok G C (heap s1) v) ClassCast x1 |] ==>
    50             x2 = raise_if (\<not> cast_ok G C (heap s1) v) ClassCast x1 |] ==>
    51          G\<turnstile>Norm s0 -Cast C e\<succ>v-> (x2,s1)"
    51          G\<turnstile>Norm s0 -Cast C e\<succ>v-> (x2,s1)"
    52 
    52 
    53   -- "cf. 15.7.1"
    53   \<comment> "cf. 15.7.1"
    54 | Lit:  "G\<turnstile>Norm s -Lit v\<succ>v-> Norm s"
    54 | Lit:  "G\<turnstile>Norm s -Lit v\<succ>v-> Norm s"
    55 
    55 
    56 | BinOp:"[| G\<turnstile>Norm s -e1\<succ>v1-> s1;
    56 | BinOp:"[| G\<turnstile>Norm s -e1\<succ>v1-> s1;
    57             G\<turnstile>s1     -e2\<succ>v2-> s2;
    57             G\<turnstile>s1     -e2\<succ>v2-> s2;
    58             v = (case bop of Eq  => Bool (v1 = v2)
    58             v = (case bop of Eq  => Bool (v1 = v2)
    59                            | Add => Intg (the_Intg v1 + the_Intg v2)) |] ==>
    59                            | Add => Intg (the_Intg v1 + the_Intg v2)) |] ==>
    60          G\<turnstile>Norm s -BinOp bop e1 e2\<succ>v-> s2"
    60          G\<turnstile>Norm s -BinOp bop e1 e2\<succ>v-> s2"
    61 
    61 
    62   -- "cf. 15.13.1, 15.2"
    62   \<comment> "cf. 15.13.1, 15.2"
    63 | LAcc: "G\<turnstile>Norm s -LAcc v\<succ>the (locals s v)-> Norm s"
    63 | LAcc: "G\<turnstile>Norm s -LAcc v\<succ>the (locals s v)-> Norm s"
    64 
    64 
    65   -- "cf. 15.25.1"
    65   \<comment> "cf. 15.25.1"
    66 | LAss: "[| G\<turnstile>Norm s -e\<succ>v-> (x,(h,l));
    66 | LAss: "[| G\<turnstile>Norm s -e\<succ>v-> (x,(h,l));
    67             l' = (if x = None then l(va\<mapsto>v) else l) |] ==>
    67             l' = (if x = None then l(va\<mapsto>v) else l) |] ==>
    68          G\<turnstile>Norm s -va::=e\<succ>v-> (x,(h,l'))"
    68          G\<turnstile>Norm s -va::=e\<succ>v-> (x,(h,l'))"
    69 
    69 
    70   -- "cf. 15.10.1, 15.2"
    70   \<comment> "cf. 15.10.1, 15.2"
    71 | FAcc: "[| G\<turnstile>Norm s0 -e\<succ>a'-> (x1,s1); 
    71 | FAcc: "[| G\<turnstile>Norm s0 -e\<succ>a'-> (x1,s1); 
    72             v = the (snd (the (heap s1 (the_Addr a'))) (fn,T)) |] ==>
    72             v = the (snd (the (heap s1 (the_Addr a'))) (fn,T)) |] ==>
    73          G\<turnstile>Norm s0 -{T}e..fn\<succ>v-> (np a' x1,s1)"
    73          G\<turnstile>Norm s0 -{T}e..fn\<succ>v-> (np a' x1,s1)"
    74 
    74 
    75   -- "cf. 15.25.1"
    75   \<comment> "cf. 15.25.1"
    76 | FAss: "[| G\<turnstile>     Norm s0  -e1\<succ>a'-> (x1,s1); a = the_Addr a';
    76 | FAss: "[| G\<turnstile>     Norm s0  -e1\<succ>a'-> (x1,s1); a = the_Addr a';
    77             G\<turnstile>(np a' x1,s1) -e2\<succ>v -> (x2,s2);
    77             G\<turnstile>(np a' x1,s1) -e2\<succ>v -> (x2,s2);
    78             h  = heap s2; (c,fs) = the (h a);
    78             h  = heap s2; (c,fs) = the (h a);
    79             h' = h(a\<mapsto>(c,(fs((fn,T)\<mapsto>v)))) |] ==>
    79             h' = h(a\<mapsto>(c,(fs((fn,T)\<mapsto>v)))) |] ==>
    80          G\<turnstile>Norm s0 -{T}e1..fn:=e2\<succ>v-> c_hupd h' (x2,s2)"
    80          G\<turnstile>Norm s0 -{T}e1..fn:=e2\<succ>v-> c_hupd h' (x2,s2)"
    81 
    81 
    82   -- "cf. 15.11.4.1, 15.11.4.2, 15.11.4.4, 15.11.4.5, 14.15"
    82   \<comment> "cf. 15.11.4.1, 15.11.4.2, 15.11.4.4, 15.11.4.5, 14.15"
    83 | Call: "[| G\<turnstile>Norm s0 -e\<succ>a'-> s1; a = the_Addr a';
    83 | Call: "[| G\<turnstile>Norm s0 -e\<succ>a'-> s1; a = the_Addr a';
    84             G\<turnstile>s1 -ps[\<succ>]pvs-> (x,(h,l)); dynT = fst (the (h a));
    84             G\<turnstile>s1 -ps[\<succ>]pvs-> (x,(h,l)); dynT = fst (the (h a));
    85             (md,rT,pns,lvars,blk,res) = the (method (G,dynT) (mn,pTs));
    85             (md,rT,pns,lvars,blk,res) = the (method (G,dynT) (mn,pTs));
    86             G\<turnstile>(np a' x,(h,(init_vars lvars)(pns[\<mapsto>]pvs)(This\<mapsto>a'))) -blk-> s3;
    86             G\<turnstile>(np a' x,(h,(init_vars lvars)(pns[\<mapsto>]pvs)(This\<mapsto>a'))) -blk-> s3;
    87             G\<turnstile> s3 -res\<succ>v -> (x4,s4) |] ==>
    87             G\<turnstile> s3 -res\<succ>v -> (x4,s4) |] ==>
    88          G\<turnstile>Norm s0 -{C}e..mn({pTs}ps)\<succ>v-> (x4,(heap s4,l))"
    88          G\<turnstile>Norm s0 -{C}e..mn({pTs}ps)\<succ>v-> (x4,(heap s4,l))"
    89 
    89 
    90 
    90 
    91   -- "evaluation of expression lists"
    91   \<comment> "evaluation of expression lists"
    92 
    92 
    93   -- "cf. 15.5"
    93   \<comment> "cf. 15.5"
    94 | XcptEs:"G\<turnstile>(Some xc,s) -e[\<succ>]undefined-> (Some xc,s)"
    94 | XcptEs:"G\<turnstile>(Some xc,s) -e[\<succ>]undefined-> (Some xc,s)"
    95 
    95 
    96   -- "cf. 15.11.???"
    96   \<comment> "cf. 15.11.???"
    97 | Nil:  "G\<turnstile>Norm s0 -[][\<succ>][]-> Norm s0"
    97 | Nil:  "G\<turnstile>Norm s0 -[][\<succ>][]-> Norm s0"
    98 
    98 
    99   -- "cf. 15.6.4"
    99   \<comment> "cf. 15.6.4"
   100 | Cons: "[| G\<turnstile>Norm s0 -e  \<succ> v -> s1;
   100 | Cons: "[| G\<turnstile>Norm s0 -e  \<succ> v -> s1;
   101             G\<turnstile>     s1 -es[\<succ>]vs-> s2 |] ==>
   101             G\<turnstile>     s1 -es[\<succ>]vs-> s2 |] ==>
   102          G\<turnstile>Norm s0 -e#es[\<succ>]v#vs-> s2"
   102          G\<turnstile>Norm s0 -e#es[\<succ>]v#vs-> s2"
   103 
   103 
   104 
   104 
   105   -- "execution of statements"
   105   \<comment> "execution of statements"
   106 
   106 
   107   -- "cf. 14.1"
   107   \<comment> "cf. 14.1"
   108 | XcptS:"G\<turnstile>(Some xc,s) -c-> (Some xc,s)"
   108 | XcptS:"G\<turnstile>(Some xc,s) -c-> (Some xc,s)"
   109 
   109 
   110   -- "cf. 14.5"
   110   \<comment> "cf. 14.5"
   111 | Skip: "G\<turnstile>Norm s -Skip-> Norm s"
   111 | Skip: "G\<turnstile>Norm s -Skip-> Norm s"
   112 
   112 
   113   -- "cf. 14.7"
   113   \<comment> "cf. 14.7"
   114 | Expr: "[| G\<turnstile>Norm s0 -e\<succ>v-> s1 |] ==>
   114 | Expr: "[| G\<turnstile>Norm s0 -e\<succ>v-> s1 |] ==>
   115          G\<turnstile>Norm s0 -Expr e-> s1"
   115          G\<turnstile>Norm s0 -Expr e-> s1"
   116 
   116 
   117   -- "cf. 14.2"
   117   \<comment> "cf. 14.2"
   118 | Comp: "[| G\<turnstile>Norm s0 -c1-> s1;
   118 | Comp: "[| G\<turnstile>Norm s0 -c1-> s1;
   119             G\<turnstile>     s1 -c2-> s2|] ==>
   119             G\<turnstile>     s1 -c2-> s2|] ==>
   120          G\<turnstile>Norm s0 -c1;; c2-> s2"
   120          G\<turnstile>Norm s0 -c1;; c2-> s2"
   121 
   121 
   122   -- "cf. 14.8.2"
   122   \<comment> "cf. 14.8.2"
   123 | Cond: "[| G\<turnstile>Norm s0  -e\<succ>v-> s1;
   123 | Cond: "[| G\<turnstile>Norm s0  -e\<succ>v-> s1;
   124             G\<turnstile> s1 -(if the_Bool v then c1 else c2)-> s2|] ==>
   124             G\<turnstile> s1 -(if the_Bool v then c1 else c2)-> s2|] ==>
   125          G\<turnstile>Norm s0 -If(e) c1 Else c2-> s2"
   125          G\<turnstile>Norm s0 -If(e) c1 Else c2-> s2"
   126 
   126 
   127   -- "cf. 14.10, 14.10.1"
   127   \<comment> "cf. 14.10, 14.10.1"
   128 | LoopF:"[| G\<turnstile>Norm s0 -e\<succ>v-> s1; \<not>the_Bool v |] ==>
   128 | LoopF:"[| G\<turnstile>Norm s0 -e\<succ>v-> s1; \<not>the_Bool v |] ==>
   129          G\<turnstile>Norm s0 -While(e) c-> s1"
   129          G\<turnstile>Norm s0 -While(e) c-> s1"
   130 | LoopT:"[| G\<turnstile>Norm s0 -e\<succ>v-> s1;  the_Bool v;
   130 | LoopT:"[| G\<turnstile>Norm s0 -e\<succ>v-> s1;  the_Bool v;
   131       G\<turnstile>s1 -c-> s2; G\<turnstile>s2 -While(e) c-> s3 |] ==>
   131       G\<turnstile>s1 -c-> s2; G\<turnstile>s2 -While(e) c-> s3 |] ==>
   132          G\<turnstile>Norm s0 -While(e) c-> s3"
   132          G\<turnstile>Norm s0 -While(e) c-> s3"