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