Nominal/activities/tphols09/IDW/SB-Calc.thy
author Christian Urban <christian.urban@kcl.ac.uk>
Tue, 07 Jan 2025 12:42:42 +0000
changeset 653 2807ec31d144
parent 415 f1be8028a4a9
permissions -rw-r--r--
updated

theory Calc
imports Main
begin


section {* A Simple Calculator *}

text {*

Task: Write a calculator which behaves as follows:

  calc "ADD 1 2"   = 3

  calc "SUM 1 2 3"   = 6

  cal "LET x 3 IN ADD x 1"   = 4

*}






subsection {* A Parser Primer *}


(*  'a -> 'b  *)

(*  'a -> 'b * 'c  *)



(*  string list -> 'a * string list  *)



(*  FAIL  *)

(*  MORE  *)

(*  ABORT *)





subsection {* Lexer: From a String to a Token List *}



ML {*
datatype token =
  Keyword of string |
  Number of int |
  Variable of string |
  Space |
  End
*}


ML {* Scan.many1 *}
ML {*
val keyword = Scan.many1 Symbol.is_ascii_upper
*}

ML {* keyword (explode "AND 1 2") *}

ML {* try keyword (explode "AND") *}

ML {* try keyword (explode " AND") *}



ML {*
val keyword = Scan.many1 Symbol.is_ascii_upper >>
  (Keyword o implode)
*}

ML {* op >> *}


ML {* keyword (explode "AND 1 2") *}



ML {*
val variable = Scan.many1 Symbol.is_ascii_lower >>
  (Variable o implode)

val space = Scan.many1 Symbol.is_ascii_blank >> K Space
*}



ML {*
fun int_of d = ord d - ord "0"

fun join_digits ds =
  fold (fn d => fn i => 10 * i + int_of d) ds 0

val number = Scan.many1 Symbol.is_ascii_digit >>
  (Number o join_digits)
*}



ML {*
val tokens = Scan.bulk
  (keyword || variable || number || space)
*}

ML {* op || *}

ML {* Scan.bulk *}


ML {* tokens (explode "ADD 1 2") *}



ML {*
fun lex str =
  explode str
  |> these o Scan.read Symbol.stopper tokens
  |> filter_out (fn Space => true | _ => false)
*}

ML {* Scan.read *}


ML {* lex "ADD 1 2" *}

ML {* lex "SUM 1 2 3" *}

ML {* lex "LET x 3 IN ADD x 1" *}

ML {* lex "ADD 1 (ADD 2 3)" *}

ML {* tokens (explode "ADD (ADD 1 2) 3") *}





subsection {* Lexer: Error Handling *}

ML {* Scan.!! *}


ML {*
val token = keyword || variable || number || space

fun bad_input s = "Unexpected symbol: " ^ quote s
fun lex_err (xs, _) = bad_input (hd xs)

val tokens = Scan.!! lex_err (Scan.bulk token)
*}


ML {* tokens (explode "ADD (ADD 1 2) 3") *}

ML {*
let
  fun lex1 str =
    explode str
    |> Scan.read Symbol.stopper tokens
in
  lex1 "ADD (ADD 1 2) 3"
end
*}
ML {*
let
  fun lex2 str =
    explode str
    |> tokens
in
  lex2 "ADD 1 2"
end
*}



ML {*
fun lex str =
  Source.of_string str
  |> Source.source Symbol.stopper tokens NONE
  |> Source.exhaust
  |> filter_out (fn Space => true | _ => false)
*}

ML {* lex "ADD 1 2" *}

ML {* lex "SUM 1 2 3" *}

ML {* lex "LET x 3 IN ADD x 1" *}

ML {* try lex "ADD 1 (ADD 2 3" *}

ML {* lex "Add 1 2" *}





subsection {* Calculator: Basic Expressions *}



ML {*
fun keyword name = Scan.one
  (fn Keyword n => n = name | _ => false)
*}

ML {*
val number = Scan.some
  (fn Number n => SOME n | _ => NONE)
*}

ML {*
val stop = Scan.one (fn t => t = End)
*}

ML {*
fun expr st =
 (number ||
  keyword "ADD" |-- expr -- expr >> (op +) ||
  keyword "SUM" |-- 
    Scan.unless stop (Scan.repeat expr) >>
      (fn is => fold (curry (op +)) is 0)) st
*}

ML {*
val f1 = op  --
val f2 = op |--
val f3 = op  --|
*}

ML {* expr (lex "ADD 1 2" @ [End]) *}

ML {* expr (lex "SUM 1 2 3" @ [End]) *}



ML {*
fun calc str = 
  lex str @ [End]
  |> expr --| stop
  |> fst
*}

ML {* calc "ADD 1 2" *}

ML {* calc "SUM 1 2 3" *}

ML {* calc "ADD 1 ADD 2 3" *}

ML {* calc "ADD 1 SUM 2 3 4" *}

ML {* calc "SUM ADD 1 2 3" *}

ML {* try calc "ADD SUM 1 2 3" *}

ML {* try calc "ADD 1 2 3" *}





subsection {* Calculator: Error Handling *}

ML {*
fun string_of_token (Keyword s) = s
  | string_of_token (Variable s) = s
  | string_of_token (Number n) = string_of_int n
  | string_of_token _ = "end"

fun bad_token (xs, _) =
  "Unexpected token: " ^ quote (string_of_token (hd xs))

fun calc str =
  lex str @ [End]
  |> Scan.error (expr --| Scan.!! bad_token stop)
  |> fst
*}

ML {* try calc "ADD 1 2 3" *}

ML {* try calc "ADD SUM 1 2 3" *}

ML {* try calc "Add 0 0" *}



ML {*
fun append_token msg xs =
  msg ^ quote (string_of_token (hd xs))

val expr' =
  expr ||
  Scan.fail_with (append_token ("Expected number, " ^
    "ADD, or SUM, instead of "))

val stop' =
  stop ||
  Scan.fail_with (append_token "Unexpected token ")
*}

ML {*
fun calc str =
  lex str @ [End]
  |> Scan.catch (expr' --| stop')
  |> fst
*}

ML {* try calc "ADD 0 0 1" *}

ML {* try calc "Add 0 0" *}

ML {* try calc "ADD SUM 1 2 3" *}





subsection {* Calculator: Expressions with Variables *}

ML {*
fun keyword name = Scan.lift (Scan.one
  (fn Keyword n => n = name | _ => false))

fun number st = Scan.lift (Scan.some
  (fn Number n => SOME n | _ => NONE)) st

fun stop st = Scan.lift (Scan.one 
  (fn t => t = End orelse t = Keyword "IN")) st
*}

ML {* Scan.lift *}



ML {*
val variable = Scan.some
  (fn Variable n => SOME n | _ => NONE)
*}



ML {*
fun lookup env v =
  (case Symtab.lookup env v of
    SOME i => i
  | _ => error ("Unbound variable " ^ quote v))
*}



ML {*
fun expr st =
 (number ||

  Scan.peek (fn env => variable >> lookup env) ||

  keyword "LET" |-- Scan.lift variable -- expr --|
    keyword "IN" :|--
      (fn b => apfst (Symtab.update b) #> expr) ||

  keyword "ADD" |-- expr -- expr >> (op +) ||
  keyword "SUM" |-- 
    Scan.unless stop (Scan.repeat expr) >>
      (fn is => fold (curry (op +)) is 0)) st
*}

ML {* Scan.peek *}

ML {* op :|-- *}

ML {* expr (Symtab.empty, lex "LET x 1 IN x") *}



ML {*
fun calc str =
  lex str @ [End]
  |> Scan.pass Symtab.empty (expr --| stop)
  |> fst
*}

ML {* calc "ADD 1 2" *}

ML {* calc "ADD 1 ADD 2 3" *}

ML {* calc "SUM 1 2 3" *}

ML {* calc "LET x 3 IN ADD x 1" *}

ML {* calc "LET x SUM 1 1 1 IN ADD x 7" *}





subsection {* Possible extensions *}

text {*
Errors with source positions:

  see Pure/General/source_pos.ML
*}

end