| 13673 |      1 | (*  Title:      HOL/MicroJava/Comp/TranslCompTp.thy
 | 
|  |      2 |     ID:         $Id$
 | 
|  |      3 |     Author:     Martin Strecker
 | 
|  |      4 |     Copyright   GPL 2002
 | 
|  |      5 | *)
 | 
|  |      6 | 
 | 
|  |      7 | theory TranslCompTp =  JVMType + Index:
 | 
|  |      8 | 
 | 
|  |      9 | 
 | 
|  |     10 | (**********************************************************************)
 | 
|  |     11 | 
 | 
|  |     12 | 
 | 
|  |     13 | constdefs
 | 
|  |     14 |   comb :: "['a \<Rightarrow> 'b list \<times> 'c, 'c \<Rightarrow> 'b list \<times> 'd, 'a] \<Rightarrow> 'b list \<times> 'd" 
 | 
|  |     15 |   "comb == (\<lambda> f1 f2 x0. let (xs1, x1) = f1 x0; 
 | 
|  |     16 |                             (xs2, x2) = f2 x1 
 | 
|  |     17 |                         in  (xs1 @ xs2, x2))"
 | 
|  |     18 |   comb_nil :: "'a \<Rightarrow> 'b list \<times> 'a"
 | 
|  |     19 |   "comb_nil a == ([], a)"
 | 
|  |     20 | 
 | 
|  |     21 | syntax (xsymbols)
 | 
|  |     22 |   "comb" :: "['a \<Rightarrow> 'b list \<times> 'c, 'c \<Rightarrow> 'b list \<times> 'd, 'a] \<Rightarrow> 'b list \<times> 'd"    
 | 
|  |     23 |             (infixr "\<box>" 55)
 | 
|  |     24 | 
 | 
|  |     25 | lemma comb_nil_left [simp]: "comb_nil \<box> f = f"
 | 
|  |     26 | by (simp add: comb_def comb_nil_def split_beta)
 | 
|  |     27 | 
 | 
|  |     28 | lemma comb_nil_right [simp]: "f \<box> comb_nil = f"
 | 
|  |     29 | by (simp add: comb_def comb_nil_def split_beta)
 | 
|  |     30 | 
 | 
|  |     31 | lemma comb_assoc [simp]: "(fa \<box> fb) \<box> fc = fa \<box> (fb \<box> fc)"
 | 
|  |     32 | by (simp add: comb_def split_beta)
 | 
|  |     33 | 
 | 
|  |     34 | lemma comb_inv: "(xs', x') = (f1 \<box> f2) x0 \<Longrightarrow>
 | 
|  |     35 |   \<exists> xs1 x1 xs2 x2. (xs1, x1) = (f1 x0) \<and> (xs2, x2) = f2 x1 \<and> xs'= xs1 @ xs2 \<and> x'=x2"
 | 
|  |     36 | apply (case_tac "f1 x0")
 | 
|  |     37 | apply (case_tac "f2 x1")
 | 
|  |     38 | apply (simp add: comb_def split_beta)
 | 
|  |     39 | done
 | 
|  |     40 | 
 | 
|  |     41 | (**********************************************************************)
 | 
|  |     42 | 
 | 
|  |     43 | syntax
 | 
|  |     44 |   mt_of		:: "method_type \<times> state_type \<Rightarrow> method_type"
 | 
|  |     45 |   sttp_of	:: "method_type \<times> state_type \<Rightarrow> state_type"
 | 
|  |     46 | 
 | 
|  |     47 | translations
 | 
|  |     48 |   "mt_of"     => "fst"
 | 
|  |     49 |   "sttp_of"   => "snd"
 | 
|  |     50 | 
 | 
|  |     51 | consts
 | 
|  |     52 |   compTpExpr  :: "java_mb \<Rightarrow> java_mb prog \<Rightarrow> expr
 | 
|  |     53 |                    \<Rightarrow> state_type \<Rightarrow> method_type \<times> state_type"
 | 
|  |     54 | 
 | 
|  |     55 |   compTpExprs :: "java_mb \<Rightarrow> java_mb prog \<Rightarrow> expr list
 | 
|  |     56 |                    \<Rightarrow> state_type \<Rightarrow> method_type \<times> state_type"
 | 
|  |     57 | 
 | 
|  |     58 |   compTpStmt  :: "java_mb \<Rightarrow> java_mb prog \<Rightarrow> stmt
 | 
|  |     59 |                    \<Rightarrow> state_type \<Rightarrow> method_type \<times> state_type"
 | 
|  |     60 | 
 | 
|  |     61 | 
 | 
|  |     62 | constdefs
 | 
|  |     63 |   nochangeST :: "state_type \<Rightarrow> method_type \<times> state_type"
 | 
|  |     64 |   "nochangeST sttp == ([Some sttp], sttp)"
 | 
|  |     65 |   pushST :: "[ty list, state_type] \<Rightarrow> method_type \<times> state_type"
 | 
|  |     66 |   "pushST tps == (\<lambda> (ST, LT). ([Some (ST, LT)], (tps @ ST, LT)))"
 | 
|  |     67 |   dupST :: "state_type \<Rightarrow> method_type \<times> state_type"
 | 
|  |     68 |   "dupST == (\<lambda> (ST, LT). ([Some (ST, LT)], (hd ST # ST, LT)))"
 | 
|  |     69 |   dup_x1ST :: "state_type \<Rightarrow> method_type \<times> state_type"
 | 
|  |     70 |   "dup_x1ST == (\<lambda> (ST, LT). ([Some (ST, LT)], 
 | 
|  |     71 |                              (hd ST # hd (tl ST) # hd ST # (tl (tl ST)), LT)))"
 | 
|  |     72 |   popST :: "[nat, state_type] \<Rightarrow> method_type \<times> state_type"
 | 
|  |     73 |   "popST n == (\<lambda> (ST, LT). ([Some (ST, LT)], (drop n ST, LT)))"
 | 
|  |     74 |   replST :: "[nat, ty, state_type] \<Rightarrow> method_type \<times> state_type"
 | 
|  |     75 |   "replST n tp == (\<lambda> (ST, LT). ([Some (ST, LT)], (tp # (drop n ST), LT)))"
 | 
|  |     76 | 
 | 
|  |     77 |   storeST :: "[nat, ty, state_type] \<Rightarrow> method_type \<times> state_type"
 | 
|  |     78 |   "storeST i tp == (\<lambda> (ST, LT). ([Some (ST, LT)], (tl ST, LT [i:= OK tp])))"
 | 
|  |     79 | 
 | 
|  |     80 | 
 | 
|  |     81 | (* Expressions *)
 | 
|  |     82 | 
 | 
|  |     83 | primrec
 | 
|  |     84 | 
 | 
|  |     85 |   "compTpExpr jmb G (NewC c) = pushST [Class c]"
 | 
|  |     86 | 
 | 
|  |     87 |   "compTpExpr jmb G (Cast c e) = 
 | 
|  |     88 |   (compTpExpr jmb G e) \<box> (replST 1 (Class c))"
 | 
|  |     89 | 
 | 
|  |     90 |   "compTpExpr jmb G (Lit val) = pushST [the (typeof (\<lambda>v. None) val)]"
 | 
|  |     91 | 
 | 
|  |     92 |   "compTpExpr jmb G (BinOp bo e1 e2) = 
 | 
|  |     93 |      (compTpExpr jmb G e1) \<box> (compTpExpr jmb G e2) \<box> 
 | 
|  |     94 |      (case bo of 
 | 
|  |     95 |        Eq => popST 2 \<box> pushST [PrimT Boolean] \<box> popST 1 \<box> pushST [PrimT Boolean]
 | 
|  |     96 |      | Add => replST 2 (PrimT Integer))"
 | 
|  |     97 | 
 | 
|  |     98 |   "compTpExpr jmb G (LAcc vn) = (\<lambda> (ST, LT). 
 | 
|  |     99 |       pushST [ok_val (LT ! (index jmb vn))] (ST, LT))"
 | 
|  |    100 | 
 | 
|  |    101 |   "compTpExpr jmb G (vn::=e) = 
 | 
|  |    102 |       (compTpExpr jmb G e) \<box> dupST \<box> (popST 1)"
 | 
|  |    103 | 
 | 
|  |    104 | 
 | 
|  |    105 |   "compTpExpr jmb G ( {cn}e..fn ) = 
 | 
|  |    106 |       (compTpExpr jmb G e) \<box> replST 1 (snd (the (field (G,cn) fn)))"
 | 
|  |    107 | 
 | 
|  |    108 |   "compTpExpr jmb G (FAss cn e1 fn e2 ) = 
 | 
|  |    109 |       (compTpExpr jmb G e1) \<box> (compTpExpr jmb G e2) \<box> dup_x1ST \<box> (popST 2)"
 | 
|  |    110 | 
 | 
|  |    111 | 
 | 
|  |    112 |   "compTpExpr jmb G ({C}a..mn({fpTs}ps)) =
 | 
|  |    113 |        (compTpExpr jmb G a) \<box> (compTpExprs jmb G ps) \<box> 
 | 
|  |    114 |        (replST ((length ps) + 1) (method_rT (the (method (G,C) (mn,fpTs)))))"
 | 
|  |    115 | 
 | 
|  |    116 | 
 | 
|  |    117 | (* Expression lists *)
 | 
|  |    118 | 
 | 
|  |    119 |   "compTpExprs jmb G [] = comb_nil"
 | 
|  |    120 | 
 | 
|  |    121 |   "compTpExprs jmb G (e#es) = (compTpExpr jmb G e) \<box> (compTpExprs jmb G es)"
 | 
|  |    122 | 
 | 
|  |    123 | 
 | 
|  |    124 | (* Statements *)
 | 
|  |    125 | primrec
 | 
|  |    126 |    "compTpStmt jmb G Skip = comb_nil"
 | 
|  |    127 | 
 | 
|  |    128 |    "compTpStmt jmb G (Expr e) =  (compTpExpr jmb G e) \<box> popST 1"
 | 
|  |    129 | 
 | 
|  |    130 |    "compTpStmt jmb G (c1;; c2) = (compTpStmt jmb G c1) \<box> (compTpStmt jmb G c2)"
 | 
|  |    131 | 
 | 
|  |    132 |    "compTpStmt jmb G (If(e) c1 Else c2) = 
 | 
|  |    133 |       (pushST [PrimT Boolean]) \<box> (compTpExpr jmb G e) \<box> popST 2 \<box>
 | 
|  |    134 |       (compTpStmt jmb G c1) \<box> nochangeST \<box> (compTpStmt jmb G c2)"
 | 
|  |    135 | 
 | 
|  |    136 |    "compTpStmt jmb G (While(e) c) = 
 | 
|  |    137 |       (pushST [PrimT Boolean]) \<box> (compTpExpr jmb G e) \<box> popST 2 \<box>
 | 
|  |    138 |       (compTpStmt jmb G c) \<box> nochangeST"
 | 
|  |    139 | 
 | 
|  |    140 | constdefs
 | 
|  |    141 |   compTpInit  :: "java_mb \<Rightarrow> (vname * ty)
 | 
|  |    142 |                    \<Rightarrow> state_type \<Rightarrow> method_type \<times> state_type"
 | 
|  |    143 |   "compTpInit jmb == (\<lambda> (vn,ty). (pushST [ty]) \<box>  (storeST (index jmb vn) ty))"
 | 
|  |    144 | 
 | 
|  |    145 | consts
 | 
|  |    146 |   compTpInitLvars :: "[java_mb, (vname \<times> ty) list] 
 | 
|  |    147 |                    \<Rightarrow> state_type \<Rightarrow> method_type \<times> state_type"
 | 
|  |    148 | 
 | 
|  |    149 | primrec 
 | 
|  |    150 |   "compTpInitLvars jmb [] = comb_nil"
 | 
|  |    151 |   "compTpInitLvars jmb (lv#lvars) = (compTpInit jmb lv) \<box> (compTpInitLvars jmb lvars)"
 | 
|  |    152 | 
 | 
|  |    153 | constdefs
 | 
|  |    154 |    start_ST :: "opstack_type"
 | 
|  |    155 |   "start_ST == []"
 | 
|  |    156 | 
 | 
|  |    157 |    start_LT :: "cname \<Rightarrow> ty list \<Rightarrow> nat \<Rightarrow> locvars_type"
 | 
|  |    158 |   "start_LT C pTs n ==  (OK (Class C))#((map OK pTs))@(replicate n Err)"
 | 
|  |    159 | 
 | 
|  |    160 |   compTpMethod  :: "[java_mb prog, cname, java_mb mdecl] \<Rightarrow> method_type"
 | 
|  |    161 |   "compTpMethod G C == \<lambda> ((mn,pTs),rT, jmb). 
 | 
|  |    162 |                          let (pns,lvars,blk,res) = jmb
 | 
|  |    163 |                          in (mt_of
 | 
|  |    164 |                             ((compTpInitLvars jmb lvars \<box> 
 | 
|  |    165 |                               compTpStmt jmb G blk \<box> 
 | 
|  |    166 |                               compTpExpr jmb G res \<box>
 | 
|  |    167 |                               nochangeST)
 | 
|  |    168 |                                 (start_ST, start_LT C pTs (length lvars))))"
 | 
|  |    169 | 
 | 
|  |    170 |   compTp :: "java_mb prog => prog_type"
 | 
|  |    171 |   "compTp G C sig == let (D, rT, jmb) = (the (method (G, C) sig))
 | 
|  |    172 |                       in compTpMethod G C (sig, rT, jmb)"
 | 
|  |    173 | 
 | 
|  |    174 | 
 | 
|  |    175 | 
 | 
|  |    176 | (**********************************************************************)
 | 
|  |    177 |   (* Computing the maximum stack size from the method_type *)
 | 
|  |    178 | 
 | 
|  |    179 | constdefs
 | 
|  |    180 |   ssize_sto :: "(state_type option) \<Rightarrow> nat"
 | 
|  |    181 |   "ssize_sto sto ==  case sto of None \<Rightarrow> 0 | (Some (ST, LT)) \<Rightarrow> length ST"
 | 
|  |    182 | 
 | 
|  |    183 |   max_of_list :: "nat list \<Rightarrow> nat"
 | 
|  |    184 |   "max_of_list xs == foldr max xs 0"
 | 
|  |    185 | 
 | 
|  |    186 |   max_ssize :: "method_type \<Rightarrow> nat"
 | 
|  |    187 |   "max_ssize mt == max_of_list (map ssize_sto mt)"
 | 
|  |    188 | 
 | 
|  |    189 | 
 | 
|  |    190 | end
 | 
|  |    191 | 
 |