| author | Christian Urban <christian.urban@kcl.ac.uk> | 
| Sat, 07 Nov 2020 00:07:28 +0000 | |
| changeset 800 | bdd731e4edbf | 
| parent 789 | 966c9fd84693 | 
| child 813 | 553cd0c5e983 | 
| permissions | -rw-r--r-- | 
| 654 | 1 | // A Small LLVM Compiler for a Simple Functional Language | 
| 644 | 2 | // (includes an external lexer and parser) | 
| 645 | 3 | // | 
| 4 | // call with | |
| 5 | // | |
| 789 | 6 | // amm fun_llvm.sc write fact.fun | 
| 7 | // | |
| 8 | // amm fun_llvm.sc write defs.fun | |
| 645 | 9 | // | 
| 789 | 10 | // this will generate a .ll file. Other options are compile and run. | 
| 654 | 11 | // | 
| 789 | 12 | // You can interpret an .ll file using lli. | 
| 655 | 13 | // | 
| 14 | // The optimiser can be invoked as | |
| 15 | // | |
| 16 | // opt -O1 -S in_file.ll > out_file.ll | |
| 17 | // opt -O3 -S in_file.ll > out_file.ll | |
| 18 | // | |
| 19 | // The code produced for the various architectures can be obtains with | |
| 20 | // | |
| 21 | // llc -march=x86 -filetype=asm in_file.ll -o - | |
| 22 | // llc -march=arm -filetype=asm in_file.ll -o - | |
| 23 | // | |
| 24 | // Producing an executable can be achieved by | |
| 25 | // | |
| 26 | // llc -filetype=obj in_file.ll | |
| 27 | // gcc in_file.o -o a.out | |
| 28 | // ./a.out | |
| 29 | ||
| 645 | 30 | |
| 789 | 31 | import $file.fun_tokens, fun_tokens._ | 
| 32 | import $file.fun_parser, fun_parser._ | |
| 645 | 33 | import scala.util._ | 
| 626 | 34 | |
| 35 | ||
| 220 
141041fc76b5
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 36 | // for generating new labels | 
| 
141041fc76b5
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 37 | var counter = -1 | 
| 
141041fc76b5
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 38 | |
| 
141041fc76b5
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 39 | def Fresh(x: String) = {
 | 
| 
141041fc76b5
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 40 | counter += 1 | 
| 
141041fc76b5
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 41 | x ++ "_" ++ counter.toString() | 
| 
141041fc76b5
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 42 | } | 
| 
141041fc76b5
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 43 | |
| 678 | 44 | // Internal CPS language for FUN | 
| 648 | 45 | abstract class KExp | 
| 653 | 46 | abstract class KVal | 
| 648 | 47 | |
| 653 | 48 | case class KVar(s: String) extends KVal | 
| 49 | case class KNum(i: Int) extends KVal | |
| 656 | 50 | case class Kop(o: String, v1: KVal, v2: KVal) extends KVal | 
| 653 | 51 | case class KCall(o: String, vrs: List[KVal]) extends KVal | 
| 655 | 52 | case class KWrite(v: KVal) extends KVal | 
| 649 | 53 | |
| 653 | 54 | case class KIf(x1: String, e1: KExp, e2: KExp) extends KExp {
 | 
| 55 | override def toString = s"KIf $x1\nIF\n$e1\nELSE\n$e2" | |
| 649 | 56 | } | 
| 653 | 57 | case class KLet(x: String, e1: KVal, e2: KExp) extends KExp {
 | 
| 648 | 58 | override def toString = s"let $x = $e1 in \n$e2" | 
| 59 | } | |
| 653 | 60 | case class KReturn(v: KVal) extends KExp | 
| 648 | 61 | |
| 654 | 62 | |
| 655 | 63 | // CPS translation from Exps to KExps using a | 
| 654 | 64 | // continuation k. | 
| 653 | 65 | def CPS(e: Exp)(k: KVal => KExp) : KExp = e match {
 | 
| 66 | case Var(s) => k(KVar(s)) | |
| 67 | case Num(i) => k(KNum(i)) | |
| 68 |   case Aop(o, e1, e2) => {
 | |
| 69 |     val z = Fresh("tmp")
 | |
| 70 | CPS(e1)(y1 => | |
| 656 | 71 | CPS(e2)(y2 => KLet(z, Kop(o, y1, y2), k(KVar(z))))) | 
| 653 | 72 | } | 
| 73 |   case If(Bop(o, b1, b2), e1, e2) => {
 | |
| 74 |     val z = Fresh("tmp")
 | |
| 75 | CPS(b1)(y1 => | |
| 655 | 76 | CPS(b2)(y2 => | 
| 656 | 77 | KLet(z, Kop(o, y1, y2), KIf(z, CPS(e1)(k), CPS(e2)(k))))) | 
| 653 | 78 | } | 
| 79 |   case Call(name, args) => {
 | |
| 80 |     def aux(args: List[Exp], vs: List[KVal]) : KExp = args match {
 | |
| 81 |       case Nil => {
 | |
| 82 |           val z = Fresh("tmp")
 | |
| 83 | KLet(z, KCall(name, vs), k(KVar(z))) | |
| 84 | } | |
| 85 | case e::es => CPS(e)(y => aux(es, vs ::: List(y))) | |
| 648 | 86 | } | 
| 653 | 87 | aux(args, Nil) | 
| 88 | } | |
| 656 | 89 | case Sequence(e1, e2) => | 
| 679 | 90 | CPS(e1)(_ => CPS(e2)(y2 => k(y2))) | 
| 655 | 91 |   case Write(e) => {
 | 
| 92 |     val z = Fresh("tmp")
 | |
| 93 | CPS(e)(y => KLet(z, KWrite(y), k(KVar(z)))) | |
| 94 | } | |
| 653 | 95 | } | 
| 96 | ||
| 679 | 97 | //initial continuation | 
| 653 | 98 | def CPSi(e: Exp) = CPS(e)(KReturn) | 
| 99 | ||
| 654 | 100 | // some testcases | 
| 653 | 101 | val e1 = Aop("*", Var("a"), Num(3))
 | 
| 654 | 102 | CPSi(e1) | 
| 653 | 103 | |
| 104 | val e2 = Aop("+", Aop("*", Var("a"), Num(3)), Num(4))
 | |
| 654 | 105 | CPSi(e2) | 
| 653 | 106 | |
| 107 | val e3 = Aop("+", Num(2), Aop("*", Var("a"), Num(3)))
 | |
| 654 | 108 | CPSi(e3) | 
| 648 | 109 | |
| 653 | 110 | val e4 = Aop("+", Aop("-", Num(1), Num(2)), Aop("*", Var("a"), Num(3)))
 | 
| 654 | 111 | CPSi(e4) | 
| 653 | 112 | |
| 113 | val e5 = If(Bop("==", Num(1), Num(1)), Num(3), Num(4))
 | |
| 654 | 114 | CPSi(e5) | 
| 653 | 115 | |
| 116 | val e6 = If(Bop("!=", Num(10), Num(10)), e5, Num(40))
 | |
| 654 | 117 | CPSi(e6) | 
| 648 | 118 | |
| 653 | 119 | val e7 = Call("foo", List(Num(3)))
 | 
| 654 | 120 | CPSi(e7) | 
| 653 | 121 | |
| 705 | 122 | val e8 = Call("foo", List(Aop("*", Num(3), Num(1)), Num(4), Aop("+", Num(5), Num(6))))
 | 
| 654 | 123 | CPSi(e8) | 
| 653 | 124 | |
| 125 | val e9 = Sequence(Aop("*", Var("a"), Num(3)), Aop("+", Var("b"), Num(6)))
 | |
| 654 | 126 | CPSi(e9) | 
| 649 | 127 | |
| 128 | val e = Aop("*", Aop("+", Num(1), Call("foo", List(Var("a"), Num(3)))), Num(4))
 | |
| 654 | 129 | CPSi(e) | 
| 653 | 130 | |
| 648 | 131 | |
| 132 | ||
| 133 | ||
| 625 | 134 | // convenient string interpolations | 
| 135 | // for instructions, labels and methods | |
| 136 | import scala.language.implicitConversions | |
| 137 | import scala.language.reflectiveCalls | |
| 138 | ||
| 139 | implicit def sring_inters(sc: StringContext) = new {
 | |
| 140 | def i(args: Any*): String = " " ++ sc.s(args:_*) ++ "\n" | |
| 141 | def l(args: Any*): String = sc.s(args:_*) ++ ":\n" | |
| 142 | def m(args: Any*): String = sc.s(args:_*) ++ "\n" | |
| 143 | } | |
| 144 | ||
| 656 | 145 | // mathematical and boolean operations | 
| 653 | 146 | def compile_op(op: String) = op match {
 | 
| 147 | case "+" => "add i32 " | |
| 148 | case "*" => "mul i32 " | |
| 149 | case "-" => "sub i32 " | |
| 656 | 150 | case "/" => "sdiv i32 " | 
| 151 | case "%" => "srem i32 " | |
| 653 | 152 | case "==" => "icmp eq i32 " | 
| 656 | 153 | case "<=" => "icmp sle i32 " // signed less or equal | 
| 154 | case "<" => "icmp slt i32 " // signed less than | |
| 653 | 155 | } | 
| 220 
141041fc76b5
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 156 | |
| 653 | 157 | def compile_val(v: KVal) : String = v match {
 | 
| 158 | case KNum(i) => s"$i" | |
| 159 | case KVar(s) => s"%$s" | |
| 656 | 160 | case Kop(op, x1, x2) => | 
| 653 | 161 |     s"${compile_op(op)} ${compile_val(x1)}, ${compile_val(x2)}"
 | 
| 162 | case KCall(x1, args) => | |
| 163 |     s"call i32 @$x1 (${args.map(compile_val).mkString("i32 ", ", i32 ", "")})"
 | |
| 655 | 164 | case KWrite(x1) => | 
| 165 |     s"call i32 @printInt (i32 ${compile_val(x1)})"
 | |
| 653 | 166 | } | 
| 648 | 167 | |
| 649 | 168 | // compile K expressions | 
| 169 | def compile_exp(a: KExp) : String = a match {
 | |
| 653 | 170 | case KReturn(v) => | 
| 171 |     i"ret i32 ${compile_val(v)}"
 | |
| 172 | case KLet(x: String, v: KVal, e: KExp) => | |
| 173 |     i"%$x = ${compile_val(v)}" ++ compile_exp(e)
 | |
| 174 |   case KIf(x, e1, e2) => {
 | |
| 679 | 175 |     val if_br = Fresh("if_branch")
 | 
| 176 |     val else_br = Fresh("else_branch")
 | |
| 649 | 177 | i"br i1 %$x, label %$if_br, label %$else_br" ++ | 
| 178 | l"\n$if_br" ++ | |
| 653 | 179 | compile_exp(e1) ++ | 
| 649 | 180 | l"\n$else_br" ++ | 
| 653 | 181 | compile_exp(e2) | 
| 220 
141041fc76b5
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 182 | } | 
| 653 | 183 | } | 
| 184 | ||
| 655 | 185 | |
| 186 | val prelude = """ | |
| 187 | @.str = private constant [4 x i8] c"%d\0A\00" | |
| 188 | ||
| 189 | declare i32 @printf(i8*, ...) | |
| 220 
141041fc76b5
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 190 | |
| 655 | 191 | define i32 @printInt(i32 %x) {
 | 
| 192 | %t0 = getelementptr [4 x i8], [4 x i8]* @.str, i32 0, i32 0 | |
| 679 | 193 | call i32 (i8*, ...) @printf(i8* %t0, i32 %x) | 
| 655 | 194 | ret i32 %x | 
| 195 | } | |
| 196 | ||
| 197 | """ | |
| 653 | 198 | |
| 220 
141041fc76b5
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 199 | |
| 625 | 200 | // compile function for declarations and main | 
| 201 | def compile_decl(d: Decl) : String = d match {
 | |
| 649 | 202 |   case Def(name, args, body) => { 
 | 
| 203 |     m"define i32 @$name (${args.mkString("i32 %", ", i32 %", "")}) {" ++
 | |
| 653 | 204 | compile_exp(CPSi(body)) ++ | 
| 649 | 205 | m"}\n" | 
| 221 
824ffbf66ab4
added fun tail
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
220diff
changeset | 206 | } | 
| 649 | 207 |   case Main(body) => {
 | 
| 208 |     m"define i32 @main() {" ++
 | |
| 789 | 209 | compile_exp(CPS(body)(_ => KReturn(KNum(0)))) ++ | 
| 649 | 210 | m"}\n" | 
| 221 
824ffbf66ab4
added fun tail
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
220diff
changeset | 211 | } | 
| 220 
141041fc76b5
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 212 | } | 
| 
141041fc76b5
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 213 | |
| 626 | 214 | // main compiler functions | 
| 220 
141041fc76b5
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 215 | |
| 789 | 216 | def compile_prog(prog: List[Decl]) : String = | 
| 217 | prelude ++ (prog.map(compile_decl).mkString) | |
| 218 | ||
| 220 
141041fc76b5
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 219 | |
| 789 | 220 | @main | 
| 221 | def compile(fname: String) = {
 | |
| 222 | val path = os.pwd / fname | |
| 223 |     val file = fname.stripSuffix("." ++ path.ext)
 | |
| 224 | val tks = tokenise(os.read(path)) | |
| 225 | val ast = parse_tks(tks) | |
| 226 | println(compile_prog(ast)) | |
| 644 | 227 | } | 
| 228 | ||
| 789 | 229 | @main | 
| 230 | def write(fname: String) = {
 | |
| 231 | val path = os.pwd / fname | |
| 232 |     val file = fname.stripSuffix("." ++ path.ext)
 | |
| 233 | val tks = tokenise(os.read(path)) | |
| 234 | val ast = parse_tks(tks) | |
| 235 | val code = compile_prog(ast) | |
| 236 | os.write.over(os.pwd / (file ++ ".ll"), code) | |
| 626 | 237 | } | 
| 238 | ||
| 789 | 239 | @main | 
| 240 | def run(fname: String) = {
 | |
| 241 | val path = os.pwd / fname | |
| 242 |     val file = fname.stripSuffix("." ++ path.ext)
 | |
| 243 | val tks = tokenise(os.read(path)) | |
| 244 | val ast = parse_tks(tks) | |
| 245 | val code = compile_prog(ast) | |
| 246 | os.write.over(os.pwd / (file ++ ".ll"), code) | |
| 247 |     os.proc("llc", "-filetype=obj", file ++ ".ll").call()
 | |
| 248 |     os.proc("gcc", file ++ ".o", "-o", file).call()
 | |
| 249 | print(os.proc(os.pwd / file).call().out.string) | |
| 657 | 250 | } | 
| 644 | 251 | |
| 657 | 252 | |
| 253 | ||
| 254 | ||
| 255 |