Theory Syntax

section ‹The $\lambda\mu$-calculus›

text‹More examples, as well as a call-by-value programming language built on
top of our formalisation, can be found in an associated Bitbucket repository~cite"bitbucket".›
  
theory Syntax
  imports Main
begin

subsection ‹Syntax›

datatype type = 
     Iota
    | Fun type type (infixr "" 200)

text‹To deal with $\alpha$-equivalence, we use De Bruijn's nameless representation wherein each bound
     variable is represented by a natural number, its index, that denotes the number of binders
     that must be traversed to arrive at the one that binds the given variable.
     Each free variable has an index that points into the top-level context, not enclosed in
     any abstractions.›
  
datatype trm =
      LVar nat    ("`_" [100] 100)
    | Lbd type trm ("λ_:_" [0, 60] 60)
    | App trm trm (infix "°" 60)
    | Mu type cmd ("μ_:_" [0, 60] 60)
and cmd = 
      MVar nat trm ("<_>_" [0, 60] 60)
      
datatype ctxt = 
  ContextEmpty ("") |
  ContextApp ctxt trm (infixl "" 75)
  
primrec ctxt_app :: "ctxt  ctxt  ctxt" (infix "." 60) where
  " . F = F" |
  "(E  t) . F = (E . F)  t"
  
fun is_val :: "trm  bool" where
  "is_val (λ T : v) = True" |
  "is_val _ = False"

  
end