Theory PromelaDatastructures
section "Data structures as used in Promela"
theory PromelaDatastructures
imports
CAVA_Base.CAVA_Base
CAVA_Base.Lexord_List
PromelaAST
"HOL-Library.IArray"
Deriving.Compare_Instances
CAVA_Base.CAVA_Code_Target
begin
subsection ‹Abstract Syntax Tree \emph{after} preprocessing›
text ‹
From the plain AST stemming from the parser, we'd like to have one containing more information while also removing duplicated constructs. This is achieved in the preprocessing step.
The additional information contains:
\begin{itemize}
\item variable type (including whether it represents a channel or not)
\item global vs local variable
\end{itemize}
Also certain constructs are expanded (like for-loops) or different nodes in the
AST are collapsed into one parametrized node (e.g.\ the different send-operations).
This preprocessing phase also tries to detect certain static errors and will bail out with an exception if such is encountered.
›