| 13673 |      1 | (*  Title:      HOL/MicroJava/Comp/Index.thy
 | 
|  |      2 |     ID:         $Id$
 | 
|  |      3 |     Author:     Martin Strecker
 | 
|  |      4 | *)
 | 
|  |      5 | 
 | 
|  |      6 | (* Index of variable in list of parameter names and local variables *)
 | 
|  |      7 | 
 | 
| 17778 |      8 | theory Index
 | 
|  |      9 | imports AuxLemmas DefsComp
 | 
|  |     10 | begin
 | 
| 13673 |     11 | 
 | 
|  |     12 | (*indexing a variable name among all variable declarations in a method body*)
 | 
|  |     13 | constdefs
 | 
|  |     14 |  index :: "java_mb => vname => nat"
 | 
|  |     15 |  "index ==  \<lambda> (pn,lv,blk,res) v.
 | 
|  |     16 |   if v = This
 | 
|  |     17 |   then 0 
 | 
|  |     18 |   else Suc (length (takeWhile (\<lambda> z. z~=v) (pn @ map fst lv)))"
 | 
|  |     19 | 
 | 
|  |     20 | 
 | 
|  |     21 | lemma index_length_pns: "
 | 
|  |     22 |   \<lbrakk> i = index (pns,lvars,blk,res) vn;
 | 
|  |     23 |   wf_java_mdecl G C ((mn,pTs),rT, (pns,lvars,blk,res)); 
 | 
|  |     24 |   vn \<in> set pns\<rbrakk> 
 | 
|  |     25 |   \<Longrightarrow> 0 < i \<and> i < Suc (length pns)"
 | 
|  |     26 | apply (simp add: wf_java_mdecl_def index_def)
 | 
|  |     27 | apply (subgoal_tac "vn \<noteq> This")
 | 
|  |     28 | apply (auto intro: length_takeWhile)
 | 
|  |     29 | done
 | 
|  |     30 | 
 | 
|  |     31 | lemma index_length_lvars: "
 | 
|  |     32 |   \<lbrakk> i = index (pns,lvars,blk,res) vn;
 | 
|  |     33 |   wf_java_mdecl G C ((mn,pTs),rT, (pns,lvars,blk,res)); 
 | 
|  |     34 |   vn \<in> set (map fst lvars)\<rbrakk> 
 | 
|  |     35 |   \<Longrightarrow> (length pns) < i \<and> i < Suc((length pns) + (length lvars))"
 | 
|  |     36 | apply (simp add: wf_java_mdecl_def index_def)
 | 
|  |     37 | apply (subgoal_tac "vn \<noteq> This")
 | 
|  |     38 | apply simp
 | 
|  |     39 | apply (subgoal_tac "\<forall> x \<in> set pns. (\<lambda>z. z \<noteq> vn) x")
 | 
|  |     40 | apply (simp add: takeWhile_append2)
 | 
|  |     41 | apply (subgoal_tac "length (takeWhile (\<lambda>z. z \<noteq> vn) (map fst lvars)) < length (map fst lvars)")
 | 
|  |     42 | apply simp
 | 
|  |     43 | apply (rule length_takeWhile)
 | 
|  |     44 | apply simp
 | 
|  |     45 | apply (simp add: map_of_in_set)
 | 
|  |     46 | apply (intro strip notI) apply simp apply blast
 | 
|  |     47 | done
 | 
|  |     48 | 
 | 
|  |     49 | 
 | 
|  |     50 | (***  index  ***)
 | 
|  |     51 | 
 | 
|  |     52 | lemma select_at_index : 
 | 
|  |     53 |   "x \<in> set (gjmb_plns (gmb G C S)) \<or> x = This 
 | 
|  |     54 |   \<Longrightarrow> (the (loc This) # glvs (gmb G C S) loc) ! (index (gmb G C S) x) = 
 | 
|  |     55 |      the (loc x)"
 | 
|  |     56 | apply (simp only: index_def gjmb_plns_def)
 | 
|  |     57 | apply (case_tac "(gmb G C S)")
 | 
|  |     58 | apply (simp add: galldefs del: set_append map_append)
 | 
|  |     59 | apply (case_tac b, simp add: gmb_def gjmb_lvs_def del: set_append map_append)
 | 
|  |     60 | apply (intro strip)
 | 
|  |     61 | apply (simp del: set_append map_append)
 | 
|  |     62 | apply (frule length_takeWhile)
 | 
|  |     63 | apply (frule_tac f = "(the \<circ> loc)" in nth_map)
 | 
|  |     64 | apply simp
 | 
|  |     65 | done
 | 
|  |     66 | 
 | 
| 14045 |     67 | lemma lift_if: "(f (if b then t else e)) = (if b then (f t) else (f e))"
 | 
|  |     68 | apply auto
 | 
|  |     69 | done
 | 
|  |     70 | 
 | 
| 13673 |     71 | lemma update_at_index: "
 | 
|  |     72 |   \<lbrakk> distinct (gjmb_plns (gmb G C S)); 
 | 
|  |     73 |   x \<in> set (gjmb_plns (gmb G C S)); x \<noteq> This \<rbrakk> \<Longrightarrow> 
 | 
|  |     74 |   locvars_xstate G C S (Norm (h, l))[index (gmb G C S) x := val] =
 | 
|  |     75 |           locvars_xstate G C S (Norm (h, l(x\<mapsto>val)))"
 | 
|  |     76 | apply (simp only: locvars_xstate_def locvars_locals_def index_def)
 | 
|  |     77 | apply (case_tac "(gmb G C S)", simp)
 | 
|  |     78 | apply (case_tac b, simp)
 | 
|  |     79 | apply (rule conjI)
 | 
|  |     80 | apply (simp add: gl_def)
 | 
|  |     81 | apply (simp add: galldefs del: set_append map_append)
 | 
|  |     82 | done
 | 
|  |     83 | 
 | 
|  |     84 | 
 | 
|  |     85 | (* !!!! incomprehensible: why can't List.takeWhile_append2 be applied the same 
 | 
|  |     86 |   way in the second case as in the first case ? *)
 | 
|  |     87 | lemma index_of_var: "\<lbrakk> xvar \<notin> set pns; xvar \<notin> set (map fst zs); xvar \<noteq> This \<rbrakk>
 | 
|  |     88 |   \<Longrightarrow> index (pns, zs @ ((xvar, xval) # xys), blk, res) xvar = Suc (length pns + length zs)"
 | 
|  |     89 | apply (simp add: index_def)
 | 
|  |     90 | apply (subgoal_tac "(\<And>x. ((x \<in> (set pns)) \<Longrightarrow> ((\<lambda>z. (z \<noteq> xvar))x)))")
 | 
|  |     91 | apply (simp add: List.takeWhile_append2)
 | 
|  |     92 | apply (subgoal_tac "(takeWhile (\<lambda>z. z \<noteq> xvar) (map fst zs @ xvar # map fst xys)) = map fst zs @ (takeWhile (\<lambda>z. z \<noteq> xvar) (xvar # map fst xys))")
 | 
|  |     93 | apply simp
 | 
|  |     94 | apply (rule List.takeWhile_append2)
 | 
|  |     95 | apply auto
 | 
|  |     96 | done
 | 
|  |     97 | 
 | 
|  |     98 | 
 | 
|  |     99 | 
 | 
|  |    100 | 
 | 
|  |    101 | (* The following def should replace the conditions in WellType.thy / wf_java_mdecl
 | 
|  |    102 | *)
 | 
|  |    103 | constdefs 
 | 
|  |    104 |   disjoint_varnames :: "[vname list, (vname \<times> ty) list] \<Rightarrow> bool"
 | 
|  |    105 | (* This corresponds to the original def in wf_java_mdecl:
 | 
|  |    106 |   "disjoint_varnames pns lvars \<equiv> 
 | 
|  |    107 |   nodups pns \<and> unique lvars \<and> This \<notin> set pns \<and> This \<notin> set (map fst lvars) \<and> 
 | 
|  |    108 | 	(\<forall>pn\<in>set pns. map_of lvars pn = None)"
 | 
|  |    109 | *)
 | 
|  |    110 | 
 | 
|  |    111 |   "disjoint_varnames pns lvars \<equiv> 
 | 
|  |    112 |   distinct pns \<and> unique lvars \<and> This \<notin> set pns \<and> This \<notin> set (map fst lvars) \<and> 
 | 
|  |    113 |   (\<forall>pn\<in>set pns. pn \<notin> set (map fst lvars))"
 | 
|  |    114 | 
 | 
|  |    115 | 
 | 
|  |    116 | lemma index_of_var2: "
 | 
|  |    117 |   disjoint_varnames pns (lvars_pre @ (vn, ty) # lvars_post)
 | 
|  |    118 |   \<Longrightarrow> index (pns, lvars_pre @ (vn, ty) # lvars_post, blk, res) vn =
 | 
|  |    119 |   Suc (length pns + length lvars_pre)"
 | 
|  |    120 | apply (simp add: disjoint_varnames_def index_def unique_def distinct_append)
 | 
|  |    121 | apply (subgoal_tac "vn \<noteq> This", simp)
 | 
|  |    122 | apply (subgoal_tac
 | 
|  |    123 |   "takeWhile (\<lambda>z. z \<noteq> vn) (map fst lvars_pre @ vn # map fst lvars_post) =
 | 
|  |    124 |   map fst lvars_pre @ takeWhile (\<lambda>z. z \<noteq> vn) (vn # map fst lvars_post)")
 | 
|  |    125 | apply simp 
 | 
|  |    126 | apply (rule List.takeWhile_append2)
 | 
|  |    127 | apply auto
 | 
|  |    128 | done
 | 
|  |    129 | 
 | 
|  |    130 | lemma wf_java_mdecl_disjoint_varnames: 
 | 
|  |    131 |   "wf_java_mdecl G C (S,rT,(pns,lvars,blk,res)) 
 | 
|  |    132 |   \<Longrightarrow> disjoint_varnames pns lvars"
 | 
|  |    133 | apply (case_tac S)
 | 
|  |    134 | apply (simp add: wf_java_mdecl_def disjoint_varnames_def  map_of_in_set)
 | 
|  |    135 | done
 | 
|  |    136 | 
 | 
|  |    137 | lemma wf_java_mdecl_length_pTs_pns: 
 | 
|  |    138 |   "wf_java_mdecl G C ((mn, pTs), rT, pns, lvars, blk, res)
 | 
|  |    139 |   \<Longrightarrow> length pTs = length pns"
 | 
|  |    140 | by (simp add: wf_java_mdecl_def)
 | 
|  |    141 | 
 | 
|  |    142 | end
 |