| 13673 |      1 | (*  Title:      HOL/MicroJava/Comp/DefsComp.thy
 | 
|  |      2 |     ID:         $Id$
 | 
|  |      3 |     Author:     Martin Strecker
 | 
|  |      4 |     Copyright   GPL 2002
 | 
|  |      5 | *)
 | 
|  |      6 | 
 | 
|  |      7 | (* Definitions for accessing parts of methods, states etc. *)
 | 
|  |      8 | 
 | 
|  |      9 | theory DefsComp = JVMExec:
 | 
|  |     10 | 
 | 
|  |     11 | 
 | 
|  |     12 | 
 | 
|  |     13 | constdefs
 | 
|  |     14 |   method_rT :: "cname \<times> ty \<times> 'c \<Rightarrow> ty"
 | 
|  |     15 |   "method_rT mtd == (fst (snd mtd))"
 | 
|  |     16 | 
 | 
|  |     17 | 
 | 
|  |     18 | constdefs
 | 
|  |     19 | (* g = get *)
 | 
|  |     20 |   gx :: "xstate \<Rightarrow> val option"  "gx \<equiv> fst"
 | 
|  |     21 |   gs :: "xstate \<Rightarrow> state"  "gs \<equiv> snd"
 | 
|  |     22 |   gh :: "xstate \<Rightarrow> aheap"        "gh \<equiv> fst\<circ>snd"
 | 
|  |     23 |   gl :: "xstate \<Rightarrow> State.locals" "gl \<equiv> snd\<circ>snd"
 | 
|  |     24 | 
 | 
|  |     25 |   gmb :: "'a prog \<Rightarrow> cname \<Rightarrow> sig \<Rightarrow> 'a"
 | 
|  |     26 |     "gmb G cn si \<equiv> snd(snd(the(method (G,cn) si)))"
 | 
|  |     27 |   gis :: "jvm_method \<Rightarrow> bytecode"
 | 
|  |     28 |     "gis \<equiv> fst \<circ> snd \<circ> snd"
 | 
|  |     29 |   glvs :: "java_mb \<Rightarrow> State.locals \<Rightarrow> locvars"
 | 
|  |     30 |     "glvs jmb loc \<equiv> map (the\<circ>loc) (gjmb_plns jmb)"
 | 
|  |     31 | 
 | 
|  |     32 | (* jmb = aus einem JavaMaethodBody *)
 | 
|  |     33 |   gjmb_pns  :: "java_mb \<Rightarrow> vname list"     "gjmb_pns \<equiv> fst"
 | 
|  |     34 |   gjmb_lvs  :: "java_mb \<Rightarrow> (vname\<times>ty)list" "gjmb_lvs \<equiv> fst\<circ>snd"
 | 
|  |     35 |   gjmb_blk  :: "java_mb \<Rightarrow> stmt"           "gjmb_blk \<equiv> fst\<circ>snd\<circ>snd"
 | 
|  |     36 |   gjmb_res  :: "java_mb \<Rightarrow> expr"           "gjmb_res \<equiv> snd\<circ>snd\<circ>snd"
 | 
|  |     37 |   gjmb_plns :: "java_mb \<Rightarrow> vname list"
 | 
|  |     38 |     "gjmb_plns \<equiv> \<lambda>jmb. gjmb_pns jmb @ map fst (gjmb_lvs jmb)"
 | 
|  |     39 |   
 | 
|  |     40 | lemmas gdefs = gx_def gh_def gl_def gmb_def gis_def glvs_def
 | 
|  |     41 | lemmas gjmbdefs = gjmb_pns_def gjmb_lvs_def gjmb_blk_def gjmb_res_def gjmb_plns_def
 | 
|  |     42 | 
 | 
|  |     43 | lemmas galldefs = gdefs gjmbdefs
 | 
|  |     44 | 
 | 
|  |     45 | 
 | 
|  |     46 | 
 | 
|  |     47 | constdefs 
 | 
|  |     48 |   locvars_locals :: "java_mb prog \<Rightarrow> cname \<Rightarrow> sig \<Rightarrow> State.locals \<Rightarrow> locvars"
 | 
|  |     49 |   "locvars_locals G C S lvs == the (lvs This) # glvs (gmb G C S) lvs"
 | 
|  |     50 | 
 | 
|  |     51 |   locals_locvars :: "java_mb prog \<Rightarrow> cname \<Rightarrow> sig \<Rightarrow> locvars \<Rightarrow> State.locals"
 | 
|  |     52 |   "locals_locvars G C S lvs == 
 | 
|  |     53 |   empty ((gjmb_plns (gmb G C S))[\<mapsto>](tl lvs)) (This\<mapsto>(hd lvs))"
 | 
|  |     54 | 
 | 
|  |     55 |   locvars_xstate :: "java_mb prog \<Rightarrow> cname \<Rightarrow> sig \<Rightarrow> xstate \<Rightarrow> locvars"
 | 
|  |     56 |   "locvars_xstate G C S xs == locvars_locals G C S (gl xs)"
 | 
|  |     57 | 
 | 
|  |     58 | 
 | 
|  |     59 | lemma locvars_xstate_par_dep: 
 | 
|  |     60 |   "lv1 = lv2 \<Longrightarrow> 
 | 
|  |     61 |   locvars_xstate G C S (xcpt1, hp1, lv1) = locvars_xstate G C S (xcpt2, hp2, lv2)"
 | 
|  |     62 | by (simp add: locvars_xstate_def gl_def)
 | 
|  |     63 | 
 | 
|  |     64 | 
 | 
|  |     65 | 
 | 
|  |     66 | 
 | 
|  |     67 | (**********************************************************************)
 | 
|  |     68 | (* Conversions *)
 | 
|  |     69 | 
 | 
|  |     70 | 
 | 
|  |     71 | lemma gx_conv [simp]: "gx (xcpt, s) = xcpt" by (simp add: gx_def)
 | 
|  |     72 | 
 | 
|  |     73 | lemma gh_conv [simp]: "gh (xcpt, h, l) = h" by (simp add: gh_def)
 | 
|  |     74 | 
 | 
|  |     75 | 
 | 
|  |     76 | end
 |