Theory Syntax
section ‹Abstract Syntax of IMP2›
theory Syntax
imports Main
begin
text ‹We define the abstract syntax of the IMP2 language,
and a minimal concrete syntax for direct use in terms.›
subsection ‹Primitives›
text ‹Variable and procedure names are strings.›
type_synonym vname = string
type_synonym pname = string
text ‹The variable names are partitioned into local and global variables.›
fun is_global :: "vname ⇒ bool" where
"is_global [] ⟷ True"
| "is_global (CHR ''G''#_) ⟷ True"
| "is_global _ ⟷ False"
abbreviation "is_local a ≡ ¬is_global a"
text ‹
Primitive values are integers,
and values are arrays modeled as functions from integers to primitive values.
Note that values and primitive values are usually part of the semantics, however,
as they occur as literals in the abstract syntax, we already define them here.
›
type_synonym pval = "int"
type_synonym val = "int ⇒ pval"
subsection ‹Arithmetic Expressions›
text ‹Arithmetic expressions consist of constants, indexed array variables,
and unary and binary operations. The operations are modeled by reflecting
arbitrary functions into the abstract syntax.›