| 8245 |      1 | (*  Title:      HOL/MicroJava/BV/BVLightSpec.thy
 | 
|  |      2 |     ID:         $Id$
 | 
|  |      3 |     Author:     Gerwin Klein
 | 
|  |      4 |     Copyright   1999 Technische Universitaet Muenchen
 | 
|  |      5 | 
 | 
|  |      6 | Specification of a lightweight bytecode verifier
 | 
|  |      7 | *)
 | 
|  |      8 | 
 | 
| 9012 |      9 | theory LBVSpec = BVSpec:;
 | 
| 8245 |     10 | 
 | 
|  |     11 | types
 | 
|  |     12 |   certificate       = "state_type option list" 
 | 
|  |     13 |   class_certificate = "sig \\<Rightarrow> certificate"
 | 
| 9012 |     14 |   prog_certificate  = "cname \\<Rightarrow> class_certificate";
 | 
| 8245 |     15 | 
 | 
|  |     16 | consts
 | 
| 9012 |     17 |  wtl_LS :: "[load_and_store,state_type,state_type,p_count,p_count] \\<Rightarrow> bool";
 | 
| 8245 |     18 | primrec
 | 
|  |     19 | "wtl_LS (Load idx) s s' max_pc pc =
 | 
|  |     20 |  (let (ST,LT) = s
 | 
|  |     21 |   in
 | 
|  |     22 |   pc+1 < max_pc \\<and>
 | 
|  |     23 |   idx < length LT \\<and> 
 | 
|  |     24 |   (\\<exists>ts. (LT ! idx) = Some ts \\<and>  
 | 
|  |     25 |    (ts # ST , LT) = s'))"
 | 
|  |     26 |   
 | 
|  |     27 | "wtl_LS (Store idx) s s' max_pc pc =
 | 
|  |     28 |  (let (ST,LT) = s
 | 
|  |     29 |   in
 | 
|  |     30 |   pc+1 < max_pc \\<and>
 | 
|  |     31 |   idx < length LT \\<and>
 | 
|  |     32 |   (\\<exists>ts ST'. ST = ts # ST' \\<and>
 | 
|  |     33 |    (ST' , LT[idx:=Some ts]) = s'))"
 | 
|  |     34 | 
 | 
|  |     35 | "wtl_LS (Bipush i) s s' max_pc pc =
 | 
|  |     36 | 	(let (ST,LT) = s
 | 
|  |     37 | 	 in
 | 
|  |     38 | 	 pc+1 < max_pc \\<and>
 | 
|  |     39 | 	 ((PrimT Integer) # ST , LT) = s')"
 | 
|  |     40 | 
 | 
|  |     41 | "wtl_LS (Aconst_null) s s' max_pc pc =
 | 
|  |     42 | 	(let (ST,LT) = s
 | 
|  |     43 | 	 in
 | 
|  |     44 | 	 pc+1 < max_pc \\<and>
 | 
| 9012 |     45 | 	 (NT # ST , LT) = s')";
 | 
| 8245 |     46 | 
 | 
|  |     47 | consts
 | 
| 9012 |     48 |  wtl_MO	:: "[manipulate_object,jvm_prog,state_type,state_type,p_count,p_count] \\<Rightarrow> bool"; 
 | 
| 8245 |     49 | primrec
 | 
|  |     50 | "wtl_MO (Getfield F C) G s s' max_pc pc =
 | 
|  |     51 | 	(let (ST,LT) = s
 | 
|  |     52 | 	 in
 | 
|  |     53 | 	 pc+1 < max_pc \\<and>
 | 
|  |     54 | 	 is_class G C \\<and>
 | 
|  |     55 | 	 (\\<exists>T oT ST'. field (G,C) F = Some(C,T) \\<and>
 | 
| 8390 |     56 |                           ST = oT # ST'  \\<and> 
 | 
|  |     57 | 		                  G \\<turnstile> oT \\<preceq> (Class C) \\<and>
 | 
|  |     58 | 		          (T # ST' , LT) = s'))"
 | 
| 8245 |     59 | 
 | 
|  |     60 | "wtl_MO (Putfield F C) G s s' max_pc pc =
 | 
|  |     61 | 	(let (ST,LT) = s
 | 
|  |     62 | 	 in
 | 
|  |     63 | 	 pc+1 < max_pc \\<and>
 | 
|  |     64 | 	 is_class G C \\<and> 
 | 
|  |     65 | 	 (\\<exists>T vT oT ST'.
 | 
|  |     66 |              field (G,C) F = Some(C,T) \\<and>
 | 
|  |     67 |              ST = vT # oT # ST' \\<and> 
 | 
|  |     68 |              G \\<turnstile> oT \\<preceq> (Class C) \\<and>
 | 
| 8390 |     69 |              G \\<turnstile> vT \\<preceq> T  \\<and>
 | 
| 9012 |     70 |              (ST' , LT) = s'))";
 | 
| 8245 |     71 | 
 | 
|  |     72 | 
 | 
|  |     73 | consts 
 | 
| 9012 |     74 |  wtl_CO	:: "[create_object,jvm_prog,state_type,state_type,p_count,p_count] \\<Rightarrow> bool"; 
 | 
| 8245 |     75 | primrec
 | 
|  |     76 | "wtl_CO (New C) G s s' max_pc pc =
 | 
|  |     77 | 	(let (ST,LT) = s
 | 
|  |     78 | 	 in
 | 
|  |     79 | 	 pc+1 < max_pc \\<and>
 | 
|  |     80 | 	 is_class G C \\<and>
 | 
| 9012 |     81 | 	 ((Class C) # ST , LT) = s')";
 | 
| 8245 |     82 | 
 | 
|  |     83 | consts
 | 
| 9012 |     84 |  wtl_CH	:: "[check_object,jvm_prog,state_type,state_type,p_count,p_count] \\<Rightarrow> bool"; 
 | 
| 8245 |     85 | primrec
 | 
|  |     86 | "wtl_CH (Checkcast C) G s s' max_pc pc = 
 | 
|  |     87 | 	(let (ST,LT) = s
 | 
|  |     88 | 	 in
 | 
|  |     89 | 	 pc+1 < max_pc \\<and>
 | 
|  |     90 | 	 is_class G C \\<and> 
 | 
|  |     91 | 	 (\\<exists>rt ST'. ST = RefT rt # ST' \\<and>
 | 
| 9012 |     92 |                    (Class C # ST' , LT) = s'))";
 | 
| 8245 |     93 | 
 | 
|  |     94 | consts 
 | 
| 9012 |     95 |  wtl_OS	:: "[op_stack,state_type,state_type,p_count,p_count] \\<Rightarrow> bool"; 
 | 
| 8245 |     96 | primrec
 | 
|  |     97 | "wtl_OS Pop s s' max_pc pc =
 | 
|  |     98 | 	(let (ST,LT) = s
 | 
|  |     99 | 	 in
 | 
|  |    100 | 	 \\<exists>ts ST'. pc+1 < max_pc \\<and>
 | 
|  |    101 | 		ST = ts # ST' \\<and> 	 
 | 
|  |    102 | 		(ST' , LT) = s')"
 | 
|  |    103 | 
 | 
|  |    104 | "wtl_OS Dup s s' max_pc pc =
 | 
|  |    105 | 	(let (ST,LT) = s
 | 
|  |    106 | 	 in
 | 
|  |    107 | 	 pc+1 < max_pc \\<and>
 | 
|  |    108 | 	 (\\<exists>ts ST'. ST = ts # ST' \\<and> 	 
 | 
|  |    109 | 		   (ts # ts # ST' , LT) = s'))"
 | 
|  |    110 | 
 | 
|  |    111 | "wtl_OS Dup_x1 s s' max_pc pc =
 | 
|  |    112 | 	(let (ST,LT) = s
 | 
|  |    113 | 	 in
 | 
|  |    114 | 	 pc+1 < max_pc \\<and>
 | 
|  |    115 | 	 (\\<exists>ts1 ts2 ST'. ST = ts1 # ts2 # ST' \\<and> 	 
 | 
|  |    116 | 		        (ts1 # ts2 # ts1 # ST' , LT) = s'))"
 | 
|  |    117 | 
 | 
|  |    118 | "wtl_OS Dup_x2 s s' max_pc pc =
 | 
|  |    119 | 	(let (ST,LT) = s
 | 
|  |    120 | 	 in
 | 
|  |    121 | 	 pc+1 < max_pc \\<and>
 | 
|  |    122 | 	 (\\<exists>ts1 ts2 ts3 ST'. ST = ts1 # ts2 # ts3 # ST' \\<and> 	 
 | 
|  |    123 | 		            (ts1 # ts2 # ts3 # ts1 # ST' , LT) = s'))"
 | 
|  |    124 | 
 | 
|  |    125 | "wtl_OS Swap s s' max_pc pc =
 | 
|  |    126 | 	(let (ST,LT) = s
 | 
|  |    127 | 	 in
 | 
|  |    128 | 	 pc+1 < max_pc \\<and>
 | 
|  |    129 | 	 (\\<exists>ts ts' ST'. ST = ts' # ts # ST' \\<and> 	 
 | 
| 9012 |    130 | 		       (ts # ts' # ST' , LT) = s'))";
 | 
| 8245 |    131 | 
 | 
|  |    132 | consts 
 | 
| 9012 |    133 |  wtl_BR	:: "[branch,jvm_prog,state_type,state_type,certificate,p_count,p_count] \\<Rightarrow> bool"; 
 | 
| 8245 |    134 | primrec
 | 
|  |    135 | "wtl_BR (Ifcmpeq branch) G s s' cert max_pc pc =
 | 
|  |    136 | 	(let (ST,LT) = s
 | 
|  |    137 | 	 in
 | 
|  |    138 | 	 pc+1 < max_pc \\<and> (nat(int pc+branch)) < max_pc \\<and> 
 | 
| 8390 |    139 | 	 (\\<exists>ts ts' ST'. ST = ts # ts' # ST' \\<and>
 | 
|  |    140 |           ((\\<exists>p. ts = PrimT p \\<and> ts' = PrimT p) \\<or>
 | 
|  |    141 |            (\\<exists>r r'. ts = RefT r \\<and> ts' = RefT r')) \\<and>
 | 
| 8245 |    142 | 		       ((ST' , LT) = s') \\<and>
 | 
|  |    143 |             cert ! (nat(int pc+branch)) \\<noteq> None \\<and>
 | 
|  |    144 | 		       G \\<turnstile> (ST' , LT) <=s the (cert ! (nat(int pc+branch)))))"
 | 
|  |    145 |   
 | 
|  |    146 | "wtl_BR (Goto branch) G s s' cert max_pc pc =
 | 
|  |    147 | 	((let (ST,LT) = s
 | 
|  |    148 | 	 in
 | 
|  |    149 | 	 (nat(int pc+branch)) < max_pc \\<and> cert ! (nat(int pc+branch)) \\<noteq> None \\<and>
 | 
|  |    150 | 	 G \\<turnstile> (ST , LT) <=s the (cert ! (nat(int pc+branch)))) \\<and>
 | 
| 9012 |    151 |    (cert ! (pc+1) = Some s'))";
 | 
| 8245 |    152 |   
 | 
|  |    153 | consts
 | 
| 9012 |    154 |  wtl_MI :: "[meth_inv,jvm_prog,state_type,state_type,certificate,p_count,p_count] \\<Rightarrow> bool"; 
 | 
| 8245 |    155 | primrec
 | 
| 8390 |    156 | "wtl_MI (Invoke mn fpTs) G s s' cert max_pc pc =
 | 
| 8245 |    157 | 	(let (ST,LT) = s
 | 
|  |    158 | 	 in
 | 
|  |    159 |          pc+1 < max_pc \\<and>
 | 
| 8390 |    160 |          (\\<exists>apTs X ST'. ST = (rev apTs) @ (X # ST') \\<and>
 | 
| 8245 |    161 |          length apTs = length fpTs \\<and>
 | 
| 8390 |    162 |          (\\<exists>s''. cert ! (pc+1) = Some s'' \\<and> 
 | 
|  |    163 |          ((s'' = s' \\<and> X = NT) \\<or>
 | 
|  |    164 |           ((G \\<turnstile> s' <=s s'') \\<and> (\\<exists>C. X = Class C \\<and>  
 | 
|  |    165 |           (\\<forall>(aT,fT)\\<in>set(zip apTs fpTs). G \\<turnstile> aT \\<preceq> fT) \\<and> 
 | 
|  |    166 |           (\\<exists>D rT b. method (G,C) (mn,fpTs) = Some(D,rT,b) \\<and>
 | 
| 9012 |    167 |           (rT # ST' , LT) = s')))))))";
 | 
| 8245 |    168 | 
 | 
|  |    169 | consts
 | 
| 9012 |    170 |  wtl_MR	:: "[meth_ret,jvm_prog,ty,state_type,state_type,certificate,p_count,p_count] \\<Rightarrow> bool";
 | 
| 8245 |    171 | primrec
 | 
|  |    172 |   "wtl_MR Return G rT s s' cert max_pc pc = 
 | 
|  |    173 | 	((let (ST,LT) = s
 | 
|  |    174 | 	 in
 | 
|  |    175 | 	 (\\<exists>T ST'. ST = T # ST' \\<and> G \\<turnstile> T \\<preceq> rT)) \\<and>
 | 
| 9012 |    176 |    (cert ! (pc+1) = Some s'))";
 | 
| 8245 |    177 | 
 | 
|  |    178 | 
 | 
|  |    179 | consts 
 | 
| 9012 |    180 |  wtl_inst :: "[instr,jvm_prog,ty,state_type,state_type,certificate,p_count,p_count] \\<Rightarrow> bool";
 | 
| 8245 |    181 |  primrec
 | 
|  |    182 |  "wtl_inst (LS ins) G rT s s' cert max_pc pc = wtl_LS ins s s' max_pc pc"
 | 
|  |    183 |  "wtl_inst (CO ins) G rT s s' cert max_pc pc = wtl_CO ins G s s' max_pc pc"
 | 
|  |    184 |  "wtl_inst (MO ins) G rT s s' cert max_pc pc = wtl_MO ins G s s' max_pc pc"
 | 
|  |    185 |  "wtl_inst (CH ins) G rT s s' cert max_pc pc = wtl_CH ins G s s' max_pc pc"
 | 
| 8390 |    186 |  "wtl_inst (MI ins) G rT s s' cert max_pc pc = wtl_MI ins G s s' cert max_pc pc"
 | 
| 8245 |    187 |  "wtl_inst (MR ins) G rT s s' cert max_pc pc = wtl_MR ins G rT s s' cert max_pc pc"
 | 
|  |    188 |  "wtl_inst (OS ins) G rT s s' cert max_pc pc = wtl_OS ins s s' max_pc pc"
 | 
| 9012 |    189 |  "wtl_inst (BR ins) G rT s s' cert max_pc pc = wtl_BR ins G s s' cert max_pc pc";
 | 
| 8245 |    190 | 
 | 
|  |    191 | 
 | 
|  |    192 | constdefs
 | 
|  |    193 | wtl_inst_option :: "[instr,jvm_prog,ty,state_type,state_type,certificate,p_count,p_count] \\<Rightarrow> bool"
 | 
|  |    194 | "wtl_inst_option i G rT s0 s1 cert max_pc pc \\<equiv>
 | 
|  |    195 |      (case cert!pc of 
 | 
|  |    196 |           None     \\<Rightarrow> wtl_inst i G rT s0 s1 cert max_pc pc
 | 
|  |    197 |         | Some s0' \\<Rightarrow> (G \\<turnstile> s0 <=s s0') \\<and>
 | 
| 9012 |    198 |                       wtl_inst i G rT s0' s1 cert max_pc pc)";
 | 
| 8245 |    199 |   
 | 
|  |    200 | consts
 | 
| 9012 |    201 |  wtl_inst_list :: "[instr list,jvm_prog,ty,state_type,state_type,certificate,p_count,p_count] \\<Rightarrow> bool";
 | 
| 8245 |    202 | primrec
 | 
|  |    203 |   "wtl_inst_list [] G rT s0 s2 cert max_pc pc = (s0 = s2)"
 | 
|  |    204 | 
 | 
|  |    205 |   "wtl_inst_list (instr#is) G rT s0 s2 cert max_pc pc =
 | 
|  |    206 |      (\\<exists>s1. wtl_inst_option instr G rT s0 s1 cert max_pc pc \\<and> 
 | 
| 9012 |    207 |            wtl_inst_list is G rT s1 s2 cert max_pc (pc+1))";
 | 
| 8245 |    208 | 
 | 
|  |    209 | constdefs
 | 
|  |    210 |  wtl_method :: "[jvm_prog,cname,ty list,ty,nat,instr list,certificate] \\<Rightarrow> bool" 
 | 
|  |    211 |  "wtl_method G C pTs rT mxl ins cert \\<equiv> 
 | 
|  |    212 | 	let max_pc = length ins 
 | 
|  |    213 |         in 
 | 
|  |    214 | 	0 < max_pc \\<and>  
 | 
|  |    215 |         (\\<exists>s2. wtl_inst_list ins G rT 
 | 
|  |    216 |                             ([],(Some(Class C))#((map Some pTs))@(replicate mxl None))
 | 
|  |    217 |                             s2 cert max_pc 0)"
 | 
|  |    218 | 
 | 
| 9012 |    219 |  wtl_jvm_prog :: "[jvm_prog,prog_certificate] \\<Rightarrow> bool"
 | 
| 8245 |    220 |  "wtl_jvm_prog G cert \\<equiv> 
 | 
|  |    221 |     wf_prog (\\<lambda>G C (sig,rT,maxl,b). 
 | 
| 9012 |    222 |                wtl_method G C (snd sig) rT maxl b (cert C sig)) G"; 
 | 
|  |    223 | 
 | 
|  |    224 | lemma rev_eq: "\\<lbrakk>length a = n; length x = n; rev a @ b # c = rev x @ y # z\\<rbrakk> \\<Longrightarrow> a = x \\<and> b = y \\<and> c = z";
 | 
|  |    225 | proof auto;
 | 
|  |    226 | qed;
 | 
|  |    227 | 
 | 
|  |    228 | lemma wtl_inst_unique: 
 | 
|  |    229 | "wtl_inst i G rT s0 s1 cert max_pc pc \\<longrightarrow>
 | 
|  |    230 |  wtl_inst i G rT s0 s1' cert max_pc pc \\<longrightarrow> s1 = s1'" (is "?P i");
 | 
|  |    231 | proof (induct i);
 | 
|  |    232 | case LS;
 | 
|  |    233 |  show "?P (LS load_and_store)"; by (induct load_and_store) auto;
 | 
|  |    234 | case CO;
 | 
|  |    235 |  show "?P (CO create_object)"; by (induct create_object) auto;
 | 
|  |    236 | case MO;
 | 
|  |    237 |  show "?P (MO manipulate_object)"; by (induct manipulate_object) auto;
 | 
|  |    238 | case CH;
 | 
|  |    239 |  show "?P (CH check_object)"; by (induct check_object) auto;
 | 
|  |    240 | case MI;
 | 
|  |    241 |  show "?P (MI meth_inv)"; proof (induct meth_inv);
 | 
|  |    242 |  case Invoke;
 | 
|  |    243 |   have "\\<exists>x y. s0 = (x,y)"; by (simp);
 | 
|  |    244 |   thus "wtl_inst (MI (Invoke mname list)) G rT s0 s1 cert max_pc pc \\<longrightarrow>
 | 
|  |    245 |         wtl_inst (MI (Invoke mname list)) G rT s0 s1' cert max_pc pc \\<longrightarrow>
 | 
|  |    246 |         s1 = s1'";
 | 
|  |    247 |   proof (elim);
 | 
|  |    248 |     apply_end(clarsimp_tac, drule rev_eq, assumption+);
 | 
|  |    249 |   qed auto;
 | 
|  |    250 |  qed;
 | 
|  |    251 | case MR;
 | 
|  |    252 |  show "?P (MR meth_ret)"; by (induct meth_ret) auto;
 | 
|  |    253 | case OS; 
 | 
|  |    254 |  show "?P (OS op_stack)"; by (induct op_stack) auto;
 | 
|  |    255 | case BR;  
 | 
|  |    256 |  show "?P (BR branch)"; by (induct branch) auto;
 | 
|  |    257 | qed;
 | 
|  |    258 | 
 | 
|  |    259 | lemma wtl_inst_option_unique:
 | 
|  |    260 | "\\<lbrakk>wtl_inst_option i G rT s0 s1 cert max_pc pc; 
 | 
|  |    261 |   wtl_inst_option i G rT s0 s1' cert max_pc pc\\<rbrakk> \\<Longrightarrow> s1 = s1'";
 | 
|  |    262 | by (cases "cert!pc") (auto simp add: wtl_inst_unique wtl_inst_option_def);
 | 
|  |    263 | 
 | 
|  |    264 | 
 | 
|  |    265 | lemma wtl_inst_list_unique: 
 | 
|  |    266 | "\\<forall> s0 pc. wtl_inst_list is G rT s0 s1 cert max_pc pc \\<longrightarrow> 
 | 
|  |    267 |  wtl_inst_list is G rT s0 s1' cert max_pc pc \\<longrightarrow> s1=s1'" (is "?P is");
 | 
|  |    268 | proof (induct "?P" "is"); 
 | 
|  |    269 |   case Nil;
 | 
|  |    270 |   show "?P []"; by simp;
 | 
| 8245 |    271 | 
 | 
| 9012 |    272 |   case Cons;
 | 
|  |    273 |   show "?P (a # list)"; 
 | 
|  |    274 |   proof (intro);
 | 
|  |    275 |     fix s0; fix pc; 
 | 
|  |    276 |     let "?o s0 s1" = "wtl_inst_option a G rT s0 s1 cert max_pc pc";
 | 
|  |    277 |     let "?l l s1 s2 pc" = "wtl_inst_list l G rT s1 s2 cert max_pc pc"; 
 | 
|  |    278 |     assume a: "?l (a#list) s0 s1 pc";
 | 
|  |    279 |     assume b: "?l (a#list) s0 s1' pc";
 | 
|  |    280 |     with a;
 | 
|  |    281 |     show "s1 = s1'";
 | 
|  |    282 |       obtain s s' where   "?o s0 s" "?o s0 s'"
 | 
|  |    283 |                   and l:  "?l list s s1 (Suc pc)"
 | 
|  |    284 |                   and l': "?l list s' s1' (Suc pc)"; by auto;
 | 
|  |    285 |       have "s=s'"; by(rule wtl_inst_option_unique);
 | 
|  |    286 |       with l l' Cons;
 | 
|  |    287 |       show ?thesis; by blast;
 | 
|  |    288 |     qed;
 | 
|  |    289 |   qed;
 | 
|  |    290 | qed;
 | 
|  |    291 |         
 | 
|  |    292 | lemma wtl_partial:
 | 
|  |    293 | "\\<forall> pc' pc s. 
 | 
|  |    294 |    wtl_inst_list is G rT s s' cert mpc pc \\<longrightarrow> \
 | 
|  |    295 |    pc' < length is \\<longrightarrow> \
 | 
|  |    296 |    (\\<exists> a b s1. a @ b = is \\<and> length a = pc' \\<and> \
 | 
|  |    297 |               wtl_inst_list a G rT s  s1 cert mpc pc \\<and> \
 | 
|  |    298 |               wtl_inst_list b G rT s1 s' cert mpc (pc+length a))" (is "?P is");
 | 
|  |    299 | proof (induct "?P" "is");
 | 
|  |    300 |   case Nil;
 | 
|  |    301 |   show "?P []"; by auto;
 | 
|  |    302 |   
 | 
|  |    303 |   case Cons;
 | 
|  |    304 |   show "?P (a#list)";
 | 
|  |    305 |   proof (intro allI impI);
 | 
|  |    306 |     fix pc' pc s;
 | 
|  |    307 |     assume length: "pc' < length (a # list)";
 | 
|  |    308 |     assume wtl:    "wtl_inst_list (a # list) G rT s s' cert mpc pc";
 | 
|  |    309 |     show "\\<exists> a' b s1. 
 | 
|  |    310 |             a' @ b = a#list \\<and> length a' = pc' \\<and> \
 | 
|  |    311 |             wtl_inst_list a' G rT s  s1 cert mpc pc \\<and> \
 | 
|  |    312 |             wtl_inst_list b G rT s1 s' cert mpc (pc+length a')"
 | 
|  |    313 |         (is "\\<exists> a b s1. ?E a b s1");
 | 
|  |    314 |     proof (cases "pc'");
 | 
|  |    315 |       case 0;
 | 
|  |    316 |       with wtl;
 | 
|  |    317 |       have "?E [] (a#list) s"; by simp;
 | 
|  |    318 |       thus ?thesis; by blast;
 | 
|  |    319 |     next;
 | 
|  |    320 |       case Suc;
 | 
|  |    321 |       with wtl;
 | 
|  |    322 |       show ?thesis; 
 | 
|  |    323 |         obtain s0 where wtlSuc: "wtl_inst_list list G rT s0 s' cert mpc (Suc pc)"
 | 
|  |    324 |                   and   wtlOpt: "wtl_inst_option a G rT s s0 cert mpc pc"; by auto;
 | 
|  |    325 |         from Cons;
 | 
|  |    326 |         show ?thesis; 
 | 
|  |    327 |           obtain a' b s1' 
 | 
|  |    328 |             where "a' @ b = list" "length a' = nat"
 | 
|  |    329 |             and w:"wtl_inst_list a' G rT s0 s1' cert mpc (Suc pc)"
 | 
|  |    330 |             and   "wtl_inst_list b G rT s1' s' cert mpc (Suc pc + length a')"; 
 | 
|  |    331 |           proof (elim allE impE);
 | 
|  |    332 |             from length Suc; show "nat < length list"; by simp; 
 | 
|  |    333 |             from wtlSuc;     show "wtl_inst_list list G rT s0 s' cert mpc (Suc pc)"; .; 
 | 
|  |    334 |           qed (elim exE conjE, auto);
 | 
|  |    335 |           with Suc wtlOpt;          
 | 
|  |    336 |           have "?E (a#a') b s1'"; by (auto simp del: split_paired_Ex);   
 | 
|  |    337 |           thus ?thesis; by blast;
 | 
|  |    338 |         qed;
 | 
|  |    339 |       qed;
 | 
|  |    340 |     qed;
 | 
|  |    341 |   qed;
 | 
|  |    342 | qed;
 | 
|  |    343 | 
 | 
|  |    344 | lemma "wtl_append1":
 | 
|  |    345 | "\\<lbrakk>wtl_inst_list x G rT s0 s1 cert (length (x@y)) 0; 
 | 
|  |    346 |   wtl_inst_list y G rT s1 s2 cert (length (x@y)) (length x)\\<rbrakk> \\<Longrightarrow>
 | 
|  |    347 |   wtl_inst_list (x@y) G rT s0 s2 cert (length (x@y)) 0";
 | 
|  |    348 | proof -;
 | 
|  |    349 |   assume w:
 | 
|  |    350 |     "wtl_inst_list x G rT s0 s1 cert (length (x@y)) 0"
 | 
|  |    351 |     "wtl_inst_list y G rT s1 s2 cert (length (x@y)) (length x)";
 | 
|  |    352 | 
 | 
|  |    353 | have
 | 
|  |    354 | "\\<forall> pc s0. 
 | 
|  |    355 |  wtl_inst_list x G rT s0 s1 cert (pc+length (x@y)) pc \\<longrightarrow> 
 | 
|  |    356 |  wtl_inst_list y G rT s1 s2 cert (pc+length (x@y)) (pc+length x) \\<longrightarrow>
 | 
|  |    357 |  wtl_inst_list (x@y) G rT s0 s2 cert (pc+length (x@y)) pc" (is "?P x");
 | 
|  |    358 | proof (induct "?P" "x");
 | 
|  |    359 |   case Nil;
 | 
|  |    360 |   show "?P []"; by simp;
 | 
|  |    361 | next;
 | 
|  |    362 |   case Cons;
 | 
|  |    363 |   show "?P (a#list)"; 
 | 
|  |    364 |   proof intro;
 | 
|  |    365 |     fix pc s0;
 | 
|  |    366 |     assume y: 
 | 
|  |    367 |       "wtl_inst_list y G rT s1 s2 cert (pc + length ((a # list) @ y)) (pc + length (a # list))";
 | 
|  |    368 |     assume al: 
 | 
|  |    369 |       "wtl_inst_list (a # list) G rT s0 s1 cert (pc + length ((a # list) @ y)) pc";
 | 
|  |    370 |     thus "wtl_inst_list ((a # list) @ y) G rT s0 s2 cert (pc + length ((a # list) @ y)) pc";
 | 
|  |    371 |       obtain s' where 
 | 
|  |    372 |        a: "wtl_inst_option a G rT s0 s' cert (Suc pc + length (list@y)) pc" and
 | 
|  |    373 |        l: "wtl_inst_list list G rT s' s1 cert (Suc pc + length (list@y)) (Suc pc)"; by auto;      
 | 
|  |    374 |       with y Cons;
 | 
|  |    375 |       have "wtl_inst_list (list @ y) G rT s' s2 cert (Suc pc + length (list @ y)) (Suc pc)";
 | 
|  |    376 |         by elim (assumption, simp+);
 | 
|  |    377 |       with a;
 | 
|  |    378 |       show ?thesis; by (auto simp del: split_paired_Ex);
 | 
|  |    379 |     qed;
 | 
|  |    380 |   qed;
 | 
|  |    381 | qed;
 | 
|  |    382 | 
 | 
|  |    383 |   with w;
 | 
|  |    384 |   show ?thesis; 
 | 
|  |    385 |   proof elim;
 | 
|  |    386 |     from w; show "wtl_inst_list x G rT s0 s1 cert (0+length (x @ y)) 0"; by simp;
 | 
|  |    387 |   qed simp+;
 | 
|  |    388 | qed;
 | 
|  |    389 | 
 | 
|  |    390 | lemma wtl_cons_appendl:
 | 
|  |    391 | "\\<lbrakk>wtl_inst_list a G rT s0 s1 cert (length (a@i#b)) 0; 
 | 
|  |    392 |   wtl_inst_option i G rT s1 s2 cert (length (a@i#b)) (length a); 
 | 
|  |    393 |   wtl_inst_list b G rT s2 s3 cert (length (a@i#b)) (Suc (length a))\\<rbrakk> \\<Longrightarrow> 
 | 
|  |    394 |   wtl_inst_list (a@i#b) G rT s0 s3 cert (length (a@i#b)) 0";
 | 
|  |    395 | proof -;
 | 
|  |    396 |   assume a: "wtl_inst_list a G rT s0 s1 cert (length (a@i#b)) 0";
 | 
|  |    397 | 
 | 
|  |    398 |   assume "wtl_inst_option i G rT s1 s2 cert (length (a@i#b)) (length a)"
 | 
|  |    399 |          "wtl_inst_list b G rT s2 s3 cert (length (a@i#b)) (Suc (length a))";
 | 
|  |    400 | 
 | 
|  |    401 |   hence "wtl_inst_list (i#b) G rT s1 s3 cert (length (a@i#b)) (length a)";
 | 
|  |    402 |     by (auto simp del: split_paired_Ex);
 | 
|  |    403 | 
 | 
|  |    404 |   with a;
 | 
|  |    405 |   show ?thesis; by (rule wtl_append1);
 | 
|  |    406 | qed;
 | 
|  |    407 | 
 | 
|  |    408 | lemma "wtl_append":
 | 
|  |    409 | "\\<lbrakk>wtl_inst_list a G rT s0 s1 cert (length (a@i#b)) 0; 
 | 
|  |    410 |   wtl_inst_option i G rT s1 s2 cert (length (a@i#b)) (length a); 
 | 
|  |    411 |   wtl_inst_list b G rT s2 s3 cert (length (a@i#b)) (Suc (length a))\\<rbrakk> \\<Longrightarrow> 
 | 
|  |    412 |   wtl_inst_list (a@[i]) G rT s0 s2 cert (length (a@i#b)) 0";
 | 
|  |    413 | proof -;
 | 
|  |    414 |   assume a: "wtl_inst_list a G rT s0 s1 cert (length (a@i#b)) 0";
 | 
|  |    415 |   assume i: "wtl_inst_option i G rT s1 s2 cert (length (a@i#b)) (length a)";
 | 
|  |    416 |   assume b: "wtl_inst_list b G rT s2 s3 cert (length (a@i#b)) (Suc (length a))";
 | 
|  |    417 | 
 | 
|  |    418 |   have "\\<forall> s0 pc. wtl_inst_list a G rT s0 s1 cert (pc+length (a@i#b)) pc \\<longrightarrow> 
 | 
|  |    419 |         wtl_inst_option i G rT s1 s2 cert (pc+length (a@i#b)) (pc + length a) \\<longrightarrow> 
 | 
|  |    420 |         wtl_inst_list b G rT s2 s3 cert (pc+length (a@i#b)) (Suc pc + length a) \\<longrightarrow> 
 | 
|  |    421 |           wtl_inst_list (a@[i]) G rT s0 s2 cert (pc+length (a@i#b)) pc" (is "?P a");
 | 
|  |    422 |   proof (induct "?P" "a");
 | 
|  |    423 |     case Nil;
 | 
|  |    424 |     show "?P []"; by (simp del: split_paired_Ex);
 | 
|  |    425 |     case Cons;
 | 
|  |    426 |     show "?P (a#list)" (is "\\<forall>s0 pc. ?x s0 pc \\<longrightarrow> ?y s0 pc \\<longrightarrow> ?z s0 pc \\<longrightarrow> ?p s0 pc"); 
 | 
|  |    427 |     proof intro;
 | 
|  |    428 |       fix s0 pc;
 | 
|  |    429 |       assume y: "?y s0 pc";
 | 
|  |    430 |       assume z: "?z s0 pc";
 | 
|  |    431 |       assume "?x s0 pc";
 | 
|  |    432 |       thus "?p s0 pc";
 | 
|  |    433 |         obtain s0' where opt: "wtl_inst_option a G rT s0 s0' cert (pc + length ((a # list) @ i # b)) pc"
 | 
|  |    434 |                     and list: "wtl_inst_list list G rT s0' s1 cert (Suc pc + length (list @ i # b)) (Suc pc)";
 | 
|  |    435 |           by (auto simp del: split_paired_Ex);
 | 
|  |    436 |         with y z Cons;
 | 
|  |    437 |         have "wtl_inst_list (list @ [i]) G rT s0' s2 cert (Suc pc + length (list @ i # b)) (Suc pc)"; 
 | 
|  |    438 |         proof elim; 
 | 
|  |    439 |           from list; show "wtl_inst_list list G rT s0' s1 cert (Suc pc + length (list @ i # b)) (Suc pc)"; .;
 | 
|  |    440 |         qed auto;
 | 
|  |    441 |         with opt;
 | 
|  |    442 |         show ?thesis; by (auto simp del: split_paired_Ex);
 | 
|  |    443 |       qed;
 | 
|  |    444 |     qed;
 | 
|  |    445 |   qed;
 | 
|  |    446 |   with a i b;
 | 
|  |    447 |   show ?thesis; 
 | 
|  |    448 |   proof elim;
 | 
|  |    449 |     from a; show "wtl_inst_list a G rT s0 s1 cert (0+length (a@i#b)) 0"; by simp;
 | 
|  |    450 |   qed auto;
 | 
|  |    451 | qed;
 | 
|  |    452 | 
 | 
|  |    453 | end;
 |