Theory Small_Step
section "Extension of language IMP with inputs, outputs, and nondeterminism"
theory Small_Step
imports
"HOL-IMP.BExp"
"HOL-IMP.Star"
begin
text ‹
\null
In a previous paper of mine \<^cite>‹"Noce24"›, the notion of termination-sensitive information flow
security with respect to a level-based interference relation, as studied in \<^cite>‹"Volpano96"›,
\<^cite>‹"Volpano97"› and formalized in \<^cite>‹"Nipkow23"›, is generalized to the notion of
termination-sensitive information flow correctness with respect to an interference function mapping
program states to (generally) intransitive interference relations. Moreover, a static type system is
specified and is proven to be capable of enforcing such information flow correctness policies.
The present paper extends both the aforesaid information flow correctness criterion and the related
static type system to the case of an imperative programming language supporting inputs, outputs, and
nondeterminism. Regarding inputs and nondeterminism, \<^cite>‹"Volpano96"›, section 7.1, observes that
``if we try to extend the core language with a primitive random number generator \emph{rand( )} and
allow an assignment such as \emph{z := rand( )} to be well typed when \emph{z} is low, then the
soundness theorem no longer holds'', and from this infers that ``new security models [...] should be
explored as potential notions of type soundness for new type systems that deal with nondeterministic
programs''. The present paper shows that this difficulty can be solved by extending the inductive
definition of the programming language's operational semantics so as to reflect the fact that, even
though the input instruction \emph{z := rand( )} may set \emph{z} to an arbitrary input value, the
same program state is produced whenever the input value is the same. As shown in this paper, this
enables to apply a suitably extended information flow correctness criterion based on stateful
intransitive noninterference, as well as an extended static type system enforcing this criterion, to
such an extended programming language.
The didactic imperative programming language IMP employed in \<^cite>‹"Nipkow23"›, extended with an
input instruction, an output instruction, and a control structure allowing for nondeterministic
choice, will be used for this purpose. Yet, in the same way as in my previous paper \<^cite>‹"Noce24"›,
the introduced concepts are applicable to larger, real-world imperative programming languages, too,
by just affording the additional type system complexity arising from richer language constructs.
For further information about the formal definitions and proofs contained in this paper, refer to
Isabelle documentation, particularly \<^cite>‹"Paulson24"›, \<^cite>‹"Nipkow24-4"›, \<^cite>‹"Krauss24"›,
\<^cite>‹"Nipkow11"›, and \<^cite>‹"Ballarin24"›.
As mentioned above, the first task to be tackled, which is the subject of this section, consists of
extending the original syntax, big-step operational semantics, and small-step operational semantics
of language IMP, as formalized in \<^cite>‹"Nipkow24-1"›, \<^cite>‹"Nipkow24-2"›, and \<^cite>‹"Nipkow24-3"›,
respectively.
›
subsection "Extended syntax"
text ‹
The starting point is extending the original syntax of language IMP with the following additional
constructs.
▪ An input instruction @{text "IN x"}, which sets variable \emph{x} to an input value.
▪ An output instruction @{text "OUT x"}, which outputs the current value of variable \emph{x}.
▪ A control structure @{text "c⇩1 OR c⇩2"}, which allows for a nondeterministic choice between
commands @{text "c⇩1"} and @{text "c⇩2"}.
\null
›
declare [[syntax_ambiguity_warning = false]]