| author | wenzelm | 
| Sat, 14 May 2011 12:40:11 +0200 | |
| changeset 42800 | df2dc9406287 | 
| parent 39758 | b8a53e3a0ee2 | 
| permissions | -rw-r--r-- | 
| 13673 | 1 | (* Title: HOL/MicroJava/Comp/TranslComp.thy | 
| 2 | Author: Martin Strecker | |
| 3 | *) | |
| 4 | ||
| 5 | (* Compiling MicroJava into MicroJVM -- Translation functions *) | |
| 6 | ||
| 16417 | 7 | theory TranslComp imports TranslCompTp begin | 
| 13673 | 8 | |
| 9 | ||
| 10 | (* parameter java_mb only serves to define function index later *) | |
| 11 | ||
| 12 | ||
| 13 | (**********************************************************************) | |
| 14 | (** code generation for expressions **) | |
| 15 | ||
| 39758 | 16 | primrec compExpr :: "java_mb => expr => instr list" | 
| 17 | and compExprs :: "java_mb => expr list => instr list" | |
| 18 | where | |
| 19 | ||
| 13673 | 20 | (*class instance creation*) | 
| 39758 | 21 | "compExpr jmb (NewC c) = [New c]" | | 
| 13673 | 22 | |
| 23 | (*type cast*) | |
| 39758 | 24 | "compExpr jmb (Cast c e) = compExpr jmb e @ [Checkcast c]" | | 
| 13673 | 25 | |
| 26 | ||
| 27 | (*literals*) | |
| 39758 | 28 | "compExpr jmb (Lit val) = [LitPush val]" | | 
| 13673 | 29 | |
| 30 | ||
| 31 | (* binary operation *) | |
| 32 | "compExpr jmb (BinOp bo e1 e2) = compExpr jmb e1 @ compExpr jmb e2 @ | |
| 33 | (case bo of Eq => [Ifcmpeq 3,LitPush(Bool False),Goto 2,LitPush(Bool True)] | |
| 39758 | 34 | | Add => [IAdd])" | | 
| 13673 | 35 | |
| 36 | (*local variable*) | |
| 39758 | 37 | "compExpr jmb (LAcc vn) = [Load (index jmb vn)]" | | 
| 13673 | 38 | |
| 39 | ||
| 40 | (*assignement*) | |
| 39758 | 41 | "compExpr jmb (vn::=e) = compExpr jmb e @ [Dup , Store (index jmb vn)]" | | 
| 13673 | 42 | |
| 43 | ||
| 44 | (*field access*) | |
| 39758 | 45 | "compExpr jmb ( {cn}e..fn ) = compExpr jmb e @ [Getfield fn cn]" |
 | 
| 13673 | 46 | |
| 47 | ||
| 48 | (*field assignement - expected syntax:  {_}_.._:=_ *)
 | |
| 49 | "compExpr jmb (FAss cn e1 fn e2 ) = | |
| 39758 | 50 | compExpr jmb e1 @ compExpr jmb e2 @ [Dup_x1 , Putfield fn cn]" | | 
| 13673 | 51 | |
| 52 | ||
| 53 | (*method call*) | |
| 54 | "compExpr jmb (Call cn e1 mn X ps) = | |
| 39758 | 55 | compExpr jmb e1 @ compExprs jmb ps @ [Invoke cn mn X]" | | 
| 13673 | 56 | |
| 57 | ||
| 58 | (** code generation expression lists **) | |
| 59 | ||
| 39758 | 60 | "compExprs jmb [] = []" | | 
| 13673 | 61 | |
| 62 | "compExprs jmb (e#es) = compExpr jmb e @ compExprs jmb es" | |
| 63 | ||
| 64 | ||
| 39758 | 65 | primrec compStmt :: "java_mb => stmt => instr list" where | 
| 66 | ||
| 13673 | 67 | (** code generation for statements **) | 
| 68 | ||
| 39758 | 69 | "compStmt jmb Skip = []" | | 
| 13673 | 70 | |
| 39758 | 71 | "compStmt jmb (Expr e) = ((compExpr jmb e) @ [Pop])" | | 
| 13673 | 72 | |
| 39758 | 73 | "compStmt jmb (c1;; c2) = ((compStmt jmb c1) @ (compStmt jmb c2))" | | 
| 13673 | 74 | |
| 75 | "compStmt jmb (If(e) c1 Else c2) = | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
16417diff
changeset | 76 | (let cnstf = LitPush (Bool False); | 
| 13673 | 77 | cnd = compExpr jmb e; | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
16417diff
changeset | 78 | thn = compStmt jmb c1; | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
16417diff
changeset | 79 | els = compStmt jmb c2; | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
16417diff
changeset | 80 | test = Ifcmpeq (int(length thn +2)); | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
16417diff
changeset | 81 | thnex = Goto (int(length els +1)) | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
16417diff
changeset | 82 | in | 
| 39758 | 83 | [cnstf] @ cnd @ [test] @ thn @ [thnex] @ els)" | | 
| 13673 | 84 | |
| 85 | "compStmt jmb (While(e) c) = | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
16417diff
changeset | 86 | (let cnstf = LitPush (Bool False); | 
| 13673 | 87 | cnd = compExpr jmb e; | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
16417diff
changeset | 88 | bdy = compStmt jmb c; | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
16417diff
changeset | 89 | test = Ifcmpeq (int(length bdy +2)); | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
16417diff
changeset | 90 | loop = Goto (-(int((length bdy) + (length cnd) +2))) | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
16417diff
changeset | 91 | in | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
16417diff
changeset | 92 | [cnstf] @ cnd @ [test] @ bdy @ [loop])" | 
| 13673 | 93 | |
| 94 | (**********************************************************************) | |
| 95 | ||
| 96 | (*compiling methods, classes and programs*) | |
| 97 | ||
| 98 | (*initialising a single variable*) | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
32960diff
changeset | 99 | definition load_default_val :: "ty => instr" where | 
| 13673 | 100 | "load_default_val ty == LitPush (default_val ty)" | 
| 101 | ||
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
32960diff
changeset | 102 | definition compInit :: "java_mb => (vname * ty) => instr list" where | 
| 13673 | 103 | "compInit jmb == \<lambda> (vn,ty). [load_default_val ty, Store (index jmb vn)]" | 
| 104 | ||
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
32960diff
changeset | 105 | definition compInitLvars :: "[java_mb, (vname \<times> ty) list] \<Rightarrow> bytecode" where | 
| 13673 | 106 | "compInitLvars jmb lvars == concat (map (compInit jmb) lvars)" | 
| 107 | ||
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
32960diff
changeset | 108 | definition compMethod :: "java_mb prog \<Rightarrow> cname \<Rightarrow> java_mb mdecl \<Rightarrow> jvm_method mdecl" where | 
| 13673 | 109 | "compMethod G C jmdl == let (sig, rT, jmb) = jmdl; | 
| 110 | (pns,lvars,blk,res) = jmb; | |
| 111 | mt = (compTpMethod G C jmdl); | |
| 112 | bc = compInitLvars jmb lvars @ | |
| 113 | compStmt jmb blk @ compExpr jmb res @ | |
| 114 | [Return] | |
| 115 | in (sig, rT, max_ssize mt, length lvars, bc, [])" | |
| 116 | ||
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
32960diff
changeset | 117 | definition compClass :: "java_mb prog => java_mb cdecl=> jvm_method cdecl" where | 
| 13673 | 118 | "compClass G == \<lambda> (C,cno,fdls,jmdls). (C,cno,fdls, map (compMethod G C) jmdls)" | 
| 119 | ||
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
32960diff
changeset | 120 | definition comp :: "java_mb prog => jvm_prog" where | 
| 13673 | 121 | "comp G == map (compClass G) G" | 
| 122 | ||
| 123 | end | |
| 124 | ||
| 125 |