src/HOL/MicroJava/BV/BVSpec.thy
changeset 8023 3e5ddca613dd
parent 8011 d14c4e9e9c8e
child 8032 1eaae1a2f8ff
     1.1 --- a/src/HOL/MicroJava/BV/BVSpec.thy	Thu Nov 18 12:12:39 1999 +0100
     1.2 +++ b/src/HOL/MicroJava/BV/BVSpec.thy	Fri Nov 19 16:30:52 1999 +0100
     1.3 @@ -22,7 +22,7 @@
     1.4  	 pc+1 < max_pc \\<and>
     1.5  	 idx < length LT \\<and>
     1.6  	 (\\<exists>ts. (LT ! idx) = Some ts \\<and> 
     1.7 -	       G \\<turnstile> phi ! (pc+1) >>>= (ts # ST , LT)))"
     1.8 +	       G \\<turnstile> (ts # ST , LT) <=s phi ! (pc+1)))"
     1.9  
    1.10  "wt_LS (Store idx) G phi max_pc pc =
    1.11  	(let (ST,LT) = phi ! pc
    1.12 @@ -30,19 +30,19 @@
    1.13  	 pc+1 < max_pc \\<and>
    1.14  	 idx < length LT \\<and>
    1.15  	 (\\<exists>ts ST'. ST = ts # ST' \\<and>
    1.16 -		   G \\<turnstile> phi ! (pc+1) >>>= (ST' , LT[idx:=Some ts])))"
    1.17 +		   G \\<turnstile> (ST' , LT[idx:=Some ts]) <=s phi ! (pc+1)))"
    1.18  
    1.19  "wt_LS (Bipush i) G phi max_pc pc =
    1.20  	(let (ST,LT) = phi ! pc
    1.21  	 in
    1.22  	 pc+1 < max_pc \\<and>
    1.23 -	 G \\<turnstile> phi ! (pc+1) >>>= ((PrimT Integer) # ST , LT))"
    1.24 +	 G \\<turnstile> ((PrimT Integer) # ST , LT) <=s phi ! (pc+1))"
    1.25  
    1.26  "wt_LS (Aconst_null) G phi max_pc pc =
    1.27  	(let (ST,LT) = phi ! pc
    1.28  	 in
    1.29  	 pc+1 < max_pc \\<and>
    1.30 -	 G \\<turnstile> phi ! (pc+1) >>>= (NT # ST , LT))"
    1.31 +	 G \\<turnstile> (NT # ST , LT) <=s phi ! (pc+1))"
    1.32  
    1.33  consts
    1.34   wt_MO	:: "[manipulate_object,jvm_prog,method_type,p_count,p_count] \\<Rightarrow> bool" 
    1.35 @@ -55,7 +55,7 @@
    1.36  	 (\\<exists>T oT ST'. cfield (G,C) F = Some(C,T) \\<and>
    1.37                         ST = oT # ST' \\<and> 
    1.38  		       G \\<turnstile> oT \\<preceq> (Class C) \\<and>
    1.39 -		       G \\<turnstile> phi ! (pc+1) >>>= (T # ST' , LT)))"
    1.40 +		       G \\<turnstile> (T # ST' , LT) <=s phi ! (pc+1)))"
    1.41  
    1.42  "wt_MO (Putfield F C) G phi max_pc pc =
    1.43  	(let (ST,LT) = phi ! pc
    1.44 @@ -67,7 +67,7 @@
    1.45               ST = vT # oT # ST' \\<and> 
    1.46               G \\<turnstile> oT \\<preceq> (Class C) \\<and>
    1.47  	     G \\<turnstile> vT \\<preceq> T  \\<and>
    1.48 -             G \\<turnstile> phi ! (pc+1) >>>= (ST' , LT)))"
    1.49 +             G \\<turnstile> (ST' , LT) <=s phi ! (pc+1)))"
    1.50  
    1.51  
    1.52  consts 
    1.53 @@ -78,7 +78,7 @@
    1.54  	 in
    1.55  	 pc+1 < max_pc \\<and>
    1.56  	 is_class G C \\<and>
    1.57 -	 G \\<turnstile> phi ! (pc+1) >>>= ((Class C) # ST , LT))"
    1.58 +	 G \\<turnstile> ((Class C) # ST , LT) <=s phi ! (pc+1))"
    1.59  
    1.60  consts
    1.61   wt_CH	:: "[check_object,jvm_prog,method_type,p_count,p_count] \\<Rightarrow> bool" 
    1.62 @@ -89,7 +89,7 @@
    1.63  	 pc+1 < max_pc \\<and>
    1.64  	 is_class G C \\<and> 
    1.65  	 (\\<exists>rt ST'. ST = RefT rt # ST' \\<and>
    1.66 -                   G \\<turnstile> phi ! (pc+1) >>>= (Class C # ST' , LT)))"
    1.67 +                   G \\<turnstile> (Class C # ST' , LT) <=s phi ! (pc+1)))"
    1.68  
    1.69  consts 
    1.70   wt_OS	:: "[op_stack,jvm_prog,method_type,p_count,p_count] \\<Rightarrow> bool" 
    1.71 @@ -99,35 +99,35 @@
    1.72  	 in
    1.73  	 \\<exists>ts ST'. pc+1 < max_pc \\<and>
    1.74  		ST = ts # ST' \\<and> 	 
    1.75 -		G \\<turnstile> phi ! (pc+1) >>>= (ST' , LT))"
    1.76 +		G \\<turnstile> (ST' , LT) <=s phi ! (pc+1))"
    1.77  
    1.78  "wt_OS Dup G phi max_pc pc =
    1.79  	(let (ST,LT) = phi ! pc
    1.80  	 in
    1.81  	 pc+1 < max_pc \\<and>
    1.82  	 (\\<exists>ts ST'. ST = ts # ST' \\<and> 	 
    1.83 -		   G \\<turnstile> phi ! (pc+1) >>>= (ts # ts # ST' , LT)))"
    1.84 +		   G \\<turnstile> (ts # ts # ST' , LT) <=s phi ! (pc+1)))"
    1.85  
    1.86  "wt_OS Dup_x1 G phi max_pc pc =
    1.87  	(let (ST,LT) = phi ! pc
    1.88  	 in
    1.89  	 pc+1 < max_pc \\<and>
    1.90  	 (\\<exists>ts1 ts2 ST'. ST = ts1 # ts2 # ST' \\<and> 	 
    1.91 -		   G \\<turnstile> phi ! (pc+1) >>>= (ts1 # ts2 # ts1 # ST' , LT)))"
    1.92 +		   G \\<turnstile> (ts1 # ts2 # ts1 # ST' , LT) <=s phi ! (pc+1)))"
    1.93  
    1.94  "wt_OS Dup_x2 G phi max_pc pc =
    1.95  	(let (ST,LT) = phi ! pc
    1.96  	 in
    1.97  	 pc+1 < max_pc \\<and>
    1.98  	 (\\<exists>ts1 ts2 ts3 ST'. ST = ts1 # ts2 # ts3 # ST' \\<and> 	 
    1.99 -		   G \\<turnstile> phi ! (pc+1) >>>= (ts1 # ts2 # ts3 # ts1 # ST' , LT)))"
   1.100 +		   G \\<turnstile> (ts1 # ts2 # ts3 # ts1 # ST' , LT) <=s phi ! (pc+1)))"
   1.101  
   1.102  "wt_OS Swap G phi max_pc pc =
   1.103  	(let (ST,LT) = phi ! pc
   1.104  	 in
   1.105  	 pc+1 < max_pc \\<and>
   1.106  	 (\\<exists>ts ts' ST'. ST = ts' # ts # ST' \\<and> 	 
   1.107 -		       G \\<turnstile> phi ! (pc+1) >>>= (ts # ts' # ST' , LT)))"
   1.108 +		       G \\<turnstile> (ts # ts' # ST' , LT) <=s phi ! (pc+1)))"
   1.109  
   1.110  consts 
   1.111   wt_BR	:: "[branch,jvm_prog,method_type,p_count,p_count] \\<Rightarrow> bool" 
   1.112 @@ -137,13 +137,13 @@
   1.113  	 in
   1.114  	 pc+1 < max_pc \\<and> (nat(int pc+branch)) < max_pc \\<and> 
   1.115  	 (\\<exists>ts ts' ST'. ST = ts # ts' # ST' \\<and>  ts = ts' \\<and>
   1.116 -		       G \\<turnstile> phi ! (pc+1) >>>= (ST' , LT) \\<and>
   1.117 -		       G \\<turnstile> phi ! (nat(int pc+branch)) >>>= (ST' , LT)))"
   1.118 +		       G \\<turnstile> (ST' , LT) <=s phi ! (pc+1) \\<and>
   1.119 +		       G \\<turnstile> (ST' , LT) <=s phi ! (nat(int pc+branch))))"
   1.120  "wt_BR (Goto branch) G phi max_pc pc =
   1.121  	(let (ST,LT) = phi ! pc
   1.122  	 in
   1.123  	 (nat(int pc+branch)) < max_pc \\<and> 
   1.124 -	 G \\<turnstile> phi ! (nat(int pc+branch)) >>>= (ST , LT))"
   1.125 +	 G \\<turnstile> (ST , LT) <=s phi ! (nat(int pc+branch)))"
   1.126  
   1.127  consts
   1.128   wt_MI :: "[meth_inv,jvm_prog,method_type,p_count,p_count] \\<Rightarrow> bool" 
   1.129 @@ -156,7 +156,7 @@
   1.130           length apTs = length fpTs \\<and>
   1.131           (\\<forall>(aT,fT)\\<in>set(zip apTs fpTs). G \\<turnstile> aT \\<preceq> fT) \\<and>
   1.132           (\\<exists>D rT b. cmethd (G,C) (mn,fpTs) = Some(D,rT,b) \\<and>
   1.133 -         G \\<turnstile> phi ! (pc+1) >>>= (rT # ST' , LT))))"
   1.134 +         G \\<turnstile> (rT # ST' , LT) <=s phi ! (pc+1))))"
   1.135  
   1.136  constdefs
   1.137   wt_MR	:: "[jvm_prog,ty,method_type,p_count] \\<Rightarrow> bool"
   1.138 @@ -183,7 +183,7 @@
   1.139  constdefs
   1.140   wt_start :: "[jvm_prog,cname,ty list,nat,method_type] \\<Rightarrow> bool"
   1.141   "wt_start G C pTs mxl phi \\<equiv> 
   1.142 -    G \\<turnstile> phi!0 >>>= ([],(Some(Class C))#((map Some pTs))@(replicate mxl None))"
   1.143 +    G \\<turnstile> ([],(Some(Class C))#((map Some pTs))@(replicate mxl None)) <=s phi!0"
   1.144  
   1.145  
   1.146   wt_method :: "[jvm_prog,cname,ty list,ty,nat,instr list,method_type] \\<Rightarrow> bool"