src/HOL/MicroJava/J/Eval.thy
author nipkow
Thu, 11 Nov 1999 12:23:45 +0100
changeset 8011 d14c4e9e9c8e
child 8034 6fc37b5c5e98
permissions -rw-r--r--
*** empty log message ***
nipkow@8011
     1
(*  Title:      HOL/MicroJava/J/Eval.thy
nipkow@8011
     2
    ID:         $Id$
nipkow@8011
     3
    Author:     David von Oheimb
nipkow@8011
     4
    Copyright   1999 Technische Universitaet Muenchen
nipkow@8011
     5
nipkow@8011
     6
Operational evaluation (big-step) semantics of the 
nipkow@8011
     7
execution of Java expressions and statements
nipkow@8011
     8
*)
nipkow@8011
     9
nipkow@8011
    10
Eval = State + 
nipkow@8011
    11
nipkow@8011
    12
consts
nipkow@8011
    13
  eval  :: "javam prog \\<Rightarrow> (xstate \\<times> expr      \\<times> val      \\<times> xstate) set"
nipkow@8011
    14
  evals :: "javam prog \\<Rightarrow> (xstate \\<times> expr list \\<times> val list \\<times> xstate) set"
nipkow@8011
    15
  exec  :: "javam prog \\<Rightarrow> (xstate \\<times> stmt                 \\<times> xstate) set"
nipkow@8011
    16
nipkow@8011
    17
syntax
nipkow@8011
    18
  eval :: "[javam prog,xstate,expr,val,xstate] \\<Rightarrow> bool "("_\\<turnstile>_ -_\\<succ>_\\<rightarrow> _"[51,82,82,82,82]81)
nipkow@8011
    19
  evals:: "[javam prog,xstate,expr list,
nipkow@8011
    20
	                      val list,xstate] \\<Rightarrow> bool "("_\\<turnstile>_ -_[\\<succ>]_\\<rightarrow> _"[51,82,51,51,82]81)
nipkow@8011
    21
  exec :: "[javam prog,xstate,stmt,    xstate] \\<Rightarrow> bool "("_\\<turnstile>_ -_\\<rightarrow> _"  [51,82,82,   82]81)
nipkow@8011
    22
nipkow@8011
    23
translations
nipkow@8011
    24
  "G\\<turnstile>s -e \\<succ> v\\<rightarrow> (x,s')" == "(s, e, v, _args x s') \\<in> eval  G"
nipkow@8011
    25
  "G\\<turnstile>s -e \\<succ> v\\<rightarrow>    s' " == "(s, e, v,         s') \\<in> eval  G"
nipkow@8011
    26
  "G\\<turnstile>s -e[\\<succ>]v\\<rightarrow> (x,s')" == "(s, e, v, _args x s') \\<in> evals G"
nipkow@8011
    27
  "G\\<turnstile>s -e[\\<succ>]v\\<rightarrow>    s' " == "(s, e, v,         s') \\<in> evals G"
nipkow@8011
    28
  "G\\<turnstile>s -c     \\<rightarrow> (x,s')" == "(s, c,    _args x s') \\<in> exec  G"
nipkow@8011
    29
  "G\\<turnstile>s -c     \\<rightarrow>    s' " == "(s, c,            s') \\<in> exec  G"
nipkow@8011
    30
nipkow@8011
    31
inductive "eval G" "evals G" "exec G" intrs
nipkow@8011
    32
nipkow@8011
    33
(* evaluation of expressions *)
nipkow@8011
    34
nipkow@8011
    35
  (* cf. 15.5 *)
nipkow@8011
    36
  XcptE				  "G\\<turnstile>(Some xc,s) -e\\<succ>arbitrary\\<rightarrow> (Some xc,s)"
nipkow@8011
    37
nipkow@8011
    38
  (* cf. 15.8.1 *)
nipkow@8011
    39
  NewC	"\\<lbrakk>h = heap s; (a,x) = new_Addr h;
nipkow@8011
    40
	  h'= h(a\\<mapsto>(C,init_vars (fields (G,C))))\\<rbrakk> \\<Longrightarrow>
nipkow@8011
    41
				   G\\<turnstile>Norm s -NewC C\\<succ>Addr a\\<rightarrow> c_hupd h' (x,s)"
nipkow@8011
    42
nipkow@8011
    43
  (* cf. 15.15 *)
nipkow@8011
    44
  Cast	"\\<lbrakk>G\\<turnstile>Norm s0 -e\\<succ>v\\<rightarrow> (x1,s1);
nipkow@8011
    45
	  x2=raise_if (\\<not> cast_ok G T (heap s1) v) ClassCast x1\\<rbrakk> \\<Longrightarrow>
nipkow@8011
    46
			        G\\<turnstile>Norm s0 -Cast T e\\<succ>v\\<rightarrow> (x2,s1)"
nipkow@8011
    47
nipkow@8011
    48
  (* cf. 15.7.1 *)
nipkow@8011
    49
  Lit				   "G\\<turnstile>Norm s -Lit v\\<succ>v\\<rightarrow> Norm s"
nipkow@8011
    50
nipkow@8011
    51
  (* cf. 15.13.1, 15.2 *)
nipkow@8011
    52
  LAcc				  "G\\<turnstile>Norm s -LAcc v\\<succ>the (locals s v)\\<rightarrow> Norm s"
nipkow@8011
    53
nipkow@8011
    54
  (* cf. 15.25.1 *)
nipkow@8011
    55
  LAss  "\\<lbrakk>G\\<turnstile>Norm s -e\\<succ>v\\<rightarrow>  (x,(h,l));
nipkow@8011
    56
	  l' = (if x = None then l(va\\<mapsto>v) else l)\\<rbrakk> \\<Longrightarrow>
nipkow@8011
    57
				   G\\<turnstile>Norm s -va\\<Colon>=e\\<succ>v\\<rightarrow> (x,(h,l'))"
nipkow@8011
    58
nipkow@8011
    59
nipkow@8011
    60
  (* cf. 15.10.1, 15.2 *)
nipkow@8011
    61
  FAcc	"\\<lbrakk>G\\<turnstile>Norm s0 -e\\<succ>a'\\<rightarrow> (x1,s1); 
nipkow@8011
    62
	  v = the (snd (the (heap s1 (the_Addr a'))) (fn,T))\\<rbrakk> \\<Longrightarrow>
nipkow@8011
    63
				 G\\<turnstile>Norm s0 -{T}e..fn\\<succ>v\\<rightarrow> (np a' x1,s1)"
nipkow@8011
    64
nipkow@8011
    65
  (* cf. 15.25.1 *)
nipkow@8011
    66
  FAss  "\\<lbrakk>G\\<turnstile>     Norm s0  -e1\\<succ>a'\\<rightarrow> (x1,s1); a = the_Addr a';
nipkow@8011
    67
	  G\\<turnstile>(np a' x1,s1) -e2\\<succ>v \\<rightarrow> (x2,s2);
nipkow@8011
    68
	  h = heap s2; (c,fs) = the (h a);
nipkow@8011
    69
	  h' = h(a\\<mapsto>(c,(fs((fn,T)\\<mapsto>v))))\\<rbrakk> \\<Longrightarrow>
nipkow@8011
    70
			  G\\<turnstile>Norm s0 -{T}e1..fn\\<in>=e2\\<succ>v\\<rightarrow> c_hupd h' (x2,s2)"
nipkow@8011
    71
nipkow@8011
    72
  (* cf. 15.11.4.1, 15.11.4.2, 15.11.4.4, 15.11.4.5, 14.15 *)
nipkow@8011
    73
  Call	"\\<lbrakk>G\\<turnstile>Norm s0 -e\\<succ>a'\\<rightarrow> s1; a = the_Addr a';
nipkow@8011
    74
	   G\\<turnstile>s1 -ps[\\<succ>]pvs\\<rightarrow> (x,(h,l)); dynT = fst (the (h a));
nipkow@8011
    75
	   (md,rT,pns,lvars,blk,res) = the (cmethd (G,dynT) (mn,pTs));
nipkow@8011
    76
	   G\\<turnstile>(np a' x,(h,(init_vars lvars)(pns[\\<mapsto>]pvs)(This\\<mapsto>a'))) -blk\\<rightarrow> s3;
nipkow@8011
    77
	   G\\<turnstile>     s3 -res\\<succ>v \\<rightarrow> (x4,s4)\\<rbrakk> \\<Longrightarrow>
nipkow@8011
    78
			    G\\<turnstile>Norm s0 -e..mn({pTs}ps)\\<succ>v\\<rightarrow> (x4,(heap s4,l))"
nipkow@8011
    79
nipkow@8011
    80
nipkow@8011
    81
(* evaluation of expression lists *)
nipkow@8011
    82
nipkow@8011
    83
  (* cf. 15.5 *)
nipkow@8011
    84
  XcptEs			  "G\\<turnstile>(Some xc,s) -e[\\<succ>]arbitrary\\<rightarrow> (Some xc,s)"
nipkow@8011
    85
nipkow@8011
    86
  (* cf. 15.11.??? *)
nipkow@8011
    87
  Nil
nipkow@8011
    88
				    "G\\<turnstile>Norm s0 -[][\\<succ>][]\\<rightarrow> Norm s0"
nipkow@8011
    89
nipkow@8011
    90
  (* cf. 15.6.4 *)
nipkow@8011
    91
  Cons	"\\<lbrakk>G\\<turnstile>Norm s0 -e  \\<succ> v \\<rightarrow> s1;
nipkow@8011
    92
           G\\<turnstile>     s1 -es[\\<succ>]vs\\<rightarrow> s2\\<rbrakk> \\<Longrightarrow>
nipkow@8011
    93
				   G\\<turnstile>Norm s0 -e#es[\\<succ>]v#vs\\<rightarrow> s2"
nipkow@8011
    94
nipkow@8011
    95
(* execution of statements *)
nipkow@8011
    96
nipkow@8011
    97
  (* cf. 14.1 *)
nipkow@8011
    98
  XcptS				 "G\\<turnstile>(Some xc,s) -s0\\<rightarrow> (Some xc,s)"
nipkow@8011
    99
nipkow@8011
   100
  (* cf. 14.5 *)
nipkow@8011
   101
  Skip	 			    "G\\<turnstile>Norm s -Skip\\<rightarrow> Norm s"
nipkow@8011
   102
nipkow@8011
   103
  (* cf. 14.7 *)
nipkow@8011
   104
  Expr	"\\<lbrakk>G\\<turnstile>Norm s0 -e\\<succ>v\\<rightarrow> s1\\<rbrakk> \\<Longrightarrow>
nipkow@8011
   105
				  G\\<turnstile>Norm s0 -Expr e\\<rightarrow> s1"
nipkow@8011
   106
nipkow@8011
   107
  (* cf. 14.2 *)
nipkow@8011
   108
  Comp	"\\<lbrakk>G\\<turnstile>Norm s0 -s \\<rightarrow> s1;
nipkow@8011
   109
	  G\\<turnstile>     s1 -t \\<rightarrow> s2\\<rbrakk> \\<Longrightarrow>
nipkow@8011
   110
				 G\\<turnstile>Norm s0 -(s;; t)\\<rightarrow> s2"
nipkow@8011
   111
nipkow@8011
   112
  (* cf. 14.8.2 *)
nipkow@8011
   113
  Cond	"\\<lbrakk>G\\<turnstile>Norm s0  -e \\<succ>v\\<rightarrow> s1;
nipkow@8011
   114
	  G\\<turnstile>     s1 -(if  the_Bool v then s else t)\\<rightarrow> s2\\<rbrakk> \\<Longrightarrow>
nipkow@8011
   115
		        G\\<turnstile>Norm s0 -(If(e) s Else t)\\<rightarrow> s2"
nipkow@8011
   116
nipkow@8011
   117
  (* cf. 14.10, 14.10.1 *)
nipkow@8011
   118
  Loop	"\\<lbrakk>G\\<turnstile>Norm s0 -(If(e) (s;; While(e) s) Else Skip)\\<rightarrow> s1\\<rbrakk> \\<Longrightarrow>
nipkow@8011
   119
			    G\\<turnstile>Norm s0 -(While(e) s)\\<rightarrow> s1"
nipkow@8011
   120
nipkow@8011
   121
end