| 11229 |      1 | (*  Title:      HOL/MicroJava/BV/Typing_Framework_err.thy
 | 
|  |      2 |     ID:         $Id$
 | 
|  |      3 |     Author:     Gerwin Klein
 | 
|  |      4 |     Copyright   2000 TUM
 | 
|  |      5 | 
 | 
|  |      6 | *)
 | 
|  |      7 | 
 | 
| 13224 |      8 | header {* \isaheader{Lifting the Typing Framework to err, app, and eff} *}
 | 
| 11229 |      9 | 
 | 
| 27680 |     10 | theory Typing_Framework_err
 | 
|  |     11 | imports Typing_Framework SemilatAlg
 | 
|  |     12 | begin
 | 
| 11229 |     13 | 
 | 
|  |     14 | constdefs
 | 
|  |     15 | 
 | 
| 13066 |     16 | wt_err_step :: "'s ord \<Rightarrow> 's err step_type \<Rightarrow> 's err list \<Rightarrow> bool"
 | 
|  |     17 | "wt_err_step r step ts \<equiv> wt_step (Err.le r) Err step ts"
 | 
| 12516 |     18 | 
 | 
| 13066 |     19 | wt_app_eff :: "'s ord \<Rightarrow> (nat \<Rightarrow> 's \<Rightarrow> bool) \<Rightarrow> 's step_type \<Rightarrow> 's list \<Rightarrow> bool"
 | 
|  |     20 | "wt_app_eff r app step ts \<equiv>
 | 
| 12516 |     21 |   \<forall>p < size ts. app p (ts!p) \<and> (\<forall>(q,t) \<in> set (step p (ts!p)). t <=_r ts!q)"
 | 
|  |     22 | 
 | 
|  |     23 | map_snd :: "('b \<Rightarrow> 'c) \<Rightarrow> ('a \<times> 'b) list \<Rightarrow> ('a \<times> 'c) list"
 | 
| 13066 |     24 | "map_snd f \<equiv> map (\<lambda>(x,y). (x, f y))"
 | 
|  |     25 | 
 | 
|  |     26 | error :: "nat \<Rightarrow> (nat \<times> 'a err) list"
 | 
| 15425 |     27 | "error n \<equiv> map (\<lambda>x. (x,Err)) [0..<n]"
 | 
| 11229 |     28 | 
 | 
| 13066 |     29 | err_step :: "nat \<Rightarrow> (nat \<Rightarrow> 's \<Rightarrow> bool) \<Rightarrow> 's step_type \<Rightarrow> 's err step_type"
 | 
|  |     30 | "err_step n app step p t \<equiv> 
 | 
|  |     31 |   case t of 
 | 
|  |     32 |     Err   \<Rightarrow> error n
 | 
|  |     33 |   | OK t' \<Rightarrow> if app p t' then map_snd OK (step p t') else error n"
 | 
| 11229 |     34 | 
 | 
| 13224 |     35 | app_mono :: "'s ord \<Rightarrow> (nat \<Rightarrow> 's \<Rightarrow> bool) \<Rightarrow> nat \<Rightarrow> 's set \<Rightarrow> bool"
 | 
|  |     36 | "app_mono r app n A \<equiv>
 | 
|  |     37 |  \<forall>s p t. s \<in> A \<and> p < n \<and> s <=_r t \<longrightarrow> app p t \<longrightarrow> app p s"
 | 
|  |     38 | 
 | 
|  |     39 | 
 | 
| 13066 |     40 | lemmas err_step_defs = err_step_def map_snd_def error_def
 | 
| 12516 |     41 | 
 | 
| 13224 |     42 | 
 | 
| 13067 |     43 | lemma bounded_err_stepD:
 | 
|  |     44 |   "bounded (err_step n app step) n \<Longrightarrow> 
 | 
|  |     45 |   p < n \<Longrightarrow>  app p a \<Longrightarrow> (q,b) \<in> set (step p a) \<Longrightarrow> 
 | 
|  |     46 |   q < n"
 | 
|  |     47 |   apply (simp add: bounded_def err_step_def)
 | 
|  |     48 |   apply (erule allE, erule impE, assumption)
 | 
|  |     49 |   apply (erule_tac x = "OK a" in allE, drule bspec)
 | 
|  |     50 |    apply (simp add: map_snd_def)
 | 
|  |     51 |    apply fast
 | 
|  |     52 |   apply simp
 | 
|  |     53 |   done
 | 
|  |     54 | 
 | 
|  |     55 | 
 | 
|  |     56 | lemma in_map_sndD: "(a,b) \<in> set (map_snd f xs) \<Longrightarrow> \<exists>b'. (a,b') \<in> set xs"
 | 
|  |     57 |   apply (induct xs)
 | 
|  |     58 |   apply (auto simp add: map_snd_def)
 | 
|  |     59 |   done
 | 
| 11229 |     60 | 
 | 
| 12516 |     61 | 
 | 
| 13067 |     62 | lemma bounded_err_stepI:
 | 
|  |     63 |   "\<forall>p. p < n \<longrightarrow> (\<forall>s. ap p s \<longrightarrow> (\<forall>(q,s') \<in> set (step p s). q < n))
 | 
|  |     64 |   \<Longrightarrow> bounded (err_step n ap step) n"
 | 
|  |     65 | apply (unfold bounded_def)
 | 
|  |     66 | apply clarify
 | 
|  |     67 | apply (simp add: err_step_def split: err.splits)
 | 
|  |     68 | apply (simp add: error_def)
 | 
|  |     69 |  apply blast
 | 
|  |     70 | apply (simp split: split_if_asm) 
 | 
|  |     71 |  apply (blast dest: in_map_sndD)
 | 
|  |     72 | apply (simp add: error_def)
 | 
|  |     73 | apply blast
 | 
|  |     74 | done
 | 
|  |     75 | 
 | 
|  |     76 | 
 | 
| 13224 |     77 | lemma bounded_lift:
 | 
|  |     78 |   "bounded step n \<Longrightarrow> bounded (err_step n app step) n"
 | 
|  |     79 |   apply (unfold bounded_def err_step_def error_def)
 | 
|  |     80 |   apply clarify
 | 
|  |     81 |   apply (erule allE, erule impE, assumption)
 | 
|  |     82 |   apply (case_tac s)
 | 
|  |     83 |   apply (auto simp add: map_snd_def split: split_if_asm)
 | 
|  |     84 |   done
 | 
|  |     85 | 
 | 
|  |     86 | 
 | 
|  |     87 | lemma le_list_map_OK [simp]:
 | 
|  |     88 |   "\<And>b. map OK a <=[Err.le r] map OK b = (a <=[r] b)"
 | 
|  |     89 |   apply (induct a)
 | 
|  |     90 |    apply simp
 | 
|  |     91 |   apply simp
 | 
|  |     92 |   apply (case_tac b)
 | 
|  |     93 |    apply simp
 | 
|  |     94 |   apply simp
 | 
|  |     95 |   done
 | 
|  |     96 | 
 | 
|  |     97 | 
 | 
|  |     98 | lemma map_snd_lessI:
 | 
|  |     99 |   "x <=|r| y \<Longrightarrow> map_snd OK x <=|Err.le r| map_snd OK y"
 | 
|  |    100 |   apply (induct x)
 | 
|  |    101 |   apply (unfold lesubstep_type_def map_snd_def)
 | 
|  |    102 |   apply auto
 | 
|  |    103 |   done
 | 
|  |    104 | 
 | 
|  |    105 | 
 | 
|  |    106 | lemma mono_lift:
 | 
|  |    107 |   "order r \<Longrightarrow> app_mono r app n A \<Longrightarrow> bounded (err_step n app step) n \<Longrightarrow>
 | 
|  |    108 |   \<forall>s p t. s \<in> A \<and> p < n \<and> s <=_r t \<longrightarrow> app p t \<longrightarrow> step p s <=|r| step p t \<Longrightarrow>
 | 
|  |    109 |   mono (Err.le r) (err_step n app step) n (err A)"
 | 
|  |    110 | apply (unfold app_mono_def mono_def err_step_def)
 | 
|  |    111 | apply clarify
 | 
|  |    112 | apply (case_tac s)
 | 
|  |    113 |  apply simp 
 | 
|  |    114 | apply simp
 | 
|  |    115 | apply (case_tac t)
 | 
|  |    116 |  apply simp
 | 
|  |    117 |  apply clarify
 | 
|  |    118 |  apply (simp add: lesubstep_type_def error_def)
 | 
|  |    119 |  apply clarify
 | 
|  |    120 |  apply (drule in_map_sndD)
 | 
|  |    121 |  apply clarify
 | 
|  |    122 |  apply (drule bounded_err_stepD, assumption+)
 | 
|  |    123 |  apply (rule exI [of _ Err])
 | 
|  |    124 |  apply simp
 | 
|  |    125 | apply simp
 | 
|  |    126 | apply (erule allE, erule allE, erule allE, erule impE)
 | 
|  |    127 |  apply (rule conjI, assumption)
 | 
|  |    128 |  apply (rule conjI, assumption)
 | 
|  |    129 |  apply assumption
 | 
|  |    130 | apply (rule conjI)
 | 
|  |    131 | apply clarify
 | 
|  |    132 | apply (erule allE, erule allE, erule allE, erule impE)
 | 
|  |    133 |  apply (rule conjI, assumption)
 | 
|  |    134 |  apply (rule conjI, assumption)
 | 
|  |    135 |  apply assumption
 | 
|  |    136 | apply (erule impE, assumption)
 | 
|  |    137 | apply (rule map_snd_lessI, assumption)
 | 
|  |    138 | apply clarify
 | 
|  |    139 | apply (simp add: lesubstep_type_def error_def)
 | 
|  |    140 | apply clarify
 | 
|  |    141 | apply (drule in_map_sndD)
 | 
|  |    142 | apply clarify
 | 
|  |    143 | apply (drule bounded_err_stepD, assumption+)
 | 
|  |    144 | apply (rule exI [of _ Err])
 | 
|  |    145 | apply simp
 | 
|  |    146 | done
 | 
|  |    147 |  
 | 
|  |    148 | lemma in_errorD:
 | 
|  |    149 |   "(x,y) \<in> set (error n) \<Longrightarrow> y = Err"
 | 
|  |    150 |   by (auto simp add: error_def)
 | 
|  |    151 | 
 | 
|  |    152 | lemma pres_type_lift:
 | 
|  |    153 |   "\<forall>s\<in>A. \<forall>p. p < n \<longrightarrow> app p s \<longrightarrow> (\<forall>(q, s')\<in>set (step p s). s' \<in> A) 
 | 
|  |    154 |   \<Longrightarrow> pres_type (err_step n app step) n (err A)"  
 | 
|  |    155 | apply (unfold pres_type_def err_step_def)
 | 
|  |    156 | apply clarify
 | 
|  |    157 | apply (case_tac b)
 | 
|  |    158 |  apply simp
 | 
|  |    159 | apply (case_tac s)
 | 
|  |    160 |  apply simp
 | 
|  |    161 |  apply (drule in_errorD)
 | 
|  |    162 |  apply simp
 | 
|  |    163 | apply (simp add: map_snd_def split: split_if_asm)
 | 
|  |    164 |  apply fast
 | 
|  |    165 | apply (drule in_errorD)
 | 
|  |    166 | apply simp
 | 
|  |    167 | done
 | 
|  |    168 | 
 | 
|  |    169 | 
 | 
|  |    170 | 
 | 
| 13067 |    171 | text {*
 | 
|  |    172 |   There used to be a condition here that each instruction must have a
 | 
|  |    173 |   successor. This is not needed any more, because the definition of
 | 
|  |    174 |   @{term error} trivially ensures that there is a successor for
 | 
|  |    175 |   the critical case where @{term app} does not hold. 
 | 
|  |    176 | *}
 | 
| 13066 |    177 | lemma wt_err_imp_wt_app_eff:
 | 
|  |    178 |   assumes wt: "wt_err_step r (err_step (size ts) app step) ts"
 | 
| 13067 |    179 |   assumes b:  "bounded (err_step (size ts) app step) (size ts)"
 | 
| 13066 |    180 |   shows "wt_app_eff r app step (map ok_val ts)"
 | 
|  |    181 | proof (unfold wt_app_eff_def, intro strip, rule conjI)
 | 
|  |    182 |   fix p assume "p < size (map ok_val ts)"
 | 
|  |    183 |   hence lp: "p < size ts" by simp
 | 
| 13067 |    184 |   hence ts: "0 < size ts" by (cases p) auto
 | 
|  |    185 |   hence err: "(0,Err) \<in> set (error (size ts))" by (simp add: error_def)
 | 
| 11229 |    186 | 
 | 
|  |    187 |   from wt lp
 | 
| 13066 |    188 |   have [intro?]: "\<And>p. p < size ts \<Longrightarrow> ts ! p \<noteq> Err" 
 | 
|  |    189 |     by (unfold wt_err_step_def wt_step_def) simp
 | 
| 11229 |    190 | 
 | 
|  |    191 |   show app: "app p (map ok_val ts ! p)"
 | 
| 13067 |    192 |   proof (rule ccontr)
 | 
|  |    193 |     from wt lp obtain s where
 | 
| 12516 |    194 |       OKp:  "ts ! p = OK s" and
 | 
| 13066 |    195 |       less: "\<forall>(q,t) \<in> set (err_step (size ts) app step p (ts!p)). t <=_(Err.le r) ts!q"
 | 
|  |    196 |       by (unfold wt_err_step_def wt_step_def stable_def) 
 | 
| 12516 |    197 |          (auto iff: not_Err_eq)
 | 
| 13067 |    198 |     assume "\<not> app p (map ok_val ts ! p)"
 | 
|  |    199 |     with OKp lp have "\<not> app p s" by simp
 | 
|  |    200 |     with OKp have "err_step (size ts) app step p (ts!p) = error (size ts)" 
 | 
|  |    201 |       by (simp add: err_step_def)    
 | 
|  |    202 |     with err ts obtain q where 
 | 
|  |    203 |       "(q,Err) \<in> set (err_step (size ts) app step p (ts!p))" and
 | 
|  |    204 |       q: "q < size ts" by auto    
 | 
|  |    205 |     with less have "ts!q = Err" by auto
 | 
|  |    206 |     moreover from q have "ts!q \<noteq> Err" ..
 | 
|  |    207 |     ultimately show False by blast
 | 
| 12516 |    208 |   qed
 | 
|  |    209 |   
 | 
|  |    210 |   show "\<forall>(q,t)\<in>set(step p (map ok_val ts ! p)). t <=_r map ok_val ts ! q" 
 | 
|  |    211 |   proof clarify
 | 
|  |    212 |     fix q t assume q: "(q,t) \<in> set (step p (map ok_val ts ! p))"
 | 
| 11229 |    213 | 
 | 
|  |    214 |     from wt lp q
 | 
|  |    215 |     obtain s where
 | 
|  |    216 |       OKp:  "ts ! p = OK s" and
 | 
| 13066 |    217 |       less: "\<forall>(q,t) \<in> set (err_step (size ts) app step p (ts!p)). t <=_(Err.le r) ts!q"
 | 
|  |    218 |       by (unfold wt_err_step_def wt_step_def stable_def) 
 | 
| 11229 |    219 |          (auto iff: not_Err_eq)
 | 
|  |    220 | 
 | 
| 13067 |    221 |     from b lp app q have lq: "q < size ts" by (rule bounded_err_stepD)
 | 
| 12516 |    222 |     hence "ts!q \<noteq> Err" ..
 | 
|  |    223 |     then obtain s' where OKq: "ts ! q = OK s'" by (auto iff: not_Err_eq)
 | 
| 11229 |    224 | 
 | 
| 12516 |    225 |     from lp lq OKp OKq app less q
 | 
|  |    226 |     show "t <=_r map ok_val ts ! q"
 | 
|  |    227 |       by (auto simp add: err_step_def map_snd_def) 
 | 
| 11229 |    228 |   qed
 | 
|  |    229 | qed
 | 
|  |    230 | 
 | 
|  |    231 | 
 | 
| 13066 |    232 | lemma wt_app_eff_imp_wt_err:
 | 
|  |    233 |   assumes app_eff: "wt_app_eff r app step ts"
 | 
|  |    234 |   assumes bounded: "bounded (err_step (size ts) app step) (size ts)"
 | 
|  |    235 |   shows "wt_err_step r (err_step (size ts) app step) (map OK ts)"
 | 
|  |    236 | proof (unfold wt_err_step_def wt_step_def, intro strip, rule conjI)
 | 
|  |    237 |   fix p assume "p < size (map OK ts)" 
 | 
|  |    238 |   hence p: "p < size ts" by simp
 | 
| 11229 |    239 |   thus "map OK ts ! p \<noteq> Err" by simp
 | 
| 12516 |    240 |   { fix q t
 | 
| 13066 |    241 |     assume q: "(q,t) \<in> set (err_step (size ts) app step p (map OK ts ! p))" 
 | 
|  |    242 |     with p app_eff obtain 
 | 
| 12516 |    243 |       "app p (ts ! p)" "\<forall>(q,t) \<in> set (step p (ts!p)). t <=_r ts!q"
 | 
| 13066 |    244 |       by (unfold wt_app_eff_def) blast
 | 
| 11229 |    245 |     moreover
 | 
| 13066 |    246 |     from q p bounded have "q < size ts"
 | 
|  |    247 |       by - (rule boundedD)
 | 
| 11229 |    248 |     hence "map OK ts ! q = OK (ts!q)" by simp
 | 
|  |    249 |     moreover
 | 
|  |    250 |     have "p < size ts" by (rule p)
 | 
| 12516 |    251 |     moreover note q
 | 
| 11229 |    252 |     ultimately     
 | 
| 12516 |    253 |     have "t <=_(Err.le r) map OK ts ! q" 
 | 
|  |    254 |       by (auto simp add: err_step_def map_snd_def)
 | 
| 11229 |    255 |   }
 | 
| 13066 |    256 |   thus "stable (Err.le r) (err_step (size ts) app step) (map OK ts) p"
 | 
| 11229 |    257 |     by (unfold stable_def) blast
 | 
|  |    258 | qed
 | 
|  |    259 | 
 | 
|  |    260 | end
 | 
| 13224 |    261 | 
 |