Theory Design_UML
chapter‹Example: The Employee Design Model›
theory
Design_UML
imports
"../../../UML_Main"
begin
text ‹\label{ex:employee-design:uml}›
section‹Introduction›
text‹
For certain concepts like classes and class-types, only a generic
definition for its resulting semantics can be given. Generic means,
there is a function outside HOL that ``compiles'' a concrete,
closed-world class diagram into a ``theory'' of this data model,
consisting of a bunch of definitions for classes, accessors, method,
casts, and tests for actual types, as well as proofs for the
fundamental properties of these operations in this concrete data
model.›
text‹Such generic function or ``compiler'' can be implemented in
Isabelle on the ML level. This has been done, for a semantics
following the open-world assumption, for UML 2.0
in~\<^cite>‹"brucker.ea:extensible:2008-b" and "brucker:interactive:2007"›. In
this paper, we follow another approach for UML 2.4: we define the
concepts of the compilation informally, and present a concrete
example which is verified in Isabelle/HOL.›
subsection‹Outlining the Example›
text‹We are presenting here a ``design-model'' of the (slightly
modified) example Figure 7.3, page 20 of
the OCL standard~\<^cite>‹"omg:ocl:2012"›. To be precise, this theory contains the formalization of
the data-part covered by the UML class model (see \autoref{fig:person}):›
text‹
\begin{figure}
\centering\scalebox{.3}{\includegraphics{figures/person.png}}%
\caption{A simple UML class model drawn from Figure 7.3,
page 20 of~\<^cite>‹"omg:ocl:2012"›. \label{fig:person}}
\end{figure}
›
text‹This means that the association (attached to the association class
\inlineocl{EmployeeRanking}) with the association ends \inlineocl+boss+ and \inlineocl+employees+ is implemented
by the attribute \inlineocl+boss+ and the operation \inlineocl+employees+ (to be discussed in the OCL part
captured by the subsequent theory).
›
section‹Example Data-Universe and its Infrastructure›
text‹Ideally, the following is generated automatically from a UML class model.›
text‹Our data universe consists in the concrete class diagram just of node's,
and implicitly of the class object. Each class implies the existence of a class
type defined for the corresponding object representations as follows:›