| author | Christian Urban <christian.urban@kcl.ac.uk> | 
| Wed, 08 Oct 2025 10:42:10 +0100 | |
| changeset 1002 | 4358a7def8cb | 
| parent 976 | 4be299d9b41b | 
| permissions | -rw-r--r-- | 
| 654 | 1 | // A Small LLVM Compiler for a Simple Functional Language | 
| 976 | 2 | // (includes a lexer and parser) | 
| 645 | 3 | // | 
| 813 | 4 | // | 
| 645 | 5 | // call with | 
| 6 | // | |
| 813 | 7 | // amm fun_llvm.sc main fact.fun | 
| 8 | // amm fun_llvm.sc main defs.fun | |
| 9 | // | |
| 10 | // or | |
| 11 | // | |
| 789 | 12 | // amm fun_llvm.sc write fact.fun | 
| 13 | // amm fun_llvm.sc write defs.fun | |
| 645 | 14 | // | 
| 813 | 15 | // this will generate an .ll file. | 
| 16 | // | |
| 17 | // or | |
| 654 | 18 | // | 
| 813 | 19 | // amm fun_llvm.sc run fact.fun | 
| 20 | // amm fun_llvm.sc run defs.fun | |
| 21 | // | |
| 22 | // | |
| 23 | // You can interpret an .ll file using lli, for example | |
| 24 | // | |
| 25 | // lli fact.ll | |
| 655 | 26 | // | 
| 27 | // The optimiser can be invoked as | |
| 28 | // | |
| 29 | // opt -O1 -S in_file.ll > out_file.ll | |
| 30 | // opt -O3 -S in_file.ll > out_file.ll | |
| 31 | // | |
| 813 | 32 | // The code produced for the various architectures can be obtain with | 
| 655 | 33 | // | 
| 34 | // llc -march=x86 -filetype=asm in_file.ll -o - | |
| 35 | // llc -march=arm -filetype=asm in_file.ll -o - | |
| 36 | // | |
| 37 | // Producing an executable can be achieved by | |
| 38 | // | |
| 39 | // llc -filetype=obj in_file.ll | |
| 40 | // gcc in_file.o -o a.out | |
| 41 | // ./a.out | |
| 42 | ||
| 645 | 43 | |
| 976 | 44 | // > using toolkit 0.5.0 | 
| 960 | 45 | // > using file fun_tokens.scala | 
| 46 | // > using file fun_parser.scala | |
| 47 | ||
| 976 | 48 | import $file.fun_tokens, fun_tokens._ | 
| 49 | import $file.fun_parser, fun_parser._ | |
| 626 | 50 | |
| 51 | ||
| 220 
141041fc76b5
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 52 | // for generating new labels | 
| 
141041fc76b5
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 53 | var counter = -1 | 
| 
141041fc76b5
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 54 | |
| 
141041fc76b5
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 55 | def Fresh(x: String) = {
 | 
| 
141041fc76b5
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 56 | counter += 1 | 
| 
141041fc76b5
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 57 | x ++ "_" ++ counter.toString() | 
| 
141041fc76b5
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 58 | } | 
| 
141041fc76b5
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 59 | |
| 678 | 60 | // Internal CPS language for FUN | 
| 648 | 61 | abstract class KExp | 
| 653 | 62 | abstract class KVal | 
| 648 | 63 | |
| 653 | 64 | case class KVar(s: String) extends KVal | 
| 65 | case class KNum(i: Int) extends KVal | |
| 656 | 66 | case class Kop(o: String, v1: KVal, v2: KVal) extends KVal | 
| 653 | 67 | case class KCall(o: String, vrs: List[KVal]) extends KVal | 
| 655 | 68 | case class KWrite(v: KVal) extends KVal | 
| 649 | 69 | |
| 819 | 70 | case class KLet(x: String, e1: KVal, e2: KExp) extends KExp {
 | 
| 71 | override def toString = s"LET $x = $e1 in \n$e2" | |
| 649 | 72 | } | 
| 819 | 73 | case class KIf(x1: String, e1: KExp, e2: KExp) extends KExp {
 | 
| 74 |   def pad(e: KExp) = e.toString.replaceAll("(?m)^", "  ")
 | |
| 75 | ||
| 76 | override def toString = | |
| 77 |      s"IF $x1\nTHEN\n${pad(e1)}\nELSE\n${pad(e2)}"
 | |
| 648 | 78 | } | 
| 653 | 79 | case class KReturn(v: KVal) extends KExp | 
| 648 | 80 | |
| 957 | 81 | // some functions for drawing KVal-trees | 
| 82 | // inspired by William Bradford Larcombe | |
| 83 | ||
| 84 | def draw_vals(vs: List[KVal], prefix: String) : String = {
 | |
| 85 | val vsi = vs.iterator | |
| 86 | vsi.map(v => draw_val(v, prefix, vsi.hasNext)).mkString | |
| 87 | } | |
| 88 | ||
| 89 | def draw_val(k: KVal, prefix: String, more: Boolean) : String = {
 | |
| 90 |   val full_prefix = s"$prefix${if more then "├" else "└"}"
 | |
| 91 |   val childPrefix = s"$prefix${if more then "│" else ""}  "
 | |
| 92 |   s"\n${full_prefix}" ++ 
 | |
| 93 |   (k match {
 | |
| 94 | case KVar(x) => x | |
| 95 | case KNum(n) => n.toString | |
| 96 |     case Kop(op, v1 , v2) => s"KOp($op) ${draw_vals(List(v1, v2), childPrefix)}" 
 | |
| 97 |     case KCall(nme, as) => s"KCall($nme) ${draw_vals(as, childPrefix)}" 
 | |
| 98 |     case KWrite(v) => s"KWrite ${draw_val(v, childPrefix, false)}" 
 | |
| 99 | }) | |
| 100 | } | |
| 101 | ||
| 102 | def draw(k: KVal) = "│" ++ draw_val(k, "", false) | |
| 103 | ||
| 104 | // val k1 = KVar("foo")
 | |
| 105 | // val k2 = KNum(1) | |
| 106 | // val k3 = Kop("-", Kop("+", k1, k2), KNum(2))
 | |
| 107 | // println(draw(k3).mkString) | |
| 108 | // println(draw(KCall("bar", List(k1,k2,k3,k2,k1))).mkString)
 | |
| 109 | ||
| 654 | 110 | |
| 655 | 111 | // CPS translation from Exps to KExps using a | 
| 654 | 112 | // continuation k. | 
| 653 | 113 | def CPS(e: Exp)(k: KVal => KExp) : KExp = e match {
 | 
| 114 | case Var(s) => k(KVar(s)) | |
| 115 | case Num(i) => k(KNum(i)) | |
| 116 |   case Aop(o, e1, e2) => {
 | |
| 117 |     val z = Fresh("tmp")
 | |
| 118 | CPS(e1)(y1 => | |
| 656 | 119 | CPS(e2)(y2 => KLet(z, Kop(o, y1, y2), k(KVar(z))))) | 
| 653 | 120 | } | 
| 121 |   case If(Bop(o, b1, b2), e1, e2) => {
 | |
| 122 |     val z = Fresh("tmp")
 | |
| 123 | CPS(b1)(y1 => | |
| 655 | 124 | CPS(b2)(y2 => | 
| 656 | 125 | KLet(z, Kop(o, y1, y2), KIf(z, CPS(e1)(k), CPS(e2)(k))))) | 
| 653 | 126 | } | 
| 127 |   case Call(name, args) => {
 | |
| 128 |     def aux(args: List[Exp], vs: List[KVal]) : KExp = args match {
 | |
| 129 |       case Nil => {
 | |
| 130 |           val z = Fresh("tmp")
 | |
| 131 | KLet(z, KCall(name, vs), k(KVar(z))) | |
| 132 | } | |
| 133 | case e::es => CPS(e)(y => aux(es, vs ::: List(y))) | |
| 648 | 134 | } | 
| 653 | 135 | aux(args, Nil) | 
| 136 | } | |
| 656 | 137 | case Sequence(e1, e2) => | 
| 679 | 138 | CPS(e1)(_ => CPS(e2)(y2 => k(y2))) | 
| 655 | 139 |   case Write(e) => {
 | 
| 140 |     val z = Fresh("tmp")
 | |
| 141 | CPS(e)(y => KLet(z, KWrite(y), k(KVar(z)))) | |
| 142 | } | |
| 653 | 143 | } | 
| 144 | ||
| 679 | 145 | //initial continuation | 
| 976 | 146 | def CPSi(e: Exp) = CPS(e)(KReturn(_)) | 
| 653 | 147 | |
| 957 | 148 | |
| 149 | ||
| 819 | 150 | //some testcases: | 
| 905 | 151 | // (1 + 2) * 3 | 
| 152 | println(CPSi(Aop("*", Aop("+", Num(1), Num(2)), Num(3))).toString)
 | |
| 153 | ||
| 154 | // 3 * (1 + 2) | |
| 155 | println(CPSi(Aop("*", Num(3), Aop("+", Num(1), Num(2)))).toString)
 | |
| 156 | ||
| 157 | //some testcases: | |
| 158 | ||
| 819 | 159 | // numbers and vars | 
| 160 | println(CPSi(Num(1)).toString) | |
| 161 | println(CPSi(Var("z")).toString)
 | |
| 162 | ||
| 163 | // a * 3 | |
| 653 | 164 | val e1 = Aop("*", Var("a"), Num(3))
 | 
| 819 | 165 | println(CPSi(e1).toString) | 
| 653 | 166 | |
| 819 | 167 | // (a * 3) + 4 | 
| 653 | 168 | val e2 = Aop("+", Aop("*", Var("a"), Num(3)), Num(4))
 | 
| 819 | 169 | println(CPSi(e2).toString) | 
| 653 | 170 | |
| 819 | 171 | // 2 + (a * 3) | 
| 653 | 172 | val e3 = Aop("+", Num(2), Aop("*", Var("a"), Num(3)))
 | 
| 819 | 173 | println(CPSi(e3).toString) | 
| 648 | 174 | |
| 819 | 175 | //(1 - 2) + (a * 3) | 
| 653 | 176 | val e4 = Aop("+", Aop("-", Num(1), Num(2)), Aop("*", Var("a"), Num(3)))
 | 
| 819 | 177 | println(CPSi(e4).toString) | 
| 653 | 178 | |
| 819 | 179 | // 3 + 4 ; 1 * 7 | 
| 180 | val es = Sequence(Aop("+", Num(3), Num(4)),
 | |
| 181 |                   Aop("*", Num(1), Num(7)))
 | |
| 182 | println(CPSi(es).toString) | |
| 183 | ||
| 184 | // if (1 == 1) then 3 else 4 | |
| 653 | 185 | val e5 = If(Bop("==", Num(1), Num(1)), Num(3), Num(4))
 | 
| 819 | 186 | println(CPSi(e5).toString) | 
| 653 | 187 | |
| 819 | 188 | // if (1 == 1) then 3 + 7 else 4 * 2 | 
| 189 | val ei = If(Bop("==", Num(1), Num(1)), 
 | |
| 190 |                 Aop("+", Num(3), Num(7)),
 | |
| 191 |                 Aop("*", Num(4), Num(2)))
 | |
| 192 | println(CPSi(ei).toString) | |
| 193 | ||
| 194 | ||
| 195 | // if (10 != 10) then e5 else 40 | |
| 653 | 196 | val e6 = If(Bop("!=", Num(10), Num(10)), e5, Num(40))
 | 
| 819 | 197 | println(CPSi(e6).toString) | 
| 648 | 198 | |
| 653 | 199 | |
| 819 | 200 | // foo(3) | 
| 201 | val e7 = Call("foo", List(Num(3)))
 | |
| 202 | println(CPSi(e7).toString) | |
| 203 | ||
| 204 | // foo(3 * 1, 4, 5 + 6) | |
| 205 | val e8 = Call("foo", List(Aop("*", Num(3), Num(1)), 
 | |
| 206 | Num(4), | |
| 207 |                           Aop("+", Num(5), Num(6))))
 | |
| 208 | println(CPSi(e8).toString) | |
| 653 | 209 | |
| 819 | 210 | // a * 3 ; b + 6 | 
| 211 | val e9 = Sequence(Aop("*", Var("a"), Num(3)), 
 | |
| 212 |                   Aop("+", Var("b"), Num(6)))
 | |
| 213 | println(CPSi(e9).toString) | |
| 649 | 214 | |
| 819 | 215 | |
| 216 | val e10 = Aop("*", Aop("+", Num(1), Call("foo", List(Var("a"), Num(3)))), Num(4))
 | |
| 217 | println(CPSi(e10).toString) | |
| 218 | ||
| 653 | 219 | |
| 648 | 220 | |
| 221 | ||
| 625 | 222 | // convenient string interpolations | 
| 223 | // for instructions, labels and methods | |
| 224 | import scala.language.implicitConversions | |
| 225 | import scala.language.reflectiveCalls | |
| 226 | ||
| 957 | 227 | extension (sc: StringContext) {
 | 
| 625 | 228 | def i(args: Any*): String = " " ++ sc.s(args:_*) ++ "\n" | 
| 229 | def l(args: Any*): String = sc.s(args:_*) ++ ":\n" | |
| 230 | def m(args: Any*): String = sc.s(args:_*) ++ "\n" | |
| 231 | } | |
| 232 | ||
| 656 | 233 | // mathematical and boolean operations | 
| 653 | 234 | def compile_op(op: String) = op match {
 | 
| 235 | case "+" => "add i32 " | |
| 236 | case "*" => "mul i32 " | |
| 237 | case "-" => "sub i32 " | |
| 656 | 238 | case "/" => "sdiv i32 " | 
| 239 | case "%" => "srem i32 " | |
| 653 | 240 | case "==" => "icmp eq i32 " | 
| 813 | 241 | case "<=" => "icmp sle i32 " // signed less or equal | 
| 242 | case "<" => "icmp slt i32 " // signed less than | |
| 653 | 243 | } | 
| 220 
141041fc76b5
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 244 | |
| 813 | 245 | // compile K values | 
| 653 | 246 | def compile_val(v: KVal) : String = v match {
 | 
| 247 | case KNum(i) => s"$i" | |
| 248 | case KVar(s) => s"%$s" | |
| 656 | 249 | case Kop(op, x1, x2) => | 
| 653 | 250 |     s"${compile_op(op)} ${compile_val(x1)}, ${compile_val(x2)}"
 | 
| 251 | case KCall(x1, args) => | |
| 252 |     s"call i32 @$x1 (${args.map(compile_val).mkString("i32 ", ", i32 ", "")})"
 | |
| 655 | 253 | case KWrite(x1) => | 
| 254 |     s"call i32 @printInt (i32 ${compile_val(x1)})"
 | |
| 653 | 255 | } | 
| 648 | 256 | |
| 649 | 257 | // compile K expressions | 
| 258 | def compile_exp(a: KExp) : String = a match {
 | |
| 653 | 259 | case KReturn(v) => | 
| 260 |     i"ret i32 ${compile_val(v)}"
 | |
| 261 | case KLet(x: String, v: KVal, e: KExp) => | |
| 262 |     i"%$x = ${compile_val(v)}" ++ compile_exp(e)
 | |
| 263 |   case KIf(x, e1, e2) => {
 | |
| 679 | 264 |     val if_br = Fresh("if_branch")
 | 
| 265 |     val else_br = Fresh("else_branch")
 | |
| 649 | 266 | i"br i1 %$x, label %$if_br, label %$else_br" ++ | 
| 267 | l"\n$if_br" ++ | |
| 653 | 268 | compile_exp(e1) ++ | 
| 649 | 269 | l"\n$else_br" ++ | 
| 653 | 270 | compile_exp(e2) | 
| 220 
141041fc76b5
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 271 | } | 
| 653 | 272 | } | 
| 273 | ||
| 655 | 274 | val prelude = """ | 
| 275 | @.str = private constant [4 x i8] c"%d\0A\00" | |
| 276 | ||
| 277 | declare i32 @printf(i8*, ...) | |
| 220 
141041fc76b5
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 278 | |
| 655 | 279 | define i32 @printInt(i32 %x) {
 | 
| 280 | %t0 = getelementptr [4 x i8], [4 x i8]* @.str, i32 0, i32 0 | |
| 679 | 281 | call i32 (i8*, ...) @printf(i8* %t0, i32 %x) | 
| 655 | 282 | ret i32 %x | 
| 283 | } | |
| 284 | ||
| 285 | """ | |
| 653 | 286 | |
| 220 
141041fc76b5
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 287 | |
| 625 | 288 | // compile function for declarations and main | 
| 289 | def compile_decl(d: Decl) : String = d match {
 | |
| 649 | 290 |   case Def(name, args, body) => { 
 | 
| 291 |     m"define i32 @$name (${args.mkString("i32 %", ", i32 %", "")}) {" ++
 | |
| 653 | 292 | compile_exp(CPSi(body)) ++ | 
| 649 | 293 | m"}\n" | 
| 221 
824ffbf66ab4
added fun tail
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
220diff
changeset | 294 | } | 
| 649 | 295 |   case Main(body) => {
 | 
| 296 |     m"define i32 @main() {" ++
 | |
| 789 | 297 | compile_exp(CPS(body)(_ => KReturn(KNum(0)))) ++ | 
| 649 | 298 | m"}\n" | 
| 221 
824ffbf66ab4
added fun tail
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
220diff
changeset | 299 | } | 
| 220 
141041fc76b5
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 300 | } | 
| 
141041fc76b5
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 301 | |
| 813 | 302 | |
| 626 | 303 | // main compiler functions | 
| 813 | 304 | def compile(prog: List[Decl]) : String = | 
| 789 | 305 | prelude ++ (prog.map(compile_decl).mkString) | 
| 306 | ||
| 307 | @main | |
| 813 | 308 | def main(fname: String) = {
 | 
| 789 | 309 | val path = os.pwd / fname | 
| 310 |     val file = fname.stripSuffix("." ++ path.ext)
 | |
| 311 | val tks = tokenise(os.read(path)) | |
| 312 | val ast = parse_tks(tks) | |
| 813 | 313 | println(compile(ast)) | 
| 644 | 314 | } | 
| 315 | ||
| 789 | 316 | @main | 
| 317 | def write(fname: String) = {
 | |
| 318 | val path = os.pwd / fname | |
| 319 |     val file = fname.stripSuffix("." ++ path.ext)
 | |
| 320 | val tks = tokenise(os.read(path)) | |
| 321 | val ast = parse_tks(tks) | |
| 813 | 322 | val code = compile(ast) | 
| 789 | 323 | os.write.over(os.pwd / (file ++ ".ll"), code) | 
| 626 | 324 | } | 
| 325 | ||
| 789 | 326 | @main | 
| 327 | def run(fname: String) = {
 | |
| 328 | val path = os.pwd / fname | |
| 329 |     val file = fname.stripSuffix("." ++ path.ext)
 | |
| 813 | 330 | write(fname) | 
| 789 | 331 |     os.proc("llc", "-filetype=obj", file ++ ".ll").call()
 | 
| 813 | 332 |     os.proc("gcc", file ++ ".o", "-o", file ++ ".bin").call()
 | 
| 333 | os.proc(os.pwd / (file ++ ".bin")).call(stdout = os.Inherit) | |
| 334 | println(s"done.") | |
| 657 | 335 | } | 
| 644 | 336 | |
| 657 | 337 | |
| 338 | ||
| 339 | ||
| 819 | 340 | // CPS functions | 
| 341 | /* | |
| 657 | 342 | |
| 819 | 343 | def fact(n: Int) : Int = | 
| 344 | if (n == 0) 1 else n * fact(n - 1) | |
| 345 | ||
| 346 | fact(6) | |
| 347 | ||
| 348 | def factT(n: Int, acc: Int) : Int = | |
| 349 | if (n == 0) acc else factT(n - 1, acc * n) | |
| 350 | ||
| 351 | factT(6, 1) | |
| 352 | ||
| 353 | def factC(n: Int, ret: Int => Int) : Int = {
 | |
| 354 | if (n == 0) ret(1) | |
| 355 | else factC(n - 1, x => ret(x * n)) | |
| 356 | } | |
| 357 | ||
| 958 
6caee1c0222e
updated and added pascal.while file
 Christian Urban <christian.urban@kcl.ac.uk> parents: 
957diff
changeset | 358 | |
| 819 | 359 | |
| 958 
6caee1c0222e
updated and added pascal.while file
 Christian Urban <christian.urban@kcl.ac.uk> parents: 
957diff
changeset | 360 | |
| 819 | 361 | |
| 362 | fibC(10, x => {println(s"Result: $x") ; 1})
 | |
| 363 | ||
| 364 | ||
| 365 | */ | |
| 960 | 366 | /* | 
| 958 
6caee1c0222e
updated and added pascal.while file
 Christian Urban <christian.urban@kcl.ac.uk> parents: 
957diff
changeset | 367 | factC(6, x => x) | 
| 
6caee1c0222e
updated and added pascal.while file
 Christian Urban <christian.urban@kcl.ac.uk> parents: 
957diff
changeset | 368 | factC(6, x => {println(s"The final Result is $x") ; 0})
 | 
| 
6caee1c0222e
updated and added pascal.while file
 Christian Urban <christian.urban@kcl.ac.uk> parents: 
957diff
changeset | 369 | factC(6, _ + 1) | 
| 
6caee1c0222e
updated and added pascal.while file
 Christian Urban <christian.urban@kcl.ac.uk> parents: 
957diff
changeset | 370 | |
| 
6caee1c0222e
updated and added pascal.while file
 Christian Urban <christian.urban@kcl.ac.uk> parents: 
957diff
changeset | 371 | def fib(n: Int) : Int = {
 | 
| 
6caee1c0222e
updated and added pascal.while file
 Christian Urban <christian.urban@kcl.ac.uk> parents: 
957diff
changeset | 372 | if (n == 0 || n == 1) 1 | 
| 
6caee1c0222e
updated and added pascal.while file
 Christian Urban <christian.urban@kcl.ac.uk> parents: 
957diff
changeset | 373 | else fib(n - 1) + fib(n - 2) | 
| 
6caee1c0222e
updated and added pascal.while file
 Christian Urban <christian.urban@kcl.ac.uk> parents: 
957diff
changeset | 374 | } | 
| 
6caee1c0222e
updated and added pascal.while file
 Christian Urban <christian.urban@kcl.ac.uk> parents: 
957diff
changeset | 375 | |
| 
6caee1c0222e
updated and added pascal.while file
 Christian Urban <christian.urban@kcl.ac.uk> parents: 
957diff
changeset | 376 | |
| 
6caee1c0222e
updated and added pascal.while file
 Christian Urban <christian.urban@kcl.ac.uk> parents: 
957diff
changeset | 377 | def fibC(n: Int, ret: Int => Int) : Int = {
 | 
| 
6caee1c0222e
updated and added pascal.while file
 Christian Urban <christian.urban@kcl.ac.uk> parents: 
957diff
changeset | 378 | if (n == 0 || n == 1) ret(1) | 
| 
6caee1c0222e
updated and added pascal.while file
 Christian Urban <christian.urban@kcl.ac.uk> parents: 
957diff
changeset | 379 | else fibC(n - 1, x => fibC(n - 2, y => ret(x + y))) | 
| 960 | 380 | } | 
| 381 | */ |