Theory Design_OCL
theory
Design_OCL
imports
Design_UML
begin
text ‹\label{ex:employee-design:ocl}›
section‹OCL Part: Invariant›
text‹These recursive predicates can be defined conservatively
by greatest fix-point
constructions---automatically. See~\<^cite>‹"brucker.ea:hol-ocl-book:2006" and "brucker:interactive:2007"›
for details. For the purpose of this example, we state them as axioms
here.
\begin{ocl}
context Person
inv label : self .boss <> null implies (self .salary ≤ ((self .boss) .salary))
\end{ocl}
›
definition Person_label⇩i⇩n⇩v :: "Person ⇒ Boolean"
where "Person_label⇩i⇩n⇩v (self) ≡
(self .boss <> null implies (self .salary ≤⇩i⇩n⇩t ((self .boss) .salary)))"
definition Person_label⇩i⇩n⇩v⇩A⇩T⇩p⇩r⇩e :: "Person ⇒ Boolean"
where "Person_label⇩i⇩n⇩v⇩A⇩T⇩p⇩r⇩e (self) ≡
(self .boss@pre <> null implies (self .salary@pre ≤⇩i⇩n⇩t ((self .boss@pre) .salary@pre)))"
definition Person_label⇩g⇩l⇩o⇩b⇩a⇩l⇩i⇩n⇩v :: "Boolean"
where "Person_label⇩g⇩l⇩o⇩b⇩a⇩l⇩i⇩n⇩v ≡ (Person .allInstances()->forAll⇩S⇩e⇩t(x | Person_label⇩i⇩n⇩v (x)) and
(Person .allInstances@pre()->forAll⇩S⇩e⇩t(x | Person_label⇩i⇩n⇩v⇩A⇩T⇩p⇩r⇩e (x))))"
lemma "τ ⊨ δ (X .boss) ⟹ τ ⊨ Person .allInstances()->includes⇩S⇩e⇩t(X .boss) ∧
τ ⊨ Person .allInstances()->includes⇩S⇩e⇩t(X) "
oops
lemma REC_pre : "τ ⊨ Person_label⇩g⇩l⇩o⇩b⇩a⇩l⇩i⇩n⇩v
⟹ τ ⊨ Person .allInstances()->includes⇩S⇩e⇩t(X)
⟹ ∃ REC. τ ⊨ REC(X) ≜ (Person_label⇩i⇩n⇩v (X) and (X .boss <> null implies REC(X .boss)))"
oops
text‹This allows to state a predicate:›
axiomatization inv⇩P⇩e⇩r⇩s⇩o⇩n⇩_⇩l⇩a⇩b⇩e⇩l :: "Person ⇒ Boolean"
where inv⇩P⇩e⇩r⇩s⇩o⇩n⇩_⇩l⇩a⇩b⇩e⇩l_def:
"(τ ⊨ Person .allInstances()->includes⇩S⇩e⇩t(self)) ⟹
(τ ⊨ (inv⇩P⇩e⇩r⇩s⇩o⇩n⇩_⇩l⇩a⇩b⇩e⇩l(self) ≜ (self .boss <> null implies
(self .salary ≤⇩i⇩n⇩t ((self .boss) .salary)) and
inv⇩P⇩e⇩r⇩s⇩o⇩n⇩_⇩l⇩a⇩b⇩e⇩l(self .boss))))"
axiomatization inv⇩P⇩e⇩r⇩s⇩o⇩n⇩_⇩l⇩a⇩b⇩e⇩l⇩A⇩T⇩p⇩r⇩e :: "Person ⇒ Boolean"
where inv⇩P⇩e⇩r⇩s⇩o⇩n⇩_⇩l⇩a⇩b⇩e⇩l⇩A⇩T⇩p⇩r⇩e_def:
"(τ ⊨ Person .allInstances@pre()->includes⇩S⇩e⇩t(self)) ⟹
(τ ⊨ (inv⇩P⇩e⇩r⇩s⇩o⇩n⇩_⇩l⇩a⇩b⇩e⇩l⇩A⇩T⇩p⇩r⇩e(self) ≜ (self .boss@pre <> null implies
(self .salary@pre ≤⇩i⇩n⇩t ((self .boss@pre) .salary@pre)) and
inv⇩P⇩e⇩r⇩s⇩o⇩n⇩_⇩l⇩a⇩b⇩e⇩l⇩A⇩T⇩p⇩r⇩e(self .boss@pre))))"
lemma inv_1 :
"(τ ⊨ Person .allInstances()->includes⇩S⇩e⇩t(self)) ⟹
(τ ⊨ inv⇩P⇩e⇩r⇩s⇩o⇩n⇩_⇩l⇩a⇩b⇩e⇩l(self) = ((τ ⊨ (self .boss ≐ null)) ∨
( τ ⊨ (self .boss <> null) ∧
τ ⊨ ((self .salary) ≤⇩i⇩n⇩t (self .boss .salary)) ∧
τ ⊨ (inv⇩P⇩e⇩r⇩s⇩o⇩n⇩_⇩l⇩a⇩b⇩e⇩l(self .boss))))) "
oops
lemma inv_2 :
"(τ ⊨ Person .allInstances@pre()->includes⇩S⇩e⇩t(self)) ⟹
(τ ⊨ inv⇩P⇩e⇩r⇩s⇩o⇩n⇩_⇩l⇩a⇩b⇩e⇩l⇩A⇩T⇩p⇩r⇩e(self)) = ((τ ⊨ (self .boss@pre ≐ null)) ∨
(τ ⊨ (self .boss@pre <> null) ∧
(τ ⊨ (self .boss@pre .salary@pre ≤⇩i⇩n⇩t self .salary@pre)) ∧
(τ ⊨ (inv⇩P⇩e⇩r⇩s⇩o⇩n⇩_⇩l⇩a⇩b⇩e⇩l⇩A⇩T⇩p⇩r⇩e(self .boss@pre)))))"
oops
text‹A very first attempt to characterize the axiomatization by an inductive
definition - this can not be the last word since too weak (should be equality!)›
coinductive inv :: "Person ⇒ (𝔄)st ⇒ bool" where
"(τ ⊨ (δ self)) ⟹ ((τ ⊨ (self .boss ≐ null)) ∨
(τ ⊨ (self .boss <> null) ∧ (τ ⊨ (self .boss .salary ≤⇩i⇩n⇩t self .salary)) ∧
( (inv(self .boss))τ )))
⟹ ( inv self τ)"
section‹OCL Part: The Contract of a Recursive Query›
text‹This part is analogous to the Analysis Model and skipped here.›
end