(***************************************************************************** * Featherweight-OCL --- A Formal Semantics for UML-OCL Version OCL 2.5 * for the OMG Standard. * http://www.brucker.ch/projects/hol-testgen/ * * UML_String.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_String imports "../UML_PropertyProfiles" begin section\Basic Type String: Operations\ subsection\Fundamental Properties on Strings: Strict Equality \label{sec:string-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 :: [('\)String,('\)String] \ ('\)Boolean" begin definition StrictRefEq\<^sub>S\<^sub>t\<^sub>r\<^sub>i\<^sub>n\<^sub>g[code_unfold] : "(x::('\)String) \ 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>S\<^sub>t\<^sub>r\<^sub>i\<^sub>n\<^sub>g : 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::('\)String) \ y" by unfold_locales (auto simp: StrictRefEq\<^sub>S\<^sub>t\<^sub>r\<^sub>i\<^sub>n\<^sub>g) subsection\Basic String Constants\ text\Although the remaining part of this library reasons about integers abstractly, we provide here as example some convenient shortcuts.\ definition OclStringa ::"('\)String" ("\") where "\ = (\ _ . \\''a''\\)" definition OclStringb ::"('\)String" ("\") where "\ = (\ _ . \\''b''\\)" definition OclStringc ::"('\)String" ("\") where "\ = (\ _ . \\''c''\\)" text\Etc.\ text_raw\\isatagafp\ subsection\Validity and Definedness Properties\ lemma "\(null::('\)String) = false" by simp lemma "\(null::('\)String) = 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:OclStringa_def) lemma [simp,code_unfold]: "\ \ = true" by(simp add:OclStringa_def) text_raw\\endisatagafp\ subsection\String 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>S\<^sub>t\<^sub>r\<^sub>i\<^sub>n\<^sub>g ::"('\)String \ ('\)String \ ('\)String" (infix "+\<^sub>s\<^sub>t\<^sub>r\<^sub>i\<^sub>n\<^sub>g" 40) where "x +\<^sub>s\<^sub>t\<^sub>r\<^sub>i\<^sub>n\<^sub>g y \ \ \. if (\ x) \ = true \ \ (\ y) \ = true \ then \\concat [\\x \\\, \\y \\\]\\ else invalid \ " interpretation OclAdd\<^sub>S\<^sub>t\<^sub>r\<^sub>i\<^sub>n\<^sub>g : profile_bin\<^sub>d_\<^sub>d "(+\<^sub>s\<^sub>t\<^sub>r\<^sub>i\<^sub>n\<^sub>g)" "\ x y. \\concat [\\x\\, \\y\\]\\" by unfold_locales (auto simp:OclAdd\<^sub>S\<^sub>t\<^sub>r\<^sub>i\<^sub>n\<^sub>g_def bot_option_def null_option_def) (* TODO : size(), concat, substring(s:string) toInteger, toReal, at(i:Integer), characters() etc. *) subsubsection\Basic Properties\ lemma OclAdd\<^sub>S\<^sub>t\<^sub>r\<^sub>i\<^sub>n\<^sub>g_not_commute: "\X Y. (X +\<^sub>s\<^sub>t\<^sub>r\<^sub>i\<^sub>n\<^sub>g Y) \ (Y +\<^sub>s\<^sub>t\<^sub>r\<^sub>i\<^sub>n\<^sub>g X)" apply(rule_tac x = "\_. \\''b''\\" in exI) apply(rule_tac x = "\_. \\''a''\\" in exI) apply(simp_all add:OclAdd\<^sub>S\<^sub>t\<^sub>r\<^sub>i\<^sub>n\<^sub>g_def) by(auto, drule fun_cong, auto) 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>s\<^sub>t\<^sub>r\<^sub>i\<^sub>n\<^sub>g \\ )" Assert "\ \ (( \ +\<^sub>s\<^sub>t\<^sub>r\<^sub>i\<^sub>n\<^sub>g \ ) \\<^sub>s\<^sub>t\<^sub>r\<^sub>i\<^sub>n\<^sub>g \\ )" Assert "\ |\ (( \ +\<^sub>s\<^sub>t\<^sub>r\<^sub>i\<^sub>n\<^sub>g ( \ +\<^sub>s\<^sub>t\<^sub>r\<^sub>i\<^sub>n\<^sub>g \ )) <\<^sub>s\<^sub>t\<^sub>r\<^sub>i\<^sub>n\<^sub>g \\ )" Assert "\ \ not (\ (null +\<^sub>s\<^sub>t\<^sub>r\<^sub>i\<^sub>n\<^sub>g \)) " *) 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 String\ Assert "\ \ \ <> \" Assert "\ \ \ <> \" Assert "\ \ \ \ \" Assert "\ \ \ \" Assert "\ \ \ \" Assert "\ \ \ (null::('\)String)" Assert "\ \ (invalid \ invalid)" Assert "\ \ (null \ null)" Assert "\ \ (\ \ \)" Assert "\ |\ (\ \ \)" Assert "\ |\ (invalid \ \)" Assert "\ |\ (null \ \)" Assert "\ |\ (invalid \ (invalid::('\)String))" (* Without typeconstraint not executable.*) Assert "\ |\ \ (invalid \ (invalid::('\)String))" (* Without typeconstraint not executable.*) Assert "\ |\ (invalid <> (invalid::('\)String))" (* Without typeconstraint not executable.*) Assert "\ |\ \ (invalid <> (invalid::('\)String))" (* Without typeconstraint not executable.*) Assert "\ \ (null \ (null::('\)String) )" (* Without typeconstraint not executable.*) Assert "\ \ (null \ (null::('\)String) )" (* Without typeconstraint not executable.*) Assert "\ \ (\ \ \)" Assert "\ |\ (\ <> \)" Assert "\ |\ (\ \ \)" Assert "\ \ (\ <> \)" (*Assert "\ |\ (\ <\<^sub>s\<^sub>t\<^sub>r\<^sub>i\<^sub>n\<^sub>g null)" Assert "\ |\ (\ (\ <\<^sub>s\<^sub>t\<^sub>r\<^sub>i\<^sub>n\<^sub>g null))" *) end