(***************************************************************************** * Featherweight-OCL --- A Formal Semantics for UML-OCL Version OCL 2.5 * for the OMG Standard. * http://www.brucker.ch/projects/hol-testgen/ * * UML_Library.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_Library imports (* Basic Type Operations *) "basic_types/UML_Boolean" "basic_types/UML_Void" "basic_types/UML_Integer" "basic_types/UML_Real" "basic_types/UML_String" (* Collection Type Operations *) "collection_types/UML_Pair" "collection_types/UML_Bag" "collection_types/UML_Set" "collection_types/UML_Sequence" begin section\Miscellaneous Stuff\ subsection\Definition: asBoolean\ definition OclAsBoolean\<^sub>I\<^sub>n\<^sub>t :: "('\) Integer \ ('\) Boolean" ("(_)->oclAsType\<^sub>I\<^sub>n\<^sub>t'(Boolean')") where "OclAsBoolean\<^sub>I\<^sub>n\<^sub>t X = (\\. if (\ X) \ = true \ then \\\\X \\\ \ 0\\ else invalid \)" interpretation OclAsBoolean\<^sub>I\<^sub>n\<^sub>t : profile_mono\<^sub>d OclAsBoolean\<^sub>I\<^sub>n\<^sub>t "\x. \\\\x\\ \ 0\\" by unfold_locales (auto simp: OclAsBoolean\<^sub>I\<^sub>n\<^sub>t_def bot_option_def) definition OclAsBoolean\<^sub>R\<^sub>e\<^sub>a\<^sub>l :: "('\) Real \ ('\) Boolean" ("(_)->oclAsType\<^sub>R\<^sub>e\<^sub>a\<^sub>l'(Boolean')") where "OclAsBoolean\<^sub>R\<^sub>e\<^sub>a\<^sub>l X = (\\. if (\ X) \ = true \ then \\\\X \\\ \ 0\\ else invalid \)" interpretation OclAsBoolean\<^sub>R\<^sub>e\<^sub>a\<^sub>l : profile_mono\<^sub>d OclAsBoolean\<^sub>R\<^sub>e\<^sub>a\<^sub>l "\x. \\\\x\\ \ 0\\" by unfold_locales (auto simp: OclAsBoolean\<^sub>R\<^sub>e\<^sub>a\<^sub>l_def bot_option_def) subsection\Definition: asInteger\ definition OclAsInteger\<^sub>R\<^sub>e\<^sub>a\<^sub>l :: "('\) Real \ ('\) Integer" ("(_)->oclAsType\<^sub>R\<^sub>e\<^sub>a\<^sub>l'(Integer')") where "OclAsInteger\<^sub>R\<^sub>e\<^sub>a\<^sub>l X = (\\. if (\ X) \ = true \ then \\floor \\X \\\\\ else invalid \)" interpretation OclAsInteger\<^sub>R\<^sub>e\<^sub>a\<^sub>l : profile_mono\<^sub>d OclAsInteger\<^sub>R\<^sub>e\<^sub>a\<^sub>l "\x. \\floor \\x\\\\" by unfold_locales (auto simp: OclAsInteger\<^sub>R\<^sub>e\<^sub>a\<^sub>l_def bot_option_def) subsection\Definition: asReal\ definition OclAsReal\<^sub>I\<^sub>n\<^sub>t :: "('\) Integer \ ('\) Real" ("(_)->oclAsType\<^sub>I\<^sub>n\<^sub>t'(Real')") where "OclAsReal\<^sub>I\<^sub>n\<^sub>t X = (\\. if (\ X) \ = true \ then \\real_of_int \\X \\\\\ else invalid \)" interpretation OclAsReal\<^sub>I\<^sub>n\<^sub>t : profile_mono\<^sub>d OclAsReal\<^sub>I\<^sub>n\<^sub>t "\x. \\real_of_int \\x\\\\" by unfold_locales (auto simp: OclAsReal\<^sub>I\<^sub>n\<^sub>t_def bot_option_def) lemma Integer_subtype_of_Real: assumes "\ \ \ X" shows "\ \ X ->oclAsType\<^sub>I\<^sub>n\<^sub>t(Real) ->oclAsType\<^sub>R\<^sub>e\<^sub>a\<^sub>l(Integer) \ X" apply(insert assms, simp add: OclAsInteger\<^sub>R\<^sub>e\<^sub>a\<^sub>l_def OclAsReal\<^sub>I\<^sub>n\<^sub>t_def OclValid_def StrongEq_def) apply(subst (2 4) cp_defined, simp add: true_def) by (metis assms bot_option_def drop.elims foundation16 null_option_def) subsection\Definition: asPair\ definition OclAsPair\<^sub>S\<^sub>e\<^sub>q :: "[('\,'\::null)Sequence]\('\,'\::null,'\::null) Pair" ("(_)->asPair\<^sub>S\<^sub>e\<^sub>q'(')") where "OclAsPair\<^sub>S\<^sub>e\<^sub>q S = (if S->size\<^sub>S\<^sub>e\<^sub>q() \ \ then Pair{S->at\<^sub>S\<^sub>e\<^sub>q(\),S->at\<^sub>S\<^sub>e\<^sub>q(\)} else invalid endif)" definition OclAsPair\<^sub>S\<^sub>e\<^sub>t :: "[('\,'\::null)Set]\('\,'\::null,'\::null) Pair" ("(_)->asPair\<^sub>S\<^sub>e\<^sub>t'(')") where "OclAsPair\<^sub>S\<^sub>e\<^sub>t S = (if S->size\<^sub>S\<^sub>e\<^sub>t() \ \ then let v = S->any\<^sub>S\<^sub>e\<^sub>t() in Pair{v,S->excluding\<^sub>S\<^sub>e\<^sub>t(v)->any\<^sub>S\<^sub>e\<^sub>t()} else invalid endif)" definition OclAsPair\<^sub>B\<^sub>a\<^sub>g :: "[('\,'\::null)Bag]\('\,'\::null,'\::null) Pair" ("(_)->asPair\<^sub>B\<^sub>a\<^sub>g'(')") where "OclAsPair\<^sub>B\<^sub>a\<^sub>g S = (if S->size\<^sub>B\<^sub>a\<^sub>g() \ \ then let v = S->any\<^sub>B\<^sub>a\<^sub>g() in Pair{v,S->excluding\<^sub>B\<^sub>a\<^sub>g(v)->any\<^sub>B\<^sub>a\<^sub>g()} else invalid endif)" subsection\Definition: asSet\ definition OclAsSet\<^sub>S\<^sub>e\<^sub>q :: "[('\,'\::null)Sequence]\('\,'\)Set" ("(_)->asSet\<^sub>S\<^sub>e\<^sub>q'(')") where "OclAsSet\<^sub>S\<^sub>e\<^sub>q S = (S->iterate\<^sub>S\<^sub>e\<^sub>q(b; x = Set{} | x ->including\<^sub>S\<^sub>e\<^sub>t(b)))" definition OclAsSet\<^sub>P\<^sub>a\<^sub>i\<^sub>r :: "[('\,'\::null,'\::null) Pair]\('\,'\)Set" ("(_)->asSet\<^sub>P\<^sub>a\<^sub>i\<^sub>r'(')") where "OclAsSet\<^sub>P\<^sub>a\<^sub>i\<^sub>r S = Set{S .First(), S .Second()}" definition OclAsSet\<^sub>B\<^sub>a\<^sub>g :: "('\,'\::null) Bag\('\,'\)Set" ("(_)->asSet\<^sub>B\<^sub>a\<^sub>g'(')") where "OclAsSet\<^sub>B\<^sub>a\<^sub>g S = (\ \. if (\ S) \ = true \ then Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e\\ Rep_Set_base S \ \\ else if (\ S) \ = true \ then null \ else invalid \)" subsection\Definition: asSequence\ definition OclAsSeq\<^sub>S\<^sub>e\<^sub>t :: "[('\,'\::null)Set]\('\,'\)Sequence" ("(_)->asSequence\<^sub>S\<^sub>e\<^sub>t'(')") where "OclAsSeq\<^sub>S\<^sub>e\<^sub>t S = (S->iterate\<^sub>S\<^sub>e\<^sub>t(b; x = Sequence{} | x ->including\<^sub>S\<^sub>e\<^sub>q(b)))" definition OclAsSeq\<^sub>B\<^sub>a\<^sub>g :: "[('\,'\::null)Bag]\('\,'\)Sequence" ("(_)->asSequence\<^sub>B\<^sub>a\<^sub>g'(')") where "OclAsSeq\<^sub>B\<^sub>a\<^sub>g S = (S->iterate\<^sub>B\<^sub>a\<^sub>g(b; x = Sequence{} | x ->including\<^sub>S\<^sub>e\<^sub>q(b)))" definition OclAsSeq\<^sub>P\<^sub>a\<^sub>i\<^sub>r :: "[('\,'\::null,'\::null) Pair]\('\,'\)Sequence" ("(_)->asSequence\<^sub>P\<^sub>a\<^sub>i\<^sub>r'(')") where "OclAsSeq\<^sub>P\<^sub>a\<^sub>i\<^sub>r S = Sequence{S .First(), S .Second()}" subsection\Definition: asBag\ definition OclAsBag\<^sub>S\<^sub>e\<^sub>q :: "[('\,'\::null)Sequence]\('\,'\)Bag" ("(_)->asBag\<^sub>S\<^sub>e\<^sub>q'(')") where "OclAsBag\<^sub>S\<^sub>e\<^sub>q S = (\\. Abs_Bag\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\\s. if list_ex ((=) s) \\Rep_Sequence\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\ then 1 else 0\\)" definition OclAsBag\<^sub>S\<^sub>e\<^sub>t :: "[('\,'\::null)Set]\('\,'\)Bag" ("(_)->asBag\<^sub>S\<^sub>e\<^sub>t'(')") where "OclAsBag\<^sub>S\<^sub>e\<^sub>t S = (\\. Abs_Bag\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\\s. if s \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\ then 1 else 0\\)" lemma assumes "\ \ \ (S ->size\<^sub>S\<^sub>e\<^sub>t())" (* S is finite *) shows "OclAsBag\<^sub>S\<^sub>e\<^sub>t S = (S->iterate\<^sub>S\<^sub>e\<^sub>t(b; x = Bag{} | x ->including\<^sub>B\<^sub>a\<^sub>g(b)))" oops definition OclAsBag\<^sub>P\<^sub>a\<^sub>i\<^sub>r :: "[('\,'\::null,'\::null) Pair]\('\,'\)Bag" ("(_)->asBag\<^sub>P\<^sub>a\<^sub>i\<^sub>r'(')") where "OclAsBag\<^sub>P\<^sub>a\<^sub>i\<^sub>r S = Bag{S .First(), S .Second()}" text_raw\\isatagafp\ subsection\Collection Types\ lemmas cp_intro'' [intro!,simp,code_unfold] = cp_intro' (* cp_intro''\<^sub>P\<^sub>a\<^sub>i\<^sub>r *) cp_intro''\<^sub>S\<^sub>e\<^sub>t cp_intro''\<^sub>S\<^sub>e\<^sub>q text_raw\\endisatagafp\ subsection\Test Statements\ lemma syntax_test: "Set{\,\} = (Set{}->including\<^sub>S\<^sub>e\<^sub>t(\)->including\<^sub>S\<^sub>e\<^sub>t(\))" by (rule refl) text\Here is an example of a nested collection.\ lemma semantic_test2: assumes H:"(Set{\} \ null) = (false::('\)Boolean)" shows "(\::('\)st) \ (Set{Set{\},null}->includes\<^sub>S\<^sub>e\<^sub>t(null))" by(simp add: OclIncludes_execute\<^sub>S\<^sub>e\<^sub>t H) lemma short_cut'[simp,code_unfold]: "(\ \ \) = false" apply(rule ext) apply(simp add: StrictRefEq\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r StrongEq_def OclInt8_def OclInt6_def true_def false_def invalid_def bot_option_def) done lemma short_cut''[simp,code_unfold]: "(\ \ \) = false" apply(rule ext) apply(simp add: StrictRefEq\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r StrongEq_def OclInt2_def OclInt1_def true_def false_def invalid_def bot_option_def) done lemma short_cut'''[simp,code_unfold]: "(\ \ \) = false" apply(rule ext) apply(simp add: StrictRefEq\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r StrongEq_def OclInt2_def OclInt1_def true_def false_def invalid_def bot_option_def) done Assert "\ \ (\ <\<^sub>i\<^sub>n\<^sub>t \) and (\ <\<^sub>i\<^sub>n\<^sub>t \) " text\Elementary computations on Sets.\ declare OclSelect_body_def [simp] Assert "\ (\ \ \(invalid::('\,'\::null) Set))" Assert "\ \ \(null::('\,'\::null) Set)" Assert "\ (\ \ \(null::('\,'\::null) Set))" Assert "\ \ \(Set{})" Assert "\ \ \(Set{Set{\},null})" Assert "\ \ \(Set{Set{\},null})" Assert "\ \ (Set{\,\}->includes\<^sub>S\<^sub>e\<^sub>t(\))" Assert "\ (\ \ (Set{\}->includes\<^sub>S\<^sub>e\<^sub>t(\)))" Assert "\ (\ \ (Set{\,\}->includes\<^sub>S\<^sub>e\<^sub>t(null)))" Assert "\ \ (Set{\,null}->includes\<^sub>S\<^sub>e\<^sub>t(null))" Assert "\ \ (Set{null,\}->includes\<^sub>S\<^sub>e\<^sub>t(null))" Assert "\ \ ((Set{})->forAll\<^sub>S\<^sub>e\<^sub>t(z | \ <\<^sub>i\<^sub>n\<^sub>t z))" Assert "\ \ ((Set{\,\})->forAll\<^sub>S\<^sub>e\<^sub>t(z | \ <\<^sub>i\<^sub>n\<^sub>t z))" Assert "\ (\ \ ((Set{\,\})->exists\<^sub>S\<^sub>e\<^sub>t(z | z <\<^sub>i\<^sub>n\<^sub>t \ )))" Assert "\ (\ \ (\(Set{\,null})->forAll\<^sub>S\<^sub>e\<^sub>t(z | \ <\<^sub>i\<^sub>n\<^sub>t z)))" Assert "\ (\ \ ((Set{\,null})->forAll\<^sub>S\<^sub>e\<^sub>t(z | \ <\<^sub>i\<^sub>n\<^sub>t z)))" Assert "\ \ ((Set{\,null})->exists\<^sub>S\<^sub>e\<^sub>t(z | \ <\<^sub>i\<^sub>n\<^sub>t z))" Assert "\ (\ \ (Set{null::'a Boolean} \ Set{}))" Assert "\ (\ \ (Set{null::'a Integer} \ Set{}))" Assert "\ (\ \ (Set{true} \ Set{false}))" Assert "\ (\ \ (Set{true,true} \ Set{false}))" Assert "\ (\ \ (Set{\} \ Set{\}))" Assert "\ \ (Set{\,null,\} \ Set{null,\})" Assert "\ \ (Set{\,null,\} <> Set{null,\})" Assert "\ \ (Set{Set{\,null}} \ Set{Set{null,\}})" Assert "\ \ (Set{Set{\,null}} <> Set{Set{null,\},null})" Assert "\ \ (Set{null}->select\<^sub>S\<^sub>e\<^sub>t(x | not x) \ Set{null})" Assert "\ \ (Set{null}->reject\<^sub>S\<^sub>e\<^sub>t(x | not x) \ Set{null})" lemma "const (Set{Set{\,null}, invalid})" by(simp add: const_ss) text\Elementary computations on Sequences.\ Assert "\ (\ \ \(invalid::('\,'\::null) Sequence))" Assert "\ \ \(null::('\,'\::null) Sequence)" Assert "\ (\ \ \(null::('\,'\::null) Sequence))" Assert "\ \ \(Sequence{})" lemma "const (Sequence{Sequence{\,null}, invalid})" by(simp add: const_ss) end