Theory Eval

Up to index of Isabelle/HOL/MicroJava

theory Eval
imports State WellType
begin

(*  Title:      HOL/MicroJava/J/Eval.thy
    ID:         $Id: Eval.thy,v 1.27 2007/07/11 09:32:10 berghofe Exp $
    Author:     David von Oheimb
    Copyright   1999 Technische Universitaet Muenchen
*)

header {* \isaheader{Operational Evaluation (big step) Semantics} *}

theory Eval imports State WellType begin


  -- "Auxiliary notions"

constdefs
  fits    :: "java_mb prog => state => val => ty => bool" ("_,_\<turnstile>_ fits _"[61,61,61,61]60)
 "G,s\<turnstile>a' fits T  ≡ case T of PrimT T' => False | RefT T' => a'=Null ∨ G\<turnstile>obj_ty(lookup_obj s a')\<preceq>T"

constdefs
  catch ::"java_mb prog => xstate => cname => bool" ("_,_\<turnstile>catch _"[61,61,61]60)
 "G,s\<turnstile>catch C≡  case abrupt s of None => False | Some a => G,store s\<turnstile> a fits Class C"



constdefs
  lupd       :: "vname => val => state => state"        ("lupd'(_\<mapsto>_')"[10,10]1000)
 "lupd vn v   ≡ λ (hp,loc). (hp, (loc(vn\<mapsto>v)))"

constdefs
  new_xcpt_var :: "vname => xstate => xstate"
 "new_xcpt_var vn ≡  λ(x,s). Norm (lupd(vn\<mapsto>the x) s)"


  -- "Evaluation relations"

inductive
  eval :: "[java_mb prog,xstate,expr,val,xstate] => bool "
          ("_ \<turnstile> _ -_\<succ>_-> _" [51,82,60,82,82] 81)
  and evals :: "[java_mb prog,xstate,expr list,
                        val list,xstate] => bool "
          ("_ \<turnstile> _ -_[\<succ>]_-> _" [51,82,60,51,82] 81)
  and exec :: "[java_mb prog,xstate,stmt,    xstate] => bool "
          ("_ \<turnstile> _ -_-> _" [51,82,60,82] 81)
  for G :: "java_mb prog"
where

  -- "evaluation of expressions"

  XcptE:"G\<turnstile>(Some xc,s) -e\<succ>arbitrary-> (Some xc,s)"  -- "cf. 15.5"

  -- "cf. 15.8.1"
| NewC: "[| h = heap s; (a,x) = new_Addr h;
            h'= h(a\<mapsto>(C,init_vars (fields (G,C)))) |] ==>
         G\<turnstile>Norm s -NewC C\<succ>Addr a-> c_hupd h' (x,s)"

  -- "cf. 15.15"
| Cast: "[| G\<turnstile>Norm s0 -e\<succ>v-> (x1,s1);
            x2 = raise_if (¬ cast_ok G C (heap s1) v) ClassCast x1 |] ==>
         G\<turnstile>Norm s0 -Cast C e\<succ>v-> (x2,s1)"

  -- "cf. 15.7.1"
| Lit:  "G\<turnstile>Norm s -Lit v\<succ>v-> Norm s"

| BinOp:"[| G\<turnstile>Norm s -e1\<succ>v1-> s1;
            G\<turnstile>s1     -e2\<succ>v2-> s2;
            v = (case bop of Eq  => Bool (v1 = v2)
                           | Add => Intg (the_Intg v1 + the_Intg v2)) |] ==>
         G\<turnstile>Norm s -BinOp bop e1 e2\<succ>v-> s2"

  -- "cf. 15.13.1, 15.2"
| LAcc: "G\<turnstile>Norm s -LAcc v\<succ>the (locals s v)-> Norm s"

  -- "cf. 15.25.1"
| LAss: "[| G\<turnstile>Norm s -e\<succ>v-> (x,(h,l));
            l' = (if x = None then l(va\<mapsto>v) else l) |] ==>
         G\<turnstile>Norm s -va::=e\<succ>v-> (x,(h,l'))"

  -- "cf. 15.10.1, 15.2"
| FAcc: "[| G\<turnstile>Norm s0 -e\<succ>a'-> (x1,s1); 
            v = the (snd (the (heap s1 (the_Addr a'))) (fn,T)) |] ==>
         G\<turnstile>Norm s0 -{T}e..fn\<succ>v-> (np a' x1,s1)"

  -- "cf. 15.25.1"
| FAss: "[| G\<turnstile>     Norm s0  -e1\<succ>a'-> (x1,s1); a = the_Addr a';
            G\<turnstile>(np a' x1,s1) -e2\<succ>v -> (x2,s2);
            h  = heap s2; (c,fs) = the (h a);
            h' = h(a\<mapsto>(c,(fs((fn,T)\<mapsto>v)))) |] ==>
         G\<turnstile>Norm s0 -{T}e1..fn:=e2\<succ>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\<turnstile>Norm s0 -e\<succ>a'-> s1; a = the_Addr a';
            G\<turnstile>s1 -ps[\<succ>]pvs-> (x,(h,l)); dynT = fst (the (h a));
            (md,rT,pns,lvars,blk,res) = the (method (G,dynT) (mn,pTs));
            G\<turnstile>(np a' x,(h,(init_vars lvars)(pns[\<mapsto>]pvs)(This\<mapsto>a'))) -blk-> s3;
            G\<turnstile> s3 -res\<succ>v -> (x4,s4) |] ==>
         G\<turnstile>Norm s0 -{C}e..mn({pTs}ps)\<succ>v-> (x4,(heap s4,l))"


  -- "evaluation of expression lists"

  -- "cf. 15.5"
| XcptEs:"G\<turnstile>(Some xc,s) -e[\<succ>]arbitrary-> (Some xc,s)"

  -- "cf. 15.11.???"
| Nil:  "G\<turnstile>Norm s0 -[][\<succ>][]-> Norm s0"

  -- "cf. 15.6.4"
| Cons: "[| G\<turnstile>Norm s0 -e  \<succ> v -> s1;
            G\<turnstile>     s1 -es[\<succ>]vs-> s2 |] ==>
         G\<turnstile>Norm s0 -e#es[\<succ>]v#vs-> s2"


  -- "execution of statements"

  -- "cf. 14.1"
| XcptS:"G\<turnstile>(Some xc,s) -c-> (Some xc,s)"

  -- "cf. 14.5"
| Skip: "G\<turnstile>Norm s -Skip-> Norm s"

  -- "cf. 14.7"
| Expr: "[| G\<turnstile>Norm s0 -e\<succ>v-> s1 |] ==>
         G\<turnstile>Norm s0 -Expr e-> s1"

  -- "cf. 14.2"
| Comp: "[| G\<turnstile>Norm s0 -c1-> s1;
            G\<turnstile>     s1 -c2-> s2|] ==>
         G\<turnstile>Norm s0 -c1;; c2-> s2"

  -- "cf. 14.8.2"
| Cond: "[| G\<turnstile>Norm s0  -e\<succ>v-> s1;
            G\<turnstile> s1 -(if the_Bool v then c1 else c2)-> s2|] ==>
         G\<turnstile>Norm s0 -If(e) c1 Else c2-> s2"

  -- "cf. 14.10, 14.10.1"
| LoopF:"[| G\<turnstile>Norm s0 -e\<succ>v-> s1; ¬the_Bool v |] ==>
         G\<turnstile>Norm s0 -While(e) c-> s1"
| LoopT:"[| G\<turnstile>Norm s0 -e\<succ>v-> s1;  the_Bool v;
      G\<turnstile>s1 -c-> s2; G\<turnstile>s2 -While(e) c-> s3 |] ==>
         G\<turnstile>Norm s0 -While(e) c-> s3"


lemmas eval_evals_exec_induct = eval_evals_exec.induct [split_format (complete)]

lemma NewCI: "[|new_Addr (heap s) = (a,x);  
       s' = c_hupd (heap s(a\<mapsto>(C,init_vars (fields (G,C))))) (x,s)|] ==>  
       G\<turnstile>Norm s -NewC C\<succ>Addr a-> s'"
apply (simp (no_asm_simp))
apply (rule eval_evals_exec.NewC)
apply auto
done

lemma eval_evals_exec_no_xcpt: 
 "!!s s'. (G\<turnstile>(x,s) -e \<succ>  v -> (x',s') --> x'=None --> x=None) ∧  
          (G\<turnstile>(x,s) -es[\<succ>]vs-> (x',s') --> x'=None --> x=None) ∧  
          (G\<turnstile>(x,s) -c       -> (x',s') --> x'=None --> x=None)"
apply(simp (no_asm_simp) only: split_tupled_all)
apply(rule eval_evals_exec_induct)
apply(unfold c_hupd_def)
apply(simp_all)
done

lemma eval_no_xcpt: "G\<turnstile>(x,s) -e\<succ>v-> (None,s') ==> x=None"
apply (drule eval_evals_exec_no_xcpt [THEN conjunct1, THEN mp])
apply (fast)
done

lemma evals_no_xcpt: "G\<turnstile>(x,s) -e[\<succ>]v-> (None,s') ==> x=None"
apply (drule eval_evals_exec_no_xcpt [THEN conjunct2, THEN conjunct1, THEN mp])
apply (fast)
done

lemma exec_no_xcpt: "G \<turnstile> (x, s) -c-> (None, s')
==> x = None"
apply (drule eval_evals_exec_no_xcpt [THEN conjunct2 [THEN conjunct2], rule_format])
apply simp+
done


lemma eval_evals_exec_xcpt: 
"!!s s'. (G\<turnstile>(x,s) -e \<succ>  v -> (x',s') --> x=Some xc --> x'=Some xc ∧ s'=s) ∧  
         (G\<turnstile>(x,s) -es[\<succ>]vs-> (x',s') --> x=Some xc --> x'=Some xc ∧ s'=s) ∧  
         (G\<turnstile>(x,s) -c       -> (x',s') --> x=Some xc --> x'=Some xc ∧ s'=s)"
apply (simp (no_asm_simp) only: split_tupled_all)
apply (rule eval_evals_exec_induct)
apply (unfold c_hupd_def)
apply (simp_all)
done

lemma eval_xcpt: "G\<turnstile>(Some xc,s) -e\<succ>v-> (x',s') ==> x'=Some xc ∧  s'=s"
apply (drule eval_evals_exec_xcpt [THEN conjunct1, THEN mp])
apply (fast)
done

lemma exec_xcpt: "G\<turnstile>(Some xc,s) -s0-> (x',s') ==> x'=Some xc ∧  s'=s"
apply (drule eval_evals_exec_xcpt [THEN conjunct2, THEN conjunct2, THEN mp])
apply (fast)
done

end

lemma eval_evals_exec_induct:

  [| !!xc a b e. P1.0 (Some xc) a b e arbitrary (Some xc) a b;
     !!h a b aa x h' C.
        [| h = fst (a, b); (aa, x) = new_Addr h;
           h' = h(aa |-> (C, init_vars (fields (G, C)))) |]
        ==> (λ(v, x, y). P1.0 None a b (NewC C) (Addr aa) v x y)
             (c_hupd h' (x, a, b));
     !!a b e v x1 aa ba x2 C.
        [| G \<turnstile> Norm (a, b) -e\<succ>v-> (x1, aa, ba);
           P1.0 None a b e v x1 aa ba;
           x2 = raise_if (¬ cast_ok G C (fst (aa, ba)) v) ClassCast x1 |]
        ==> P1.0 None a b (Cast C e) v x2 aa ba;
     !!a b v. P1.0 None a b (Lit v) v None a b;
     !!a b e1 v1 aa ab ba e2 v2 ac ad bb v bop.
        [| G \<turnstile> Norm (a, b) -e1\<succ>v1-> (aa, ab, ba);
           P1.0 None a b e1 v1 aa ab ba;
           G \<turnstile> (aa, ab, ba) -e2\<succ>v2-> (ac, ad, bb);
           P1.0 aa ab ba e2 v2 ac ad bb;
           v = (case bop of Eq => Bool (v1 = v2)
                | Add => Intg (the_Intg v1 + the_Intg v2)) |]
        ==> P1.0 None a b (BinOp bop e1 e2) v ac ad bb;
     !!a b v. P1.0 None a b (LAcc v) (the (snd (a, b) v)) None a b;
     !!a b e v x h l l' va.
        [| G \<turnstile> Norm (a, b) -e\<succ>v-> (x, h, l);
           P1.0 None a b e v x h l; l' = (if x = None then l(va |-> v) else l) |]
        ==> P1.0 None a b (va::=e) v x h l';
     !!a b e a' x1 aa ba v fn T.
        [| G \<turnstile> Norm (a, b) -e\<succ>a'-> (x1, aa, ba);
           P1.0 None a b e a' x1 aa ba;
           v = the (snd (the (fst (aa, ba) (the_Addr a'))) (fn, T)) |]
        ==> P1.0 None a b ({T}e..fn) v (np a' x1) aa ba;
     !!a b e1 a' x1 aa ba ab e2 v x2 ac bb h c fs h' fn T.
        [| G \<turnstile> Norm (a, b) -e1\<succ>a'-> (x1, aa, ba);
           P1.0 None a b e1 a' x1 aa ba; ab = the_Addr a';
           G \<turnstile> (np a' x1, aa, ba) -e2\<succ>v-> (x2, ac, bb);
           P1.0 (np a' x1) aa ba e2 v x2 ac bb; h = fst (ac, bb);
           (c, fs) = the (h ab); h' = h(ab |-> (c, fs((fn, T) |-> v))) |]
        ==> (λ(va, x, y). P1.0 None a b ({T}e1..fn:=e2) v va x y)
             (c_hupd h' (x2, ac, bb));
     !!a b e a' aa ab ba ac ps pvs x h l dynT md rT pns lvars blk res mn pTs ad ae
        bb v x4 af bc C.
        [| G \<turnstile> Norm (a, b) -e\<succ>a'-> (aa, ab, ba);
           P1.0 None a b e a' aa ab ba; ac = the_Addr a';
           G \<turnstile> (aa, ab, ba) -ps[\<succ>]pvs-> (x, h, l);
           P2.0 aa ab ba ps pvs x h l; dynT = fst (the (h ac));
           (md, rT, pns, lvars, blk, res) = the (method (G, dynT) (mn, pTs));
           G \<turnstile> (np a' x, h, init_vars lvars(pns [|->] pvs, This |->
                           a')) -blk-> (ad, ae, bb);
           P3.0 (np a' x) h (init_vars lvars(pns [|->] pvs, This |-> a')) blk ad
            ae bb;
           G \<turnstile> (ad, ae, bb) -res\<succ>v-> (x4, af, bc);
           P1.0 ad ae bb res v x4 af bc |]
        ==> P1.0 None a b ({C}e..mn( {pTs}ps)) v x4 (fst (af, bc)) l;
     !!xc a b e. P2.0 (Some xc) a b e arbitrary (Some xc) a b;
     !!a b. P2.0 None a b [] [] None a b;
     !!a b e v aa ab ba es vs ac ad bb.
        [| G \<turnstile> Norm (a, b) -e\<succ>v-> (aa, ab, ba);
           P1.0 None a b e v aa ab ba;
           G \<turnstile> (aa, ab, ba) -es[\<succ>]vs-> (ac, ad, bb);
           P2.0 aa ab ba es vs ac ad bb |]
        ==> P2.0 None a b (e # es) (v # vs) ac ad bb;
     !!xc a b c. P3.0 (Some xc) a b c (Some xc) a b;
     !!a b. P3.0 None a b Skip None a b;
     !!a b e v aa ab ba.
        [| G \<turnstile> Norm (a, b) -e\<succ>v-> (aa, ab, ba);
           P1.0 None a b e v aa ab ba |]
        ==> P3.0 None a b (Expr e) aa ab ba;
     !!a b c1 aa ab ba c2 ac ad bb.
        [| G \<turnstile> Norm (a, b) -c1-> (aa, ab, ba);
           P3.0 None a b c1 aa ab ba;
           G \<turnstile> (aa, ab, ba) -c2-> (ac, ad, bb);
           P3.0 aa ab ba c2 ac ad bb |]
        ==> P3.0 None a b (c1;; c2) ac ad bb;
     !!a b e v aa ab ba c1 c2 ac ad bb.
        [| G \<turnstile> Norm (a, b) -e\<succ>v-> (aa, ab, ba);
           P1.0 None a b e v aa ab ba;
           G \<turnstile> (aa, ab,
                           ba) -(if the_Bool v then c1 else c2)-> (ac, ad, bb);
           P3.0 aa ab ba (if the_Bool v then c1 else c2) ac ad bb |]
        ==> P3.0 None a b (If (e) c1 Else c2) ac ad bb;
     !!a b e v aa ab ba c.
        [| G \<turnstile> Norm (a, b) -e\<succ>v-> (aa, ab, ba);
           P1.0 None a b e v aa ab ba; ¬ the_Bool v |]
        ==> P3.0 None a b (While (e) c) aa ab ba;
     !!a b e v aa ab ba c ac ad bb ae af bc.
        [| G \<turnstile> Norm (a, b) -e\<succ>v-> (aa, ab, ba);
           P1.0 None a b e v aa ab ba; the_Bool v;
           G \<turnstile> (aa, ab, ba) -c-> (ac, ad, bb);
           P3.0 aa ab ba c ac ad bb;
           G \<turnstile> (ac, ad, bb) -While (e) c-> (ae, af, bc);
           P3.0 ac ad bb (While (e) c) ae af bc |]
        ==> P3.0 None a b (While (e) c) ae af bc |]
  ==> (G \<turnstile> (x1a, x1b, x1c) -x2a\<succ>x3a-> (x4a, x4b, x4c) -->
       P1.0 x1a x1b x1c x2a x3a x4a x4b x4c) ∧
      (G \<turnstile> (x5a, x5b, x5c) -x6a[\<succ>]x7a-> (x8a, x8b, x8c) -->
       P2.0 x5a x5b x5c x6a x7a x8a x8b x8c) ∧
      (G \<turnstile> (x9a, x9b, x9c) -x10a-> (x11a, x11b, x11c) -->
       P3.0 x9a x9b x9c x10a x11a x11b x11c)

lemma NewCI:

  [| new_Addr (fst s) = (a, x);
     s' = c_hupd (fst s(a |-> (C, init_vars (fields (G, C))))) (x, s) |]
  ==> G \<turnstile> Norm s -NewC C\<succ>Addr a-> s'

lemma eval_evals_exec_no_xcpt:

  (G \<turnstile> (x, s) -e\<succ>v-> (x', s') --> x' = None --> x = None) ∧
  (G \<turnstile> (x, s) -es[\<succ>]vs-> (x', s') --> x' = None --> x = None) ∧
  (G \<turnstile> (x, s) -c-> (x', s') --> x' = None --> x = None)

lemma eval_no_xcpt:

  G \<turnstile> (x, s) -e\<succ>v-> Norm s' ==> x = None

lemma evals_no_xcpt:

  G \<turnstile> (x, s) -e[\<succ>]v-> Norm s' ==> x = None

lemma exec_no_xcpt:

  G \<turnstile> (x, s) -c-> Norm s' ==> x = None

lemma eval_evals_exec_xcpt:

  (G \<turnstile> (x, s) -e\<succ>v-> (x', s') -->
   x = Some xc --> x' = Some xcs' = s) ∧
  (G \<turnstile> (x, s) -es[\<succ>]vs-> (x', s') -->
   x = Some xc --> x' = Some xcs' = s) ∧
  (G \<turnstile> (x, s) -c-> (x', s') --> x = Some xc --> x' = Some xcs' = s)

lemma eval_xcpt:

  G \<turnstile> (Some xc, s) -e\<succ>v-> (x', s') ==> x' = Some xcs' = s

lemma exec_xcpt:

  G \<turnstile> (Some xc, s) -s0.0-> (x', s') ==> x' = Some xcs' = s