Theory Design_OCL

(*****************************************************************************
 * Featherweight-OCL --- A Formal Semantics for UML-OCL Version OCL 2.5
 *                       for the OMG Standard.
 *                       http://www.brucker.ch/projects/hol-testgen/
 *
 * Design_OCL.thy --- OCL Contracts and an Example.
 * This file is part of HOL-TestGen.
 *
 * Copyright (c) 2012-2015 Université Paris-Saclay, Univ. Paris-Sud, France
 *               2013-2015 IRT SystemX, France
 *
 * All rights reserved.
 *
 * Redistribution and use in source and binary forms, with or without
 * modification, are permitted provided that the following conditions are
 * met:
 *
 *     * Redistributions of source code must retain the above copyright
 *       notice, this list of conditions and the following disclaimer.
 *
 *     * Redistributions in binary form must reproduce the above
 *       copyright notice, this list of conditions and the following
 *       disclaimer in the documentation and/or other materials provided
 *       with the distribution.
 *
 *     * Neither the name of the copyright holders nor the names of its
 *       contributors may be used to endorse or promote products derived
 *       from this software without specific prior written permission.
 *
 * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS
 * "AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT
 * LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR
 * A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT
 * OWNER OR CONTRIBUTORS BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL,
 * SPECIAL, EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT
 * LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE,
 * DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY
 * THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT
 * (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
 * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
 ******************************************************************************)

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_labelinv :: "Person  Boolean" 
where     "Person_labelinv (self)   
                 (self .boss <> null implies (self .salary  int  ((self .boss) .salary)))"
                                       

definition Person_labelinvATpre :: "Person  Boolean" 
where     "Person_labelinvATpre (self)   
                 (self .boss@pre <> null implies (self .salary@pre int ((self .boss@pre) .salary@pre)))"

definition Person_labelglobalinv :: "Boolean"
where     "Person_labelglobalinv  (Person .allInstances()->forAllSet(x | Person_labelinv (x)) and 
                                  (Person .allInstances@pre()->forAllSet(x | Person_labelinvATpre (x))))"
                                  
                                  
lemma "τ  δ (X .boss)  τ  Person .allInstances()->includesSet(X .boss) 
                            τ  Person .allInstances()->includesSet(X) "
oops  
(* To be generated generically ... hard, but crucial lemma that should hold. 
   It means that X and it successor are object representation that actually
   occur in the state. *)

lemma REC_pre : "τ  Person_labelglobalinv 
        τ  Person .allInstances()->includesSet(X) ― ‹X› represented object in state›
         REC.  τ  REC(X)   (Person_labelinv (X) and (X .boss <> null implies REC(X .boss)))"
oops (* Attempt to allegiate the burden of he following axiomatizations: could be
        a witness for a constant specification ...*)       

text‹This allows to state a predicate:›
                                       
axiomatization invPerson_label :: "Person  Boolean"
where invPerson_label_def:
"(τ  Person .allInstances()->includesSet(self))  
 (τ  (invPerson_label(self)   (self .boss <> null implies  
                                  (self .salary  int  ((self .boss) .salary)) and
                                   invPerson_label(self .boss))))"

axiomatization invPerson_labelATpre :: "Person  Boolean"
where invPerson_labelATpre_def: 
"(τ  Person .allInstances@pre()->includesSet(self)) 
 (τ  (invPerson_labelATpre(self)  (self .boss@pre <> null implies 
                                   (self .salary@pre  int  ((self .boss@pre) .salary@pre)) and
                                    invPerson_labelATpre(self .boss@pre))))"


lemma inv_1 : 
"(τ  Person .allInstances()->includesSet(self)) 
    (τ  invPerson_label(self) = ((τ  (self .boss  null)) 
                               ( τ  (self .boss <> null)  
                                 τ  ((self .salary)  int  (self .boss .salary))  
                                 τ  (invPerson_label(self .boss))))) "
oops (* Let's hope that this holds ... *)


lemma inv_2 : 
"(τ  Person .allInstances@pre()->includesSet(self)) 
    (τ  invPerson_labelATpre(self)) =  ((τ  (self .boss@pre  null)) 
                                     (τ  (self .boss@pre <> null) 
                                     (τ  (self .boss@pre .salary@pre int self .salary@pre))  
                                     (τ  (invPerson_labelATpre(self .boss@pre)))))"
oops (* Let's hope that this holds ... *)

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 int 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