solutions/cw4/parser.sc
author Christian Urban <christian.urban@kcl.ac.uk>
Tue, 19 Sep 2023 09:54:41 +0100
changeset 920 7af2eea19646
parent 894 02ef5c3abc51
child 959 64ec1884d860
permissions -rw-r--r--
updated

// CW3

import $file.lexer
import lexer._ 

case class ~[+A, +B](x: A, y: B)

// parser combinators

abstract class Parser[I, T](using is: I => Seq[_])  {
  def parse(in: I): Set[(T, I)]  

  def parse_all(in: I) : Set[T] =
    for ((hd, tl) <- parse(in); 
        if is(tl).isEmpty) yield hd
}

// alternative parser
class AltParser[I, T](p: => Parser[I, T], 
                      q: => Parser[I, T])(using I => Seq[_]) extends Parser[I, T] {
  def parse(in: I) = p.parse(in) ++ q.parse(in)   
}

// sequence parser
class SeqParser[I, T, S](p: => Parser[I, T], 
                         q: => Parser[I, S])(using I => Seq[_]) extends Parser[I, ~[T, S]] {
  def parse(in: I) = 
    for ((hd1, tl1) <- p.parse(in); 
         (hd2, tl2) <- q.parse(tl1)) yield (new ~(hd1, hd2), tl2)
}

// map parser
class MapParser[I, T, S](p: => Parser[I, T], 
                         f: T => S)(using I => Seq[_]) extends Parser[I, S] {
  def parse(in: I) = for ((hd, tl) <- p.parse(in)) yield (f(hd), tl)
}

// more convenient syntax for parser combinators
extension [I, T](p: Parser[I, T])(using I => Seq[_]) {
  def ||(q : => Parser[I, T]) = new AltParser[I, T](p, q)
  def ~[S] (q : => Parser[I, S]) = new SeqParser[I, T, S](p, q)
  def map[S](f: => T => S) = new MapParser[I, T, S](p, f)
}

/*
// atomic parser for (particular) strings
case class StrParser(s: String) extends Parser[String, String] {
  def parse(sb: String) = {
    val (prefix, suffix) = sb.splitAt(s.length)
    if (prefix == s) Set((prefix, suffix)) else Set()
  }
}

extension (sc: StringContext) 
  def p(args: Any*) = StrParser(sc.s(args:_*))
*/

case class TokenParser(t: Token) extends Parser[List[Token], Token] {
    def parse(in: List[Token]) = {
      // an example of an atomic parser for characters
      if (!in.isEmpty && in.head == t) Set((t, in.tail)) else Set()
    }
}   

case class TokenListParser(ts: List[Token]) extends Parser[List[Token], List[Token]] {
    def parse(tsb: List[Token]) = {
        val (prefix, suffix) = tsb.splitAt(ts.length)
        if (prefix == ts) Set((prefix, suffix)) else Set()
    }
}

// Implicit definitions to go from a token 
// or a list of tokens to a TokenListParser
implicit def token2parser(t: Token) : Parser[List[Token], Token] = 
  TokenParser(t)

extension (t: Token) {
    def || (q : => Parser[List[Token], Token]) = 
      new AltParser[List[Token], Token](t, q)
    def ~[S](q : => Parser[List[Token], S]) = 
      new SeqParser[List[Token], Token, S](t, q)  
}



// Abstract Syntax Trees
abstract class Stmt
abstract class AExp
abstract class BExp

type Block = List[Stmt]

case object Skip extends Stmt
case class If(a: BExp, bl1: Block, bl2: Block) extends Stmt
case class While(b: BExp, bl: Block) extends Stmt
case class Assign(s: String, a: AExp) extends Stmt
case class Read(s: String) extends Stmt
case class WriteId(s: String) extends Stmt  // for printing values of variables
case class WriteString(s: String) extends Stmt  // for printing words
case class For(counter: String, lower: AExp, upper: AExp, code: Block) extends Stmt
case object Break extends Stmt


case class Var(s: String) extends AExp
case class Num(i: Int) extends AExp
case class Aop(o: String, a1: AExp, a2: AExp) extends AExp

case object True extends BExp
case object False extends BExp
case class Bop(o: String, a1: AExp, a2: AExp) extends BExp
case class And(b1: BExp, b2: BExp) extends BExp
case class Or(b1: BExp, b2: BExp) extends BExp

case class IdParser() extends Parser[List[Token], String] {
    def parse(tsb: List[Token]) = tsb match {
        case T_ID(id) :: rest => Set((id, rest))
        case _ => Set()
    }
}

case class NumParser() extends Parser[List[Token], Int] {
    def parse(tsb: List[Token]) = tsb match {
        case T_NUM(n) :: rest => Set((n, rest))
        case _ => Set()
    }
}

case class StringParser() extends Parser[List[Token], String] {
    def parse(tsb: List[Token]) = tsb match {
        case T_STRING(s) :: rest => Set((s, rest))
        case _ => Set()
    }
}

// WHILE Language Parsing

// WHILE Language Parsing
lazy val AExp: Parser[List[Token], AExp] = 
  (Te ~ T_OP("+") ~ AExp).map{ case x ~ _ ~ z => Aop("+", x, z): AExp } ||
  (Te ~ T_OP("-") ~ AExp).map{ case x ~ _ ~ z => Aop("-", x, z): AExp } || Te
lazy val Te: Parser[List[Token], AExp] = 
  (Fa ~ T_OP("*") ~ Te).map{ case x ~ _ ~ z => Aop("*", x, z): AExp } || 
  (Fa ~ T_OP("/") ~ Te).map{ case x ~ _ ~ z => Aop("/", x, z): AExp } || 
  (Fa ~ T_OP("%") ~ Te).map{ case x ~ _ ~ z => Aop("%", x, z): AExp } || Fa  
lazy val Fa: Parser[List[Token], AExp] = 
   (T_PAREN("(") ~ AExp ~ T_PAREN(")")).map{ case _ ~ y ~ _ => y } || 
   IdParser().map{Var(_)}  || 
   NumParser().map{Num(_)}

lazy val BExp: Parser[List[Token], BExp] = 
   (AExp ~ T_OP("==") ~ AExp).map{ case x ~ _ ~ z => Bop("==", x, z): BExp } || 
   (AExp ~ T_OP("!=") ~ AExp).map{ case x ~ _ ~ z => Bop("!=", x, z): BExp } || 
   (AExp ~ T_OP("<") ~ AExp).map{ case x ~ _ ~ z => Bop("<", x, z): BExp } || 
   (AExp ~ T_OP(">") ~ AExp).map{ case x ~ _ ~ z => Bop(">", x, z): BExp } ||
   (T_PAREN("(") ~ BExp ~ T_PAREN(")") ~ T_OP("&&") ~ BExp).map{ case _ ~ y ~ _ ~ _ ~ v => And(y, v): BExp } ||
   (T_PAREN("(") ~ BExp ~ T_PAREN(")") ~ T_OP("||") ~ BExp).map{ case _ ~ y ~ _ ~ _ ~ v => Or(y, v): BExp } ||
   (T_KEYWORD("true").map(_ => True: BExp )) || 
   (T_KEYWORD("false").map(_ => False: BExp )) ||
   (T_PAREN("(") ~ BExp ~ T_PAREN(")")).map{ case _ ~ x ~ _ => x }

lazy val Stmt: Parser[List[Token], Stmt] =
    T_KEYWORD("skip").map(_ => Skip: Stmt) ||
    T_KEYWORD("break").map(_ => Break: Stmt) ||
    (IdParser() ~ T_OP(":=") ~ AExp).map{ case id ~ _ ~ z => Assign(id, z): Stmt } ||
    (T_KEYWORD("if") ~ BExp ~ T_KEYWORD("then") ~ Block ~ T_KEYWORD("else") ~ Block).map{ case _ ~ y ~ _ ~ u ~ _ ~ w => If(y, u, w): Stmt } ||
    (T_KEYWORD("while") ~ BExp ~ T_KEYWORD("do") ~ Block).map{ case _ ~ y ~ _ ~ w => While(y, w) : Stmt } ||
    (T_KEYWORD("for") ~ IdParser() ~ T_OP(":=") ~ AExp ~T_KEYWORD("upto") ~ AExp ~ T_KEYWORD("do") ~ Block).map{ 
        case _ ~ id ~ _ ~ low ~ _ ~ high ~ _ ~ bl => For(id, low, high, bl) : Stmt } ||
    (T_KEYWORD("read") ~ IdParser()).map{ case _ ~ id => Read(id): Stmt} ||
    (T_KEYWORD("write") ~ IdParser()).map{ case _ ~ id => WriteId(id): Stmt} ||
    (T_KEYWORD("write") ~ StringParser()).map{ case _ ~ s => WriteString(s): Stmt} || 
    (T_KEYWORD("write") ~ T_PAREN("(") ~ IdParser() ~ T_PAREN(")")).map{ case _ ~ _ ~ id ~ _ => WriteId(id): Stmt} ||
    (T_KEYWORD("write") ~  T_PAREN("(") ~ StringParser() ~ T_PAREN(")")).map{ case _ ~ _ ~ s ~ _ => WriteString(s): Stmt}

lazy val Stmts: Parser[List[Token], Block] =
    (Stmt ~ T_SEMI ~ Stmts).map{ case x ~ _ ~ z => x :: z : Block } ||
    (Stmt.map(s => List(s) : Block))

lazy val Block: Parser[List[Token], Block] =
    (T_PAREN("{") ~ Stmts ~ T_PAREN("}")).map{ case x ~ y ~ z => y} ||
    (Stmt.map(s => List(s)))