(***************************************************************************** * Featherweight-OCL --- A Formal Semantics for UML-OCL Version OCL 2.5 * for the OMG Standard. * http://www.brucker.ch/projects/hol-testgen/ * * UML_Integer.thy --- Library definitions. * 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 UML_Integer imports "../UML_PropertyProfiles" begin section\Basic Type Integer: Operations\ subsection\Fundamental Predicates on Integers: Strict Equality \label{sec:integer-strict-eq}\ text\The last basic operation belonging to the fundamental infrastructure of a value-type in OCL is the weak equality, which is defined similar to the @{typ "('\)Boolean"}-case as strict extension of the strong equality:\ overloading StrictRefEq \ "StrictRefEq :: [('\)Integer,('\)Integer] \ ('\)Boolean" begin definition StrictRefEq\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r[code_unfold] : "(x::('\)Integer) \ y \ \ \. if (\ x) \ = true \ \ (\ y) \ = true \ then (x \ y) \ else invalid \" end text\Property proof in terms of @{term "profile_bin\<^sub>S\<^sub>t\<^sub>r\<^sub>o\<^sub>n\<^sub>g\<^sub>E\<^sub>q_\<^sub>v_\<^sub>v"}\ interpretation StrictRefEq\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r : profile_bin\<^sub>S\<^sub>t\<^sub>r\<^sub>o\<^sub>n\<^sub>g\<^sub>E\<^sub>q_\<^sub>v_\<^sub>v "\ x y. (x::('\)Integer) \ y" by unfold_locales (auto simp: StrictRefEq\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r) subsection\Basic Integer Constants\ text\Although the remaining part of this library reasons about integers abstractly, we provide here as example some convenient shortcuts.\ definition OclInt0 ::"('\)Integer" ("\") where "\ = (\ _ . \\0::int\\)" definition OclInt1 ::"('\)Integer" ("\") where "\ = (\ _ . \\1::int\\)" definition OclInt2 ::"('\)Integer" ("\") where "\ = (\ _ . \\2::int\\)" text\Etc.\ text_raw\\isatagafp\ definition OclInt3 ::"('\)Integer" ("\") where "\ = (\ _ . \\3::int\\)" definition OclInt4 ::"('\)Integer" ("\") where "\ = (\ _ . \\4::int\\)" definition OclInt5 ::"('\)Integer" ("\") where "\ = (\ _ . \\5::int\\)" definition OclInt6 ::"('\)Integer" ("\") where "\ = (\ _ . \\6::int\\)" definition OclInt7 ::"('\)Integer" ("\") where "\ = (\ _ . \\7::int\\)" definition OclInt8 ::"('\)Integer" ("\") where "\ = (\ _ . \\8::int\\)" definition OclInt9 ::"('\)Integer" ("\") where "\ = (\ _ . \\9::int\\)" definition OclInt10 ::"('\)Integer" ("\\")where "\\ = (\ _ . \\10::int\\)" subsection\Validity and Definedness Properties\ lemma "\(null::('\)Integer) = false" by simp lemma "\(null::('\)Integer) = true" by simp lemma [simp,code_unfold]: "\ (\_. \\n\\) = true" by(simp add:defined_def true_def bot_fun_def bot_option_def null_fun_def null_option_def) lemma [simp,code_unfold]: "\ (\_. \\n\\) = true" by(simp add:valid_def true_def bot_fun_def bot_option_def) (* ecclectic proofs to make examples executable *) lemma [simp,code_unfold]: "\ \ = true" by(simp add:OclInt0_def) lemma [simp,code_unfold]: "\ \ = true" by(simp add:OclInt0_def) lemma [simp,code_unfold]: "\ \ = true" by(simp add:OclInt1_def) lemma [simp,code_unfold]: "\ \ = true" by(simp add:OclInt1_def) lemma [simp,code_unfold]: "\ \ = true" by(simp add:OclInt2_def) lemma [simp,code_unfold]: "\ \ = true" by(simp add:OclInt2_def) lemma [simp,code_unfold]: "\ \ = true" by(simp add:OclInt6_def) lemma [simp,code_unfold]: "\ \ = true" by(simp add:OclInt6_def) lemma [simp,code_unfold]: "\ \ = true" by(simp add:OclInt8_def) lemma [simp,code_unfold]: "\ \ = true" by(simp add:OclInt8_def) lemma [simp,code_unfold]: "\ \ = true" by(simp add:OclInt9_def) lemma [simp,code_unfold]: "\ \ = true" by(simp add:OclInt9_def) text_raw\\endisatagafp\ subsection\Arithmetical Operations\ subsubsection\Definition\ text\Here is a common case of a built-in operation on built-in types. Note that the arguments must be both defined (non-null, non-bot).\ text\Note that we can not follow the lexis of the OCL Standard for Isabelle technical reasons; these operators are heavily overloaded in the HOL library that a further overloading would lead to heavy technical buzz in this document. \ definition OclAdd\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r ::"('\)Integer \ ('\)Integer \ ('\)Integer" (infix "+\<^sub>i\<^sub>n\<^sub>t" 40) where "x +\<^sub>i\<^sub>n\<^sub>t y \ \ \. if (\ x) \ = true \ \ (\ y) \ = true \ then \\\\x \\\ + \\y \\\\\ else invalid \ " interpretation OclAdd\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r : profile_bin\<^sub>d_\<^sub>d "(+\<^sub>i\<^sub>n\<^sub>t)" "\ x y. \\\\x\\ + \\y\\\\" by unfold_locales (auto simp:OclAdd\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r_def bot_option_def null_option_def) definition OclMinus\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r ::"('\)Integer \ ('\)Integer \ ('\)Integer" (infix "-\<^sub>i\<^sub>n\<^sub>t" 41) where "x -\<^sub>i\<^sub>n\<^sub>t y \ \ \. if (\ x) \ = true \ \ (\ y) \ = true \ then \\\\x \\\ - \\y \\\\\ else invalid \ " interpretation OclMinus\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r : profile_bin\<^sub>d_\<^sub>d "(-\<^sub>i\<^sub>n\<^sub>t)" "\ x y. \\\\x\\ - \\y\\\\" by unfold_locales (auto simp:OclMinus\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r_def bot_option_def null_option_def) definition OclMult\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r ::"('\)Integer \ ('\)Integer \ ('\)Integer" (infix "*\<^sub>i\<^sub>n\<^sub>t" 45) where "x *\<^sub>i\<^sub>n\<^sub>t y \ \ \. if (\ x) \ = true \ \ (\ y) \ = true \ then \\\\x \\\ * \\y \\\\\ else invalid \" interpretation OclMult\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r : profile_bin\<^sub>d_\<^sub>d "OclMult\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r" "\ x y. \\\\x\\ * \\y\\\\" by unfold_locales (auto simp:OclMult\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r_def bot_option_def null_option_def) text\Here is the special case of division, which is defined as invalid for division by zero.\ definition OclDivision\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r ::"('\)Integer \ ('\)Integer \ ('\)Integer" (infix "div\<^sub>i\<^sub>n\<^sub>t" 45) where "x div\<^sub>i\<^sub>n\<^sub>t y \ \ \. if (\ x) \ = true \ \ (\ y) \ = true \ then if y \ \ OclInt0 \ then \\\\x \\\ div \\y \\\\\ else invalid \ else invalid \ " (* TODO: special locale setup.*) definition OclModulus\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r ::"('\)Integer \ ('\)Integer \ ('\)Integer" (infix "mod\<^sub>i\<^sub>n\<^sub>t" 45) where "x mod\<^sub>i\<^sub>n\<^sub>t y \ \ \. if (\ x) \ = true \ \ (\ y) \ = true \ then if y \ \ OclInt0 \ then \\\\x \\\ mod \\y \\\\\ else invalid \ else invalid \ " (* TODO: special locale setup.*) definition OclLess\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r ::"('\)Integer \ ('\)Integer \ ('\)Boolean" (infix "<\<^sub>i\<^sub>n\<^sub>t" 35) where "x <\<^sub>i\<^sub>n\<^sub>t y \ \ \. if (\ x) \ = true \ \ (\ y) \ = true \ then \\\\x \\\ < \\y \\\\\ else invalid \ " interpretation OclLess\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r : profile_bin\<^sub>d_\<^sub>d "(<\<^sub>i\<^sub>n\<^sub>t)" "\ x y. \\\\x\\ < \\y\\\\" by unfold_locales (auto simp:OclLess\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r_def bot_option_def null_option_def) definition OclLe\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r ::"('\)Integer \ ('\)Integer \ ('\)Boolean" (infix "\\<^sub>i\<^sub>n\<^sub>t" 35) where "x \\<^sub>i\<^sub>n\<^sub>t y \ \ \. if (\ x) \ = true \ \ (\ y) \ = true \ then \\\\x \\\ \ \\y \\\\\ else invalid \ " interpretation OclLe\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r : profile_bin\<^sub>d_\<^sub>d "(\\<^sub>i\<^sub>n\<^sub>t)" "\ x y. \\\\x\\ \ \\y\\\\" by unfold_locales (auto simp:OclLe\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r_def bot_option_def null_option_def) subsubsection\Basic Properties\ lemma OclAdd\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r_commute: "(X +\<^sub>i\<^sub>n\<^sub>t Y) = (Y +\<^sub>i\<^sub>n\<^sub>t X)" by(rule ext,auto simp:true_def false_def OclAdd\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r_def invalid_def split: option.split option.split_asm bool.split bool.split_asm) subsubsection\Execution with Invalid or Null or Zero as Argument\ lemma OclAdd\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r_zero1[simp,code_unfold] : "(x +\<^sub>i\<^sub>n\<^sub>t \) = (if \ x and not (\ x) then invalid else x endif)" proof (rule ext, rename_tac \, case_tac "(\ x and not (\ x)) \ = true \") fix \ show "(\ x and not (\ x)) \ = true \ \ (x +\<^sub>i\<^sub>n\<^sub>t \) \ = (if \ x and not (\ x) then invalid else x endif) \" apply(subst OclIf_true', simp add: OclValid_def) by (metis OclAdd\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r_def OclNot_defargs OclValid_def foundation5 foundation9) next fix \ have A: "\\. (\ \ not (\ x and not (\ x))) = (x \ = invalid \ \ \ \ \ x)" by (metis OclNot_not OclOr_def defined5 defined6 defined_not_I foundation11 foundation18' foundation6 foundation7 foundation9 invalid_def) have B: "\ \ \ x \ \\\\x \\\\\ = x \" apply(cases "x \", metis bot_option_def foundation16) apply(rename_tac x', case_tac x', metis bot_option_def foundation16 null_option_def) by(simp) show "(x +\<^sub>i\<^sub>n\<^sub>t \) \ = (if \ x and not (\ x) then invalid else x endif) \" when "\ \ not (\ x and not (\ x))" apply(insert that, subst OclIf_false', simp, simp add: A, auto simp: OclAdd\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r_def OclInt0_def) (* *) apply(simp add: foundation16'[simplified OclValid_def]) apply(simp add: B) by(simp add: OclValid_def) qed(metis OclValid_def defined5 defined6 defined_and_I defined_not_I foundation9) lemma OclAdd\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r_zero2[simp,code_unfold] : "(\ +\<^sub>i\<^sub>n\<^sub>t x) = (if \ x and not (\ x) then invalid else x endif)" by(subst OclAdd\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r_commute, simp) (* TODO Basic proproperties for multiplication, division, modulus. *) subsection\Test Statements\ text\Here follows a list of code-examples, that explain the meanings of the above definitions by compilation to code and execution to @{term "True"}.\ Assert "\ \ ( \ \\<^sub>i\<^sub>n\<^sub>t \\ )" Assert "\ \ (( \ +\<^sub>i\<^sub>n\<^sub>t \ ) \\<^sub>i\<^sub>n\<^sub>t \\ )" Assert "\ |\ (( \ +\<^sub>i\<^sub>n\<^sub>t ( \ +\<^sub>i\<^sub>n\<^sub>t \ )) <\<^sub>i\<^sub>n\<^sub>t \\ )" Assert "\ \ not (\ (null +\<^sub>i\<^sub>n\<^sub>t \)) " Assert "\ \ (((\ *\<^sub>i\<^sub>n\<^sub>t \) div\<^sub>i\<^sub>n\<^sub>t \\) \\<^sub>i\<^sub>n\<^sub>t \) " Assert "\ \ not (\ (\ div\<^sub>i\<^sub>n\<^sub>t \)) " Assert "\ \ not (\ (\ div\<^sub>i\<^sub>n\<^sub>t \)) " lemma integer_non_null [simp]: "((\_. \\n\\) \ (null::('\)Integer)) = false" by(rule ext,auto simp: StrictRefEq\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r valid_def bot_fun_def bot_option_def null_fun_def null_option_def StrongEq_def) lemma null_non_integer [simp]: "((null::('\)Integer) \ (\_. \\n\\)) = false" by(rule ext,auto simp: StrictRefEq\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r valid_def bot_fun_def bot_option_def null_fun_def null_option_def StrongEq_def) lemma OclInt0_non_null [simp,code_unfold]: "(\ \ null) = false" by(simp add: OclInt0_def) lemma null_non_OclInt0 [simp,code_unfold]: "(null \ \) = false" by(simp add: OclInt0_def) lemma OclInt1_non_null [simp,code_unfold]: "(\ \ null) = false" by(simp add: OclInt1_def) lemma null_non_OclInt1 [simp,code_unfold]: "(null \ \) = false" by(simp add: OclInt1_def) lemma OclInt2_non_null [simp,code_unfold]: "(\ \ null) = false" by(simp add: OclInt2_def) lemma null_non_OclInt2 [simp,code_unfold]: "(null \ \) = false" by(simp add: OclInt2_def) lemma OclInt6_non_null [simp,code_unfold]: "(\ \ null) = false" by(simp add: OclInt6_def) lemma null_non_OclInt6 [simp,code_unfold]: "(null \ \) = false" by(simp add: OclInt6_def) lemma OclInt8_non_null [simp,code_unfold]: "(\ \ null) = false" by(simp add: OclInt8_def) lemma null_non_OclInt8 [simp,code_unfold]: "(null \ \) = false" by(simp add: OclInt8_def) lemma OclInt9_non_null [simp,code_unfold]: "(\ \ null) = false" by(simp add: OclInt9_def) lemma null_non_OclInt9 [simp,code_unfold]: "(null \ \) = false" by(simp add: OclInt9_def) text\Here follows a list of code-examples, that explain the meanings of the above definitions by compilation to code and execution to @{term "True"}.\ text\Elementary computations on Integer\ Assert "\ \ ((\ <\<^sub>i\<^sub>n\<^sub>t \) and (\ <\<^sub>i\<^sub>n\<^sub>t \))" Assert "\ \ \ <> \" Assert "\ \ \ <> \" Assert "\ \ \ \ \" Assert "\ \ \ \" Assert "\ \ \ \" Assert "\ \ \ (null::('\)Integer)" Assert "\ \ (invalid \ invalid)" Assert "\ \ (null \ null)" Assert "\ \ (\ \ \)" Assert "\ |\ (\ \ \\)" Assert "\ |\ (invalid \ \\)" Assert "\ |\ (null \ \\)" Assert "\ |\ (invalid \ (invalid::('\)Integer))" (* Without typeconstraint not executable.*) Assert "\ |\ \ (invalid \ (invalid::('\)Integer))" (* Without typeconstraint not executable.*) Assert "\ |\ (invalid <> (invalid::('\)Integer))" (* Without typeconstraint not executable.*) Assert "\ |\ \ (invalid <> (invalid::('\)Integer))" (* Without typeconstraint not executable.*) Assert "\ \ (null \ (null::('\)Integer) )" (* Without typeconstraint not executable.*) Assert "\ \ (null \ (null::('\)Integer) )" (* Without typeconstraint not executable.*) Assert "\ \ (\ \ \)" Assert "\ |\ (\ <> \)" Assert "\ |\ (\ \ \\)" Assert "\ \ (\ <> \\)" Assert "\ |\ (\ <\<^sub>i\<^sub>n\<^sub>t null)" Assert "\ |\ (\ (\ <\<^sub>i\<^sub>n\<^sub>t null))" end