progs/lecture3.scala
changeset 506 28f78d9081d7
parent 491 2a30c7dfe3ed
--- a/progs/lecture3.scala	Thu Dec 04 15:04:05 2025 +0000
+++ b/progs/lecture3.scala	Fri Dec 05 10:20:00 2025 +0000
@@ -4,11 +4,14 @@
 // last week:
 // higher-order functions
 // maps
+// recursion
 
-// - recursion
-// - Sudoku
-// - string interpolations
+
 // - Pattern-Matching
+// - Datatypes 
+// - Countdown Game
+// - String interpolations
+
 
 def fact(n: BigInt) : BigInt = {
   if (n == 0) 1 
@@ -21,9 +24,45 @@
   else fib(n - 1) + fib(n - 2)
 }
 
+def list_len(xs: List[Int]) : Int =
+  if (xs == Nil) 0
+  else 1 + list_len(xs.tail)
+
+def list_len(xs: List[Int]) : Int = xs match {
+  case Nil => 0
+  case _::xt => 1 + list_len(xt)
+}
+
+def list_len(xs: List[Int]) : Int =
+  if (xs == Nil) 0
+  else 1 + list_len(xs.tail)
+
+
+def list_map(xs: List[Int], f: Int => Int) : List[Int] =
+  if (xs == Nil) Nil
+  else f(xs.head) :: list_map(xs.tail, f)
 
 
 
+def altproduct(xs: List[Int]) : List[Int] = xs match {
+  case Nil => Nil
+  case (x::y::xs) => x * y :: altproduct(y::xs)
+  case (x::Nil) => List(x)
+}
+
+altproduct(List(1,2,3,4,5))
+
+
+def powerset(xs: Set[Int]) : Set[Set[Int]] = {
+  if (xs == Set()) Set(Set())
+  else {
+    val ps = powerset(xs.tail)
+    ps ++ ps.map(_ + xs.head)
+  }
+}     
+
+powerset(Set(1,2,3)).mkString("\n")
+
 
 def inc(n: Int) : Int = n + 1
 
@@ -37,69 +76,570 @@
 
 
 
+// Pattern Matching
+//==================
+
+// A powerful tool which has even landed in Java during 
+// the last few years (https://inside.java/2021/06/13/podcast-017/).
+// ...Scala already has it for many years and the concept is
+// older than your friendly lecturer, that is stone old  ;o)
+
+// The general schema:
+//
+//    expression match {
+//       case pattern1 => expression1
+//       case pattern2 => expression2
+//       ...
+//       case patternN => expressionN
+//    }
+
+def my_map(f: Int => Int, xs: List[Int]) : List[Int] =  
+  xs match {
+    case Nil => Nil 
+    case hd::tl => f(hd) ::  my_map(f, tl)
+  }
+
+my_map(x => x * x, List(1,2,3,4))
+
+
+// recall
+def len(xs: List[Int]) : Int = {
+    if (xs == Nil) 0
+    else 1 + len(xs.tail)
+}    
+
+def len(xs: List[Int]) : Int = xs match {
+    case Nil => 0
+    case hd::tail => 1 + len(tail)
+}  
+
+
+def my_map_int(lst: List[Int], f: Int => Int) : List[Int] = 
+  lst match {
+    case Nil => Nil
+    case x::xs => f(x)::my_map_int(xs, f)
+  }
+
+def my_map_option(opt: Option[Int], f: Int => Int) : Option[Int] = 
+  opt match {
+    case None => None
+    case Some(x) => Some(f(x))
+  }
+
+my_map_option(None, x => x * x)
+my_map_option(Some(8), x => x * x)
+
+
+// you can also have cases combined
+def season(month: String) : String = month match {
+  case "March" | "April" | "May" => "It's spring"
+  case "June" | "July" | "August" => "It's summer"
+  case "September" | "October" | "November" => "It's autumn"
+  case "December" => "It's winter"
+  case "January" | "February" => "It's unfortunately winter"
+  case _ => "Wrong month"
+}
+
+
+season("November")
+season("abcd")
+// pattern-match on integers
+
+def fib(n: Int) : Int = n match { 
+  case 0 | 1 => 1
+  case n => fib(n - 1) + fib(n - 2)
+}
+
+fib(10)
+
+// pattern-match on results
+
+// Silly: fizz buzz
+def fizz_buzz(n: Int) : String = (n % 3, n % 5) match {
+  case (0, 0) => "fizz buzz"
+  case (0, _) => "fizz"
+  case (_, 0) => "buzz"
+  case _ => n.toString  
+}
+
+for (n <- 1 to 20) 
+ println(fizz_buzz(n))
+
+// more interesting patterns for lists - calculate the deltas between 
+// elements
+
+List(4,3,2,1) -> List(1,1,.. )
+
+def delta(xs: List[Int]) : List[Int] = xs match {
+  case Nil => Nil
+  case x::Nil => x::Nil
+  case x::y::xs => (x - y)::delta(y::xs)
+}
+
+delta(List(10, 7, 8, 2, 5, 10))
+
+
+// guards in pattern-matching
+
+def foo(xs: List[Int]) : String = xs match {
+  case Nil => s"this list is empty"
+  case x :: xs if x % 2 == 0 
+     => s"the first elemnt is even"
+  case x :: y :: rest if x == y
+     => s"this has two elemnts that are the same"
+  case hd :: tl => s"this list is standard $hd::$tl"
+}
+
+foo(Nil)
+foo(List(1,2,3))
+foo(List(1,2))
+foo(List(1,1,2,3))
+foo(List(2,2,2,3))
+
+
+// Trees
+
+enum Tree {
+  case Leaf(x: Int) 
+  case Node(s: String, left: Tree, right: Tree)  
+}
+import Tree._
+
+val lf = Leaf(20)
+val tr = Node("foo", Leaf(10), Leaf(23))
+
+def tree_size(t: Tree) : Int = t match {
+  case Leaf(_) => 1
+  case Node(_, left , right) => 
+    1 + tree_size(left) + tree_size(right)
+}
+
+tree_size(tr)
+
+val lst : List[Tree] = List(lf, tr)
+
+
+abstract class Colour
+case object Red extends Colour 
+case object Green extends Colour 
+case object Blue extends Colour
+case object Yellow extends Colour
+
+
+def fav_colour(c: Colour) : Boolean = c match {
+  case Green => true
+  case _  => false 
+}
+
+fav_colour(Blue)
+
+enum ChessPiece:
+  case Queen, Rook, Bishop, Knight, Pawn
+  def value = this match
+    case Queen  => 9
+    case Rook   => 5
+    case Bishop => 3
+    case Knight => 3
+    case Pawn   => 1
+
+
+
+// ... a tiny bit more useful: Roman Numerals
+
+sealed abstract class RomanDigit 
+case object I extends RomanDigit 
+case object V extends RomanDigit 
+case object X extends RomanDigit 
+case object L extends RomanDigit 
+case object C extends RomanDigit 
+case object D extends RomanDigit 
+case object M extends RomanDigit 
+
+type RomanNumeral = List[RomanDigit] 
+
+List(X,I,M,A)
+
+/*
+I    -> 1
+II   -> 2
+III  -> 3
+IV   -> 4
+V    -> 5
+VI   -> 6
+VII  -> 7
+VIII -> 8
+IX   -> 9
+X    -> 10
+*/
+
+def RomanNumeral2Int(rs: RomanNumeral): Int = rs match { 
+  case Nil => 0
+  case M::r    => 1000 + RomanNumeral2Int(r)  
+  case C::M::r => 900 + RomanNumeral2Int(r)
+  case D::r    => 500 + RomanNumeral2Int(r)
+  case C::D::r => 400 + RomanNumeral2Int(r)
+  case C::r    => 100 + RomanNumeral2Int(r)
+  case X::C::r => 90 + RomanNumeral2Int(r)
+  case L::r    => 50 + RomanNumeral2Int(r)
+  case X::L::r => 40 + RomanNumeral2Int(r)
+  case X::r    => 10 + RomanNumeral2Int(r)
+  case I::X::r => 9 + RomanNumeral2Int(r)
+  case V::r    => 5 + RomanNumeral2Int(r)
+  case I::V::r => 4 + RomanNumeral2Int(r)
+  case I::r    => 1 + RomanNumeral2Int(r)
+}
+
+RomanNumeral2Int(List(I,V))             // 4
+RomanNumeral2Int(List(I,I,I,I))         // 4 (invalid Roman number)
+RomanNumeral2Int(List(V,I))             // 6
+RomanNumeral2Int(List(I,X))             // 9
+RomanNumeral2Int(List(M,C,M,L,X,X,I,X)) // 1979
+RomanNumeral2Int(List(M,M,X,V,I,I))     // 2017
+
+
+abstract class Rexp
+case object ZERO extends Rexp                      // matches nothing
+case object ONE extends Rexp                       // matches the empty string
+case class CHAR(c: Char) extends Rexp              // matches a character c
+case class ALT(r1: Rexp, r2: Rexp) extends Rexp    // alternative
+case class SEQ(r1: Rexp, r2: Rexp) extends Rexp    // sequence
+case class STAR(r: Rexp) extends Rexp              // star
+
+def depth(r: Rexp) : Int = r match {
+  case ZERO => 1
+  case ONE => 1
+  case CHAR(_) => 1
+  case ALT(r1, r2) => 1 + List(depth(r1), depth(r2)).max
+  case SEQ(r1, r2) => 1 + List(depth(r1), depth(r2)).max
+  case STAR(r1) => 1 + depth(r1)
+}
+
+
 
 
 
-// A Recursive Web Crawler / Email Harvester
-//===========================================
-//
-// the idea is to look for links using the
-// regular expression "https?://[^"]*" and for
-// email addresses using another regex.
+// expressions (essentially trees)
+
+abstract class Exp
+case class N(n: Int) extends Exp                  // for numbers
+case class Plus(e1: Exp, e2: Exp) extends Exp
+case class Times(e1: Exp, e2: Exp) extends Exp
+
+def string(e: Exp) : String = e match {
+  case N(n) => s"$n"
+  case Plus(e1, e2) => s"(${string(e1)} + ${string(e2)})" 
+  case Times(e1, e2) => s"(${string(e1)} * ${string(e2)})"
+}
+
+val e = Plus(N(9), Times(N(3), N(4)))
+e.toString
+println(string(e))
+
+def eval(e: Exp) : Int = e match {
+  case N(n) => n
+  case Plus(e1, e2) => eval(e1) + eval(e2) 
+  case Times(e1, e2) => eval(e1) * eval(e2) 
+}
 
-import io.Source
-import scala.util._
+println(eval(e))
+
+// simplification rules:
+// e + 0, 0 + e => e 
+// e * 0, 0 * e => 0
+// e * 1, 1 * e => e
+//
+// (....9 ....)
 
-// gets the first 10K of a web-page
-def get_page(url: String) : String = {
-  Try(Source.fromURL(url)("ISO-8859-1").take(10000).mkString).
-    getOrElse { println(s"  Problem with: $url"); ""}
+def simp(e: Exp) : Exp = e match {
+  case N(n) => N(n)
+  case Plus(e1, e2) => (simp(e1), simp(e2)) match {
+    case (N(0), e2s) => e2s
+    case (e1s, N(0)) => e1s
+    case (e1s, e2s) => Plus(e1s, e2s)
+  }  
+  case Times(e1, e2) => (simp(e1), simp(e2)) match {
+    case (N(0), _) => N(0)
+    case (_, N(0)) => N(0)
+    case (N(1), e2s) => e2s
+    case (e1s, N(1)) => e1s
+    case (e1s, e2s) => Times(e1s, e2s)
+  }  
 }
 
-// regex for URLs and emails
-val http_pattern = """"https?://[^"]*"""".r
-val email_pattern = """([a-z0-9_\.-]+)@([\da-z\.-]+)\.([a-z\.]{2,6})""".r
+
+val e2 = Times(Plus(N(0), N(1)), Plus(N(0), N(9)))
+println(string(e2))
+println(string(simp(e2)))
+
+
+
+// String interpolations as patterns
+
+val date = "2019-11-26"
+val s"$year-$month-$day" = date
+
+def parse_date(date: String) : Option[(Int, Int, Int)]= date match {
+  case s"$year-$month-$day" => Some((day.toInt, month.toInt, year.toInt))
+  case s"$day/$month/$year" => Some((day.toInt, month.toInt, year.toInt))
+  case s"$day.$month.$year" => Some((day.toInt, month.toInt, year.toInt))
+  case _ => None
+} 
+
+parse_date("2019-11-26")
+parse_date("26/11/2019")
+parse_date("26.11.2019")
+
 
-//test case:
-//email_pattern.findAllIn
-//  ("foo bla christian@kcl.ac.uk 1234567").toList
+// Countdown Game using Powerset
+//===============================
+
+
+def powerset(xs: Set[Int]) : Set[Set[Int]] = {
+  if (xs == Set()) Set(Set())
+  else {
+    val ps = powerset(xs.tail)  
+    ps ++ ps.map(_ + xs.head)
+  }
+}  
+
+powerset(Set(1,2,3)).mkString("\n")
+
+// proper subsets
+def psubsets(xs: Set[Int]) = 
+  powerset(xs) -- Set(Set(), xs) 
+
+psubsets(Set(1,2,3)).mkString("\n")
+
+def splits(xs: Set[Int]) : Set[(Set[Int], Set[Int])] =
+  psubsets(xs).map(s => (s, xs -- s))
+
+splits(Set(1,2,3,4)).mkString("\n")
 
 
-// drops the first and last character from a string
-def unquote(s: String) = s.drop(1).dropRight(1)
+enum Tree {
+  case Num(i: Int)
+  case Add(l: Tree, r: Tree)
+  case Mul(l: Tree, r: Tree)
+  case Sub(l: Tree, r: Tree)
+  case Div(l: Tree, r: Tree)
+}
+import Tree._
 
-def get_all_URLs(page: String): Set[String] = 
-  http_pattern.findAllIn(page).map(unquote).toSet
+//pretty printing
+def pp(tr: Tree) : String = tr match {
+  case Num(n) => s"$n"
+  case Add(l, r) => s"(${pp(l)} + ${pp(r)})"
+  case Mul(l, r) => s"(${pp(l)} * ${pp(r)})"
+  case Sub(l, r) => s"(${pp(l)} - ${pp(r)})"
+  case Div(l, r) => s"(${pp(l)} / ${pp(r)})"
+}
+
 
-// naive version of crawl - searches until a given depth,
-// visits pages potentially more than once
-def crawl(url: String, n: Int) : Unit = {
-  if (n == 0) ()
-  else {
-    println(s"  Visiting: $n $url")
-    for (u <- get_all_URLs(get_page(url))) crawl(u, n - 1)
+def eval(tr: Tree) : Option[Int] = tr match {
+  case Num(n) => Some(n)
+  case Add(l, r) => 
+    for (ln <- eval(l); rn <- eval(r)) yield ln + rn
+  case Mul(l, r) => 
+    for (ln <- eval(l); rn <- eval(r)) yield ln * rn 
+  case Sub(l, r) => 
+    for (ln <- eval(l); rn <- eval(r)) yield ln - rn
+  case Div(l, r) => 
+    for (ln <- eval(l); rn <- eval(r); if rn != 0) 
+      yield ln / rn 
+}
+
+
+
+def search(nums: Set[Int]) : Set[Tree] = nums.size match {
+  case 0 => Set()
+  case 1 => Set(Num(nums.head))
+  case 2 => {
+    val ln = Num(nums.head)
+    val rn = Num(nums.tail.head)
+    Set(Add(ln, rn), Mul(ln, rn),
+        Sub(ln, rn), Sub(rn, ln),
+        Div(ln, rn), Div(rn, ln))
+  }
+  case n => {
+    val spls = splits(nums)
+    val res = 
+      for ((ls, rs) <- spls) yield {
+      for (lt <- search(ls);
+           rt <- search(rs)) yield {
+            Set(Add(lt, rt), Mul(lt, rt),
+                Sub(lt, rt), Sub(rt, lt),
+                Div(lt, rt), Div(rt, lt))
+           }
+    }
+    res.flatten.flatten
   }
 }
 
-// some starting URLs for the crawler
-val startURL = """https://nms.kcl.ac.uk/christian.urban/"""
+Set(Set(1,2), Set(2,3), Set(4,5)).flatten
+
+search(Set(1))
+search(Set(1, 2)).mkString("\n")
+search(Set(1, 2,3)).map(pp).mkString("\n")
+search(Set(1, 2,3)).map(pl).mkString("\n")
+search(Set(1, 2,3)).map(tr => s"${pp(tr)} = ${eval(tr)}").mkString("\n")
+
 
-crawl(startURL, 2)
+def search(nums: Set[Int]) : Set[Tree] = nums.size match {
+  case 0 => Set()
+  case 1 => Set(Num(nums.head))
+  case xs => {
+    val spls = splits(nums)
+    val subtrs = 
+      for ((lspls, rspls) <- spls;
+            lt <- search(lspls); 
+            rt <- search(rspls)) yield {
+        Set(Add(lt, rt), Mul(lt, rt))
+    } 
+    subtrs.flatten
+  }
+}
+
+println(search(Set(1,2,3,4)).mkString("\n"))
+
+def pp(tr: Tree) : String = tr match {
+  case Num(n) => s"$n"
+  case Add(l, r) => s"(${pp(l)} + ${pp(r)})"
+  case Mul(l, r) => s"(${pp(l)} * ${pp(r)})"
+}
+
 
 
-// a primitive email harvester
-def emails(url: String, n: Int) : Set[String] = {
-  if (n == 0) Set()
-  else {
-    println(s"  Visiting: $n $url")
-    val page = get_page(url)
-    val new_emails = email_pattern.findAllIn(page).toSet
-    new_emails ++ (for (u <- get_all_URLs(page)) yield emails(u, n - 1)).flatten
+def eval(tr: Tree) : Int = tr match {
+  case Num(n) => n
+  case Add(l, r) => eval(l) + eval(r) 
+  case Mul(l, r) => eval(l) * eval(r)
+}
+
+
+
+
+
+def search(nums: Set[Int]) : Set[Tree] =  nums.size match {
+  case 0 => Set()
+  case 1 => Set(Num(nums.head))
+  case 2 => {
+    val l = nums.head
+    val r = nums.tail.head
+    Set(Add(Num(l), Num(r)), 
+        Mul(Num(l), Num(r)))
+        ++ Option.when(l <= r)(Sub(Num(r), Num(l)))
+        ++ Option.when(l > r)(Sub(Num(l), Num(r)))
+        ++ Option.when(r > 0 && l % r == 0)(Div(Num(l), Num(r)))
+        ++ Option.when(l > 0 && r % l == 0)(Div(Num(r), Num(l)))
+  }
+  case xs => {
+    val spls = splits(nums)
+    val subtrs = 
+      for ((lspls, rspls) <- spls;
+           lt <- search(lspls); 
+          rt <- search(rspls)) yield {
+        Set(Add(lt, rt), Sub(lt, rt),
+            Mul(lt, rt), Div(lt, rt))
+    } 
+    subtrs.flatten
   }
 }
 
-emails(startURL, 2)
+println(search(Set(1,2,3,4)).mkString("\n"))
+
+def eval(tr: Tree) : Option[Int] = tr match {
+  case Num(n) => Some(n)
+  case Add(l, r) => 
+    for (rl <- eval(l); rr <- eval(r)) yield rl + rr
+  case Mul(l, r) => 
+    for (rl <- eval(l); rr <- eval(r)) yield rl * rr  
+  case Sub(l, r) => 
+    for (rl <- eval(l); rr <- eval(r);
+         if 0 <= rl - rr) yield rl - rr   
+  case Div(l, r) => 
+    for (rl <- eval(l); rr <- eval(r);
+         if rr > 0 && rl % rr == 0) yield rl / rr          
+}
+
+eval(Add(Num(1), Num(2)))
+eval(Mul(Add(Num(1), Num(2)), Num(4)))
+eval(Sub(Num(3), Num(2)))
+eval(Sub(Num(3), Num(6)))
+eval(Div(Num(6), Num(2)))
+eval(Div(Num(6), Num(4)))
+
+def time_needed[T](n: Int, code: => T) = {
+  val start = System.nanoTime()
+  for (i <- (0 to n)) code
+  val end = System.nanoTime()
+  (end - start) / 1.0e9
+}
+
+def check(xs: Set[Int], target: Int) =
+  search(xs).find(eval(_) == Some(target))
+
+for (sol <- check(Set(50, 5, 4, 9, 10, 8), 560)) {
+  println(s"${pp(sol)} => ${eval(sol)}")
+}
+
+
+
+time_needed(1, check(Set(50, 5, 4, 9, 10, 8), 560))
+
+
+println(check(Set(25, 5, 2, 10, 7, 1), 986).mkString("\n"))
+
+for (sol <- check(Set(25, 5, 2, 10, 7, 1), 986)) {
+  println(s"${pp(sol)} => ${eval(sol)}")
+}
 
+for (sol <- check(Set(25, 5, 2, 10, 7, 1), -1)) {
+  println(s"${pp(sol)} => ${eval(sol)}")
+}
+
+for (sol <- check(Set(100, 25, 75, 50, 7, 10), 360)) {
+  println(s"${pp(sol)} => ${eval(sol)}")
+}
+time_needed(1, check(Set(100, 25, 75, 50, 7, 10), 360))
+
+
+
+time_needed(1, check(Set(25, 5, 2, 10, 7, 1), 986))
+time_needed(1, check(Set(25, 5, 2, 10, 7, 1), -1))
+
+
+def generate(nums: Set[Int]) : Set[(Tree, Int)] =  nums.size match {
+  case 0 => Set()
+  case 1 => Set((Num(nums.head), nums.head))
+  case xs => {
+    val spls = splits(nums)
+    val subtrs =
+      for ((lspls, rspls) <- spls;
+           (lt, ln) <- generate(lspls); 
+           (rt, rn) <- generate(rspls)) yield {
+        Set((Add(lt, rt), ln + rn),
+            (Mul(lt, rt), ln * rn))
+        ++ Option.when(ln <= rn)((Sub(rt, lt), rn - ln)) 
+        ++ Option.when(ln > rn)((Sub(lt, rt), ln - rn))
+        ++ Option.when(rn > 0 && ln % rn == 0)((Div(lt, rt), ln / rn))
+        ++ Option.when(ln > 0 && rn % ln == 0)((Div(rt, lt), rn / ln))
+    } 
+    subtrs.flatten
+  }
+}
+
+def check2(xs: Set[Int], target: Int) =
+  generate(xs).find(_._2 == target)
+
+for ((sol, ev) <- check2(Set(50, 5, 4, 9, 10, 8), 560)) {
+  println(s"${pp(sol)} => ${eval(sol)} / $ev")
+}
+
+time_needed(1, check(Set(50, 5, 4, 9, 10, 8), 560))
+time_needed(1, check2(Set(50, 5, 4, 9, 10, 8), 560))
+
+time_needed(1, check(Set(50, 5, 4, 9, 10, 8), -1))
+time_needed(1, check2(Set(50, 5, 4, 9, 10, 8), -1))
 
 
 // Sudoku 
@@ -305,312 +845,6 @@
 
 
 
-// Pattern Matching
-//==================
-
-// A powerful tool which has even landed in Java during 
-// the last few years (https://inside.java/2021/06/13/podcast-017/).
-// ...Scala already has it for many years and the concept is
-// older than your friendly lecturer, that is stone old  ;o)
-
-// The general schema:
-//
-//    expression match {
-//       case pattern1 => expression1
-//       case pattern2 => expression2
-//       ...
-//       case patternN => expressionN
-//    }
-
-def my_map(f: Int => Int, xs: List[Int]) : List[Int] =  
-  xs match {
-    case Nil => Nil 
-    case hd::tl => f(hd) ::  my_map(f, tl)
-  }
-
-my_map(x => x * x, List(1,2,3,4))
-
-
-// recall
-def len(xs: List[Int]) : Int = {
-    if (xs == Nil) 0
-    else 1 + len(xs.tail)
-}    
-
-def len(xs: List[Int]) : Int = xs match {
-    case Nil => 0
-    case hd::tail => 1 + len(tail)
-}  
-
-
-def my_map_int(lst: List[Int], f: Int => Int) : List[Int] = 
-  lst match {
-    case Nil => Nil
-    case x::xs => f(x)::my_map_int(xs, f)
-  }
-
-def my_map_option(opt: Option[Int], f: Int => Int) : Option[Int] = 
-  opt match {
-    case None => None
-    case Some(x) => Some(f(x))
-  }
-
-my_map_option(None, x => x * x)
-my_map_option(Some(8), x => x * x)
-
-
-// you can also have cases combined
-def season(month: String) : String = month match {
-  case "March" | "April" | "May" => "It's spring"
-  case "June" | "July" | "August" => "It's summer"
-  case "September" | "October" | "November" => "It's autumn"
-  case "December" => "It's winter"
-  case "January" | "February" => "It's unfortunately winter"
-  case _ => "Wrong month"
-}
-
-// pattern-match on integers
-
-def fib(n: Int) : Int = n match { 
-  case 0 | 1 => 1
-  case n => fib(n - 1) + fib(n - 2)
-}
-
-fib(10)
-
-// pattern-match on results
-
-// Silly: fizz buzz
-def fizz_buzz(n: Int) : String = (n % 3, n % 5) match {
-  case (0, 0) => "fizz buzz"
-  case (0, _) => "fizz"
-  case (_, 0) => "buzz"
-  case _ => n.toString  
-}
-
-for (n <- 1 to 20) 
- println(fizz_buzz(n))
-
-// more interesting patterns for lists - calculate the deltas between 
-// elements
-
-List(4,3,2,1) -> List(1,1,.. )
-
-def delta(xs: List[Int]) : List[Int] = xs match {
-  case Nil => Nil
-  case x::Nil => x::Nil
-  case x::y::xs => (x - y)::delta(y::xs)
-}
-
-delta(List(10, 7, 8, 2, 5, 10))
-
-
-// guards in pattern-matching
-
-def foo(xs: List[Int]) : String = xs match {
-  case Nil => s"this list is empty"
-  case x :: xs if x % 2 == 0 
-     => s"the first elemnt is even"
-  case x :: y :: rest if x == y
-     => s"this has two elemnts that are the same"
-  case hd :: tl => s"this list is standard $hd::$tl"
-}
-
-foo(Nil)
-foo(List(1,2,3))
-foo(List(1,2))
-foo(List(1,1,2,3))
-foo(List(2,2,2,3))
-
-
-// Trees
-
-abstract class Tree
-case class Leaf(x: Int) extends Tree
-case class Node(s: String, left: Tree, right: Tree) extends Tree 
-
-val lf = Leaf(20)
-val tr = Node("foo", Leaf(10), Leaf(23))
-
-def sizet(t: Tree) : Int = t match {
-  case Leaf(_) => 1
-  case Node(_, left , right) => 1 + sizet(left) + sizet(right)
-}
-
-sizet(tr)
-
-val lst : List[Tree] = List(lf, tr)
-
-
-abstract class Colour
-case object Red extends Colour 
-case object Green extends Colour 
-case object Blue extends Colour
-case object Yellow extends Colour
-
-
-def fav_colour(c: Colour) : Boolean = c match {
-  case Green => true
-  case _  => false 
-}
-
-fav_colour(Blue)
-
-enum ChessPiece:
-  case Queen, Rook, Bishop, Knight, Pawn
-  def value = this match
-    case Queen  => 9
-    case Rook   => 5
-    case Bishop => 3
-    case Knight => 3
-    case Pawn   => 1
-
-
-
-// ... a tiny bit more useful: Roman Numerals
-
-sealed abstract class RomanDigit 
-case object I extends RomanDigit 
-case object V extends RomanDigit 
-case object X extends RomanDigit 
-case object L extends RomanDigit 
-case object C extends RomanDigit 
-case object D extends RomanDigit 
-case object M extends RomanDigit 
-
-type RomanNumeral = List[RomanDigit] 
-
-List(X,I,M,A)
-
-/*
-I    -> 1
-II   -> 2
-III  -> 3
-IV   -> 4
-V    -> 5
-VI   -> 6
-VII  -> 7
-VIII -> 8
-IX   -> 9
-X    -> 10
-*/
-
-def RomanNumeral2Int(rs: RomanNumeral): Int = rs match { 
-  case Nil => 0
-  case M::r    => 1000 + RomanNumeral2Int(r)  
-  case C::M::r => 900 + RomanNumeral2Int(r)
-  case D::r    => 500 + RomanNumeral2Int(r)
-  case C::D::r => 400 + RomanNumeral2Int(r)
-  case C::r    => 100 + RomanNumeral2Int(r)
-  case X::C::r => 90 + RomanNumeral2Int(r)
-  case L::r    => 50 + RomanNumeral2Int(r)
-  case X::L::r => 40 + RomanNumeral2Int(r)
-  case X::r    => 10 + RomanNumeral2Int(r)
-  case I::X::r => 9 + RomanNumeral2Int(r)
-  case V::r    => 5 + RomanNumeral2Int(r)
-  case I::V::r => 4 + RomanNumeral2Int(r)
-  case I::r    => 1 + RomanNumeral2Int(r)
-}
-
-RomanNumeral2Int(List(I,V))             // 4
-RomanNumeral2Int(List(I,I,I,I))         // 4 (invalid Roman number)
-RomanNumeral2Int(List(V,I))             // 6
-RomanNumeral2Int(List(I,X))             // 9
-RomanNumeral2Int(List(M,C,M,L,X,X,I,X)) // 1979
-RomanNumeral2Int(List(M,M,X,V,I,I))     // 2017
-
-
-abstract class Rexp
-case object ZERO extends Rexp                      // matches nothing
-case object ONE extends Rexp                       // matches the empty string
-case class CHAR(c: Char) extends Rexp              // matches a character c
-case class ALT(r1: Rexp, r2: Rexp) extends Rexp    // alternative
-case class SEQ(r1: Rexp, r2: Rexp) extends Rexp    // sequence
-case class STAR(r: Rexp) extends Rexp              // star
-
-def depth(r: Rexp) : Int = r match {
-  case ZERO => 1
-  case ONE => 1
-  case CHAR(_) => 1
-  case ALT(r1, r2) => 1 + List(depth(r1), depth(r2)).max
-  case SEQ(r1, r2) => 1 + List(depth(r1), depth(r2)).max
-  case STAR(r1) => 1 + depth(r1)
-}
-
-
-
-
-
-// expressions (essentially trees)
-
-abstract class Exp
-case class N(n: Int) extends Exp                  // for numbers
-case class Plus(e1: Exp, e2: Exp) extends Exp
-case class Times(e1: Exp, e2: Exp) extends Exp
-
-def string(e: Exp) : String = e match {
-  case N(n) => s"$n"
-  case Plus(e1, e2) => s"(${string(e1)} + ${string(e2)})" 
-  case Times(e1, e2) => s"(${string(e1)} * ${string(e2)})"
-}
-
-val e = Plus(N(9), Times(N(3), N(4)))
-e.toString
-println(string(e))
-
-def eval(e: Exp) : Int = e match {
-  case N(n) => n
-  case Plus(e1, e2) => eval(e1) + eval(e2) 
-  case Times(e1, e2) => eval(e1) * eval(e2) 
-}
-
-println(eval(e))
-
-// simplification rules:
-// e + 0, 0 + e => e 
-// e * 0, 0 * e => 0
-// e * 1, 1 * e => e
-//
-// (....9 ....)
-
-def simp(e: Exp) : Exp = e match {
-  case N(n) => N(n)
-  case Plus(e1, e2) => (simp(e1), simp(e2)) match {
-    case (N(0), e2s) => e2s
-    case (e1s, N(0)) => e1s
-    case (e1s, e2s) => Plus(e1s, e2s)
-  }  
-  case Times(e1, e2) => (simp(e1), simp(e2)) match {
-    case (N(0), _) => N(0)
-    case (_, N(0)) => N(0)
-    case (N(1), e2s) => e2s
-    case (e1s, N(1)) => e1s
-    case (e1s, e2s) => Times(e1s, e2s)
-  }  
-}
-
-
-val e2 = Times(Plus(N(0), N(1)), Plus(N(0), N(9)))
-println(string(e2))
-println(string(simp(e2)))
-
-
-
-// String interpolations as patterns
-
-val date = "2019-11-26"
-val s"$year-$month-$day" = date
-
-def parse_date(date: String) : Option[(Int, Int, Int)]= date match {
-  case s"$year-$month-$day" => Some((day.toInt, month.toInt, year.toInt))
-  case s"$day/$month/$year" => Some((day.toInt, month.toInt, year.toInt))
-  case s"$day.$month.$year" => Some((day.toInt, month.toInt, year.toInt))
-  case _ => None
-} 
-
-parse_date("2019-11-26")
-parse_date("26/11/2019")
-parse_date("26.11.2019")