| author | wenzelm | 
| Thu, 09 Sep 2010 18:32:21 +0200 | |
| changeset 39239 | 47273e5b1441 | 
| parent 35417 | 47ee18b6ae32 | 
| child 39758 | b8a53e3a0ee2 | 
| permissions | -rw-r--r-- | 
| 13673 | 1 | (* Title: HOL/MicroJava/Comp/TranslCompTp.thy | 
| 2 | Author: Martin Strecker | |
| 3 | *) | |
| 4 | ||
| 15481 | 5 | theory TranslCompTp | 
| 6 | imports Index "../BV/JVMType" | |
| 7 | begin | |
| 8 | ||
| 13673 | 9 | (**********************************************************************) | 
| 10 | ||
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
35102diff
changeset | 11 | definition comb :: "['a \<Rightarrow> 'b list \<times> 'c, 'c \<Rightarrow> 'b list \<times> 'd, 'a] \<Rightarrow> 'b list \<times> 'd" where | 
| 13673 | 12 | "comb == (\<lambda> f1 f2 x0. let (xs1, x1) = f1 x0; | 
| 13 | (xs2, x2) = f2 x1 | |
| 14 | in (xs1 @ xs2, x2))" | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
35102diff
changeset | 15 | |
| 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
35102diff
changeset | 16 | definition comb_nil :: "'a \<Rightarrow> 'b list \<times> 'a" where | 
| 13673 | 17 | "comb_nil a == ([], a)" | 
| 18 | ||
| 35355 
613e133966ea
modernized syntax declarations, and make them actually work with authentic syntax;
 wenzelm parents: 
35102diff
changeset | 19 | notation (xsymbols) | 
| 
613e133966ea
modernized syntax declarations, and make them actually work with authentic syntax;
 wenzelm parents: 
35102diff
changeset | 20 | comb (infixr "\<box>" 55) | 
| 13673 | 21 | |
| 22 | lemma comb_nil_left [simp]: "comb_nil \<box> f = f" | |
| 23 | by (simp add: comb_def comb_nil_def split_beta) | |
| 24 | ||
| 25 | lemma comb_nil_right [simp]: "f \<box> comb_nil = f" | |
| 26 | by (simp add: comb_def comb_nil_def split_beta) | |
| 27 | ||
| 28 | lemma comb_assoc [simp]: "(fa \<box> fb) \<box> fc = fa \<box> (fb \<box> fc)" | |
| 29 | by (simp add: comb_def split_beta) | |
| 30 | ||
| 31 | lemma comb_inv: "(xs', x') = (f1 \<box> f2) x0 \<Longrightarrow> | |
| 32 | \<exists> xs1 x1 xs2 x2. (xs1, x1) = (f1 x0) \<and> (xs2, x2) = f2 x1 \<and> xs'= xs1 @ xs2 \<and> x'=x2" | |
| 33 | apply (case_tac "f1 x0") | |
| 34 | apply (case_tac "f2 x1") | |
| 35 | apply (simp add: comb_def split_beta) | |
| 36 | done | |
| 37 | ||
| 38 | (**********************************************************************) | |
| 39 | ||
| 35102 | 40 | abbreviation (input) | 
| 41 | mt_of :: "method_type \<times> state_type \<Rightarrow> method_type" | |
| 42 | where "mt_of == fst" | |
| 13673 | 43 | |
| 35102 | 44 | abbreviation (input) | 
| 45 | sttp_of :: "method_type \<times> state_type \<Rightarrow> state_type" | |
| 46 | where "sttp_of == snd" | |
| 13673 | 47 | |
| 48 | consts | |
| 49 | compTpExpr :: "java_mb \<Rightarrow> java_mb prog \<Rightarrow> expr | |
| 50 | \<Rightarrow> state_type \<Rightarrow> method_type \<times> state_type" | |
| 51 | ||
| 52 | compTpExprs :: "java_mb \<Rightarrow> java_mb prog \<Rightarrow> expr list | |
| 53 | \<Rightarrow> state_type \<Rightarrow> method_type \<times> state_type" | |
| 54 | ||
| 55 | compTpStmt :: "java_mb \<Rightarrow> java_mb prog \<Rightarrow> stmt | |
| 56 | \<Rightarrow> state_type \<Rightarrow> method_type \<times> state_type" | |
| 57 | ||
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
35102diff
changeset | 58 | definition nochangeST :: "state_type \<Rightarrow> method_type \<times> state_type" where | 
| 13673 | 59 | "nochangeST sttp == ([Some sttp], sttp)" | 
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
35102diff
changeset | 60 | |
| 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
35102diff
changeset | 61 | definition pushST :: "[ty list, state_type] \<Rightarrow> method_type \<times> state_type" where | 
| 13673 | 62 | "pushST tps == (\<lambda> (ST, LT). ([Some (ST, LT)], (tps @ ST, LT)))" | 
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
35102diff
changeset | 63 | |
| 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
35102diff
changeset | 64 | definition dupST :: "state_type \<Rightarrow> method_type \<times> state_type" where | 
| 13673 | 65 | "dupST == (\<lambda> (ST, LT). ([Some (ST, LT)], (hd ST # ST, LT)))" | 
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
35102diff
changeset | 66 | |
| 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
35102diff
changeset | 67 | definition dup_x1ST :: "state_type \<Rightarrow> method_type \<times> state_type" where | 
| 13673 | 68 | "dup_x1ST == (\<lambda> (ST, LT). ([Some (ST, LT)], | 
| 69 | (hd ST # hd (tl ST) # hd ST # (tl (tl ST)), LT)))" | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
35102diff
changeset | 70 | |
| 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
35102diff
changeset | 71 | definition popST :: "[nat, state_type] \<Rightarrow> method_type \<times> state_type" where | 
| 13673 | 72 | "popST n == (\<lambda> (ST, LT). ([Some (ST, LT)], (drop n ST, LT)))" | 
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
35102diff
changeset | 73 | |
| 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
35102diff
changeset | 74 | definition replST :: "[nat, ty, state_type] \<Rightarrow> method_type \<times> state_type" where | 
| 13673 | 75 | "replST n tp == (\<lambda> (ST, LT). ([Some (ST, LT)], (tp # (drop n ST), LT)))" | 
| 76 | ||
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
35102diff
changeset | 77 | definition storeST :: "[nat, ty, state_type] \<Rightarrow> method_type \<times> state_type" where | 
| 13673 | 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 | ||
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
35102diff
changeset | 140 | definition compTpInit :: "java_mb \<Rightarrow> (vname * ty) | 
| 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
35102diff
changeset | 141 | \<Rightarrow> state_type \<Rightarrow> method_type \<times> state_type" where | 
| 13673 | 142 | "compTpInit jmb == (\<lambda> (vn,ty). (pushST [ty]) \<box> (storeST (index jmb vn) ty))" | 
| 143 | ||
| 144 | consts | |
| 145 | compTpInitLvars :: "[java_mb, (vname \<times> ty) list] | |
| 146 | \<Rightarrow> state_type \<Rightarrow> method_type \<times> state_type" | |
| 147 | ||
| 148 | primrec | |
| 149 | "compTpInitLvars jmb [] = comb_nil" | |
| 150 | "compTpInitLvars jmb (lv#lvars) = (compTpInit jmb lv) \<box> (compTpInitLvars jmb lvars)" | |
| 151 | ||
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
35102diff
changeset | 152 | definition start_ST :: "opstack_type" where | 
| 13673 | 153 | "start_ST == []" | 
| 154 | ||
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
35102diff
changeset | 155 | definition start_LT :: "cname \<Rightarrow> ty list \<Rightarrow> nat \<Rightarrow> locvars_type" where | 
| 13673 | 156 | "start_LT C pTs n == (OK (Class C))#((map OK pTs))@(replicate n Err)" | 
| 157 | ||
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
35102diff
changeset | 158 | definition compTpMethod :: "[java_mb prog, cname, java_mb mdecl] \<Rightarrow> method_type" where | 
| 13673 | 159 | "compTpMethod G C == \<lambda> ((mn,pTs),rT, jmb). | 
| 160 | let (pns,lvars,blk,res) = jmb | |
| 161 | in (mt_of | |
| 162 | ((compTpInitLvars jmb lvars \<box> | |
| 163 | compTpStmt jmb G blk \<box> | |
| 164 | compTpExpr jmb G res \<box> | |
| 165 | nochangeST) | |
| 166 | (start_ST, start_LT C pTs (length lvars))))" | |
| 167 | ||
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
35102diff
changeset | 168 | definition compTp :: "java_mb prog => prog_type" where | 
| 13673 | 169 | "compTp G C sig == let (D, rT, jmb) = (the (method (G, C) sig)) | 
| 170 | in compTpMethod G C (sig, rT, jmb)" | |
| 171 | ||
| 172 | ||
| 173 | ||
| 174 | (**********************************************************************) | |
| 175 | (* Computing the maximum stack size from the method_type *) | |
| 176 | ||
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
35102diff
changeset | 177 | definition ssize_sto :: "(state_type option) \<Rightarrow> nat" where | 
| 13673 | 178 | "ssize_sto sto == case sto of None \<Rightarrow> 0 | (Some (ST, LT)) \<Rightarrow> length ST" | 
| 179 | ||
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
35102diff
changeset | 180 | definition max_of_list :: "nat list \<Rightarrow> nat" where | 
| 13673 | 181 | "max_of_list xs == foldr max xs 0" | 
| 182 | ||
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
35102diff
changeset | 183 | definition max_ssize :: "method_type \<Rightarrow> nat" where | 
| 13673 | 184 | "max_ssize mt == max_of_list (map ssize_sto mt)" | 
| 185 | ||
| 186 | end |