Import of official AFP entry for Isabelle 2019.
afp-mirror/Featherweight_OCL/master This commit looks good Details

This commit is contained in:
Achim D. Brucker 2019-06-22 23:44:04 +01:00
parent 81e5b1a2af
commit 98df203945
24 changed files with 853 additions and 853 deletions

2
.ci/Jenkinsfile vendored
View File

@ -3,7 +3,7 @@ pipeline {
stages { stages {
stage('Build') { stage('Build') {
steps { steps {
sh 'docker run -v $PWD/Featherweight_OCL:/Featherweight_OCL logicalhacking:isabelle2018 isabelle build -D /Featherweight_OCL' sh 'docker run -v $PWD/Featherweight_OCL:/Featherweight_OCL logicalhacking:isabelle2019 isabelle build -D /Featherweight_OCL'
} }
} }
} }

View File

@ -1,7 +1,7 @@
chapter AFP chapter AFP
session "Featherweight_OCL-devel" (AFP) = HOL + session "Featherweight_OCL-devel" (AFP) = HOL +
description {* Featherweight-OCL *} description "Featherweight-OCL"
options [document_variants = "annex-a=annexa,-theory,-afp,-proof,-ML:document=afp,-annexa:outline=-annexa,afp,/proof,/ML", options [document_variants = "annex-a=annexa,-theory,-afp,-proof,-ML:document=afp,-annexa:outline=-annexa,afp,/proof,/ML",
show_question_marks = false, timeout = 600] show_question_marks = false, timeout = 600]
theories theories

View File

@ -44,8 +44,8 @@ theory UML_Contracts
imports UML_State imports UML_State
begin begin
text{* Modeling of an operation contract for an operation with 2 arguments, text\<open>Modeling of an operation contract for an operation with 2 arguments,
(so depending on three parameters if one takes "self" into account). *} (so depending on three parameters if one takes "self" into account).\<close>
locale contract_scheme = locale contract_scheme =
fixes f_\<upsilon> fixes f_\<upsilon>

View File

@ -56,9 +56,9 @@ imports (* Basic Type Operations *)
"collection_types/UML_Sequence" "collection_types/UML_Sequence"
begin begin
section{* Miscellaneous Stuff*} section\<open>Miscellaneous Stuff\<close>
subsection{* Definition: asBoolean *} subsection\<open>Definition: asBoolean\<close>
definition OclAsBoolean\<^sub>I\<^sub>n\<^sub>t :: "('\<AA>) Integer \<Rightarrow> ('\<AA>) Boolean" ("(_)->oclAsType\<^sub>I\<^sub>n\<^sub>t'(Boolean')") definition OclAsBoolean\<^sub>I\<^sub>n\<^sub>t :: "('\<AA>) Integer \<Rightarrow> ('\<AA>) Boolean" ("(_)->oclAsType\<^sub>I\<^sub>n\<^sub>t'(Boolean')")
where "OclAsBoolean\<^sub>I\<^sub>n\<^sub>t X = (\<lambda>\<tau>. if (\<delta> X) \<tau> = true \<tau> where "OclAsBoolean\<^sub>I\<^sub>n\<^sub>t X = (\<lambda>\<tau>. if (\<delta> X) \<tau> = true \<tau>
@ -76,7 +76,7 @@ where "OclAsBoolean\<^sub>R\<^sub>e\<^sub>a\<^sub>l X = (\<lambda>\<tau>. if
interpretation OclAsBoolean\<^sub>R\<^sub>e\<^sub>a\<^sub>l : profile_mono\<^sub>d OclAsBoolean\<^sub>R\<^sub>e\<^sub>a\<^sub>l "\<lambda>x. \<lfloor>\<lfloor>\<lceil>\<lceil>x\<rceil>\<rceil> \<noteq> 0\<rfloor>\<rfloor>" interpretation OclAsBoolean\<^sub>R\<^sub>e\<^sub>a\<^sub>l : profile_mono\<^sub>d OclAsBoolean\<^sub>R\<^sub>e\<^sub>a\<^sub>l "\<lambda>x. \<lfloor>\<lfloor>\<lceil>\<lceil>x\<rceil>\<rceil> \<noteq> 0\<rfloor>\<rfloor>"
by unfold_locales (auto simp: OclAsBoolean\<^sub>R\<^sub>e\<^sub>a\<^sub>l_def bot_option_def) by unfold_locales (auto simp: OclAsBoolean\<^sub>R\<^sub>e\<^sub>a\<^sub>l_def bot_option_def)
subsection{* Definition: asInteger *} subsection\<open>Definition: asInteger\<close>
definition OclAsInteger\<^sub>R\<^sub>e\<^sub>a\<^sub>l :: "('\<AA>) Real \<Rightarrow> ('\<AA>) Integer" ("(_)->oclAsType\<^sub>R\<^sub>e\<^sub>a\<^sub>l'(Integer')") definition OclAsInteger\<^sub>R\<^sub>e\<^sub>a\<^sub>l :: "('\<AA>) Real \<Rightarrow> ('\<AA>) Integer" ("(_)->oclAsType\<^sub>R\<^sub>e\<^sub>a\<^sub>l'(Integer')")
where "OclAsInteger\<^sub>R\<^sub>e\<^sub>a\<^sub>l X = (\<lambda>\<tau>. if (\<delta> X) \<tau> = true \<tau> where "OclAsInteger\<^sub>R\<^sub>e\<^sub>a\<^sub>l X = (\<lambda>\<tau>. if (\<delta> X) \<tau> = true \<tau>
@ -86,7 +86,7 @@ where "OclAsInteger\<^sub>R\<^sub>e\<^sub>a\<^sub>l X = (\<lambda>\<tau>. if
interpretation OclAsInteger\<^sub>R\<^sub>e\<^sub>a\<^sub>l : profile_mono\<^sub>d OclAsInteger\<^sub>R\<^sub>e\<^sub>a\<^sub>l "\<lambda>x. \<lfloor>\<lfloor>floor \<lceil>\<lceil>x\<rceil>\<rceil>\<rfloor>\<rfloor>" interpretation OclAsInteger\<^sub>R\<^sub>e\<^sub>a\<^sub>l : profile_mono\<^sub>d OclAsInteger\<^sub>R\<^sub>e\<^sub>a\<^sub>l "\<lambda>x. \<lfloor>\<lfloor>floor \<lceil>\<lceil>x\<rceil>\<rceil>\<rfloor>\<rfloor>"
by unfold_locales (auto simp: OclAsInteger\<^sub>R\<^sub>e\<^sub>a\<^sub>l_def bot_option_def) by unfold_locales (auto simp: OclAsInteger\<^sub>R\<^sub>e\<^sub>a\<^sub>l_def bot_option_def)
subsection{* Definition: asReal *} subsection\<open>Definition: asReal\<close>
definition OclAsReal\<^sub>I\<^sub>n\<^sub>t :: "('\<AA>) Integer \<Rightarrow> ('\<AA>) Real" ("(_)->oclAsType\<^sub>I\<^sub>n\<^sub>t'(Real')") definition OclAsReal\<^sub>I\<^sub>n\<^sub>t :: "('\<AA>) Integer \<Rightarrow> ('\<AA>) Real" ("(_)->oclAsType\<^sub>I\<^sub>n\<^sub>t'(Real')")
where "OclAsReal\<^sub>I\<^sub>n\<^sub>t X = (\<lambda>\<tau>. if (\<delta> X) \<tau> = true \<tau> where "OclAsReal\<^sub>I\<^sub>n\<^sub>t X = (\<lambda>\<tau>. if (\<delta> X) \<tau> = true \<tau>
@ -103,7 +103,7 @@ lemma Integer_subtype_of_Real:
apply(subst (2 4) cp_defined, simp add: true_def) apply(subst (2 4) cp_defined, simp add: true_def)
by (metis assms bot_option_def drop.elims foundation16 null_option_def) by (metis assms bot_option_def drop.elims foundation16 null_option_def)
subsection{* Definition: asPair *} subsection\<open>Definition: asPair\<close>
definition OclAsPair\<^sub>S\<^sub>e\<^sub>q :: "[('\<AA>,'\<alpha>::null)Sequence]\<Rightarrow>('\<AA>,'\<alpha>::null,'\<alpha>::null) Pair" ("(_)->asPair\<^sub>S\<^sub>e\<^sub>q'(')") definition OclAsPair\<^sub>S\<^sub>e\<^sub>q :: "[('\<AA>,'\<alpha>::null)Sequence]\<Rightarrow>('\<AA>,'\<alpha>::null,'\<alpha>::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() \<doteq> \<two> where "OclAsPair\<^sub>S\<^sub>e\<^sub>q S = (if S->size\<^sub>S\<^sub>e\<^sub>q() \<doteq> \<two>
@ -125,7 +125,7 @@ where "OclAsPair\<^sub>B\<^sub>a\<^sub>g S = (if S->size\<^sub>B\<^sub>a\<^s
else invalid else invalid
endif)" endif)"
subsection{* Definition: asSet *} subsection\<open>Definition: asSet\<close>
definition OclAsSet\<^sub>S\<^sub>e\<^sub>q :: "[('\<AA>,'\<alpha>::null)Sequence]\<Rightarrow>('\<AA>,'\<alpha>)Set" ("(_)->asSet\<^sub>S\<^sub>e\<^sub>q'(')") definition OclAsSet\<^sub>S\<^sub>e\<^sub>q :: "[('\<AA>,'\<alpha>::null)Sequence]\<Rightarrow>('\<AA>,'\<alpha>)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)))" 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)))"
@ -139,7 +139,7 @@ where "OclAsSet\<^sub>B\<^sub>a\<^sub>g S = (\<lambda> \<tau>. if (\<delta>
else if (\<upsilon> S) \<tau> = true \<tau> then null \<tau> else if (\<upsilon> S) \<tau> = true \<tau> then null \<tau>
else invalid \<tau>)" else invalid \<tau>)"
subsection{* Definition: asSequence *} subsection\<open>Definition: asSequence\<close>
definition OclAsSeq\<^sub>S\<^sub>e\<^sub>t :: "[('\<AA>,'\<alpha>::null)Set]\<Rightarrow>('\<AA>,'\<alpha>)Sequence" ("(_)->asSequence\<^sub>S\<^sub>e\<^sub>t'(')") definition OclAsSeq\<^sub>S\<^sub>e\<^sub>t :: "[('\<AA>,'\<alpha>::null)Set]\<Rightarrow>('\<AA>,'\<alpha>)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)))" 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)))"
@ -150,7 +150,7 @@ where "OclAsSeq\<^sub>B\<^sub>a\<^sub>g S = (S->iterate\<^sub>B\<^sub>a\<^su
definition OclAsSeq\<^sub>P\<^sub>a\<^sub>i\<^sub>r :: "[('\<AA>,'\<alpha>::null,'\<alpha>::null) Pair]\<Rightarrow>('\<AA>,'\<alpha>)Sequence" ("(_)->asSequence\<^sub>P\<^sub>a\<^sub>i\<^sub>r'(')") definition OclAsSeq\<^sub>P\<^sub>a\<^sub>i\<^sub>r :: "[('\<AA>,'\<alpha>::null,'\<alpha>::null) Pair]\<Rightarrow>('\<AA>,'\<alpha>)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()}" where "OclAsSeq\<^sub>P\<^sub>a\<^sub>i\<^sub>r S = Sequence{S .First(), S .Second()}"
subsection{* Definition: asBag *} subsection\<open>Definition: asBag\<close>
definition OclAsBag\<^sub>S\<^sub>e\<^sub>q :: "[('\<AA>,'\<alpha>::null)Sequence]\<Rightarrow>('\<AA>,'\<alpha>)Bag" ("(_)->asBag\<^sub>S\<^sub>e\<^sub>q'(')") definition OclAsBag\<^sub>S\<^sub>e\<^sub>q :: "[('\<AA>,'\<alpha>::null)Sequence]\<Rightarrow>('\<AA>,'\<alpha>)Bag" ("(_)->asBag\<^sub>S\<^sub>e\<^sub>q'(')")
where "OclAsBag\<^sub>S\<^sub>e\<^sub>q S = (\<lambda>\<tau>. Abs_Bag\<^sub>b\<^sub>a\<^sub>s\<^sub>e \<lfloor>\<lfloor>\<lambda>s. if list_ex ((=) s) \<lceil>\<lceil>Rep_Sequence\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \<tau>)\<rceil>\<rceil> then 1 else 0\<rfloor>\<rfloor>)" where "OclAsBag\<^sub>S\<^sub>e\<^sub>q S = (\<lambda>\<tau>. Abs_Bag\<^sub>b\<^sub>a\<^sub>s\<^sub>e \<lfloor>\<lfloor>\<lambda>s. if list_ex ((=) s) \<lceil>\<lceil>Rep_Sequence\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \<tau>)\<rceil>\<rceil> then 1 else 0\<rfloor>\<rfloor>)"
@ -165,21 +165,21 @@ oops
definition OclAsBag\<^sub>P\<^sub>a\<^sub>i\<^sub>r :: "[('\<AA>,'\<alpha>::null,'\<alpha>::null) Pair]\<Rightarrow>('\<AA>,'\<alpha>)Bag" ("(_)->asBag\<^sub>P\<^sub>a\<^sub>i\<^sub>r'(')") definition OclAsBag\<^sub>P\<^sub>a\<^sub>i\<^sub>r :: "[('\<AA>,'\<alpha>::null,'\<alpha>::null) Pair]\<Rightarrow>('\<AA>,'\<alpha>)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()}" where "OclAsBag\<^sub>P\<^sub>a\<^sub>i\<^sub>r S = Bag{S .First(), S .Second()}"
text_raw{* \isatagafp *} text_raw\<open>\isatagafp\<close>
subsection{* Collection Types *} subsection\<open>Collection Types\<close>
lemmas cp_intro'' [intro!,simp,code_unfold] = lemmas cp_intro'' [intro!,simp,code_unfold] =
cp_intro' cp_intro'
(* cp_intro''\<^sub>P\<^sub>a\<^sub>i\<^sub>r *) (* 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>t
cp_intro''\<^sub>S\<^sub>e\<^sub>q cp_intro''\<^sub>S\<^sub>e\<^sub>q
text_raw{* \endisatagafp *} text_raw\<open>\endisatagafp\<close>
subsection{* Test Statements *} subsection\<open>Test Statements\<close>
lemma syntax_test: "Set{\<two>,\<one>} = (Set{}->including\<^sub>S\<^sub>e\<^sub>t(\<one>)->including\<^sub>S\<^sub>e\<^sub>t(\<two>))" lemma syntax_test: "Set{\<two>,\<one>} = (Set{}->including\<^sub>S\<^sub>e\<^sub>t(\<one>)->including\<^sub>S\<^sub>e\<^sub>t(\<two>))"
by (rule refl) by (rule refl)
text{* Here is an example of a nested collection. *} text\<open>Here is an example of a nested collection.\<close>
lemma semantic_test2: lemma semantic_test2:
assumes H:"(Set{\<two>} \<doteq> null) = (false::('\<AA>)Boolean)" assumes H:"(Set{\<two>} \<doteq> null) = (false::('\<AA>)Boolean)"
shows "(\<tau>::('\<AA>)st) \<Turnstile> (Set{Set{\<two>},null}->includes\<^sub>S\<^sub>e\<^sub>t(null))" shows "(\<tau>::('\<AA>)st) \<Turnstile> (Set{Set{\<two>},null}->includes\<^sub>S\<^sub>e\<^sub>t(null))"
@ -207,7 +207,7 @@ done
Assert "\<tau> \<Turnstile> (\<zero> <\<^sub>i\<^sub>n\<^sub>t \<two>) and (\<zero> <\<^sub>i\<^sub>n\<^sub>t \<one>) " Assert "\<tau> \<Turnstile> (\<zero> <\<^sub>i\<^sub>n\<^sub>t \<two>) and (\<zero> <\<^sub>i\<^sub>n\<^sub>t \<one>) "
text{* Elementary computations on Sets.*} text\<open>Elementary computations on Sets.\<close>
declare OclSelect_body_def [simp] declare OclSelect_body_def [simp]
@ -248,7 +248,7 @@ Assert "\<tau> \<Turnstile> (Set{null}->reject\<^sub>S\<^sub>e\<^sub>t(x | no
lemma "const (Set{Set{\<two>,null}, invalid})" by(simp add: const_ss) lemma "const (Set{Set{\<two>,null}, invalid})" by(simp add: const_ss)
text{* Elementary computations on Sequences.*} text\<open>Elementary computations on Sequences.\<close>
Assert "\<not> (\<tau> \<Turnstile> \<upsilon>(invalid::('\<AA>,'\<alpha>::null) Sequence))" Assert "\<not> (\<tau> \<Turnstile> \<upsilon>(invalid::('\<AA>,'\<alpha>::null) Sequence))"
Assert "\<tau> \<Turnstile> \<upsilon>(null::('\<AA>,'\<alpha>::null) Sequence)" Assert "\<tau> \<Turnstile> \<upsilon>(null::('\<AA>,'\<alpha>::null) Sequence)"

View File

@ -40,15 +40,15 @@
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE. * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
******************************************************************************) ******************************************************************************)
chapter{* Formalization II: OCL Terms and Library Operations *} chapter\<open>Formalization II: OCL Terms and Library Operations\<close>
theory UML_Logic theory UML_Logic
imports UML_Types imports UML_Types
begin begin
section{* The Operations of the Boolean Type and the OCL Logic*} section\<open>The Operations of the Boolean Type and the OCL Logic\<close>
subsection{* Basic Constants *} subsection\<open>Basic Constants\<close>
lemma bot_Boolean_def : "(bot::('\<AA>)Boolean) = (\<lambda> \<tau>. \<bottom>)" lemma bot_Boolean_def : "(bot::('\<AA>)Boolean) = (\<lambda> \<tau>. \<bottom>)"
by(simp add: bot_fun_def bot_option_def) by(simp add: bot_fun_def bot_option_def)
@ -86,7 +86,7 @@ by(simp add: Sem_def true_def)
lemma textbook_false: "I\<lbrakk>false\<rbrakk> \<tau> = \<lfloor>\<lfloor>False\<rfloor>\<rfloor>" lemma textbook_false: "I\<lbrakk>false\<rbrakk> \<tau> = \<lfloor>\<lfloor>False\<rfloor>\<rfloor>"
by(simp add: Sem_def false_def) by(simp add: Sem_def false_def)
text {* text \<open>
\begin{table}[htbp] \begin{table}[htbp]
\centering \centering
\begin{tabu}{lX[,c,]} \begin{tabu}{lX[,c,]}
@ -102,12 +102,12 @@ text {*
\caption{Basic semantic constant definitions of the logic} \caption{Basic semantic constant definitions of the logic}
\label{tab:sem_basic_constants} \label{tab:sem_basic_constants}
\end{table} \end{table}
*} \<close>
subsection{* Validity and Definedness *} subsection\<open>Validity and Definedness\<close>
text{* However, this has also the consequence that core concepts like definedness, text\<open>However, this has also the consequence that core concepts like definedness,
validity and even cp have to be redefined on this type class:*} validity and even cp have to be redefined on this type class:\<close>
definition valid :: "('\<AA>,'a::null)val \<Rightarrow> ('\<AA>)Boolean" ("\<upsilon> _" [100]100) definition valid :: "('\<AA>,'a::null)val \<Rightarrow> ('\<AA>)Boolean" ("\<upsilon> _" [100]100)
where "\<upsilon> X \<equiv> \<lambda> \<tau> . if X \<tau> = bot \<tau> then false \<tau> else true \<tau>" where "\<upsilon> X \<equiv> \<lambda> \<tau> . if X \<tau> = bot \<tau> then false \<tau> else true \<tau>"
@ -124,16 +124,16 @@ lemma valid3[simp]: "\<upsilon> true = true"
lemma valid4[simp]: "\<upsilon> false = true" lemma valid4[simp]: "\<upsilon> false = true"
by(rule ext,simp add: valid_def bot_fun_def bot_option_def null_is_valid by(rule ext,simp add: valid_def bot_fun_def bot_option_def null_is_valid
null_fun_def invalid_def true_def false_def) null_fun_def invalid_def true_def false_def)
text_raw{* \isatagafp *} text_raw\<open>\isatagafp\<close>
lemma cp_valid: "(\<upsilon> X) \<tau> = (\<upsilon> (\<lambda> _. X \<tau>)) \<tau>" lemma cp_valid: "(\<upsilon> X) \<tau> = (\<upsilon> (\<lambda> _. X \<tau>)) \<tau>"
by(simp add: valid_def) by(simp add: valid_def)
text_raw{* \endisatagafp *} text_raw\<open>\endisatagafp\<close>
definition defined :: "('\<AA>,'a::null)val \<Rightarrow> ('\<AA>)Boolean" ("\<delta> _" [100]100) definition defined :: "('\<AA>,'a::null)val \<Rightarrow> ('\<AA>)Boolean" ("\<delta> _" [100]100)
where "\<delta> X \<equiv> \<lambda> \<tau> . if X \<tau> = bot \<tau> \<or> X \<tau> = null \<tau> then false \<tau> else true \<tau>" where "\<delta> X \<equiv> \<lambda> \<tau> . if X \<tau> = bot \<tau> \<or> X \<tau> = null \<tau> then false \<tau> else true \<tau>"
text{* The generalized definitions of invalid and definedness have the same text\<open>The generalized definitions of invalid and definedness have the same
properties as the old ones : *} properties as the old ones :\<close>
lemma defined1[simp]: "\<delta> invalid = false" lemma defined1[simp]: "\<delta> invalid = false"
by(rule ext,simp add: defined_def bot_fun_def bot_option_def by(rule ext,simp add: defined_def bot_fun_def bot_option_def
null_def invalid_def true_def false_def) null_def invalid_def true_def false_def)
@ -162,13 +162,13 @@ lemma valid6[simp]: "\<upsilon> \<delta> X = true"
by(rule ext, by(rule ext,
auto simp: valid_def defined_def true_def false_def auto simp: valid_def defined_def true_def false_def
bot_fun_def bot_option_def null_option_def null_fun_def) bot_fun_def bot_option_def null_option_def null_fun_def)
text_raw{* \isatagafp *} text_raw\<open>\isatagafp\<close>
lemma cp_defined:"(\<delta> X)\<tau> = (\<delta> (\<lambda> _. X \<tau>)) \<tau>" lemma cp_defined:"(\<delta> X)\<tau> = (\<delta> (\<lambda> _. X \<tau>)) \<tau>"
by(simp add: defined_def) by(simp add: defined_def)
text_raw{* \endisatagafp *} text_raw\<open>\endisatagafp\<close>
text{* The definitions above for the constants @{const defined} and @{const valid} text\<open>The definitions above for the constants @{const defined} and @{const valid}
can be rewritten into the conventional semantic "textbook" format as follows: *} can be rewritten into the conventional semantic "textbook" format as follows:\<close>
lemma textbook_defined: "I\<lbrakk>\<delta>(X)\<rbrakk> \<tau> = (if I\<lbrakk>X\<rbrakk> \<tau> = I\<lbrakk>bot\<rbrakk> \<tau> \<or> I\<lbrakk>X\<rbrakk> \<tau> = I\<lbrakk>null\<rbrakk> \<tau> lemma textbook_defined: "I\<lbrakk>\<delta>(X)\<rbrakk> \<tau> = (if I\<lbrakk>X\<rbrakk> \<tau> = I\<lbrakk>bot\<rbrakk> \<tau> \<or> I\<lbrakk>X\<rbrakk> \<tau> = I\<lbrakk>null\<rbrakk> \<tau>
then I\<lbrakk>false\<rbrakk> \<tau> then I\<lbrakk>false\<rbrakk> \<tau>
@ -181,7 +181,7 @@ lemma textbook_valid: "I\<lbrakk>\<upsilon>(X)\<rbrakk> \<tau> = (if I\<lbrakk>X
by(simp add: Sem_def valid_def) by(simp add: Sem_def valid_def)
text {* text \<open>
\autoref{tab:sem_definedness} and \autoref{tab:alglaws_definedness} \autoref{tab:sem_definedness} and \autoref{tab:alglaws_definedness}
summarize the results of this section. summarize the results of this section.
\begin{table}[htbp] \begin{table}[htbp]
@ -214,10 +214,10 @@ summarize the results of this section.
\caption{Laws of the basic predicates of the logic.} \caption{Laws of the basic predicates of the logic.}
\label{tab:alglaws_definedness} \label{tab:alglaws_definedness}
\end{table} \end{table}
*} \<close>
subsection{* The Equalities of OCL \label{sec:equality}*} subsection\<open>The Equalities of OCL \label{sec:equality}\<close>
text{* text\<open>
The OCL contains a particular version of equality, written in The OCL contains a particular version of equality, written in
Standard documents \inlineocl+_ = _+ and \inlineocl+_ <> _+ for its Standard documents \inlineocl+_ = _+ and \inlineocl+_ <> _+ for its
negation, which is referred as \emph{weak referential equality} negation, which is referred as \emph{weak referential equality}
@ -236,7 +236,7 @@ text{*
Standard. The purpose of strong equality is to define and reason Standard. The purpose of strong equality is to define and reason
over OCL. It is therefore a natural task in Featherweight OCL to over OCL. It is therefore a natural task in Featherweight OCL to
formally investigate the somewhat quite complex relationship between formally investigate the somewhat quite complex relationship between
these two. *} text{* Strong equality has two motivations: a these two.\<close> text\<open>Strong equality has two motivations: a
pragmatic one and a fundamental one. pragmatic one and a fundamental one.
\begin{enumerate} \begin{enumerate}
\item The pragmatic reason is fairly simple: users of object-oriented languages want \item The pragmatic reason is fairly simple: users of object-oriented languages want
@ -279,8 +279,8 @@ text{*
sub-expression $s$ in a term $P$ by $t$ and we have that the sub-expression $s$ in a term $P$ by $t$ and we have that the
replacement is equal to the original.'' replacement is equal to the original.''
\end{enumerate} \end{enumerate}
*} \<close>
text{* text\<open>
While weak referential equality is defined to be strict in the OCL While weak referential equality is defined to be strict in the OCL
standard, we will define strong equality as non-strict. It is quite standard, we will define strong equality as non-strict. It is quite
nasty (but not impossible) to define the logical equality in a nasty (but not impossible) to define the logical equality in a
@ -288,31 +288,31 @@ text{*
however, whenever references were used, strong equality is needed however, whenever references were used, strong equality is needed
since references refer to particular states (pre or post), and that since references refer to particular states (pre or post), and that
they mean the same thing can therefore not be taken for granted. they mean the same thing can therefore not be taken for granted.
*} \<close>
subsubsection{* Definition *} subsubsection\<open>Definition\<close>
text{* text\<open>
The strict equality on basic types (actually on all types) must be The strict equality on basic types (actually on all types) must be
exceptionally defined on @{term "null"}---otherwise the entire exceptionally defined on @{term "null"}---otherwise the entire
concept of null in the language does not make much sense. This is an concept of null in the language does not make much sense. This is an
important exception from the general rule that null important exception from the general rule that null
arguments---especially if passed as ``self''-argument---lead to arguments---especially if passed as ``self''-argument---lead to
invalid results. invalid results.
*} \<close>
text{* text\<open>
We define strong equality extremely generic, even for types that We define strong equality extremely generic, even for types that
contain a @{text "null"} or @{text "\<bottom>"} element. Strong contain a \<open>null\<close> or \<open>\<bottom>\<close> element. Strong
equality is simply polymorphic in Featherweight OCL, \ie, is equality is simply polymorphic in Featherweight OCL, \ie, is
defined identical for all types in OCL and HOL. defined identical for all types in OCL and HOL.
*} \<close>
definition StrongEq::"['\<AA> st \<Rightarrow> '\<alpha>,'\<AA> st \<Rightarrow> '\<alpha>] \<Rightarrow> ('\<AA>)Boolean" (infixl "\<triangleq>" 30) definition StrongEq::"['\<AA> st \<Rightarrow> '\<alpha>,'\<AA> st \<Rightarrow> '\<alpha>] \<Rightarrow> ('\<AA>)Boolean" (infixl "\<triangleq>" 30)
where "X \<triangleq> Y \<equiv> \<lambda> \<tau>. \<lfloor>\<lfloor>X \<tau> = Y \<tau> \<rfloor>\<rfloor>" where "X \<triangleq> Y \<equiv> \<lambda> \<tau>. \<lfloor>\<lfloor>X \<tau> = Y \<tau> \<rfloor>\<rfloor>"
text{* text\<open>
From this follow already elementary properties like: From this follow already elementary properties like:
*} \<close>
lemma [simp,code_unfold]: "(true \<triangleq> false) = false" lemma [simp,code_unfold]: "(true \<triangleq> false) = false"
by(rule ext, auto simp: StrongEq_def) by(rule ext, auto simp: StrongEq_def)
@ -320,10 +320,10 @@ lemma [simp,code_unfold]: "(false \<triangleq> true) = false"
by(rule ext, auto simp: StrongEq_def) by(rule ext, auto simp: StrongEq_def)
subsubsection{* Fundamental Predicates on Strong Equality *} subsubsection\<open>Fundamental Predicates on Strong Equality\<close>
text{* Equality reasoning in OCL is not humpty dumpty. While strong equality text\<open>Equality reasoning in OCL is not humpty dumpty. While strong equality
is clearly an equivalence: *} is clearly an equivalence:\<close>
lemma StrongEq_refl [simp]: "(X \<triangleq> X) = true" lemma StrongEq_refl [simp]: "(X \<triangleq> X) = true"
by(rule ext, simp add: null_def invalid_def true_def false_def StrongEq_def) by(rule ext, simp add: null_def invalid_def true_def false_def StrongEq_def)
@ -339,7 +339,7 @@ lemma StrongEq_trans_strong [simp]:
apply(drule_tac x=x in fun_cong)+ apply(drule_tac x=x in fun_cong)+
by auto by auto
text{* text\<open>
it is only in a limited sense a congruence, at least from the it is only in a limited sense a congruence, at least from the
point of view of this semantic theory. The point is that it is point of view of this semantic theory. The point is that it is
only a congruence on OCL expressions, not arbitrary HOL only a congruence on OCL expressions, not arbitrary HOL
@ -351,7 +351,7 @@ text{*
the sub-expressions, \ie, all sub-expressions inside an OCL the sub-expressions, \ie, all sub-expressions inside an OCL
expression refer to the same context. Expressed formally, this expression refer to the same context. Expressed formally, this
boils down to: boils down to:
*} \<close>
lemma StrongEq_subst : lemma StrongEq_subst :
assumes cp: "\<And>X. P(X)\<tau> = P(\<lambda> _. X \<tau>)\<tau>" assumes cp: "\<And>X. P(X)\<tau> = P(\<lambda> _. X \<tau>)\<tau>"
and eq: "(X \<triangleq> Y)\<tau> = true \<tau>" and eq: "(X \<triangleq> Y)\<tau> = true \<tau>"
@ -375,8 +375,8 @@ lemma valid7[simp]: "\<upsilon> (X \<triangleq> Y) = true"
lemma cp_StrongEq: "(X \<triangleq> Y) \<tau> = ((\<lambda> _. X \<tau>) \<triangleq> (\<lambda> _. Y \<tau>)) \<tau>" lemma cp_StrongEq: "(X \<triangleq> Y) \<tau> = ((\<lambda> _. X \<tau>) \<triangleq> (\<lambda> _. Y \<tau>)) \<tau>"
by(simp add: StrongEq_def) by(simp add: StrongEq_def)
subsection{* Logical Connectives and their Universal Properties *} subsection\<open>Logical Connectives and their Universal Properties\<close>
text{* text\<open>
It is a design goal to give OCL a semantics that is as closely as It is a design goal to give OCL a semantics that is as closely as
possible to a ``logical system'' in a known sense; a specification possible to a ``logical system'' in a known sense; a specification
logic where the logical connectives can not be understood other that logic where the logical connectives can not be understood other that
@ -397,7 +397,7 @@ text{*
Kleene-Logics with @{term "invalid"} as least, @{term "null"} as Kleene-Logics with @{term "invalid"} as least, @{term "null"} as
middle and @{term "true"} resp. @{term "false"} as unrelated middle and @{term "true"} resp. @{term "false"} as unrelated
top-elements. top-elements.
*} \<close>
definition OclNot :: "('\<AA>)Boolean \<Rightarrow> ('\<AA>)Boolean" ("not") definition OclNot :: "('\<AA>)Boolean \<Rightarrow> ('\<AA>)Boolean" ("not")
@ -447,16 +447,16 @@ where "X and Y \<equiv> (\<lambda> \<tau> . case X \<tau> of
| \<lfloor>\<lfloor>True\<rfloor>\<rfloor> \<Rightarrow> Y \<tau>)" | \<lfloor>\<lfloor>True\<rfloor>\<rfloor> \<Rightarrow> Y \<tau>)"
text{* text\<open>
Note that @{term "not"} is \emph{not} defined as a strict function; Note that @{term "not"} is \emph{not} defined as a strict function;
proximity to lattice laws implies that we \emph{need} a definition proximity to lattice laws implies that we \emph{need} a definition
of @{term "not"} that satisfies @{text "not(not(x))=x"}. of @{term "not"} that satisfies \<open>not(not(x))=x\<close>.
*} \<close>
text{* text\<open>
In textbook notation, the logical core constructs @{const In textbook notation, the logical core constructs @{const
"OclNot"} and @{const "OclAnd"} were represented as follows: "OclNot"} and @{const "OclAnd"} were represented as follows:
*} \<close>
lemma textbook_OclNot: lemma textbook_OclNot:
"I\<lbrakk>not(X)\<rbrakk> \<tau> = (case I\<lbrakk>X\<rbrakk> \<tau> of \<bottom> \<Rightarrow> \<bottom> "I\<lbrakk>not(X)\<rbrakk> \<tau> = (case I\<lbrakk>X\<rbrakk> \<tau> of \<bottom> \<Rightarrow> \<bottom>
| \<lfloor> \<bottom> \<rfloor> \<Rightarrow> \<lfloor> \<bottom> \<rfloor> | \<lfloor> \<bottom> \<rfloor> \<Rightarrow> \<lfloor> \<bottom> \<rfloor>
@ -685,7 +685,7 @@ lemma OclImplies_true2[simp]: "(X implies true) = true"
lemma OclImplies_false1[simp]:"(false implies X) = true" lemma OclImplies_false1[simp]:"(false implies X) = true"
by(simp add: OclImplies_def) by(simp add: OclImplies_def)
subsection{* A Standard Logical Calculus for OCL *} subsection\<open>A Standard Logical Calculus for OCL\<close>
definition OclValid :: "[('\<AA>)st, ('\<AA>)Boolean] \<Rightarrow> bool" ("(1(_)/ \<Turnstile> (_))" 50) definition OclValid :: "[('\<AA>)st, ('\<AA>)Boolean] \<Rightarrow> bool" ("(1(_)/ \<Turnstile> (_))" 50)
where "\<tau> \<Turnstile> P \<equiv> ((P \<tau>) = true \<tau>)" where "\<tau> \<Turnstile> P \<equiv> ((P \<tau>) = true \<tau>)"
@ -694,7 +694,7 @@ syntax OclNonValid :: "[('\<AA>)st, ('\<AA>)Boolean] \<Rightarrow> bool" ("(1(_
translations "\<tau> |\<noteq> P" == "\<not>(\<tau> \<Turnstile> P)" translations "\<tau> |\<noteq> P" == "\<not>(\<tau> \<Turnstile> P)"
subsubsection{* Global vs. Local Judgements*} subsubsection\<open>Global vs. Local Judgements\<close>
lemma transform1: "P = true \<Longrightarrow> \<tau> \<Turnstile> P" lemma transform1: "P = true \<Longrightarrow> \<tau> \<Turnstile> P"
by(simp add: OclValid_def) by(simp add: OclValid_def)
@ -713,9 +713,9 @@ apply(auto simp: false_def true_def defined_def bot_Boolean_def null_Boolean_def
split: option.split_asm HOL.if_split_asm) split: option.split_asm HOL.if_split_asm)
done done
text{* However, certain properties (like transitivity) can not text\<open>However, certain properties (like transitivity) can not
be \emph{transformed} from the global level to the local one, be \emph{transformed} from the global level to the local one,
they have to be re-proven on the local level. *} they have to be re-proven on the local level.\<close>
lemma (*transform3:*) lemma (*transform3:*)
assumes H : "P = true \<Longrightarrow> Q = true" assumes H : "P = true \<Longrightarrow> Q = true"
@ -725,8 +725,8 @@ apply(rule H[THEN fun_cong])
apply(rule ext) apply(rule ext)
oops oops
subsubsection{* Local Validity and Meta-logic*} subsubsection\<open>Local Validity and Meta-logic\<close>
text{* \label{sec:localVal} *} text\<open>\label{sec:localVal}\<close>
lemma foundation1[simp]: "\<tau> \<Turnstile> true" lemma foundation1[simp]: "\<tau> \<Turnstile> true"
by(auto simp: OclValid_def) by(auto simp: OclValid_def)
@ -784,12 +784,12 @@ by(simp add: OclNot_def OclValid_def true_def false_def valid_def
split: option.split option.split_asm) split: option.split option.split_asm)
text{* text\<open>
Key theorem for the $\delta$-closure: either an expression is Key theorem for the $\delta$-closure: either an expression is
defined, or it can be replaced (substituted via @{text "StrongEq_L_subst2"}; defined, or it can be replaced (substituted via \<open>StrongEq_L_subst2\<close>;
see below) by @{text invalid} or @{text null}. Strictness-reduction rules will see below) by \<open>invalid\<close> or \<open>null\<close>. Strictness-reduction rules will
usually reduce these substituted terms drastically. usually reduce these substituted terms drastically.
*} \<close>
lemma foundation8: lemma foundation8:
@ -952,7 +952,7 @@ by(simp add: OclOr_def defined_and_I defined_not_I)
lemma valid_or_I : "\<tau> \<Turnstile> \<upsilon> (x) \<Longrightarrow> \<tau> \<Turnstile> \<upsilon> (y) \<Longrightarrow> \<tau> \<Turnstile> \<upsilon> (x or y)" lemma valid_or_I : "\<tau> \<Turnstile> \<upsilon> (x) \<Longrightarrow> \<tau> \<Turnstile> \<upsilon> (y) \<Longrightarrow> \<tau> \<Turnstile> \<upsilon> (x or y)"
by(simp add: OclOr_def valid_and_I valid_not_I) by(simp add: OclOr_def valid_and_I valid_not_I)
subsubsection{* Local Judgements and Strong Equality *} subsubsection\<open>Local Judgements and Strong Equality\<close>
lemma StrongEq_L_refl: "\<tau> \<Turnstile> (x \<triangleq> x)" lemma StrongEq_L_refl: "\<tau> \<Turnstile> (x \<triangleq> x)"
by(simp add: OclValid_def StrongEq_def) by(simp add: OclValid_def StrongEq_def)
@ -966,18 +966,18 @@ by(simp add: OclValid_def StrongEq_def true_def)
text{* In order to establish substitutivity (which does not text\<open>In order to establish substitutivity (which does not
hold in general HOL formulas) we introduce the following hold in general HOL formulas) we introduce the following
predicate that allows for a calculus of the necessary side-conditions.*} predicate that allows for a calculus of the necessary side-conditions.\<close>
definition cp :: "(('\<AA>,'\<alpha>) val \<Rightarrow> ('\<AA>,'\<beta>) val) \<Rightarrow> bool" definition cp :: "(('\<AA>,'\<alpha>) val \<Rightarrow> ('\<AA>,'\<beta>) val) \<Rightarrow> bool"
where "cp P \<equiv> (\<exists> f. \<forall> X \<tau>. P X \<tau> = f (X \<tau>) \<tau>)" where "cp P \<equiv> (\<exists> f. \<forall> X \<tau>. P X \<tau> = f (X \<tau>) \<tau>)"
text{* The rule of substitutivity in Featherweight OCL holds only text\<open>The rule of substitutivity in Featherweight OCL holds only
for context-passing expressions, \ie those that pass for context-passing expressions, \ie those that pass
the context @{text "\<tau>"} without changing it. Fortunately, all the context \<open>\<tau>\<close> without changing it. Fortunately, all
operators of the OCL language satisfy this property operators of the OCL language satisfy this property
(but not all HOL operators).*} (but not all HOL operators).\<close>
lemma StrongEq_L_subst1: "\<And> \<tau>. cp P \<Longrightarrow> \<tau> \<Turnstile> (x \<triangleq> y) \<Longrightarrow> \<tau> \<Turnstile> (P x \<triangleq> P y)" lemma StrongEq_L_subst1: "\<And> \<tau>. cp P \<Longrightarrow> \<tau> \<Turnstile> (x \<triangleq> y) \<Longrightarrow> \<tau> \<Turnstile> (P x \<triangleq> P y)"
by(auto simp: OclValid_def StrongEq_def true_def cp_def) by(auto simp: OclValid_def StrongEq_def true_def cp_def)
@ -1055,7 +1055,7 @@ lemma cp_const : "cp(\<lambda>_. c)"
lemma cp_id : "cp(\<lambda>X. X)" lemma cp_id : "cp(\<lambda>X. X)"
by (simp add: cp_def, fast) by (simp add: cp_def, fast)
text_raw{* \isatagafp *} text_raw\<open>\isatagafp\<close>
lemmas cp_intro[intro!,simp,code_unfold] = lemmas cp_intro[intro!,simp,code_unfold] =
cp_const cp_const
@ -1069,10 +1069,10 @@ lemmas cp_intro[intro!,simp,code_unfold] =
cp_StrongEq[THEN allI[THEN allI[THEN allI[THEN cpI2]], cp_StrongEq[THEN allI[THEN allI[THEN allI[THEN cpI2]],
of "StrongEq"]] of "StrongEq"]]
text_raw{* \endisatagafp *} text_raw\<open>\endisatagafp\<close>
subsection{* OCL's if then else endif *} subsection\<open>OCL's if then else endif\<close>
definition OclIf :: "[('\<AA>)Boolean , ('\<AA>,'\<alpha>::null) val, ('\<AA>,'\<alpha>) val] \<Rightarrow> ('\<AA>,'\<alpha>) val" definition OclIf :: "[('\<AA>)Boolean , ('\<AA>,'\<alpha>::null) val, ('\<AA>,'\<alpha>) val] \<Rightarrow> ('\<AA>,'\<alpha>) val"
("if (_) then (_) else (_) endif" [10,10,10]50) ("if (_) then (_) else (_) endif" [10,10,10]50)
@ -1086,12 +1086,12 @@ where "(if C then B\<^sub>1 else B\<^sub>2 endif) = (\<lambda> \<tau>. if (\<del
lemma cp_OclIf:"((if C then B\<^sub>1 else B\<^sub>2 endif) \<tau> = lemma cp_OclIf:"((if C then B\<^sub>1 else B\<^sub>2 endif) \<tau> =
(if (\<lambda> _. C \<tau>) then (\<lambda> _. B\<^sub>1 \<tau>) else (\<lambda> _. B\<^sub>2 \<tau>) endif) \<tau>)" (if (\<lambda> _. C \<tau>) then (\<lambda> _. B\<^sub>1 \<tau>) else (\<lambda> _. B\<^sub>2 \<tau>) endif) \<tau>)"
by(simp only: OclIf_def, subst cp_defined, rule refl) by(simp only: OclIf_def, subst cp_defined, rule refl)
text_raw{* \isatagafp *} text_raw\<open>\isatagafp\<close>
lemmas cp_intro'[intro!,simp,code_unfold] = lemmas cp_intro'[intro!,simp,code_unfold] =
cp_intro cp_intro
cp_OclIf[THEN allI[THEN allI[THEN allI[THEN allI[THEN cpI3]]], of "OclIf"]] cp_OclIf[THEN allI[THEN allI[THEN allI[THEN allI[THEN cpI3]]], of "OclIf"]]
text_raw{* \endisatagafp *} text_raw\<open>\endisatagafp\<close>
lemma OclIf_invalid [simp]: "(if invalid then B\<^sub>1 else B\<^sub>2 endif) = invalid" lemma OclIf_invalid [simp]: "(if invalid then B\<^sub>1 else B\<^sub>2 endif) = invalid"
by(rule ext, auto simp: OclIf_def) by(rule ext, auto simp: OclIf_def)
@ -1135,9 +1135,9 @@ by simp
subsection{* Fundamental Predicates on Basic Types: Strict (Referential) Equality *} subsection\<open>Fundamental Predicates on Basic Types: Strict (Referential) Equality\<close>
text{* text\<open>
In contrast to logical equality, the OCL standard defines an equality operation In contrast to logical equality, the OCL standard defines an equality operation
which we call ``strict referential equality''. It behaves differently for all which we call ``strict referential equality''. It behaves differently for all
types---on value types, it is basically a strict version of strong equality, types---on value types, it is basically a strict version of strong equality,
@ -1145,22 +1145,22 @@ text{*
their references within the store. We introduce strict referential equality their references within the store. We introduce strict referential equality
as an \emph{overloaded} concept and will handle it for as an \emph{overloaded} concept and will handle it for
each type instance individually. each type instance individually.
*} \<close>
consts StrictRefEq :: "[('\<AA>,'a)val,('\<AA>,'a)val] \<Rightarrow> ('\<AA>)Boolean" (infixl "\<doteq>" 30) consts StrictRefEq :: "[('\<AA>,'a)val,('\<AA>,'a)val] \<Rightarrow> ('\<AA>)Boolean" (infixl "\<doteq>" 30)
text{* with {term "not"} we can express the notation:*} text\<open>with {term "not"} we can express the notation:\<close>
syntax syntax
"notequal" :: "('\<AA>)Boolean \<Rightarrow> ('\<AA>)Boolean \<Rightarrow> ('\<AA>)Boolean" (infix "<>" 40) "notequal" :: "('\<AA>)Boolean \<Rightarrow> ('\<AA>)Boolean \<Rightarrow> ('\<AA>)Boolean" (infix "<>" 40)
translations translations
"a <> b" == "CONST OclNot(a \<doteq> b)" "a <> b" == "CONST OclNot(a \<doteq> b)"
text{* We will define instances of this equality in a case-by-case basis.*} text\<open>We will define instances of this equality in a case-by-case basis.\<close>
subsection{* Laws to Establish Definedness (\texorpdfstring{$\delta$}{d}-closure) *} subsection\<open>Laws to Establish Definedness (\texorpdfstring{$\delta$}{d}-closure)\<close>
text{* For the logical connectives, we have --- beyond text\<open>For the logical connectives, we have --- beyond
@{thm foundation6} --- the following facts: *} @{thm foundation6} --- the following facts:\<close>
lemma OclNot_defargs: lemma OclNot_defargs:
"\<tau> \<Turnstile> (not P) \<Longrightarrow> \<tau> \<Turnstile> \<delta> P" "\<tau> \<Turnstile> (not P) \<Longrightarrow> \<tau> \<Turnstile> \<delta> P"
by(auto simp: OclNot_def OclValid_def true_def invalid_def defined_def false_def by(auto simp: OclNot_def OclValid_def true_def invalid_def defined_def false_def
@ -1181,7 +1181,7 @@ proof -
qed qed
subsection{* A Side-calculus for Constant Terms *} subsection\<open>A Side-calculus for Constant Terms\<close>
definition "const X \<equiv> \<forall> \<tau> \<tau>'. X \<tau> = X \<tau>'" definition "const X \<equiv> \<forall> \<tau> \<tau>'. X \<tau> = X \<tau>'"
@ -1354,7 +1354,7 @@ lemmas const_ss = const_bot const_null const_invalid const_false const_true
const_HOL_if const_HOL_and const_HOL_eq const_HOL_if const_HOL_and const_HOL_eq
text{* Miscellaneous: Overloading the syntax of ``bottom'' *} text\<open>Miscellaneous: Overloading the syntax of ``bottom''\<close>
notation bot ("\<bottom>") notation bot ("\<bottom>")

View File

@ -47,17 +47,17 @@ theory UML_PropertyProfiles
imports UML_Logic imports UML_Logic
begin begin
section{* Property Profiles for OCL Operators via Isabelle Locales *} section\<open>Property Profiles for OCL Operators via Isabelle Locales\<close>
text{* We use the Isabelle mechanism of a \emph{Locale} to generate the text\<open>We use the Isabelle mechanism of a \emph{Locale} to generate the
common lemmas for each type and operator; Locales can be seen as a common lemmas for each type and operator; Locales can be seen as a
functor that takes a local theory and generates a number of theorems. functor that takes a local theory and generates a number of theorems.
In our case, we will instantiate later these locales by the local theory In our case, we will instantiate later these locales by the local theory
of an operator definition and obtain the common rules for strictness, definedness of an operator definition and obtain the common rules for strictness, definedness
propagation, context-passingness and constance in a systematic way. propagation, context-passingness and constance in a systematic way.
*} \<close>
subsection{* Property Profiles for Monadic Operators *} subsection\<open>Property Profiles for Monadic Operators\<close>
locale profile_mono_scheme_defined = locale profile_mono_scheme_defined =
fixes f :: "('\<AA>,'\<alpha>::null)val \<Rightarrow> ('\<AA>,'\<beta>::null)val" fixes f :: "('\<AA>,'\<alpha>::null)val \<Rightarrow> ('\<AA>,'\<beta>::null)val"
@ -140,7 +140,7 @@ begin
split: if_split_asm) split: if_split_asm)
end end
subsection{* Property Profiles for Single *} subsection\<open>Property Profiles for Single\<close>
locale profile_single = locale profile_single =
fixes d:: "('\<AA>,'a::null)val \<Rightarrow> '\<AA> Boolean" fixes d:: "('\<AA>,'a::null)val \<Rightarrow> '\<AA> Boolean"
@ -148,7 +148,7 @@ locale profile_single =
assumes d_cp0: "d X \<tau> = d (\<lambda> _. X \<tau>) \<tau>" assumes d_cp0: "d X \<tau> = d (\<lambda> _. X \<tau>) \<tau>"
assumes d_const[simp,code_unfold]: "const X \<Longrightarrow> const (d X)" assumes d_const[simp,code_unfold]: "const X \<Longrightarrow> const (d X)"
subsection{* Property Profiles for Binary Operators *} subsection\<open>Property Profiles for Binary Operators\<close>
definition "bin' f g d\<^sub>x d\<^sub>y X Y = definition "bin' f g d\<^sub>x d\<^sub>y X Y =
(f X Y = (\<lambda> \<tau>. if (d\<^sub>x X) \<tau> = true \<tau> \<and> (d\<^sub>y Y) \<tau> = true \<tau> (f X Y = (\<lambda> \<tau>. if (d\<^sub>x X) \<tau> = true \<tau> \<and> (d\<^sub>y Y) \<tau> = true \<tau>
@ -223,14 +223,14 @@ begin
end end
text{* text\<open>
In our context, we will use Locales as ``Property Profiles'' for OCL operators; In our context, we will use Locales as ``Property Profiles'' for OCL operators;
if an operator @{term "f"} is of profile @{term "profile_bin_scheme defined f g"} we know if an operator @{term "f"} is of profile @{term "profile_bin_scheme defined f g"} we know
that it satisfies a number of properties like @{text "strict1"} or @{text "strict2"} that it satisfies a number of properties like \<open>strict1\<close> or \<open>strict2\<close>
\ie{} @{term "f invalid y = invalid"} and @{term "f null y = invalid"}. \ie{} @{term "f invalid y = invalid"} and @{term "f null y = invalid"}.
Since some of the more advanced Locales come with 10 - 15 theorems, property profiles Since some of the more advanced Locales come with 10 - 15 theorems, property profiles
represent a major structuring mechanism for the OCL library. represent a major structuring mechanism for the OCL library.
*} \<close>
locale profile_bin_scheme_defined = locale profile_bin_scheme_defined =

View File

@ -40,31 +40,31 @@
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE. * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
******************************************************************************) ******************************************************************************)
chapter{* Formalization III: UML/OCL constructs: State Operations and Objects *} chapter\<open>Formalization III: UML/OCL constructs: State Operations and Objects\<close>
theory UML_State theory UML_State
imports UML_Library imports UML_Library
begin begin
no_notation None ("\<bottom>") no_notation None ("\<bottom>")
section{* Introduction: States over Typed Object Universes *} section\<open>Introduction: States over Typed Object Universes\<close>
text{* \label{sec:universe} text\<open>\label{sec:universe}
In the following, we will refine the concepts of a user-defined In the following, we will refine the concepts of a user-defined
data-model (implied by a class-diagram) as well as the notion of data-model (implied by a class-diagram) as well as the notion of
$\state{}$ used in the previous section to much more detail. $\state{}$ used in the previous section to much more detail.
Surprisingly, even without a concrete notion of an objects and a Surprisingly, even without a concrete notion of an objects and a
universe of object representation, the generic infrastructure of universe of object representation, the generic infrastructure of
state-related operations is fairly rich. state-related operations is fairly rich.
*} \<close>
subsection{* Fundamental Properties on Objects: Core Referential Equality *} subsection\<open>Fundamental Properties on Objects: Core Referential Equality\<close>
subsubsection{* Definition *} subsubsection\<open>Definition\<close>
text{* Generic referential equality - to be used for instantiations text\<open>Generic referential equality - to be used for instantiations
with concrete object types ... *} with concrete object types ...\<close>
definition StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t :: "('\<AA>,'a::{object,null})val \<Rightarrow> ('\<AA>,'a)val \<Rightarrow> ('\<AA>)Boolean" definition StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t :: "('\<AA>,'a::{object,null})val \<Rightarrow> ('\<AA>,'a)val \<Rightarrow> ('\<AA>)Boolean"
where "StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t x y where "StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t x y
\<equiv> \<lambda> \<tau>. if (\<upsilon> x) \<tau> = true \<tau> \<and> (\<upsilon> y) \<tau> = true \<tau> \<equiv> \<lambda> \<tau>. if (\<upsilon> x) \<tau> = true \<tau> \<and> (\<upsilon> y) \<tau> = true \<tau>
@ -73,7 +73,7 @@ where "StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t x y
else \<lfloor>\<lfloor>(oid_of (x \<tau>)) = (oid_of (y \<tau>)) \<rfloor>\<rfloor> else \<lfloor>\<lfloor>(oid_of (x \<tau>)) = (oid_of (y \<tau>)) \<rfloor>\<rfloor>
else invalid \<tau>" else invalid \<tau>"
subsubsection{* Strictness and context passing *} subsubsection\<open>Strictness and context passing\<close>
lemma StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_strict1[simp,code_unfold] : lemma StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_strict1[simp,code_unfold] :
"(StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t x invalid) = invalid" "(StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t x invalid) = invalid"
@ -88,7 +88,7 @@ lemma cp_StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t:
"(StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t x y \<tau>) = (StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t (\<lambda>_. x \<tau>) (\<lambda>_. y \<tau>)) \<tau>" "(StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t x y \<tau>) = (StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t (\<lambda>_. x \<tau>) (\<lambda>_. y \<tau>)) \<tau>"
by(auto simp: StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def cp_valid[symmetric]) by(auto simp: StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def cp_valid[symmetric])
text_raw{* \isatagafp *} text_raw\<open>\isatagafp\<close>
lemmas cp0_StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t= cp_StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t[THEN allI[THEN allI[THEN allI[THEN cpI2]], lemmas cp0_StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t= cp_StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t[THEN allI[THEN allI[THEN allI[THEN cpI2]],
of "StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t"]] of "StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t"]]
@ -97,12 +97,12 @@ lemmas cp_intro''[intro!,simp,code_unfold] =
cp_StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t[THEN allI[THEN allI[THEN allI[THEN cpI2]], cp_StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t[THEN allI[THEN allI[THEN allI[THEN cpI2]],
of "StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t"]] of "StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t"]]
text_raw{* \endisatagafp *} text_raw\<open>\endisatagafp\<close>
subsection{* Logic and Algebraic Layer on Object *} subsection\<open>Logic and Algebraic Layer on Object\<close>
subsubsection{* Validity and Definedness Properties *} subsubsection\<open>Validity and Definedness Properties\<close>
text{* We derive the usual laws on definedness for (generic) object equality:*} text\<open>We derive the usual laws on definedness for (generic) object equality:\<close>
lemma StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_defargs: lemma StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_defargs:
"\<tau> \<Turnstile> (StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t x (y::('\<AA>,'a::{null,object})val))\<Longrightarrow> (\<tau> \<Turnstile>(\<upsilon> x)) \<and> (\<tau> \<Turnstile>(\<upsilon> y))" "\<tau> \<Turnstile> (StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t x (y::('\<AA>,'a::{null,object})val))\<Longrightarrow> (\<tau> \<Turnstile>(\<upsilon> x)) \<and> (\<tau> \<Turnstile>(\<upsilon> y))"
by(simp add: StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def OclValid_def true_def invalid_def bot_option_def by(simp add: StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def OclValid_def true_def invalid_def bot_option_def
@ -119,7 +119,7 @@ lemma StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def_homo :
"\<delta>(StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t x (y::('\<AA>,'a::{null,object})val)) = ((\<upsilon> x) and (\<upsilon> y))" "\<delta>(StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t x (y::('\<AA>,'a::{null,object})val)) = ((\<upsilon> x) and (\<upsilon> y))"
oops (* sorry *) oops (* sorry *)
subsubsection{* Symmetry *} subsubsection\<open>Symmetry\<close>
lemma StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_sym : lemma StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_sym :
assumes x_val : "\<tau> \<Turnstile> \<upsilon> x" assumes x_val : "\<tau> \<Turnstile> \<upsilon> x"
@ -128,9 +128,9 @@ by(simp add: StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def tru
x_val[simplified OclValid_def]) x_val[simplified OclValid_def])
subsubsection{* Behavior vs StrongEq *} subsubsection\<open>Behavior vs StrongEq\<close>
text{* It remains to clarify the role of the state invariant text\<open>It remains to clarify the role of the state invariant
$\inv_\sigma(\sigma)$ mentioned above that states the condition that $\inv_\sigma(\sigma)$ mentioned above that states the condition that
there is a ``one-to-one'' correspondence between object there is a ``one-to-one'' correspondence between object
representations and $\oid$'s: $\forall \mathit{oid} \in \dom\ap representations and $\oid$'s: $\forall \mathit{oid} \in \dom\ap
@ -140,33 +140,33 @@ back to \citet{richters:precise:2002}; however, we state this
condition as an invariant on states rather than a global axiom. It condition as an invariant on states rather than a global axiom. It
can, therefore, not be taken for granted that an $\oid$ makes sense can, therefore, not be taken for granted that an $\oid$ makes sense
both in pre- and post-states of OCL expressions. both in pre- and post-states of OCL expressions.
*} \<close>
text{* We capture this invariant in the predicate WFF :*} text\<open>We capture this invariant in the predicate WFF :\<close>
definition WFF :: "('\<AA>::object)st \<Rightarrow> bool" definition WFF :: "('\<AA>::object)st \<Rightarrow> bool"
where "WFF \<tau> = ((\<forall> x \<in> ran(heap(fst \<tau>)). \<lceil>heap(fst \<tau>) (oid_of x)\<rceil> = x) \<and> where "WFF \<tau> = ((\<forall> x \<in> ran(heap(fst \<tau>)). \<lceil>heap(fst \<tau>) (oid_of x)\<rceil> = x) \<and>
(\<forall> x \<in> ran(heap(snd \<tau>)). \<lceil>heap(snd \<tau>) (oid_of x)\<rceil> = x))" (\<forall> x \<in> ran(heap(snd \<tau>)). \<lceil>heap(snd \<tau>) (oid_of x)\<rceil> = x))"
text{* It turns out that WFF is a key-concept for linking strict referential equality to text\<open>It turns out that WFF is a key-concept for linking strict referential equality to
logical equality: in well-formed states (i.e. those states where the self (oid-of) field contains logical equality: in well-formed states (i.e. those states where the self (oid-of) field contains
the pointer to which the object is associated to in the state), referential equality coincides the pointer to which the object is associated to in the state), referential equality coincides
with logical equality. *} with logical equality.\<close>
text{* We turn now to the generic definition of referential equality on objects: text\<open>We turn now to the generic definition of referential equality on objects:
Equality on objects in a state is reduced to equality on the Equality on objects in a state is reduced to equality on the
references to these objects. As in HOL-OCL~\cite{brucker.ea:hol-ocl:2008,brucker.ea:hol-ocl-book:2006}, references to these objects. As in HOL-OCL~\cite{brucker.ea:hol-ocl:2008,brucker.ea:hol-ocl-book:2006},
we will store the reference of an object inside the object in a (ghost) field. we will store the reference of an object inside the object in a (ghost) field.
By establishing certain invariants (``consistent state''), it can By establishing certain invariants (``consistent state''), it can
be assured that there is a ``one-to-one-correspondence'' of objects be assured that there is a ``one-to-one-correspondence'' of objects
to their references---and therefore the definition below to their references---and therefore the definition below
behaves as we expect. *} behaves as we expect.\<close>
text{* Generic Referential Equality enjoys the usual properties: text\<open>Generic Referential Equality enjoys the usual properties:
(quasi) reflexivity, symmetry, transitivity, substitutivity for (quasi) reflexivity, symmetry, transitivity, substitutivity for
defined values. For type-technical reasons, for each concrete defined values. For type-technical reasons, for each concrete
object type, the equality @{text "\<doteq>"} is defined by generic referential object type, the equality \<open>\<doteq>\<close> is defined by generic referential
equality. *} equality.\<close>
theorem StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_vs_StrongEq: theorem StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_vs_StrongEq:
assumes WFF: "WFF \<tau>" assumes WFF: "WFF \<tau>"
@ -198,21 +198,21 @@ shows "(\<tau> \<Turnstile> (StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c
apply(auto simp: StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def OclValid_def WFF_def StrongEq_def true_def Ball_def) apply(auto simp: StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def OclValid_def WFF_def StrongEq_def true_def Ball_def)
by (metis foundation18' oid_preserve valid_x valid_y)+ by (metis foundation18' oid_preserve valid_x valid_y)+
text{* So, if two object descriptions live in the same state (both pre or post), the referential text\<open>So, if two object descriptions live in the same state (both pre or post), the referential
equality on objects implies in a WFF state the logical equality. *} equality on objects implies in a WFF state the logical equality.\<close>
section{* Operations on Object *} section\<open>Operations on Object\<close>
subsection{* Initial States (for testing and code generation) *} subsection\<open>Initial States (for testing and code generation)\<close>
definition \<tau>\<^sub>0 :: "('\<AA>)st" definition \<tau>\<^sub>0 :: "('\<AA>)st"
where "\<tau>\<^sub>0 \<equiv> (\<lparr>heap=Map.empty, assocs = Map.empty\<rparr>, where "\<tau>\<^sub>0 \<equiv> (\<lparr>heap=Map.empty, assocs = Map.empty\<rparr>,
\<lparr>heap=Map.empty, assocs = Map.empty\<rparr>)" \<lparr>heap=Map.empty, assocs = Map.empty\<rparr>)"
subsection{* OclAllInstances *} subsection\<open>OclAllInstances\<close>
text{* To denote OCL types occurring in OCL expressions syntactically---as, for example, text\<open>To denote OCL types occurring in OCL expressions syntactically---as, for example,
as ``argument'' of \inlineocl{oclAllInstances()}---we use the inverses of the injection functions into the object as ``argument'' of \inlineocl{oclAllInstances()}---we use the inverses of the injection functions into the object
universes; we show that this is a sufficient ``characterization.'' *} universes; we show that this is a sufficient ``characterization.''\<close>
definition OclAllInstances_generic :: "(('\<AA>::object) st \<Rightarrow> '\<AA> state) \<Rightarrow> ('\<AA>::object \<rightharpoonup> '\<alpha>) \<Rightarrow> definition OclAllInstances_generic :: "(('\<AA>::object) st \<Rightarrow> '\<AA> state) \<Rightarrow> ('\<AA>::object \<rightharpoonup> '\<alpha>) \<Rightarrow>
('\<AA>, '\<alpha> option option) Set" ('\<AA>, '\<alpha> option option) Set"
@ -264,7 +264,7 @@ by (metis OclAllInstances_generic_defined
A foundation16' foundation18 foundation24 foundation6) A foundation16' foundation18 foundation24 foundation6)
text{* One way to establish the actual presence of an object representation in a state is:*} text\<open>One way to establish the actual presence of an object representation in a state is:\<close>
definition "is_represented_in_state fst_snd x H \<tau> = (x \<tau> \<in> (Some o H) ` ran (heap (fst_snd \<tau>)))" definition "is_represented_in_state fst_snd x H \<tau> = (x \<tau> \<in> (Some o H) ` ran (heap (fst_snd \<tau>)))"
@ -305,12 +305,12 @@ proof -
by(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject, (simp add: bot_option_def true_def)+) by(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject, (simp add: bot_option_def true_def)+)
qed qed
text{* Here comes a couple of operational rules that allow to infer the value text\<open>Here comes a couple of operational rules that allow to infer the value
of \inlineisar+oclAllInstances+ from the context $\tau$. These rules are a special-case of \inlineisar+oclAllInstances+ from the context $\tau$. These rules are a special-case
in the sense that they are the only rules that relate statements with \emph{different} in the sense that they are the only rules that relate statements with \emph{different}
$\tau$'s. For that reason, new concepts like ``constant contexts P'' are necessary $\tau$'s. For that reason, new concepts like ``constant contexts P'' are necessary
(for which we do not elaborate an own theory for reasons of space limitations; (for which we do not elaborate an own theory for reasons of space limitations;
in examples, we will prove resulting constraints straight forward by hand). *} in examples, we will prove resulting constraints straight forward by hand).\<close>
lemma state_update_vs_allInstances_generic_including': lemma state_update_vs_allInstances_generic_including':
@ -475,7 +475,7 @@ qed
declare OclAllInstances_generic_def [simp] declare OclAllInstances_generic_def [simp]
subsubsection{* OclAllInstances (@post) *} subsubsection\<open>OclAllInstances (@post)\<close>
definition OclAllInstances_at_post :: "('\<AA> :: object \<rightharpoonup> '\<alpha>) \<Rightarrow> ('\<AA>, '\<alpha> option option) Set" definition OclAllInstances_at_post :: "('\<AA> :: object \<rightharpoonup> '\<alpha>) \<Rightarrow> ('\<AA>, '\<alpha> option option) Set"
("_ .allInstances'(')") ("_ .allInstances'(')")
@ -503,7 +503,7 @@ unfolding OclAllInstances_at_post_def
by(rule represented_generic_objects_defined[OF A[simplified OclAllInstances_at_post_def]]) by(rule represented_generic_objects_defined[OF A[simplified OclAllInstances_at_post_def]])
text{* One way to establish the actual presence of an object representation in a state is:*} text\<open>One way to establish the actual presence of an object representation in a state is:\<close>
lemma lemma
assumes A: "\<tau> \<Turnstile> H .allInstances()->includes\<^sub>S\<^sub>e\<^sub>t(x)" assumes A: "\<tau> \<Turnstile> H .allInstances()->includes\<^sub>S\<^sub>e\<^sub>t(x)"
@ -515,12 +515,12 @@ shows "(\<sigma>, \<lparr>heap=Map.empty, assocs=A\<rparr>) \<Turnstile> Type
unfolding OclAllInstances_at_post_def unfolding OclAllInstances_at_post_def
by(rule state_update_vs_allInstances_generic_empty[OF snd_conv]) by(rule state_update_vs_allInstances_generic_empty[OF snd_conv])
text{* Here comes a couple of operational rules that allow to infer the value text\<open>Here comes a couple of operational rules that allow to infer the value
of \inlineisar+oclAllInstances+ from the context $\tau$. These rules are a special-case of \inlineisar+oclAllInstances+ from the context $\tau$. These rules are a special-case
in the sense that they are the only rules that relate statements with \emph{different} in the sense that they are the only rules that relate statements with \emph{different}
$\tau$'s. For that reason, new concepts like ``constant contexts P'' are necessary $\tau$'s. For that reason, new concepts like ``constant contexts P'' are necessary
(for which we do not elaborate an own theory for reasons of space limitations; (for which we do not elaborate an own theory for reasons of space limitations;
in examples, we will prove resulting constraints straight forward by hand). *} in examples, we will prove resulting constraints straight forward by hand).\<close>
lemma state_update_vs_allInstances_at_post_including': lemma state_update_vs_allInstances_at_post_including':
@ -581,7 +581,7 @@ shows "((\<sigma>, \<lparr>heap=\<sigma>'(oid\<mapsto>Object),assocs=A\<rparr>
unfolding OclAllInstances_at_post_def unfolding OclAllInstances_at_post_def
by(rule state_update_vs_allInstances_generic_tc[OF snd_conv], insert assms) by(rule state_update_vs_allInstances_generic_tc[OF snd_conv], insert assms)
subsubsection{* OclAllInstances (@pre) *} subsubsection\<open>OclAllInstances (@pre)\<close>
definition OclAllInstances_at_pre :: "('\<AA> :: object \<rightharpoonup> '\<alpha>) \<Rightarrow> ('\<AA>, '\<alpha> option option) Set" definition OclAllInstances_at_pre :: "('\<AA> :: object \<rightharpoonup> '\<alpha>) \<Rightarrow> ('\<AA>, '\<alpha> option option) Set"
("_ .allInstances@pre'(')") ("_ .allInstances@pre'(')")
@ -608,7 +608,7 @@ unfolding OclAllInstances_at_pre_def
by(rule represented_generic_objects_defined[OF A[simplified OclAllInstances_at_pre_def]]) by(rule represented_generic_objects_defined[OF A[simplified OclAllInstances_at_pre_def]])
text{* One way to establish the actual presence of an object representation in a state is:*} text\<open>One way to establish the actual presence of an object representation in a state is:\<close>
lemma lemma
assumes A: "\<tau> \<Turnstile> H .allInstances@pre()->includes\<^sub>S\<^sub>e\<^sub>t(x)" assumes A: "\<tau> \<Turnstile> H .allInstances@pre()->includes\<^sub>S\<^sub>e\<^sub>t(x)"
@ -621,12 +621,12 @@ shows "(\<lparr>heap=Map.empty, assocs=A\<rparr>, \<sigma>) \<Turnstile> Type
unfolding OclAllInstances_at_pre_def unfolding OclAllInstances_at_pre_def
by(rule state_update_vs_allInstances_generic_empty[OF fst_conv]) by(rule state_update_vs_allInstances_generic_empty[OF fst_conv])
text{* Here comes a couple of operational rules that allow to infer the value text\<open>Here comes a couple of operational rules that allow to infer the value
of \inlineisar+oclAllInstances@pre+ from the context $\tau$. These rules are a special-case of \inlineisar+oclAllInstances@pre+ from the context $\tau$. These rules are a special-case
in the sense that they are the only rules that relate statements with \emph{different} in the sense that they are the only rules that relate statements with \emph{different}
$\tau$'s. For that reason, new concepts like ``constant contexts P'' are necessary $\tau$'s. For that reason, new concepts like ``constant contexts P'' are necessary
(for which we do not elaborate an own theory for reasons of space limitations; (for which we do not elaborate an own theory for reasons of space limitations;
in examples, we will prove resulting constraints straight forward by hand). *} in examples, we will prove resulting constraints straight forward by hand).\<close>
lemma state_update_vs_allInstances_at_pre_including': lemma state_update_vs_allInstances_at_pre_including':
@ -687,7 +687,7 @@ shows "((\<lparr>heap=\<sigma>'(oid\<mapsto>Object),assocs=A\<rparr>, \<sigma>
unfolding OclAllInstances_at_pre_def unfolding OclAllInstances_at_pre_def
by(rule state_update_vs_allInstances_generic_tc[OF fst_conv], insert assms) by(rule state_update_vs_allInstances_generic_tc[OF fst_conv], insert assms)
subsubsection{* @post or @pre *} subsubsection\<open>@post or @pre\<close>
theorem StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_vs_StrongEq'': theorem StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_vs_StrongEq'':
assumes WFF: "WFF \<tau>" assumes WFF: "WFF \<tau>"
@ -730,7 +730,7 @@ proof -
simp add: comp_def image_def, fastforce)+ simp add: comp_def image_def, fastforce)+
qed qed
subsection{* OclIsNew, OclIsDeleted, OclIsMaintained, OclIsAbsent *} subsection\<open>OclIsNew, OclIsDeleted, OclIsMaintained, OclIsAbsent\<close>
definition OclIsNew:: "('\<AA>, '\<alpha>::{null,object})val \<Rightarrow> ('\<AA>)Boolean" ("(_).oclIsNew'(')") definition OclIsNew:: "('\<AA>, '\<alpha>::{null,object})val \<Rightarrow> ('\<AA>)Boolean" ("(_).oclIsNew'(')")
where "X .oclIsNew() \<equiv> (\<lambda>\<tau> . if (\<delta> X) \<tau> = true \<tau> where "X .oclIsNew() \<equiv> (\<lambda>\<tau> . if (\<delta> X) \<tau> = true \<tau>
@ -738,9 +738,9 @@ where "X .oclIsNew() \<equiv> (\<lambda>\<tau> . if (\<delta> X) \<tau> = true \
oid_of (X \<tau>) \<in> dom(heap(snd \<tau>))\<rfloor>\<rfloor> oid_of (X \<tau>) \<in> dom(heap(snd \<tau>))\<rfloor>\<rfloor>
else invalid \<tau>)" else invalid \<tau>)"
text{* The following predicates --- which are not part of the OCL standard descriptions --- text\<open>The following predicates --- which are not part of the OCL standard descriptions ---
complete the goal of \inlineisar+oclIsNew+ by describing where an object belongs. complete the goal of \inlineisar+oclIsNew+ by describing where an object belongs.
*} \<close>
definition OclIsDeleted:: "('\<AA>, '\<alpha>::{null,object})val \<Rightarrow> ('\<AA>)Boolean" ("(_).oclIsDeleted'(')") definition OclIsDeleted:: "('\<AA>, '\<alpha>::{null,object})val \<Rightarrow> ('\<AA>)Boolean" ("(_).oclIsDeleted'(')")
where "X .oclIsDeleted() \<equiv> (\<lambda>\<tau> . if (\<delta> X) \<tau> = true \<tau> where "X .oclIsDeleted() \<equiv> (\<lambda>\<tau> . if (\<delta> X) \<tau> = true \<tau>
@ -772,17 +772,17 @@ lemma notNew_vs_others : "\<tau> \<Turnstile> \<delta> X \<Longrightarrow>
by(simp add: OclIsDeleted_def OclIsNew_def OclIsMaintained_def OclIsAbsent_def by(simp add: OclIsDeleted_def OclIsNew_def OclIsMaintained_def OclIsAbsent_def
OclNot_def OclValid_def true_def, blast) OclNot_def OclValid_def true_def, blast)
subsection{* OclIsModifiedOnly *} subsection\<open>OclIsModifiedOnly\<close>
subsubsection{* Definition *} subsubsection\<open>Definition\<close>
text{* The following predicate---which is not part of the OCL text\<open>The following predicate---which is not part of the OCL
standard---provides a simple, but powerful means to describe framing standard---provides a simple, but powerful means to describe framing
conditions. For any formal approach, be it animation of OCL contracts, conditions. For any formal approach, be it animation of OCL contracts,
test-case generation or die-hard theorem proving, the specification of test-case generation or die-hard theorem proving, the specification of
the part of a system transition that \emph{does not change} is of the part of a system transition that \emph{does not change} is of
primordial importance. The following operator establishes the equality primordial importance. The following operator establishes the equality
between old and new objects in the state (provided that they exist in between old and new objects in the state (provided that they exist in
both states), with the exception of those objects. *} both states), with the exception of those objects.\<close>
definition OclIsModifiedOnly ::"('\<AA>::object,'\<alpha>::{null,object})Set \<Rightarrow> '\<AA> Boolean" definition OclIsModifiedOnly ::"('\<AA>::object,'\<alpha>::{null,object})Set \<Rightarrow> '\<AA> Boolean"
("_->oclIsModifiedOnly'(')") ("_->oclIsModifiedOnly'(')")
@ -793,7 +793,7 @@ where "X->oclIsModifiedOnly() \<equiv> (\<lambda>(\<sigma>,\<sigma>').
then \<lfloor>\<lfloor>\<forall> x \<in> S. (heap \<sigma>) x = (heap \<sigma>') x\<rfloor>\<rfloor> then \<lfloor>\<lfloor>\<forall> x \<in> S. (heap \<sigma>) x = (heap \<sigma>') x\<rfloor>\<rfloor>
else invalid (\<sigma>,\<sigma>'))" else invalid (\<sigma>,\<sigma>'))"
subsubsection{* Execution with Invalid or Null or Null Element as Argument *} subsubsection\<open>Execution with Invalid or Null or Null Element as Argument\<close>
lemma "invalid->oclIsModifiedOnly() = invalid" lemma "invalid->oclIsModifiedOnly() = invalid"
by(simp add: OclIsModifiedOnly_def) by(simp add: OclIsModifiedOnly_def)
@ -810,16 +810,16 @@ lemma
apply(simp split: if_split_asm) apply(simp split: if_split_asm)
by(simp add: null_fun_def, blast) by(simp add: null_fun_def, blast)
subsubsection{* Context Passing *} subsubsection\<open>Context Passing\<close>
lemma cp_OclIsModifiedOnly : "X->oclIsModifiedOnly() \<tau> = (\<lambda>_. X \<tau>)->oclIsModifiedOnly() \<tau>" lemma cp_OclIsModifiedOnly : "X->oclIsModifiedOnly() \<tau> = (\<lambda>_. X \<tau>)->oclIsModifiedOnly() \<tau>"
by(simp only: OclIsModifiedOnly_def, case_tac \<tau>, simp only:, subst cp_defined, simp) by(simp only: OclIsModifiedOnly_def, case_tac \<tau>, simp only:, subst cp_defined, simp)
subsection{* OclSelf *} subsection\<open>OclSelf\<close>
text{* The following predicate---which is not part of the OCL text\<open>The following predicate---which is not part of the OCL
standard---explicitly retrieves in the pre or post state the original OCL expression standard---explicitly retrieves in the pre or post state the original OCL expression
given as argument. *} given as argument.\<close>
definition [simp]: "OclSelf x H fst_snd = (\<lambda>\<tau> . if (\<delta> x) \<tau> = true \<tau> definition [simp]: "OclSelf x H fst_snd = (\<lambda>\<tau> . if (\<delta> x) \<tau> = true \<tau>
then if oid_of (x \<tau>) \<in> dom(heap(fst \<tau>)) \<and> oid_of (x \<tau>) \<in> dom(heap (snd \<tau>)) then if oid_of (x \<tau>) \<in> dom(heap(fst \<tau>)) \<and> oid_of (x \<tau>) \<in> dom(heap (snd \<tau>))
@ -837,7 +837,7 @@ definition OclSelf_at_post :: "('\<AA>::object,'\<alpha>::{null,object})val \<Ri
('\<AA>::object,'\<alpha>::{null,object})val" ("(_)@post(_)") ('\<AA>::object,'\<alpha>::{null,object})val" ("(_)@post(_)")
where "x @post H = OclSelf x H snd" where "x @post H = OclSelf x H snd"
subsection{* Framing Theorem *} subsection\<open>Framing Theorem\<close>
lemma all_oid_diff: lemma all_oid_diff:
assumes def_x : "\<tau> \<Turnstile> \<delta> x" assumes def_x : "\<tau> \<Turnstile> \<delta> x"
@ -956,8 +956,8 @@ theorem framing:
qed(simp add: OclSelf_at_post_def OclSelf_at_pre_def OclValid_def StrongEq_def true_def)+ qed(simp add: OclSelf_at_post_def OclSelf_at_pre_def OclValid_def StrongEq_def true_def)+
text{* As corollary, the framing property can be expressed with only the strong equality text\<open>As corollary, the framing property can be expressed with only the strong equality
as comparison operator. *} as comparison operator.\<close>
theorem framing': theorem framing':
assumes wff : "WFF \<tau>" assumes wff : "WFF \<tau>"
@ -985,7 +985,7 @@ proof -
by(simp add: OclSelf_at_post_def OclSelf_at_pre_def OclValid_def StrongEq_def true_def)+ by(simp add: OclSelf_at_post_def OclSelf_at_pre_def OclValid_def StrongEq_def true_def)+
qed qed
subsection{* Miscellaneous *} subsection\<open>Miscellaneous\<close>
lemma pre_post_new: "\<tau> \<Turnstile> (x .oclIsNew()) \<Longrightarrow> \<not> (\<tau> \<Turnstile> \<upsilon>(x @pre H1)) \<and> \<not> (\<tau> \<Turnstile> \<upsilon>(x @post H2))" lemma pre_post_new: "\<tau> \<Turnstile> (x .oclIsNew()) \<Longrightarrow> \<not> (\<tau> \<Turnstile> \<upsilon>(x @pre H1)) \<and> \<not> (\<tau> \<Turnstile> \<upsilon>(x @post H2))"
by(simp add: OclIsNew_def OclSelf_at_pre_def OclSelf_at_post_def by(simp add: OclIsNew_def OclSelf_at_pre_def OclSelf_at_post_def
@ -1021,17 +1021,17 @@ by(simp add: OclIsMaintained_def OclSelf_at_pre_def OclSelf_at_post_def
lemma framing_same_state: "(\<sigma>, \<sigma>) \<Turnstile> (x @pre H \<triangleq> (x @post H))" lemma framing_same_state: "(\<sigma>, \<sigma>) \<Turnstile> (x @pre H \<triangleq> (x @post H))"
by(simp add: OclSelf_at_pre_def OclSelf_at_post_def OclValid_def StrongEq_def) by(simp add: OclSelf_at_pre_def OclSelf_at_post_def OclValid_def StrongEq_def)
section{* Accessors on Object *} section\<open>Accessors on Object\<close>
subsection{* Definition *} subsection\<open>Definition\<close>
definition "select_object mt incl smash deref l = smash (foldl incl mt (map deref l)) definition "select_object mt incl smash deref l = smash (foldl incl mt (map deref l))
\<comment> \<open>smash returns null with \<open>mt\<close> in input (in this case, object contains null pointer)\<close>" \<comment> \<open>smash returns null with \<open>mt\<close> in input (in this case, object contains null pointer)\<close>"
text{* The continuation @{text f} is usually instantiated with a smashing text\<open>The continuation \<open>f\<close> is usually instantiated with a smashing
function which is either the identity @{term id} or, for \inlineocl{0..1} cardinalities function which is either the identity @{term id} or, for \inlineocl{0..1} cardinalities
of associations, the @{term OclANY}-selector which also handles the of associations, the @{term OclANY}-selector which also handles the
@{term null}-cases appropriately. A standard use-case for this combinator @{term null}-cases appropriately. A standard use-case for this combinator
is for example: *} is for example:\<close>
term "(select_object mtSet UML_Set.OclIncluding UML_Set.OclANY f l oid )::('\<AA>, 'a::null)val" term "(select_object mtSet UML_Set.OclIncluding UML_Set.OclANY f l oid )::('\<AA>, 'a::null)val"
definition "select_object\<^sub>S\<^sub>e\<^sub>t = select_object mtSet UML_Set.OclIncluding id" definition "select_object\<^sub>S\<^sub>e\<^sub>t = select_object mtSet UML_Set.OclIncluding id"
@ -1047,7 +1047,7 @@ definition "select_object\<^sub>S\<^sub>e\<^sub>q = select_object mtSequence UML
definition "select_object_any\<^sub>S\<^sub>e\<^sub>q f s_set = UML_Sequence.OclANY (select_object\<^sub>S\<^sub>e\<^sub>q f s_set)" definition "select_object_any\<^sub>S\<^sub>e\<^sub>q f s_set = UML_Sequence.OclANY (select_object\<^sub>S\<^sub>e\<^sub>q f s_set)"
definition "select_object\<^sub>P\<^sub>a\<^sub>i\<^sub>r f1 f2 = (\<lambda>(a,b). OclPair (f1 a) (f2 b))" definition "select_object\<^sub>P\<^sub>a\<^sub>i\<^sub>r f1 f2 = (\<lambda>(a,b). OclPair (f1 a) (f2 b))"
subsection{* Validity and Definedness Properties *} subsection\<open>Validity and Definedness Properties\<close>
lemma select_fold_exec\<^sub>S\<^sub>e\<^sub>q: lemma select_fold_exec\<^sub>S\<^sub>e\<^sub>q:
assumes "list_all (\<lambda>f. (\<tau> \<Turnstile> \<upsilon> f)) l" assumes "list_all (\<lambda>f. (\<tau> \<Turnstile> \<upsilon> f)) l"

View File

@ -70,7 +70,7 @@ lemmas substs4 = StrongEq_L_subst4_rev
lemmas substs = substs1 substs2 substs4 [THEN iffD2] substs4 lemmas substs = substs1 substs2 substs4 [THEN iffD2] substs4
thm substs thm substs
ML{* ML\<open>
fun ocl_subst_asm_tac ctxt = FIRST'(map (fn C => (eresolve0_tac [C]) THEN' (simp_tac ctxt)) fun ocl_subst_asm_tac ctxt = FIRST'(map (fn C => (eresolve0_tac [C]) THEN' (simp_tac ctxt))
@{thms "substs"}) @{thms "substs"})
@ -81,7 +81,7 @@ val _ = Theory.setup
(Scan.succeed (ocl_subst_asm)) (Scan.succeed (ocl_subst_asm))
"ocl substition step") "ocl substition step")
*} \<close>
lemma test1 : "\<tau> \<Turnstile> A \<Longrightarrow> \<tau> \<Turnstile> (A and B \<triangleq> B)" lemma test1 : "\<tau> \<Turnstile> A \<Longrightarrow> \<tau> \<Turnstile> (A and B \<triangleq> B)"
apply(tactic "ocl_subst_asm_tac @{context} 1") apply(tactic "ocl_subst_asm_tac @{context} 1")

View File

@ -40,7 +40,7 @@
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE. * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
******************************************************************************) ******************************************************************************)
chapter{* Formalization I: OCL Types and Core Definitions \label{sec:focl-types}*} chapter\<open>Formalization I: OCL Types and Core Definitions \label{sec:focl-types}\<close>
theory UML_Types theory UML_Types
imports HOL.Transcendental imports HOL.Transcendental
@ -48,14 +48,14 @@ keywords "Assert" :: thy_decl
and "Assert_local" :: thy_decl and "Assert_local" :: thy_decl
begin begin
section{* Preliminaries *} section\<open>Preliminaries\<close>
subsection{* Notations for the Option Type *} subsection\<open>Notations for the Option Type\<close>
text{* text\<open>
First of all, we will use a more compact notation for the library First of all, we will use a more compact notation for the library
option type which occur all over in our definitions and which will make option type which occur all over in our definitions and which will make
the presentation more like a textbook: the presentation more like a textbook:
*} \<close>
no_notation ceiling ("\<lceil>_\<rceil>") (* For Real Numbers only ... Otherwise has unfortunate side-effects on syntax. *) no_notation ceiling ("\<lceil>_\<rceil>") (* For Real Numbers only ... Otherwise has unfortunate side-effects on syntax. *)
no_notation floor ("\<lfloor>_\<rfloor>") (* For Real Numbers only ... Otherwise has unfortunate side-effects on syntax. *) no_notation floor ("\<lfloor>_\<rfloor>") (* For Real Numbers only ... Otherwise has unfortunate side-effects on syntax. *)
@ -64,65 +64,65 @@ type_notation option ("\<langle>_\<rangle>\<^sub>\<bottom>") (* NOTE: "_\<^sub>\
notation Some ("\<lfloor>(_)\<rfloor>") notation Some ("\<lfloor>(_)\<rfloor>")
notation None ("\<bottom>") notation None ("\<bottom>")
text{* These commands introduce an alternative, more compact notation for the type constructor text\<open>These commands introduce an alternative, more compact notation for the type constructor
@{typ "'\<alpha> option"}, namely @{typ "\<langle>'\<alpha>\<rangle>\<^sub>\<bottom>"}. Furthermore, the constructors @{term "Some X"} and @{typ "'\<alpha> option"}, namely @{typ "\<langle>'\<alpha>\<rangle>\<^sub>\<bottom>"}. Furthermore, the constructors @{term "Some X"} and
@{term "None"} of the type @{typ "'\<alpha> option"}, namely @{term "\<lfloor>X\<rfloor>"} and @{term "\<bottom>"}. *} @{term "None"} of the type @{typ "'\<alpha> option"}, namely @{term "\<lfloor>X\<rfloor>"} and @{term "\<bottom>"}.\<close>
text{* text\<open>
The following function (corresponding to @{term the} in the Isabelle/HOL library) The following function (corresponding to @{term the} in the Isabelle/HOL library)
is defined as the inverse of the injection @{term Some}. is defined as the inverse of the injection @{term Some}.
*} \<close>
fun drop :: "'\<alpha> option \<Rightarrow> '\<alpha>" ("\<lceil>(_)\<rceil>") fun drop :: "'\<alpha> option \<Rightarrow> '\<alpha>" ("\<lceil>(_)\<rceil>")
where drop_lift[simp]: "\<lceil>\<lfloor>v\<rfloor>\<rceil> = v" where drop_lift[simp]: "\<lceil>\<lfloor>v\<rfloor>\<rceil> = v"
text{* The definitions for the constants and operations based on functions text\<open>The definitions for the constants and operations based on functions
will be geared towards a format that Isabelle can check to be a ``conservative'' will be geared towards a format that Isabelle can check to be a ``conservative''
(\ie, logically safe) axiomatic definition. By introducing an explicit (\ie, logically safe) axiomatic definition. By introducing an explicit
interpretation function (which happens to be defined just as the identity interpretation function (which happens to be defined just as the identity
since we are using a shallow embedding of OCL into HOL), all these definitions since we are using a shallow embedding of OCL into HOL), all these definitions
can be rewritten into the conventional semantic textbook format. can be rewritten into the conventional semantic textbook format.
To say it in other words: The interpretation function @{text Sem} as defined To say it in other words: The interpretation function \<open>Sem\<close> as defined
below is just a textual marker for presentation purposes, i.e. intended for readers below is just a textual marker for presentation purposes, i.e. intended for readers
used to conventional textbook notations on semantics. Since we use a ``shallow embedding'', used to conventional textbook notations on semantics. Since we use a ``shallow embedding'',
i.e. since we represent the syntax of OCL directly by HOL constants, the interpretation function i.e. since we represent the syntax of OCL directly by HOL constants, the interpretation function
is semantically not only superfluous, but from an Isabelle perspective strictly in is semantically not only superfluous, but from an Isabelle perspective strictly in
the way for certain consistency checks performed by the definitional packages. the way for certain consistency checks performed by the definitional packages.
*} \<close>
definition Sem :: "'a \<Rightarrow> 'a" ("I\<lbrakk>_\<rbrakk>") definition Sem :: "'a \<Rightarrow> 'a" ("I\<lbrakk>_\<rbrakk>")
where "I\<lbrakk>x\<rbrakk> \<equiv> x" where "I\<lbrakk>x\<rbrakk> \<equiv> x"
subsection{* Common Infrastructure for all OCL Types \label{sec:focl-common-types}*} subsection\<open>Common Infrastructure for all OCL Types \label{sec:focl-common-types}\<close>
text {* In order to have the possibility to nest collection types, text \<open>In order to have the possibility to nest collection types,
such that we can give semantics to expressions like @{text "Set{Set{\<two>},null}"}, such that we can give semantics to expressions like \<open>Set{Set{\<two>},null}\<close>,
it is necessary to introduce a uniform interface for types having it is necessary to introduce a uniform interface for types having
the @{text "invalid"} (= bottom) element. The reason is that we impose the \<open>invalid\<close> (= bottom) element. The reason is that we impose
a data-invariant on raw-collection \inlineisar|types_code| which assures a data-invariant on raw-collection \inlineisar|types_code| which assures
that the @{text "invalid"} element is not allowed inside the collection; that the \<open>invalid\<close> element is not allowed inside the collection;
all raw-collections of this form were identified with the @{text "invalid"} element all raw-collections of this form were identified with the \<open>invalid\<close> element
itself. The construction requires that the new collection type is itself. The construction requires that the new collection type is
not comparable with the raw-types (consisting of nested option type constructions), not comparable with the raw-types (consisting of nested option type constructions),
such that the data-invariant must be expressed in terms of the interface. such that the data-invariant must be expressed in terms of the interface.
In a second step, our base-types will be shown to be instances of this interface. In a second step, our base-types will be shown to be instances of this interface.
*} \<close>
text{* text\<open>
This uniform interface consists in a type class requiring the existence This uniform interface consists in a type class requiring the existence
of a bot and a null element. The construction proceeds by of a bot and a null element. The construction proceeds by
abstracting the null (defined by @{text "\<lfloor> \<bottom> \<rfloor>"} on abstracting the null (defined by \<open>\<lfloor> \<bottom> \<rfloor>\<close> on
@{text "'a option option"}) to a @{text null} element, which may \<open>'a option option\<close>) to a \<open>null\<close> element, which may
have an arbitrary semantic structure, and an undefinedness element @{text "\<bottom>"} have an arbitrary semantic structure, and an undefinedness element \<open>\<bottom>\<close>
to an abstract undefinedness element @{text "bot"} (also written to an abstract undefinedness element \<open>bot\<close> (also written
@{text "\<bottom>"} whenever no confusion arises). As a consequence, it is necessary \<open>\<bottom>\<close> whenever no confusion arises). As a consequence, it is necessary
to redefine the notions of invalid, defined, valuation etc. to redefine the notions of invalid, defined, valuation etc.
on top of this interface. *} on top of this interface.\<close>
text{* text\<open>
This interface consists in two abstract type classes @{text bot} This interface consists in two abstract type classes \<open>bot\<close>
and @{text null} for the class of all types comprising a bot and a and \<open>null\<close> for the class of all types comprising a bot and a
distinct null element. *} distinct null element.\<close>
class bot = class bot =
fixes bot :: "'a" fixes bot :: "'a"
@ -134,15 +134,15 @@ class null = bot +
assumes null_is_valid : "null \<noteq> bot" assumes null_is_valid : "null \<noteq> bot"
subsection{* Accommodation of Basic Types to the Abstract Interface *} subsection\<open>Accommodation of Basic Types to the Abstract Interface\<close>
text{* text\<open>
In the following it is shown that the ``option-option'' type is In the following it is shown that the ``option-option'' type is
in fact in the @{text null} class and that function spaces over these in fact in the \<open>null\<close> class and that function spaces over these
classes again ``live'' in these classes. This motivates the default construction classes again ``live'' in these classes. This motivates the default construction
of the semantic domain for the basic types (\inlineocl{Boolean}, of the semantic domain for the basic types (\inlineocl{Boolean},
\inlineocl{Integer}, \inlineocl{Real}, \ldots). \inlineocl{Integer}, \inlineocl{Real}, \ldots).
*} \<close>
instantiation option :: (type)bot instantiation option :: (type)bot
begin begin
@ -188,32 +188,32 @@ begin
qed qed
end end
text{* A trivial consequence of this adaption of the interface is that text\<open>A trivial consequence of this adaption of the interface is that
abstract and concrete versions of null are the same on base types abstract and concrete versions of null are the same on base types
(as could be expected). *} (as could be expected).\<close>
subsection{* The Common Infrastructure of Object Types (Class Types) and States. *} subsection\<open>The Common Infrastructure of Object Types (Class Types) and States.\<close>
text{* Recall that OCL is a textual extension of the UML; in particular, we use OCL as means to text\<open>Recall that OCL is a textual extension of the UML; in particular, we use OCL as means to
annotate UML class models. Thus, OCL inherits a notion of \emph{data} in the UML: UML class annotate UML class models. Thus, OCL inherits a notion of \emph{data} in the UML: UML class
models provide classes, inheritance, types of objects, and subtypes connecting them along models provide classes, inheritance, types of objects, and subtypes connecting them along
the inheritance hierarchie. the inheritance hierarchie.
*} \<close>
text{* For the moment, we formalize the most common notions of objects, in particular text\<open>For the moment, we formalize the most common notions of objects, in particular
the existance of object-identifiers (oid) for each object under which it can the existance of object-identifiers (oid) for each object under which it can
be referenced in a \emph{state}. *} be referenced in a \emph{state}.\<close>
type_synonym oid = nat type_synonym oid = nat
text{* We refrained from the alternative: text\<open>We refrained from the alternative:
\begin{isar}[mathescape] \begin{isar}[mathescape]
$\text{\textbf{type-synonym}}$ $\mathit{oid = ind}$ $\text{\textbf{type-synonym}}$ $\mathit{oid = ind}$
\end{isar} \end{isar}
which is slightly more abstract but non-executable. which is slightly more abstract but non-executable.
*} \<close>
text{* \emph{States} in UML/OCL are a pair of text\<open>\emph{States} in UML/OCL are a pair of
\begin{itemize} \begin{itemize}
\item a partial map from oid's to elements of an \emph{object universe}, \item a partial map from oid's to elements of an \emph{object universe},
\ie{} the set of all possible object representations. \ie{} the set of all possible object representations.
@ -221,32 +221,32 @@ text{* \emph{States} in UML/OCL are a pair of
objects living in a state. These relations can be n-ary which we model by nested lists. objects living in a state. These relations can be n-ary which we model by nested lists.
\end{itemize} \end{itemize}
For the moment we do not have to describe the concrete structure of the object universe and denote For the moment we do not have to describe the concrete structure of the object universe and denote
it by the polymorphic variable @{text "'\<AA>"}.*} it by the polymorphic variable \<open>'\<AA>\<close>.\<close>
record ('\<AA>)state = record ('\<AA>)state =
heap :: "oid \<rightharpoonup> '\<AA> " heap :: "oid \<rightharpoonup> '\<AA> "
assocs :: "oid \<rightharpoonup> ((oid list) list) list" assocs :: "oid \<rightharpoonup> ((oid list) list) list"
text{* In general, OCL operations are functions implicitly depending on a pair text\<open>In general, OCL operations are functions implicitly depending on a pair
of pre- and post-state, \ie{} \emph{state transitions}. Since this will be reflected in our of pre- and post-state, \ie{} \emph{state transitions}. Since this will be reflected in our
representation of OCL Types within HOL, we need to introduce the foundational concept of an representation of OCL Types within HOL, we need to introduce the foundational concept of an
object id (oid), which is just some infinite set, and some abstract notion of state. *} object id (oid), which is just some infinite set, and some abstract notion of state.\<close>
type_synonym ('\<AA>)st = "'\<AA> state \<times> '\<AA> state" type_synonym ('\<AA>)st = "'\<AA> state \<times> '\<AA> state"
text{* We will require for all objects that there is a function that text\<open>We will require for all objects that there is a function that
projects the oid of an object in the state (we will settle the question how to define projects the oid of an object in the state (we will settle the question how to define
this function later). We will use the Isabelle type class mechanism~\cite{haftmann.ea:constructive:2006} this function later). We will use the Isabelle type class mechanism~\cite{haftmann.ea:constructive:2006}
to capture this: *} to capture this:\<close>
class object = fixes oid_of :: "'a \<Rightarrow> oid" class object = fixes oid_of :: "'a \<Rightarrow> oid"
text{* Thus, if needed, we can constrain the object universe to objects by adding text\<open>Thus, if needed, we can constrain the object universe to objects by adding
the following type class constraint:*} the following type class constraint:\<close>
typ "'\<AA> :: object" typ "'\<AA> :: object"
text{* The major instance needed are instances constructed over options: once an object, text\<open>The major instance needed are instances constructed over options: once an object,
options of objects are also objects. *} options of objects are also objects.\<close>
instantiation option :: (object)object instantiation option :: (object)object
begin begin
definition oid_of_option_def: "oid_of x = oid_of (the x)" definition oid_of_option_def: "oid_of x = oid_of (the x)"
@ -254,43 +254,43 @@ begin
end end
subsection{* Common Infrastructure for all OCL Types (II): Valuations as OCL Types *} subsection\<open>Common Infrastructure for all OCL Types (II): Valuations as OCL Types\<close>
text{* Since OCL operations in general depend on pre- and post-states, we will text\<open>Since OCL operations in general depend on pre- and post-states, we will
represent OCL types as \emph{functions} from pre- and post-state to some represent OCL types as \emph{functions} from pre- and post-state to some
HOL raw-type that contains exactly the data in the OCL type --- see below. HOL raw-type that contains exactly the data in the OCL type --- see below.
This gives rise to the idea that we represent OCL types by \emph{Valuations}. This gives rise to the idea that we represent OCL types by \emph{Valuations}.
*} \<close>
text{* Valuations are functions from a state pair (built upon text\<open>Valuations are functions from a state pair (built upon
data universe @{typ "'\<AA>"}) to an arbitrary null-type (\ie, containing data universe @{typ "'\<AA>"}) to an arbitrary null-type (\ie, containing
at least a destinguished @{text "null"} and @{text "invalid"} element). *} at least a destinguished \<open>null\<close> and \<open>invalid\<close> element).\<close>
type_synonym ('\<AA>,'\<alpha>) val = "'\<AA> st \<Rightarrow> '\<alpha>::null" type_synonym ('\<AA>,'\<alpha>) val = "'\<AA> st \<Rightarrow> '\<alpha>::null"
text{* The definitions for the constants and operations based on valuations text\<open>The definitions for the constants and operations based on valuations
will be geared towards a format that Isabelle can check to be a ``conservative'' will be geared towards a format that Isabelle can check to be a ``conservative''
(\ie, logically safe) axiomatic definition. By introducing an explicit (\ie, logically safe) axiomatic definition. By introducing an explicit
interpretation function (which happens to be defined just as the identity interpretation function (which happens to be defined just as the identity
since we are using a shallow embedding of OCL into HOL), all these definitions since we are using a shallow embedding of OCL into HOL), all these definitions
can be rewritten into the conventional semantic textbook format as follows: *} can be rewritten into the conventional semantic textbook format as follows:\<close>
subsection{* The fundamental constants 'invalid' and 'null' in all OCL Types *} subsection\<open>The fundamental constants 'invalid' and 'null' in all OCL Types\<close>
text{* As a consequence of semantic domain definition, any OCL type will text\<open>As a consequence of semantic domain definition, any OCL type will
have the two semantic constants @{text "invalid"} (for exceptional, aborted have the two semantic constants \<open>invalid\<close> (for exceptional, aborted
computation) and @{text "null"}: computation) and \<open>null\<close>:
*} \<close>
definition invalid :: "('\<AA>,'\<alpha>::bot) val" definition invalid :: "('\<AA>,'\<alpha>::bot) val"
where "invalid \<equiv> \<lambda> \<tau>. bot" where "invalid \<equiv> \<lambda> \<tau>. bot"
text{* This conservative Isabelle definition of the polymorphic constant text\<open>This conservative Isabelle definition of the polymorphic constant
@{const invalid} is equivalent with the textbook definition: *} @{const invalid} is equivalent with the textbook definition:\<close>
lemma textbook_invalid: "I\<lbrakk>invalid\<rbrakk>\<tau> = bot" lemma textbook_invalid: "I\<lbrakk>invalid\<rbrakk>\<tau> = bot"
by(simp add: invalid_def Sem_def) by(simp add: invalid_def Sem_def)
text {* Note that the definition : text \<open>Note that the definition :
{\small {\small
\begin{isar}[mathescape] \begin{isar}[mathescape]
definition null :: "('$\mathfrak{A}$,'\<alpha>::null) val" definition null :: "('$\mathfrak{A}$,'\<alpha>::null) val"
@ -302,27 +302,27 @@ Thus, the polymorphic constant @{const null} is simply the result of
a general type class construction. Nevertheless, we can derive the a general type class construction. Nevertheless, we can derive the
semantic textbook definition for the OCL null constant based on the semantic textbook definition for the OCL null constant based on the
abstract null: abstract null:
*} \<close>
lemma textbook_null_fun: "I\<lbrakk>null::('\<AA>,'\<alpha>::null) val\<rbrakk> \<tau> = (null::('\<alpha>::null))" lemma textbook_null_fun: "I\<lbrakk>null::('\<AA>,'\<alpha>::null) val\<rbrakk> \<tau> = (null::('\<alpha>::null))"
by(simp add: null_fun_def Sem_def) by(simp add: null_fun_def Sem_def)
section{* Basic OCL Value Types *} section\<open>Basic OCL Value Types\<close>
text {* The structure of this section roughly follows the structure of Chapter text \<open>The structure of this section roughly follows the structure of Chapter
11 of the OCL standard~\cite{omg:ocl:2012}, which introduces the OCL 11 of the OCL standard~\cite{omg:ocl:2012}, which introduces the OCL
Library. *} Library.\<close>
text{* The semantic domain of the (basic) boolean type is now defined as the Standard: text\<open>The semantic domain of the (basic) boolean type is now defined as the Standard:
the space of valuation to @{typ "bool option option"}, \ie{} the Boolean base type:*} the space of valuation to @{typ "bool option option"}, \ie{} the Boolean base type:\<close>
type_synonym Boolean\<^sub>b\<^sub>a\<^sub>s\<^sub>e = "bool option option" type_synonym Boolean\<^sub>b\<^sub>a\<^sub>s\<^sub>e = "bool option option"
type_synonym ('\<AA>)Boolean = "('\<AA>,Boolean\<^sub>b\<^sub>a\<^sub>s\<^sub>e) val" type_synonym ('\<AA>)Boolean = "('\<AA>,Boolean\<^sub>b\<^sub>a\<^sub>s\<^sub>e) val"
text{* Because of the previous class definitions, Isabelle type-inference establishes that text\<open>Because of the previous class definitions, Isabelle type-inference establishes that
@{typ "('\<AA>)Boolean"} lives actually both in the type class @{term bot} and @{term null}; @{typ "('\<AA>)Boolean"} lives actually both in the type class @{term bot} and @{term null};
this type is sufficiently rich to contain at least these two elements. this type is sufficiently rich to contain at least these two elements.
Analogously we build: *} Analogously we build:\<close>
type_synonym Integer\<^sub>b\<^sub>a\<^sub>s\<^sub>e = "int option option" type_synonym Integer\<^sub>b\<^sub>a\<^sub>s\<^sub>e = "int option option"
type_synonym ('\<AA>)Integer = "('\<AA>,Integer\<^sub>b\<^sub>a\<^sub>s\<^sub>e) val" type_synonym ('\<AA>)Integer = "('\<AA>,Integer\<^sub>b\<^sub>a\<^sub>s\<^sub>e) val"
@ -332,19 +332,19 @@ type_synonym ('\<AA>)String = "('\<AA>,String\<^sub>b\<^sub>a\<^sub>s\<^sub>e) v
type_synonym Real\<^sub>b\<^sub>a\<^sub>s\<^sub>e = "real option option" type_synonym Real\<^sub>b\<^sub>a\<^sub>s\<^sub>e = "real option option"
type_synonym ('\<AA>)Real = "('\<AA>,Real\<^sub>b\<^sub>a\<^sub>s\<^sub>e) val" type_synonym ('\<AA>)Real = "('\<AA>,Real\<^sub>b\<^sub>a\<^sub>s\<^sub>e) val"
text{* Since @{term "Real"} is again a basic type, we define its semantic domain text\<open>Since @{term "Real"} is again a basic type, we define its semantic domain
as the valuations over @{text "real option option"} --- i.e. the mathematical type of real numbers. as the valuations over \<open>real option option\<close> --- i.e. the mathematical type of real numbers.
The HOL-theory for @{text real} ``Real'' transcendental numbers such as $\pi$ and $e$ as well as The HOL-theory for \<open>real\<close> ``Real'' transcendental numbers such as $\pi$ and $e$ as well as
infrastructure to reason over infinite convergent Cauchy-sequences (it is thus possible, in principle, infrastructure to reason over infinite convergent Cauchy-sequences (it is thus possible, in principle,
to reason in Featherweight OCL that the sum of inverted two-s exponentials is actually 2. to reason in Featherweight OCL that the sum of inverted two-s exponentials is actually 2.
If needed, a code-generator to compile @{text "Real"} to floating-point If needed, a code-generator to compile \<open>Real\<close> to floating-point
numbers can be added; this allows for mapping reals to an efficient machine representation; numbers can be added; this allows for mapping reals to an efficient machine representation;
of course, this feature would be logically unsafe.*} of course, this feature would be logically unsafe.\<close>
text{* For technical reasons related to the Isabelle type inference for type-classes text\<open>For technical reasons related to the Isabelle type inference for type-classes
(we don't get the properties in the right order that class instantiation provides them, (we don't get the properties in the right order that class instantiation provides them,
if we would follow the previous scheme), we give a slightly atypic definition:*} if we would follow the previous scheme), we give a slightly atypic definition:\<close>
typedef Void\<^sub>b\<^sub>a\<^sub>s\<^sub>e = "{X::unit option option. X = bot \<or> X = null }" by(rule_tac x="bot" in exI, simp) typedef Void\<^sub>b\<^sub>a\<^sub>s\<^sub>e = "{X::unit option option. X = bot \<or> X = null }" by(rule_tac x="bot" in exI, simp)
@ -353,41 +353,41 @@ type_synonym ('\<AA>)Void = "('\<AA>,Void\<^sub>b\<^sub>a\<^sub>s\<^sub>e) val"
section{* Some OCL Collection Types *} section\<open>Some OCL Collection Types\<close>
text{* For the semantic construction of the collection types, we have two goals: text\<open>For the semantic construction of the collection types, we have two goals:
\begin{enumerate} \begin{enumerate}
\item we want the types to be \emph{fully abstract}, \ie, the type should not \item we want the types to be \emph{fully abstract}, \ie, the type should not
contain junk-elements that are not representable by OCL expressions, and contain junk-elements that are not representable by OCL expressions, and
\item we want a possibility to nest collection types (so, we want the \item we want a possibility to nest collection types (so, we want the
potential of talking about @{text "Set(Set(Sequences(Pairs(X,Y))))"}). potential of talking about \<open>Set(Set(Sequences(Pairs(X,Y))))\<close>).
\end{enumerate} \end{enumerate}
The former principle rules out the option to define @{text "'\<alpha> Set"} just by The former principle rules out the option to define \<open>'\<alpha> Set\<close> just by
@{text "('\<AA>, ('\<alpha> option option) set) val"}. This would allow sets to contain \<open>('\<AA>, ('\<alpha> option option) set) val\<close>. This would allow sets to contain
junk elements such as @{text "{\<bottom>}"} which we need to identify with undefinedness junk elements such as \<open>{\<bottom>}\<close> which we need to identify with undefinedness
itself. Abandoning fully abstractness of rules would later on produce all sorts itself. Abandoning fully abstractness of rules would later on produce all sorts
of problems when quantifying over the elements of a type. of problems when quantifying over the elements of a type.
However, if we build an own type, then it must conform to our abstract interface However, if we build an own type, then it must conform to our abstract interface
in order to have nested types: arguments of type-constructors must conform to our in order to have nested types: arguments of type-constructors must conform to our
abstract interface, and the result type too. abstract interface, and the result type too.
*} \<close>
subsection{* The Construction of the Pair Type (Tuples) *} subsection\<open>The Construction of the Pair Type (Tuples)\<close>
text{* The core of an own type construction is done via a type text\<open>The core of an own type construction is done via a type
definition which provides the base-type @{text "('\<alpha>, '\<beta>) Pair\<^sub>b\<^sub>a\<^sub>s\<^sub>e"}. It definition which provides the base-type \<open>('\<alpha>, '\<beta>) Pair\<^sub>b\<^sub>a\<^sub>s\<^sub>e\<close>. It
is shown that this type ``fits'' indeed into the abstract type is shown that this type ``fits'' indeed into the abstract type
interface discussed in the previous section. *} interface discussed in the previous section.\<close>
typedef (overloaded) ('\<alpha>, '\<beta>) Pair\<^sub>b\<^sub>a\<^sub>s\<^sub>e = "{X::('\<alpha>::null \<times> '\<beta>::null) option option. typedef (overloaded) ('\<alpha>, '\<beta>) Pair\<^sub>b\<^sub>a\<^sub>s\<^sub>e = "{X::('\<alpha>::null \<times> '\<beta>::null) option option.
X = bot \<or> X = null \<or> (fst\<lceil>\<lceil>X\<rceil>\<rceil> \<noteq> bot \<and> snd\<lceil>\<lceil>X\<rceil>\<rceil> \<noteq> bot)}" X = bot \<or> X = null \<or> (fst\<lceil>\<lceil>X\<rceil>\<rceil> \<noteq> bot \<and> snd\<lceil>\<lceil>X\<rceil>\<rceil> \<noteq> bot)}"
by (rule_tac x="bot" in exI, simp) by (rule_tac x="bot" in exI, simp)
text{* We ``carve'' out from the concrete type @{typ "('\<alpha>::null \<times> '\<beta>::null) option option"} text\<open>We ``carve'' out from the concrete type @{typ "('\<alpha>::null \<times> '\<beta>::null) option option"}
the new fully abstract type, which will not contain representations like @{term "\<lfloor>\<lfloor>(\<bottom>,a)\<rfloor>\<rfloor>"} the new fully abstract type, which will not contain representations like @{term "\<lfloor>\<lfloor>(\<bottom>,a)\<rfloor>\<rfloor>"}
or @{term "\<lfloor>\<lfloor>(b,\<bottom>)\<rfloor>\<rfloor>"}. The type constuctor @{text "Pair{x,y}"} to be defined later will or @{term "\<lfloor>\<lfloor>(b,\<bottom>)\<rfloor>\<rfloor>"}. The type constuctor \<open>Pair{x,y}\<close> to be defined later will
identify these with @{term "invalid"}. identify these with @{term "invalid"}.
*} \<close>
instantiation Pair\<^sub>b\<^sub>a\<^sub>s\<^sub>e :: (null,null)bot instantiation Pair\<^sub>b\<^sub>a\<^sub>s\<^sub>e :: (null,null)bot
begin begin
@ -410,20 +410,20 @@ begin
end end
text{* ... and lifting this type to the format of a valuation gives us:*} text\<open>... and lifting this type to the format of a valuation gives us:\<close>
type_synonym ('\<AA>,'\<alpha>,'\<beta>) Pair = "('\<AA>, ('\<alpha>,'\<beta>) Pair\<^sub>b\<^sub>a\<^sub>s\<^sub>e) val" type_synonym ('\<AA>,'\<alpha>,'\<beta>) Pair = "('\<AA>, ('\<alpha>,'\<beta>) Pair\<^sub>b\<^sub>a\<^sub>s\<^sub>e) val"
type_notation Pair\<^sub>b\<^sub>a\<^sub>s\<^sub>e ("Pair'(_,_')") type_notation Pair\<^sub>b\<^sub>a\<^sub>s\<^sub>e ("Pair'(_,_')")
subsection{* The Construction of the Set Type *} subsection\<open>The Construction of the Set Type\<close>
text{* The core of an own type construction is done via a type text\<open>The core of an own type construction is done via a type
definition which provides the raw-type @{text "'\<alpha> Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e"}. It definition which provides the raw-type \<open>'\<alpha> Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e\<close>. It
is shown that this type ``fits'' indeed into the abstract type is shown that this type ``fits'' indeed into the abstract type
interface discussed in the previous section. Note that we make interface discussed in the previous section. Note that we make
no restriction whatsoever to \emph{finite} sets; while with no restriction whatsoever to \emph{finite} sets; while with
the standards type-constructors only finite sets can be denoted, the standards type-constructors only finite sets can be denoted,
there is the possibility to define in fact infinite there is the possibility to define in fact infinite
type constructors in \FOCL (c.f. \autoref{sec:type-extensions}). *} type constructors in \FOCL (c.f. \autoref{sec:type-extensions}).\<close>
typedef (overloaded) '\<alpha> Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e ="{X::('\<alpha>::null) set option option. X = bot \<or> X = null \<or> (\<forall>x\<in>\<lceil>\<lceil>X\<rceil>\<rceil>. x \<noteq> bot)}" typedef (overloaded) '\<alpha> Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e ="{X::('\<alpha>::null) set option option. X = bot \<or> X = null \<or> (\<forall>x\<in>\<lceil>\<lceil>X\<rceil>\<rceil>. x \<noteq> bot)}"
by (rule_tac x="bot" in exI, simp) by (rule_tac x="bot" in exI, simp)
@ -450,13 +450,13 @@ begin
qed qed
end end
text{* ... and lifting this type to the format of a valuation gives us:*} text\<open>... and lifting this type to the format of a valuation gives us:\<close>
type_synonym ('\<AA>,'\<alpha>) Set = "('\<AA>, '\<alpha> Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e) val" type_synonym ('\<AA>,'\<alpha>) Set = "('\<AA>, '\<alpha> Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e) val"
type_notation Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e ("Set'(_')") type_notation Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e ("Set'(_')")
subsection{* The Construction of the Bag Type *} subsection\<open>The Construction of the Bag Type\<close>
text{* The core of an own type construction is done via a type text\<open>The core of an own type construction is done via a type
definition which provides the raw-type @{text "'\<alpha> Bag\<^sub>b\<^sub>a\<^sub>s\<^sub>e"} definition which provides the raw-type \<open>'\<alpha> Bag\<^sub>b\<^sub>a\<^sub>s\<^sub>e\<close>
based on multi-sets from the \HOL library. As in Sets, it based on multi-sets from the \HOL library. As in Sets, it
is shown that this type ``fits'' indeed into the abstract type is shown that this type ``fits'' indeed into the abstract type
interface discussed in the previous section, and as in sets, we make interface discussed in the previous section, and as in sets, we make
@ -464,9 +464,9 @@ text{* The core of an own type construction is done via a type
the standards type-constructors only finite sets can be denoted, the standards type-constructors only finite sets can be denoted,
there is the possibility to define in fact infinite there is the possibility to define in fact infinite
type constructors in \FOCL (c.f. \autoref{sec:type-extensions}). type constructors in \FOCL (c.f. \autoref{sec:type-extensions}).
However, while several @{text null} elements are possible in a Bag, there However, while several \<open>null\<close> elements are possible in a Bag, there
can't be no bottom (invalid) element in them. can't be no bottom (invalid) element in them.
*} \<close>
typedef (overloaded) '\<alpha> Bag\<^sub>b\<^sub>a\<^sub>s\<^sub>e ="{X::('\<alpha>::null \<Rightarrow> nat) option option. X = bot \<or> X = null \<or> \<lceil>\<lceil>X\<rceil>\<rceil> bot = 0 }" typedef (overloaded) '\<alpha> Bag\<^sub>b\<^sub>a\<^sub>s\<^sub>e ="{X::('\<alpha>::null \<Rightarrow> nat) option option. X = bot \<or> X = null \<or> \<lceil>\<lceil>X\<rceil>\<rceil> bot = 0 }"
by (rule_tac x="bot" in exI, simp) by (rule_tac x="bot" in exI, simp)
@ -494,16 +494,16 @@ begin
qed qed
end end
text{* ... and lifting this type to the format of a valuation gives us:*} text\<open>... and lifting this type to the format of a valuation gives us:\<close>
type_synonym ('\<AA>,'\<alpha>) Bag = "('\<AA>, '\<alpha> Bag\<^sub>b\<^sub>a\<^sub>s\<^sub>e) val" type_synonym ('\<AA>,'\<alpha>) Bag = "('\<AA>, '\<alpha> Bag\<^sub>b\<^sub>a\<^sub>s\<^sub>e) val"
type_notation Bag\<^sub>b\<^sub>a\<^sub>s\<^sub>e ("Bag'(_')") type_notation Bag\<^sub>b\<^sub>a\<^sub>s\<^sub>e ("Bag'(_')")
subsection{* The Construction of the Sequence Type *} subsection\<open>The Construction of the Sequence Type\<close>
text{* The core of an own type construction is done via a type text\<open>The core of an own type construction is done via a type
definition which provides the base-type @{text "'\<alpha> Sequence\<^sub>b\<^sub>a\<^sub>s\<^sub>e"}. It definition which provides the base-type \<open>'\<alpha> Sequence\<^sub>b\<^sub>a\<^sub>s\<^sub>e\<close>. It
is shown that this type ``fits'' indeed into the abstract type is shown that this type ``fits'' indeed into the abstract type
interface discussed in the previous section. *} interface discussed in the previous section.\<close>
typedef (overloaded) '\<alpha> Sequence\<^sub>b\<^sub>a\<^sub>s\<^sub>e ="{X::('\<alpha>::null) list option option. typedef (overloaded) '\<alpha> Sequence\<^sub>b\<^sub>a\<^sub>s\<^sub>e ="{X::('\<alpha>::null) list option option.
X = bot \<or> X = null \<or> (\<forall>x\<in>set \<lceil>\<lceil>X\<rceil>\<rceil>. x \<noteq> bot)}" X = bot \<or> X = null \<or> (\<forall>x\<in>set \<lceil>\<lceil>X\<rceil>\<rceil>. x \<noteq> bot)}"
@ -534,12 +534,12 @@ begin
end end
text{* ... and lifting this type to the format of a valuation gives us:*} text\<open>... and lifting this type to the format of a valuation gives us:\<close>
type_synonym ('\<AA>,'\<alpha>) Sequence = "('\<AA>, '\<alpha> Sequence\<^sub>b\<^sub>a\<^sub>s\<^sub>e) val" type_synonym ('\<AA>,'\<alpha>) Sequence = "('\<AA>, '\<alpha> Sequence\<^sub>b\<^sub>a\<^sub>s\<^sub>e) val"
type_notation Sequence\<^sub>b\<^sub>a\<^sub>s\<^sub>e ("Sequence'(_')") type_notation Sequence\<^sub>b\<^sub>a\<^sub>s\<^sub>e ("Sequence'(_')")
subsection{* Discussion: The Representation of UML/OCL Types in Featherweight OCL *} subsection\<open>Discussion: The Representation of UML/OCL Types in Featherweight OCL\<close>
text{* In the introduction, we mentioned that there is an ``injective representation text\<open>In the introduction, we mentioned that there is an ``injective representation
mapping'' between the types of OCL and the types of Featherweight OCL (and its mapping'' between the types of OCL and the types of Featherweight OCL (and its
meta-language: HOL). This injectivity is at the heart of our representation technique meta-language: HOL). This injectivity is at the heart of our representation technique
--- a so-called \emph{shallow embedding} --- and means: OCL types were mapped one-to-one --- a so-called \emph{shallow embedding} --- and means: OCL types were mapped one-to-one
@ -547,9 +547,9 @@ to types in HOL, ruling out a resentation where
everything is mapped on some common HOL-type, say ``OCL-expression'', in which we everything is mapped on some common HOL-type, say ``OCL-expression'', in which we
would have to sort out the typing of OCL and its impact on the semantic representation would have to sort out the typing of OCL and its impact on the semantic representation
function in an own, quite heavy side-calculus. function in an own, quite heavy side-calculus.
*} \<close>
text{* After the previous sections, we are now able to exemplify this representation as follows: text\<open>After the previous sections, we are now able to exemplify this representation as follows:
\begin{table}[htbp] \begin{table}[htbp]
\centering \centering
@ -579,22 +579,22 @@ We do not formalize the representation map here; however, its principles are qui
\item the arguments of type constructors \inlineocl{Set(T)} remain corresponding HOL base-types. \item the arguments of type constructors \inlineocl{Set(T)} remain corresponding HOL base-types.
\end{enumerate} \end{enumerate}
*} \<close>
text{* Note, furthermore, that our construction of ``fully abstract types'' (no junk, no confusion) text\<open>Note, furthermore, that our construction of ``fully abstract types'' (no junk, no confusion)
assures that the logical equality to be defined in the next section works correctly and comes assures that the logical equality to be defined in the next section works correctly and comes
as element of the ``lingua franca'', \ie{} HOL. *} as element of the ``lingua franca'', \ie{} HOL.\<close>
(*<*) (*<*)
section{* Miscelleaneous: ML assertions *} section\<open>Miscelleaneous: ML assertions\<close>
text{* We introduce here a new command \emph{Assert} similar as \emph{value} for proving text\<open>We introduce here a new command \emph{Assert} similar as \emph{value} for proving
that the given term in argument is a true proposition. The difference with \emph{value} is that that the given term in argument is a true proposition. The difference with \emph{value} is that
\emph{Assert} fails if the normal form of the term evaluated is not equal to @{term True}. \emph{Assert} fails if the normal form of the term evaluated is not equal to @{term True}.
Moreover, in case \emph{value} could not normalize the given term, as another strategy of reduction Moreover, in case \emph{value} could not normalize the given term, as another strategy of reduction
we try to prove it with a single ``simp'' tactic. *} we try to prove it with a single ``simp'' tactic.\<close>
ML{* ML\<open>
fun disp_msg title msg status = title ^ ": '" ^ msg ^ "' " ^ status fun disp_msg title msg status = title ^ ": '" ^ msg ^ "' " ^ status
fun lemma msg specification_theorem concl in_local thy = fun lemma msg specification_theorem concl in_local thy =
@ -639,7 +639,7 @@ fun outer_syntax_command command_spec theory in_local =
val () = outer_syntax_command @{command_keyword Assert} Toplevel.theory Named_Target.theory_map val () = outer_syntax_command @{command_keyword Assert} Toplevel.theory Named_Target.theory_map
val () = outer_syntax_command @{command_keyword Assert_local} (Toplevel.local_theory NONE NONE) I val () = outer_syntax_command @{command_keyword Assert_local} (Toplevel.local_theory NONE NONE) I
*} \<close>
(*>*) (*>*)

View File

@ -45,13 +45,13 @@ imports "../UML_PropertyProfiles"
begin begin
subsection{* Fundamental Predicates on Basic Types: Strict (Referential) Equality *} subsection\<open>Fundamental Predicates on Basic Types: Strict (Referential) Equality\<close>
text{* text\<open>
Here is a first instance of a definition of strict value equality---for Here is a first instance of a definition of strict value equality---for
the special case of the type @{typ "('\<AA>)Boolean"}, it is just the special case of the type @{typ "('\<AA>)Boolean"}, it is just
the strict extension of the logical the strict extension of the logical
equality: equality:
*} \<close>
overloading StrictRefEq \<equiv> "StrictRefEq :: [('\<AA>)Boolean,('\<AA>)Boolean] \<Rightarrow> ('\<AA>)Boolean" overloading StrictRefEq \<equiv> "StrictRefEq :: [('\<AA>)Boolean,('\<AA>)Boolean] \<Rightarrow> ('\<AA>)Boolean"
begin begin
definition StrictRefEq\<^sub>B\<^sub>o\<^sub>o\<^sub>l\<^sub>e\<^sub>a\<^sub>n[code_unfold] : definition StrictRefEq\<^sub>B\<^sub>o\<^sub>o\<^sub>l\<^sub>e\<^sub>a\<^sub>n[code_unfold] :
@ -60,7 +60,7 @@ begin
else invalid \<tau>" else invalid \<tau>"
end end
text{* which implies elementary properties like: *} text\<open>which implies elementary properties like:\<close>
lemma [simp,code_unfold] : "(true \<doteq> false) = false" lemma [simp,code_unfold] : "(true \<doteq> false) = false"
by(simp add:StrictRefEq\<^sub>B\<^sub>o\<^sub>o\<^sub>l\<^sub>e\<^sub>a\<^sub>n) by(simp add:StrictRefEq\<^sub>B\<^sub>o\<^sub>o\<^sub>l\<^sub>e\<^sub>a\<^sub>n)
lemma [simp,code_unfold] : "(false \<doteq> true) = false" lemma [simp,code_unfold] : "(false \<doteq> true) = false"
@ -84,28 +84,28 @@ lemma true_non_null [simp,code_unfold]:"(true \<doteq> null) = false"
apply(rule ext, simp add: StrictRefEq\<^sub>B\<^sub>o\<^sub>o\<^sub>l\<^sub>e\<^sub>a\<^sub>n StrongEq_def false_def) apply(rule ext, simp add: StrictRefEq\<^sub>B\<^sub>o\<^sub>o\<^sub>l\<^sub>e\<^sub>a\<^sub>n StrongEq_def false_def)
by(simp add: true_def bot_option_def null_fun_def null_option_def) by(simp add: true_def bot_option_def null_fun_def null_option_def)
text{* With respect to strictness properties and miscelleaneous side-calculi, text\<open>With respect to strictness properties and miscelleaneous side-calculi,
strict referential equality behaves on booleans as described in the strict referential equality behaves on booleans as described in the
@{term "profile_bin\<^sub>S\<^sub>t\<^sub>r\<^sub>o\<^sub>n\<^sub>g\<^sub>E\<^sub>q_\<^sub>v_\<^sub>v"}:*} @{term "profile_bin\<^sub>S\<^sub>t\<^sub>r\<^sub>o\<^sub>n\<^sub>g\<^sub>E\<^sub>q_\<^sub>v_\<^sub>v"}:\<close>
interpretation StrictRefEq\<^sub>B\<^sub>o\<^sub>o\<^sub>l\<^sub>e\<^sub>a\<^sub>n : profile_bin\<^sub>S\<^sub>t\<^sub>r\<^sub>o\<^sub>n\<^sub>g\<^sub>E\<^sub>q_\<^sub>v_\<^sub>v "\<lambda> x y. (x::('\<AA>)Boolean) \<doteq> y" interpretation StrictRefEq\<^sub>B\<^sub>o\<^sub>o\<^sub>l\<^sub>e\<^sub>a\<^sub>n : profile_bin\<^sub>S\<^sub>t\<^sub>r\<^sub>o\<^sub>n\<^sub>g\<^sub>E\<^sub>q_\<^sub>v_\<^sub>v "\<lambda> x y. (x::('\<AA>)Boolean) \<doteq> y"
by unfold_locales (auto simp:StrictRefEq\<^sub>B\<^sub>o\<^sub>o\<^sub>l\<^sub>e\<^sub>a\<^sub>n) by unfold_locales (auto simp:StrictRefEq\<^sub>B\<^sub>o\<^sub>o\<^sub>l\<^sub>e\<^sub>a\<^sub>n)
text{* In particular, it is strict, cp-preserving and const-preserving. In particular, text\<open>In particular, it is strict, cp-preserving and const-preserving. In particular,
it generates the simplifier rules for terms like:*} it generates the simplifier rules for terms like:\<close>
lemma "(invalid \<doteq> false) = invalid" by(simp) lemma "(invalid \<doteq> false) = invalid" by(simp)
lemma "(invalid \<doteq> true) = invalid" by(simp) lemma "(invalid \<doteq> true) = invalid" by(simp)
lemma "(false \<doteq> invalid) = invalid" by(simp) lemma "(false \<doteq> invalid) = invalid" by(simp)
lemma "(true \<doteq> invalid) = invalid" by(simp) lemma "(true \<doteq> invalid) = invalid" by(simp)
lemma "((invalid::('\<AA>)Boolean) \<doteq> invalid) = invalid" by(simp) lemma "((invalid::('\<AA>)Boolean) \<doteq> invalid) = invalid" by(simp)
text{* Thus, the weak equality is \emph{not} reflexive. *} text\<open>Thus, the weak equality is \emph{not} reflexive.\<close>
subsection{* Test Statements on Boolean Operations. *} subsection\<open>Test Statements on Boolean Operations.\<close>
text{* Here follows a list of code-examples, that explain the meanings text\<open>Here follows a list of code-examples, that explain the meanings
of the above definitions by compilation to code and execution to @{term "True"}.*} of the above definitions by compilation to code and execution to @{term "True"}.\<close>
text{* Elementary computations on Boolean *} text\<open>Elementary computations on Boolean\<close>
Assert "\<tau> \<Turnstile> \<upsilon>(true)" Assert "\<tau> \<Turnstile> \<upsilon>(true)"
Assert "\<tau> \<Turnstile> \<delta>(false)" Assert "\<tau> \<Turnstile> \<delta>(false)"
Assert "\<tau> |\<noteq> \<delta>(null)" Assert "\<tau> |\<noteq> \<delta>(null)"

View File

@ -44,13 +44,13 @@ theory UML_Integer
imports "../UML_PropertyProfiles" imports "../UML_PropertyProfiles"
begin begin
section{* Basic Type Integer: Operations *} section\<open>Basic Type Integer: Operations\<close>
subsection{* Fundamental Predicates on Integers: Strict Equality \label{sec:integer-strict-eq}*} subsection\<open>Fundamental Predicates on Integers: Strict Equality \label{sec:integer-strict-eq}\<close>
text{* The last basic operation belonging to the fundamental infrastructure text\<open>The last basic operation belonging to the fundamental infrastructure
of a value-type in OCL is the weak equality, which is defined similar of a value-type in OCL is the weak equality, which is defined similar
to the @{typ "('\<AA>)Boolean"}-case as strict extension of the strong equality:*} to the @{typ "('\<AA>)Boolean"}-case as strict extension of the strong equality:\<close>
overloading StrictRefEq \<equiv> "StrictRefEq :: [('\<AA>)Integer,('\<AA>)Integer] \<Rightarrow> ('\<AA>)Boolean" overloading StrictRefEq \<equiv> "StrictRefEq :: [('\<AA>)Integer,('\<AA>)Integer] \<Rightarrow> ('\<AA>)Boolean"
begin begin
definition StrictRefEq\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r[code_unfold] : definition StrictRefEq\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r[code_unfold] :
@ -59,20 +59,20 @@ begin
else invalid \<tau>" else invalid \<tau>"
end 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"}*} text\<open>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"}\<close>
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 "\<lambda> x y. (x::('\<AA>)Integer) \<doteq> y" 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 "\<lambda> x y. (x::('\<AA>)Integer) \<doteq> y"
by unfold_locales (auto simp: StrictRefEq\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r) by unfold_locales (auto simp: StrictRefEq\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r)
subsection{* Basic Integer Constants *} subsection\<open>Basic Integer Constants\<close>
text{* Although the remaining part of this library reasons about text\<open>Although the remaining part of this library reasons about
integers abstractly, we provide here as example some convenient shortcuts. *} integers abstractly, we provide here as example some convenient shortcuts.\<close>
definition OclInt0 ::"('\<AA>)Integer" ("\<zero>") where "\<zero> = (\<lambda> _ . \<lfloor>\<lfloor>0::int\<rfloor>\<rfloor>)" definition OclInt0 ::"('\<AA>)Integer" ("\<zero>") where "\<zero> = (\<lambda> _ . \<lfloor>\<lfloor>0::int\<rfloor>\<rfloor>)"
definition OclInt1 ::"('\<AA>)Integer" ("\<one>") where "\<one> = (\<lambda> _ . \<lfloor>\<lfloor>1::int\<rfloor>\<rfloor>)" definition OclInt1 ::"('\<AA>)Integer" ("\<one>") where "\<one> = (\<lambda> _ . \<lfloor>\<lfloor>1::int\<rfloor>\<rfloor>)"
definition OclInt2 ::"('\<AA>)Integer" ("\<two>") where "\<two> = (\<lambda> _ . \<lfloor>\<lfloor>2::int\<rfloor>\<rfloor>)" definition OclInt2 ::"('\<AA>)Integer" ("\<two>") where "\<two> = (\<lambda> _ . \<lfloor>\<lfloor>2::int\<rfloor>\<rfloor>)"
text{* Etc. *} text\<open>Etc.\<close>
text_raw{* \isatagafp *} text_raw\<open>\isatagafp\<close>
definition OclInt3 ::"('\<AA>)Integer" ("\<three>") where "\<three> = (\<lambda> _ . \<lfloor>\<lfloor>3::int\<rfloor>\<rfloor>)" definition OclInt3 ::"('\<AA>)Integer" ("\<three>") where "\<three> = (\<lambda> _ . \<lfloor>\<lfloor>3::int\<rfloor>\<rfloor>)"
definition OclInt4 ::"('\<AA>)Integer" ("\<four>") where "\<four> = (\<lambda> _ . \<lfloor>\<lfloor>4::int\<rfloor>\<rfloor>)" definition OclInt4 ::"('\<AA>)Integer" ("\<four>") where "\<four> = (\<lambda> _ . \<lfloor>\<lfloor>4::int\<rfloor>\<rfloor>)"
definition OclInt5 ::"('\<AA>)Integer" ("\<five>") where "\<five> = (\<lambda> _ . \<lfloor>\<lfloor>5::int\<rfloor>\<rfloor>)" definition OclInt5 ::"('\<AA>)Integer" ("\<five>") where "\<five> = (\<lambda> _ . \<lfloor>\<lfloor>5::int\<rfloor>\<rfloor>)"
@ -82,7 +82,7 @@ definition OclInt8 ::"('\<AA>)Integer" ("\<eight>") where "\<eight> = (\<l
definition OclInt9 ::"('\<AA>)Integer" ("\<nine>") where "\<nine> = (\<lambda> _ . \<lfloor>\<lfloor>9::int\<rfloor>\<rfloor>)" definition OclInt9 ::"('\<AA>)Integer" ("\<nine>") where "\<nine> = (\<lambda> _ . \<lfloor>\<lfloor>9::int\<rfloor>\<rfloor>)"
definition OclInt10 ::"('\<AA>)Integer" ("\<one>\<zero>")where "\<one>\<zero> = (\<lambda> _ . \<lfloor>\<lfloor>10::int\<rfloor>\<rfloor>)" definition OclInt10 ::"('\<AA>)Integer" ("\<one>\<zero>")where "\<one>\<zero> = (\<lambda> _ . \<lfloor>\<lfloor>10::int\<rfloor>\<rfloor>)"
subsection{* Validity and Definedness Properties *} subsection\<open>Validity and Definedness Properties\<close>
lemma "\<delta>(null::('\<AA>)Integer) = false" by simp lemma "\<delta>(null::('\<AA>)Integer) = false" by simp
lemma "\<upsilon>(null::('\<AA>)Integer) = true" by simp lemma "\<upsilon>(null::('\<AA>)Integer) = true" by simp
@ -109,18 +109,18 @@ lemma [simp,code_unfold]: "\<upsilon> \<eight> = true" by(simp add:OclInt8_def)
lemma [simp,code_unfold]: "\<delta> \<nine> = true" by(simp add:OclInt9_def) lemma [simp,code_unfold]: "\<delta> \<nine> = true" by(simp add:OclInt9_def)
lemma [simp,code_unfold]: "\<upsilon> \<nine> = true" by(simp add:OclInt9_def) lemma [simp,code_unfold]: "\<upsilon> \<nine> = true" by(simp add:OclInt9_def)
text_raw{* \endisatagafp *} text_raw\<open>\endisatagafp\<close>
subsection{* Arithmetical Operations *} subsection\<open>Arithmetical Operations\<close>
subsubsection{* Definition *} subsubsection\<open>Definition\<close>
text{* Here is a common case of a built-in operation on built-in types. text\<open>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). *} Note that the arguments must be both defined (non-null, non-bot).\<close>
text{* Note that we can not follow the lexis of the OCL Standard for Isabelle text\<open>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 technical reasons; these operators are heavily overloaded in the HOL library
that a further overloading would lead to heavy technical buzz in this that a further overloading would lead to heavy technical buzz in this
document. document.
*} \<close>
definition OclAdd\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r ::"('\<AA>)Integer \<Rightarrow> ('\<AA>)Integer \<Rightarrow> ('\<AA>)Integer" (infix "+\<^sub>i\<^sub>n\<^sub>t" 40) definition OclAdd\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r ::"('\<AA>)Integer \<Rightarrow> ('\<AA>)Integer \<Rightarrow> ('\<AA>)Integer" (infix "+\<^sub>i\<^sub>n\<^sub>t" 40)
where "x +\<^sub>i\<^sub>n\<^sub>t y \<equiv> \<lambda> \<tau>. if (\<delta> x) \<tau> = true \<tau> \<and> (\<delta> y) \<tau> = true \<tau> where "x +\<^sub>i\<^sub>n\<^sub>t y \<equiv> \<lambda> \<tau>. if (\<delta> x) \<tau> = true \<tau> \<and> (\<delta> y) \<tau> = true \<tau>
then \<lfloor>\<lfloor>\<lceil>\<lceil>x \<tau>\<rceil>\<rceil> + \<lceil>\<lceil>y \<tau>\<rceil>\<rceil>\<rfloor>\<rfloor> then \<lfloor>\<lfloor>\<lceil>\<lceil>x \<tau>\<rceil>\<rceil> + \<lceil>\<lceil>y \<tau>\<rceil>\<rceil>\<rfloor>\<rfloor>
@ -144,8 +144,8 @@ where "x *\<^sub>i\<^sub>n\<^sub>t y \<equiv> \<lambda> \<tau>. if (\<delta> x)
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" "\<lambda> x y. \<lfloor>\<lfloor>\<lceil>\<lceil>x\<rceil>\<rceil> * \<lceil>\<lceil>y\<rceil>\<rceil>\<rfloor>\<rfloor>" 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" "\<lambda> x y. \<lfloor>\<lfloor>\<lceil>\<lceil>x\<rceil>\<rceil> * \<lceil>\<lceil>y\<rceil>\<rceil>\<rfloor>\<rfloor>"
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) 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 text\<open>Here is the special case of division, which is defined as invalid for division
by zero. *} by zero.\<close>
definition OclDivision\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r ::"('\<AA>)Integer \<Rightarrow> ('\<AA>)Integer \<Rightarrow> ('\<AA>)Integer" (infix "div\<^sub>i\<^sub>n\<^sub>t" 45) definition OclDivision\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r ::"('\<AA>)Integer \<Rightarrow> ('\<AA>)Integer \<Rightarrow> ('\<AA>)Integer" (infix "div\<^sub>i\<^sub>n\<^sub>t" 45)
where "x div\<^sub>i\<^sub>n\<^sub>t y \<equiv> \<lambda> \<tau>. if (\<delta> x) \<tau> = true \<tau> \<and> (\<delta> y) \<tau> = true \<tau> where "x div\<^sub>i\<^sub>n\<^sub>t y \<equiv> \<lambda> \<tau>. if (\<delta> x) \<tau> = true \<tau> \<and> (\<delta> y) \<tau> = true \<tau>
then if y \<tau> \<noteq> OclInt0 \<tau> then \<lfloor>\<lfloor>\<lceil>\<lceil>x \<tau>\<rceil>\<rceil> div \<lceil>\<lceil>y \<tau>\<rceil>\<rceil>\<rfloor>\<rfloor> else invalid \<tau> then if y \<tau> \<noteq> OclInt0 \<tau> then \<lfloor>\<lfloor>\<lceil>\<lceil>x \<tau>\<rceil>\<rceil> div \<lceil>\<lceil>y \<tau>\<rceil>\<rceil>\<rfloor>\<rfloor> else invalid \<tau>
@ -173,14 +173,14 @@ where "x \<le>\<^sub>i\<^sub>n\<^sub>t y \<equiv> \<lambda> \<tau>. if (\<delta>
interpretation OclLe\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r : profile_bin\<^sub>d_\<^sub>d "(\<le>\<^sub>i\<^sub>n\<^sub>t)" "\<lambda> x y. \<lfloor>\<lfloor>\<lceil>\<lceil>x\<rceil>\<rceil> \<le> \<lceil>\<lceil>y\<rceil>\<rceil>\<rfloor>\<rfloor>" interpretation OclLe\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r : profile_bin\<^sub>d_\<^sub>d "(\<le>\<^sub>i\<^sub>n\<^sub>t)" "\<lambda> x y. \<lfloor>\<lfloor>\<lceil>\<lceil>x\<rceil>\<rceil> \<le> \<lceil>\<lceil>y\<rceil>\<rceil>\<rfloor>\<rfloor>"
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) 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 *} subsubsection\<open>Basic Properties\<close>
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)" 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 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 split: option.split option.split_asm
bool.split bool.split_asm) bool.split bool.split_asm)
subsubsection{* Execution with Invalid or Null or Zero as Argument *} subsubsection\<open>Execution with Invalid or Null or Zero as Argument\<close>
lemma OclAdd\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r_zero1[simp,code_unfold] : 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 \<zero>) = (if \<upsilon> x and not (\<delta> x) then invalid else x endif)" "(x +\<^sub>i\<^sub>n\<^sub>t \<zero>) = (if \<upsilon> x and not (\<delta> x) then invalid else x endif)"
@ -214,9 +214,9 @@ by(subst OclAdd\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r_commute,
subsection{* Test Statements *} subsection\<open>Test Statements\<close>
text{* Here follows a list of code-examples, that explain the meanings text\<open>Here follows a list of code-examples, that explain the meanings
of the above definitions by compilation to code and execution to @{term "True"}.*} of the above definitions by compilation to code and execution to @{term "True"}.\<close>
Assert "\<tau> \<Turnstile> ( \<nine> \<le>\<^sub>i\<^sub>n\<^sub>t \<one>\<zero> )" Assert "\<tau> \<Turnstile> ( \<nine> \<le>\<^sub>i\<^sub>n\<^sub>t \<one>\<zero> )"
Assert "\<tau> \<Turnstile> (( \<four> +\<^sub>i\<^sub>n\<^sub>t \<four> ) \<le>\<^sub>i\<^sub>n\<^sub>t \<one>\<zero> )" Assert "\<tau> \<Turnstile> (( \<four> +\<^sub>i\<^sub>n\<^sub>t \<four> ) \<le>\<^sub>i\<^sub>n\<^sub>t \<one>\<zero> )"
@ -250,11 +250,11 @@ lemma OclInt9_non_null [simp,code_unfold]: "(\<nine> \<doteq> null) = false" by(
lemma null_non_OclInt9 [simp,code_unfold]: "(null \<doteq> \<nine>) = false" by(simp add: OclInt9_def) lemma null_non_OclInt9 [simp,code_unfold]: "(null \<doteq> \<nine>) = false" by(simp add: OclInt9_def)
text{* Here follows a list of code-examples, that explain the meanings text\<open>Here follows a list of code-examples, that explain the meanings
of the above definitions by compilation to code and execution to @{term "True"}.*} of the above definitions by compilation to code and execution to @{term "True"}.\<close>
text{* Elementary computations on Integer *} text\<open>Elementary computations on Integer\<close>
Assert "\<tau> \<Turnstile> ((\<zero> <\<^sub>i\<^sub>n\<^sub>t \<two>) and (\<zero> <\<^sub>i\<^sub>n\<^sub>t \<one>))" Assert "\<tau> \<Turnstile> ((\<zero> <\<^sub>i\<^sub>n\<^sub>t \<two>) and (\<zero> <\<^sub>i\<^sub>n\<^sub>t \<one>))"

View File

@ -44,13 +44,13 @@ theory UML_Real
imports "../UML_PropertyProfiles" imports "../UML_PropertyProfiles"
begin begin
section{* Basic Type Real: Operations *} section\<open>Basic Type Real: Operations\<close>
subsection{* Fundamental Predicates on Reals: Strict Equality \label{sec:real-strict-eq}*} subsection\<open>Fundamental Predicates on Reals: Strict Equality \label{sec:real-strict-eq}\<close>
text{* The last basic operation belonging to the fundamental infrastructure text\<open>The last basic operation belonging to the fundamental infrastructure
of a value-type in OCL is the weak equality, which is defined similar of a value-type in OCL is the weak equality, which is defined similar
to the @{typ "('\<AA>)Boolean"}-case as strict extension of the strong equality:*} to the @{typ "('\<AA>)Boolean"}-case as strict extension of the strong equality:\<close>
overloading StrictRefEq \<equiv> "StrictRefEq :: [('\<AA>)Real,('\<AA>)Real] \<Rightarrow> ('\<AA>)Boolean" overloading StrictRefEq \<equiv> "StrictRefEq :: [('\<AA>)Real,('\<AA>)Real] \<Rightarrow> ('\<AA>)Boolean"
begin begin
definition StrictRefEq\<^sub>R\<^sub>e\<^sub>a\<^sub>l [code_unfold] : definition StrictRefEq\<^sub>R\<^sub>e\<^sub>a\<^sub>l [code_unfold] :
@ -59,20 +59,20 @@ begin
else invalid \<tau>" else invalid \<tau>"
end 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"}*} text\<open>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"}\<close>
interpretation StrictRefEq\<^sub>R\<^sub>e\<^sub>a\<^sub>l : profile_bin\<^sub>S\<^sub>t\<^sub>r\<^sub>o\<^sub>n\<^sub>g\<^sub>E\<^sub>q_\<^sub>v_\<^sub>v "\<lambda> x y. (x::('\<AA>)Real) \<doteq> y" interpretation StrictRefEq\<^sub>R\<^sub>e\<^sub>a\<^sub>l : profile_bin\<^sub>S\<^sub>t\<^sub>r\<^sub>o\<^sub>n\<^sub>g\<^sub>E\<^sub>q_\<^sub>v_\<^sub>v "\<lambda> x y. (x::('\<AA>)Real) \<doteq> y"
by unfold_locales (auto simp: StrictRefEq\<^sub>R\<^sub>e\<^sub>a\<^sub>l) by unfold_locales (auto simp: StrictRefEq\<^sub>R\<^sub>e\<^sub>a\<^sub>l)
subsection{* Basic Real Constants *} subsection\<open>Basic Real Constants\<close>
text{* Although the remaining part of this library reasons about text\<open>Although the remaining part of this library reasons about
reals abstractly, we provide here as example some convenient shortcuts. *} reals abstractly, we provide here as example some convenient shortcuts.\<close>
definition OclReal0 ::"('\<AA>)Real" ("\<zero>.\<zero>") where "\<zero>.\<zero> = (\<lambda> _ . \<lfloor>\<lfloor>0::real\<rfloor>\<rfloor>)" definition OclReal0 ::"('\<AA>)Real" ("\<zero>.\<zero>") where "\<zero>.\<zero> = (\<lambda> _ . \<lfloor>\<lfloor>0::real\<rfloor>\<rfloor>)"
definition OclReal1 ::"('\<AA>)Real" ("\<one>.\<zero>") where "\<one>.\<zero> = (\<lambda> _ . \<lfloor>\<lfloor>1::real\<rfloor>\<rfloor>)" definition OclReal1 ::"('\<AA>)Real" ("\<one>.\<zero>") where "\<one>.\<zero> = (\<lambda> _ . \<lfloor>\<lfloor>1::real\<rfloor>\<rfloor>)"
definition OclReal2 ::"('\<AA>)Real" ("\<two>.\<zero>") where "\<two>.\<zero> = (\<lambda> _ . \<lfloor>\<lfloor>2::real\<rfloor>\<rfloor>)" definition OclReal2 ::"('\<AA>)Real" ("\<two>.\<zero>") where "\<two>.\<zero> = (\<lambda> _ . \<lfloor>\<lfloor>2::real\<rfloor>\<rfloor>)"
text{* Etc. *} text\<open>Etc.\<close>
text_raw{* \isatagafp *} text_raw\<open>\isatagafp\<close>
definition OclReal3 ::"('\<AA>)Real" ("\<three>.\<zero>") where "\<three>.\<zero> = (\<lambda> _ . \<lfloor>\<lfloor>3::real\<rfloor>\<rfloor>)" definition OclReal3 ::"('\<AA>)Real" ("\<three>.\<zero>") where "\<three>.\<zero> = (\<lambda> _ . \<lfloor>\<lfloor>3::real\<rfloor>\<rfloor>)"
definition OclReal4 ::"('\<AA>)Real" ("\<four>.\<zero>") where "\<four>.\<zero> = (\<lambda> _ . \<lfloor>\<lfloor>4::real\<rfloor>\<rfloor>)" definition OclReal4 ::"('\<AA>)Real" ("\<four>.\<zero>") where "\<four>.\<zero> = (\<lambda> _ . \<lfloor>\<lfloor>4::real\<rfloor>\<rfloor>)"
definition OclReal5 ::"('\<AA>)Real" ("\<five>.\<zero>") where "\<five>.\<zero> = (\<lambda> _ . \<lfloor>\<lfloor>5::real\<rfloor>\<rfloor>)" definition OclReal5 ::"('\<AA>)Real" ("\<five>.\<zero>") where "\<five>.\<zero> = (\<lambda> _ . \<lfloor>\<lfloor>5::real\<rfloor>\<rfloor>)"
@ -83,7 +83,7 @@ definition OclReal9 ::"('\<AA>)Real" ("\<nine>.\<zero>") where "\<nine>.\
definition OclReal10 ::"('\<AA>)Real" ("\<one>\<zero>.\<zero>") where "\<one>\<zero>.\<zero> = (\<lambda> _ . \<lfloor>\<lfloor>10::real\<rfloor>\<rfloor>)" definition OclReal10 ::"('\<AA>)Real" ("\<one>\<zero>.\<zero>") where "\<one>\<zero>.\<zero> = (\<lambda> _ . \<lfloor>\<lfloor>10::real\<rfloor>\<rfloor>)"
definition OclRealpi ::"('\<AA>)Real" ("\<pi>") where "\<pi> = (\<lambda> _ . \<lfloor>\<lfloor>pi\<rfloor>\<rfloor>)" definition OclRealpi ::"('\<AA>)Real" ("\<pi>") where "\<pi> = (\<lambda> _ . \<lfloor>\<lfloor>pi\<rfloor>\<rfloor>)"
subsection{* Validity and Definedness Properties *} subsection\<open>Validity and Definedness Properties\<close>
lemma "\<delta>(null::('\<AA>)Real) = false" by simp lemma "\<delta>(null::('\<AA>)Real) = false" by simp
lemma "\<upsilon>(null::('\<AA>)Real) = true" by simp lemma "\<upsilon>(null::('\<AA>)Real) = true" by simp
@ -109,18 +109,18 @@ lemma [simp,code_unfold]: "\<delta> \<eight>.\<zero> = true" by(simp add:OclReal
lemma [simp,code_unfold]: "\<upsilon> \<eight>.\<zero> = true" by(simp add:OclReal8_def) lemma [simp,code_unfold]: "\<upsilon> \<eight>.\<zero> = true" by(simp add:OclReal8_def)
lemma [simp,code_unfold]: "\<delta> \<nine>.\<zero> = true" by(simp add:OclReal9_def) lemma [simp,code_unfold]: "\<delta> \<nine>.\<zero> = true" by(simp add:OclReal9_def)
lemma [simp,code_unfold]: "\<upsilon> \<nine>.\<zero> = true" by(simp add:OclReal9_def) lemma [simp,code_unfold]: "\<upsilon> \<nine>.\<zero> = true" by(simp add:OclReal9_def)
text_raw{* \endisatagafp *} text_raw\<open>\endisatagafp\<close>
subsection{* Arithmetical Operations *} subsection\<open>Arithmetical Operations\<close>
subsubsection{* Definition *} subsubsection\<open>Definition\<close>
text{* Here is a common case of a built-in operation on built-in types. text\<open>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). *} Note that the arguments must be both defined (non-null, non-bot).\<close>
text{* Note that we can not follow the lexis of the OCL Standard for Isabelle text\<open>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 technical reasons; these operators are heavily overloaded in the HOL library
that a further overloading would lead to heavy technical buzz in this that a further overloading would lead to heavy technical buzz in this
document. document.
*} \<close>
definition OclAdd\<^sub>R\<^sub>e\<^sub>a\<^sub>l ::"('\<AA>)Real \<Rightarrow> ('\<AA>)Real \<Rightarrow> ('\<AA>)Real" (infix "+\<^sub>r\<^sub>e\<^sub>a\<^sub>l" 40) definition OclAdd\<^sub>R\<^sub>e\<^sub>a\<^sub>l ::"('\<AA>)Real \<Rightarrow> ('\<AA>)Real \<Rightarrow> ('\<AA>)Real" (infix "+\<^sub>r\<^sub>e\<^sub>a\<^sub>l" 40)
where "x +\<^sub>r\<^sub>e\<^sub>a\<^sub>l y \<equiv> \<lambda> \<tau>. if (\<delta> x) \<tau> = true \<tau> \<and> (\<delta> y) \<tau> = true \<tau> where "x +\<^sub>r\<^sub>e\<^sub>a\<^sub>l y \<equiv> \<lambda> \<tau>. if (\<delta> x) \<tau> = true \<tau> \<and> (\<delta> y) \<tau> = true \<tau>
then \<lfloor>\<lfloor>\<lceil>\<lceil>x \<tau>\<rceil>\<rceil> + \<lceil>\<lceil>y \<tau>\<rceil>\<rceil>\<rfloor>\<rfloor> then \<lfloor>\<lfloor>\<lceil>\<lceil>x \<tau>\<rceil>\<rceil> + \<lceil>\<lceil>y \<tau>\<rceil>\<rceil>\<rfloor>\<rfloor>
@ -144,8 +144,8 @@ where "x *\<^sub>r\<^sub>e\<^sub>a\<^sub>l y \<equiv> \<lambda> \<tau>. if (\<de
interpretation OclMult\<^sub>R\<^sub>e\<^sub>a\<^sub>l : profile_bin\<^sub>d_\<^sub>d "OclMult\<^sub>R\<^sub>e\<^sub>a\<^sub>l" "\<lambda> x y. \<lfloor>\<lfloor>\<lceil>\<lceil>x\<rceil>\<rceil> * \<lceil>\<lceil>y\<rceil>\<rceil>\<rfloor>\<rfloor>" interpretation OclMult\<^sub>R\<^sub>e\<^sub>a\<^sub>l : profile_bin\<^sub>d_\<^sub>d "OclMult\<^sub>R\<^sub>e\<^sub>a\<^sub>l" "\<lambda> x y. \<lfloor>\<lfloor>\<lceil>\<lceil>x\<rceil>\<rceil> * \<lceil>\<lceil>y\<rceil>\<rceil>\<rfloor>\<rfloor>"
by unfold_locales (auto simp:OclMult\<^sub>R\<^sub>e\<^sub>a\<^sub>l_def bot_option_def null_option_def) by unfold_locales (auto simp:OclMult\<^sub>R\<^sub>e\<^sub>a\<^sub>l_def bot_option_def null_option_def)
text{* Here is the special case of division, which is defined as invalid for division text\<open>Here is the special case of division, which is defined as invalid for division
by zero. *} by zero.\<close>
definition OclDivision\<^sub>R\<^sub>e\<^sub>a\<^sub>l ::"('\<AA>)Real \<Rightarrow> ('\<AA>)Real \<Rightarrow> ('\<AA>)Real" (infix "div\<^sub>r\<^sub>e\<^sub>a\<^sub>l" 45) definition OclDivision\<^sub>R\<^sub>e\<^sub>a\<^sub>l ::"('\<AA>)Real \<Rightarrow> ('\<AA>)Real \<Rightarrow> ('\<AA>)Real" (infix "div\<^sub>r\<^sub>e\<^sub>a\<^sub>l" 45)
where "x div\<^sub>r\<^sub>e\<^sub>a\<^sub>l y \<equiv> \<lambda> \<tau>. if (\<delta> x) \<tau> = true \<tau> \<and> (\<delta> y) \<tau> = true \<tau> where "x div\<^sub>r\<^sub>e\<^sub>a\<^sub>l y \<equiv> \<lambda> \<tau>. if (\<delta> x) \<tau> = true \<tau> \<and> (\<delta> y) \<tau> = true \<tau>
then if y \<tau> \<noteq> OclReal0 \<tau> then \<lfloor>\<lfloor>\<lceil>\<lceil>x \<tau>\<rceil>\<rceil> / \<lceil>\<lceil>y \<tau>\<rceil>\<rceil>\<rfloor>\<rfloor> else invalid \<tau> then if y \<tau> \<noteq> OclReal0 \<tau> then \<lfloor>\<lfloor>\<lceil>\<lceil>x \<tau>\<rceil>\<rceil> / \<lceil>\<lceil>y \<tau>\<rceil>\<rceil>\<rfloor>\<rfloor> else invalid \<tau>
@ -174,14 +174,14 @@ where "x \<le>\<^sub>r\<^sub>e\<^sub>a\<^sub>l y \<equiv> \<lambda> \<tau>. if (
interpretation OclLe\<^sub>R\<^sub>e\<^sub>a\<^sub>l : profile_bin\<^sub>d_\<^sub>d "(\<le>\<^sub>r\<^sub>e\<^sub>a\<^sub>l)" "\<lambda> x y. \<lfloor>\<lfloor>\<lceil>\<lceil>x\<rceil>\<rceil> \<le> \<lceil>\<lceil>y\<rceil>\<rceil>\<rfloor>\<rfloor>" interpretation OclLe\<^sub>R\<^sub>e\<^sub>a\<^sub>l : profile_bin\<^sub>d_\<^sub>d "(\<le>\<^sub>r\<^sub>e\<^sub>a\<^sub>l)" "\<lambda> x y. \<lfloor>\<lfloor>\<lceil>\<lceil>x\<rceil>\<rceil> \<le> \<lceil>\<lceil>y\<rceil>\<rceil>\<rfloor>\<rfloor>"
by unfold_locales (auto simp:OclLe\<^sub>R\<^sub>e\<^sub>a\<^sub>l_def bot_option_def null_option_def) by unfold_locales (auto simp:OclLe\<^sub>R\<^sub>e\<^sub>a\<^sub>l_def bot_option_def null_option_def)
subsubsection{* Basic Properties *} subsubsection\<open>Basic Properties\<close>
lemma OclAdd\<^sub>R\<^sub>e\<^sub>a\<^sub>l_commute: "(X +\<^sub>r\<^sub>e\<^sub>a\<^sub>l Y) = (Y +\<^sub>r\<^sub>e\<^sub>a\<^sub>l X)" lemma OclAdd\<^sub>R\<^sub>e\<^sub>a\<^sub>l_commute: "(X +\<^sub>r\<^sub>e\<^sub>a\<^sub>l Y) = (Y +\<^sub>r\<^sub>e\<^sub>a\<^sub>l X)"
by(rule ext,auto simp:true_def false_def OclAdd\<^sub>R\<^sub>e\<^sub>a\<^sub>l_def invalid_def by(rule ext,auto simp:true_def false_def OclAdd\<^sub>R\<^sub>e\<^sub>a\<^sub>l_def invalid_def
split: option.split option.split_asm split: option.split option.split_asm
bool.split bool.split_asm) bool.split bool.split_asm)
subsubsection{* Execution with Invalid or Null or Zero as Argument *} subsubsection\<open>Execution with Invalid or Null or Zero as Argument\<close>
lemma OclAdd\<^sub>R\<^sub>e\<^sub>a\<^sub>l_zero1[simp,code_unfold] : lemma OclAdd\<^sub>R\<^sub>e\<^sub>a\<^sub>l_zero1[simp,code_unfold] :
"(x +\<^sub>r\<^sub>e\<^sub>a\<^sub>l \<zero>.\<zero>) = (if \<upsilon> x and not (\<delta> x) then invalid else x endif)" "(x +\<^sub>r\<^sub>e\<^sub>a\<^sub>l \<zero>.\<zero>) = (if \<upsilon> x and not (\<delta> x) then invalid else x endif)"
@ -215,9 +215,9 @@ by(subst OclAdd\<^sub>R\<^sub>e\<^sub>a\<^sub>l_commute, simp)
subsection{* Test Statements *} subsection\<open>Test Statements\<close>
text{* Here follows a list of code-examples, that explain the meanings text\<open>Here follows a list of code-examples, that explain the meanings
of the above definitions by compilation to code and execution to @{term "True"}.*} of the above definitions by compilation to code and execution to @{term "True"}.\<close>
Assert "\<tau> \<Turnstile> ( \<nine>.\<zero> \<le>\<^sub>r\<^sub>e\<^sub>a\<^sub>l \<one>\<zero>.\<zero> )" Assert "\<tau> \<Turnstile> ( \<nine>.\<zero> \<le>\<^sub>r\<^sub>e\<^sub>a\<^sub>l \<one>\<zero>.\<zero> )"
Assert "\<tau> \<Turnstile> (( \<four>.\<zero> +\<^sub>r\<^sub>e\<^sub>a\<^sub>l \<four>.\<zero> ) \<le>\<^sub>r\<^sub>e\<^sub>a\<^sub>l \<one>\<zero>.\<zero> )" Assert "\<tau> \<Turnstile> (( \<four>.\<zero> +\<^sub>r\<^sub>e\<^sub>a\<^sub>l \<four>.\<zero> ) \<le>\<^sub>r\<^sub>e\<^sub>a\<^sub>l \<one>\<zero>.\<zero> )"
@ -251,11 +251,11 @@ lemma OclReal9_non_null [simp,code_unfold]: "(\<nine>.\<zero> \<doteq> null) = f
lemma null_non_OclReal9 [simp,code_unfold]: "(null \<doteq> \<nine>.\<zero>) = false" by(simp add: OclReal9_def) lemma null_non_OclReal9 [simp,code_unfold]: "(null \<doteq> \<nine>.\<zero>) = false" by(simp add: OclReal9_def)
text{* Here follows a list of code-examples, that explain the meanings text\<open>Here follows a list of code-examples, that explain the meanings
of the above definitions by compilation to code and execution to @{term "True"}.*} of the above definitions by compilation to code and execution to @{term "True"}.\<close>
text{* Elementary computations on Real *} text\<open>Elementary computations on Real\<close>
Assert "\<tau> \<Turnstile> \<one>.\<zero> <> \<two>.\<zero>" Assert "\<tau> \<Turnstile> \<one>.\<zero> <> \<two>.\<zero>"
Assert "\<tau> \<Turnstile> \<two>.\<zero> <> \<one>.\<zero>" Assert "\<tau> \<Turnstile> \<two>.\<zero> <> \<one>.\<zero>"

View File

@ -44,13 +44,13 @@ theory UML_String
imports "../UML_PropertyProfiles" imports "../UML_PropertyProfiles"
begin begin
section{* Basic Type String: Operations *} section\<open>Basic Type String: Operations\<close>
subsection{* Fundamental Properties on Strings: Strict Equality \label{sec:string-strict-eq}*} subsection\<open>Fundamental Properties on Strings: Strict Equality \label{sec:string-strict-eq}\<close>
text{* The last basic operation belonging to the fundamental infrastructure text\<open>The last basic operation belonging to the fundamental infrastructure
of a value-type in OCL is the weak equality, which is defined similar of a value-type in OCL is the weak equality, which is defined similar
to the @{typ "('\<AA>)Boolean"}-case as strict extension of the strong equality:*} to the @{typ "('\<AA>)Boolean"}-case as strict extension of the strong equality:\<close>
overloading StrictRefEq \<equiv> "StrictRefEq :: [('\<AA>)String,('\<AA>)String] \<Rightarrow> ('\<AA>)Boolean" overloading StrictRefEq \<equiv> "StrictRefEq :: [('\<AA>)String,('\<AA>)String] \<Rightarrow> ('\<AA>)Boolean"
begin begin
definition StrictRefEq\<^sub>S\<^sub>t\<^sub>r\<^sub>i\<^sub>n\<^sub>g[code_unfold] : definition StrictRefEq\<^sub>S\<^sub>t\<^sub>r\<^sub>i\<^sub>n\<^sub>g[code_unfold] :
@ -59,22 +59,22 @@ begin
else invalid \<tau>" else invalid \<tau>"
end 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"}*} text\<open>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"}\<close>
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 "\<lambda> x y. (x::('\<AA>)String) \<doteq> y" 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 "\<lambda> x y. (x::('\<AA>)String) \<doteq> y"
by unfold_locales (auto simp: StrictRefEq\<^sub>S\<^sub>t\<^sub>r\<^sub>i\<^sub>n\<^sub>g) by unfold_locales (auto simp: StrictRefEq\<^sub>S\<^sub>t\<^sub>r\<^sub>i\<^sub>n\<^sub>g)
subsection{* Basic String Constants *} subsection\<open>Basic String Constants\<close>
text{* Although the remaining part of this library reasons about text\<open>Although the remaining part of this library reasons about
integers abstractly, we provide here as example some convenient shortcuts. *} integers abstractly, we provide here as example some convenient shortcuts.\<close>
definition OclStringa ::"('\<AA>)String" ("\<a>") where "\<a> = (\<lambda> _ . \<lfloor>\<lfloor>''a''\<rfloor>\<rfloor>)" definition OclStringa ::"('\<AA>)String" ("\<a>") where "\<a> = (\<lambda> _ . \<lfloor>\<lfloor>''a''\<rfloor>\<rfloor>)"
definition OclStringb ::"('\<AA>)String" ("\<b>") where "\<b> = (\<lambda> _ . \<lfloor>\<lfloor>''b''\<rfloor>\<rfloor>)" definition OclStringb ::"('\<AA>)String" ("\<b>") where "\<b> = (\<lambda> _ . \<lfloor>\<lfloor>''b''\<rfloor>\<rfloor>)"
definition OclStringc ::"('\<AA>)String" ("\<c>") where "\<c> = (\<lambda> _ . \<lfloor>\<lfloor>''c''\<rfloor>\<rfloor>)" definition OclStringc ::"('\<AA>)String" ("\<c>") where "\<c> = (\<lambda> _ . \<lfloor>\<lfloor>''c''\<rfloor>\<rfloor>)"
text{* Etc.*} text\<open>Etc.\<close>
text_raw{* \isatagafp *} text_raw\<open>\isatagafp\<close>
subsection{* Validity and Definedness Properties *} subsection\<open>Validity and Definedness Properties\<close>
lemma "\<delta>(null::('\<AA>)String) = false" by simp lemma "\<delta>(null::('\<AA>)String) = false" by simp
lemma "\<upsilon>(null::('\<AA>)String) = true" by simp lemma "\<upsilon>(null::('\<AA>)String) = true" by simp
@ -90,18 +90,18 @@ by(simp add:valid_def true_def
(* ecclectic proofs to make examples executable *) (* ecclectic proofs to make examples executable *)
lemma [simp,code_unfold]: "\<delta> \<a> = true" by(simp add:OclStringa_def) lemma [simp,code_unfold]: "\<delta> \<a> = true" by(simp add:OclStringa_def)
lemma [simp,code_unfold]: "\<upsilon> \<a> = true" by(simp add:OclStringa_def) lemma [simp,code_unfold]: "\<upsilon> \<a> = true" by(simp add:OclStringa_def)
text_raw{* \endisatagafp *} text_raw\<open>\endisatagafp\<close>
subsection{* String Operations *} subsection\<open>String Operations\<close>
subsubsection{* Definition *} subsubsection\<open>Definition\<close>
text{* Here is a common case of a built-in operation on built-in types. text\<open>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). *} Note that the arguments must be both defined (non-null, non-bot).\<close>
text{* Note that we can not follow the lexis of the OCL Standard for Isabelle text\<open>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 technical reasons; these operators are heavily overloaded in the HOL library
that a further overloading would lead to heavy technical buzz in this that a further overloading would lead to heavy technical buzz in this
document. document.
*} \<close>
definition OclAdd\<^sub>S\<^sub>t\<^sub>r\<^sub>i\<^sub>n\<^sub>g ::"('\<AA>)String \<Rightarrow> ('\<AA>)String \<Rightarrow> ('\<AA>)String" (infix "+\<^sub>s\<^sub>t\<^sub>r\<^sub>i\<^sub>n\<^sub>g" 40) definition OclAdd\<^sub>S\<^sub>t\<^sub>r\<^sub>i\<^sub>n\<^sub>g ::"('\<AA>)String \<Rightarrow> ('\<AA>)String \<Rightarrow> ('\<AA>)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 \<equiv> \<lambda> \<tau>. if (\<delta> x) \<tau> = true \<tau> \<and> (\<delta> y) \<tau> = true \<tau> where "x +\<^sub>s\<^sub>t\<^sub>r\<^sub>i\<^sub>n\<^sub>g y \<equiv> \<lambda> \<tau>. if (\<delta> x) \<tau> = true \<tau> \<and> (\<delta> y) \<tau> = true \<tau>
then \<lfloor>\<lfloor>concat [\<lceil>\<lceil>x \<tau>\<rceil>\<rceil>, \<lceil>\<lceil>y \<tau>\<rceil>\<rceil>]\<rfloor>\<rfloor> then \<lfloor>\<lfloor>concat [\<lceil>\<lceil>x \<tau>\<rceil>\<rceil>, \<lceil>\<lceil>y \<tau>\<rceil>\<rceil>]\<rfloor>\<rfloor>
@ -112,7 +112,7 @@ interpretation OclAdd\<^sub>S\<^sub>t\<^sub>r\<^sub>i\<^sub>n\<^sub>g : profile_
(* TODO : size(), concat, substring(s:string) toInteger, toReal, at(i:Integer), characters() etc. *) (* TODO : size(), concat, substring(s:string) toInteger, toReal, at(i:Integer), characters() etc. *)
subsubsection{* Basic Properties *} subsubsection\<open>Basic Properties\<close>
lemma OclAdd\<^sub>S\<^sub>t\<^sub>r\<^sub>i\<^sub>n\<^sub>g_not_commute: "\<exists>X Y. (X +\<^sub>s\<^sub>t\<^sub>r\<^sub>i\<^sub>n\<^sub>g Y) \<noteq> (Y +\<^sub>s\<^sub>t\<^sub>r\<^sub>i\<^sub>n\<^sub>g X)" lemma OclAdd\<^sub>S\<^sub>t\<^sub>r\<^sub>i\<^sub>n\<^sub>g_not_commute: "\<exists>X Y. (X +\<^sub>s\<^sub>t\<^sub>r\<^sub>i\<^sub>n\<^sub>g Y) \<noteq> (Y +\<^sub>s\<^sub>t\<^sub>r\<^sub>i\<^sub>n\<^sub>g X)"
apply(rule_tac x = "\<lambda>_. \<lfloor>\<lfloor>''b''\<rfloor>\<rfloor>" in exI) apply(rule_tac x = "\<lambda>_. \<lfloor>\<lfloor>''b''\<rfloor>\<rfloor>" in exI)
@ -121,9 +121,9 @@ lemma OclAdd\<^sub>S\<^sub>t\<^sub>r\<^sub>i\<^sub>n\<^sub>g_not_commute: "\<exi
by(auto, drule fun_cong, auto) by(auto, drule fun_cong, auto)
subsection{* Test Statements *} subsection\<open>Test Statements\<close>
text{* Here follows a list of code-examples, that explain the meanings text\<open>Here follows a list of code-examples, that explain the meanings
of the above definitions by compilation to code and execution to @{term "True"}.*} of the above definitions by compilation to code and execution to @{term "True"}.\<close>
(* (*
Assert "\<tau> \<Turnstile> ( \<nine> \<le>\<^sub>s\<^sub>t\<^sub>r\<^sub>i\<^sub>n\<^sub>g \<one>\<zero> )" Assert "\<tau> \<Turnstile> ( \<nine> \<le>\<^sub>s\<^sub>t\<^sub>r\<^sub>i\<^sub>n\<^sub>g \<one>\<zero> )"
Assert "\<tau> \<Turnstile> (( \<four> +\<^sub>s\<^sub>t\<^sub>r\<^sub>i\<^sub>n\<^sub>g \<four> ) \<le>\<^sub>s\<^sub>t\<^sub>r\<^sub>i\<^sub>n\<^sub>g \<one>\<zero> )" Assert "\<tau> \<Turnstile> (( \<four> +\<^sub>s\<^sub>t\<^sub>r\<^sub>i\<^sub>n\<^sub>g \<four> ) \<le>\<^sub>s\<^sub>t\<^sub>r\<^sub>i\<^sub>n\<^sub>g \<one>\<zero> )"
@ -131,11 +131,11 @@ Assert "\<tau> |\<noteq> (( \<four> +\<^sub>s\<^sub>t\<^sub>r\<^sub>i\<^sub>n\<^
Assert "\<tau> \<Turnstile> not (\<upsilon> (null +\<^sub>s\<^sub>t\<^sub>r\<^sub>i\<^sub>n\<^sub>g \<one>)) " Assert "\<tau> \<Turnstile> not (\<upsilon> (null +\<^sub>s\<^sub>t\<^sub>r\<^sub>i\<^sub>n\<^sub>g \<one>)) "
*) *)
text{* Here follows a list of code-examples, that explain the meanings text\<open>Here follows a list of code-examples, that explain the meanings
of the above definitions by compilation to code and execution to @{term "True"}.*} of the above definitions by compilation to code and execution to @{term "True"}.\<close>
text{* Elementary computations on String *} text\<open>Elementary computations on String\<close>
Assert "\<tau> \<Turnstile> \<a> <> \<b>" Assert "\<tau> \<Turnstile> \<a> <> \<b>"
Assert "\<tau> \<Turnstile> \<b> <> \<a>" Assert "\<tau> \<Turnstile> \<b> <> \<a>"

View File

@ -44,18 +44,18 @@ theory UML_Void
imports "../UML_PropertyProfiles" imports "../UML_PropertyProfiles"
begin begin
section{* Basic Type Void: Operations *} section\<open>Basic Type Void: Operations\<close>
(* For technical reasons, the type does not contain to the null-class yet. *) (* For technical reasons, the type does not contain to the null-class yet. *)
text {* This \emph{minimal} OCL type contains only two elements: text \<open>This \emph{minimal} OCL type contains only two elements:
@{term "invalid"} and @{term "null"}. @{term "invalid"} and @{term "null"}.
@{term "Void"} could initially be defined as @{typ "unit option option"}, @{term "Void"} could initially be defined as @{typ "unit option option"},
however the cardinal of this type is more than two, so it would have the cost to consider however the cardinal of this type is more than two, so it would have the cost to consider
@{text "Some None"} and @{text "Some (Some ())"} seemingly everywhere.*} \<open>Some None\<close> and \<open>Some (Some ())\<close> seemingly everywhere.\<close>
subsection{* Fundamental Properties on Voids: Strict Equality *} subsection\<open>Fundamental Properties on Voids: Strict Equality\<close>
subsubsection{* Definition *} subsubsection\<open>Definition\<close>
instantiation Void\<^sub>b\<^sub>a\<^sub>s\<^sub>e :: bot instantiation Void\<^sub>b\<^sub>a\<^sub>s\<^sub>e :: bot
begin begin
@ -81,9 +81,9 @@ begin
end end
text{* The last basic operation belonging to the fundamental infrastructure text\<open>The last basic operation belonging to the fundamental infrastructure
of a value-type in OCL is the weak equality, which is defined similar of a value-type in OCL is the weak equality, which is defined similar
to the @{typ "('\<AA>)Void"}-case as strict extension of the strong equality:*} to the @{typ "('\<AA>)Void"}-case as strict extension of the strong equality:\<close>
overloading StrictRefEq \<equiv> "StrictRefEq :: [('\<AA>)Void,('\<AA>)Void] \<Rightarrow> ('\<AA>)Boolean" overloading StrictRefEq \<equiv> "StrictRefEq :: [('\<AA>)Void,('\<AA>)Void] \<Rightarrow> ('\<AA>)Boolean"
begin begin
definition StrictRefEq\<^sub>V\<^sub>o\<^sub>i\<^sub>d[code_unfold] : definition StrictRefEq\<^sub>V\<^sub>o\<^sub>i\<^sub>d[code_unfold] :
@ -92,15 +92,15 @@ begin
else invalid \<tau>" else invalid \<tau>"
end 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"}*} text\<open>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"}\<close>
interpretation StrictRefEq\<^sub>V\<^sub>o\<^sub>i\<^sub>d : profile_bin\<^sub>S\<^sub>t\<^sub>r\<^sub>o\<^sub>n\<^sub>g\<^sub>E\<^sub>q_\<^sub>v_\<^sub>v "\<lambda> x y. (x::('\<AA>)Void) \<doteq> y" interpretation StrictRefEq\<^sub>V\<^sub>o\<^sub>i\<^sub>d : profile_bin\<^sub>S\<^sub>t\<^sub>r\<^sub>o\<^sub>n\<^sub>g\<^sub>E\<^sub>q_\<^sub>v_\<^sub>v "\<lambda> x y. (x::('\<AA>)Void) \<doteq> y"
by unfold_locales (auto simp: StrictRefEq\<^sub>V\<^sub>o\<^sub>i\<^sub>d) by unfold_locales (auto simp: StrictRefEq\<^sub>V\<^sub>o\<^sub>i\<^sub>d)
subsection{* Basic Void Constants *} subsection\<open>Basic Void Constants\<close>
subsection{* Validity and Definedness Properties *} subsection\<open>Validity and Definedness Properties\<close>
lemma "\<delta>(null::('\<AA>)Void) = false" by simp lemma "\<delta>(null::('\<AA>)Void) = false" by simp
lemma "\<upsilon>(null::('\<AA>)Void) = true" by simp lemma "\<upsilon>(null::('\<AA>)Void) = true" by simp
@ -130,7 +130,7 @@ apply(rule ext, simp split:, intro conjI impI)
by(metis null_Void_def null_is_valid, simp add: true_def) by(metis null_Void_def null_is_valid, simp add: true_def)
subsection{* Test Statements *} subsection\<open>Test Statements\<close>
Assert "\<tau> \<Turnstile> ((null::('\<AA>)Void) \<doteq> null)" Assert "\<tau> \<Turnstile> ((null::('\<AA>)Void) \<doteq> null)"

View File

@ -50,7 +50,7 @@ imports "../basic_types/UML_Void"
begin begin
no_notation None ("\<bottom>") no_notation None ("\<bottom>")
section{* Collection Type Bag: Operations *} section\<open>Collection Type Bag: Operations\<close>
definition "Rep_Bag_base' x = {(x0, y). y < \<lceil>\<lceil>Rep_Bag\<^sub>b\<^sub>a\<^sub>s\<^sub>e x\<rceil>\<rceil> x0 }" definition "Rep_Bag_base' x = {(x0, y). y < \<lceil>\<lceil>Rep_Bag\<^sub>b\<^sub>a\<^sub>s\<^sub>e x\<rceil>\<rceil> x0 }"
definition "Rep_Bag_base x \<tau> = {(x0, y). y < \<lceil>\<lceil>Rep_Bag\<^sub>b\<^sub>a\<^sub>s\<^sub>e (x \<tau>)\<rceil>\<rceil> x0 }" definition "Rep_Bag_base x \<tau> = {(x0, y). y < \<lceil>\<lceil>Rep_Bag\<^sub>b\<^sub>a\<^sub>s\<^sub>e (x \<tau>)\<rceil>\<rceil> x0 }"
@ -60,10 +60,10 @@ definition ApproxEq (infixl "\<cong>" 30)
where "X \<cong> Y \<equiv> \<lambda> \<tau>. \<lfloor>\<lfloor>Rep_Set_base X \<tau> = Rep_Set_base Y \<tau> \<rfloor>\<rfloor>" where "X \<cong> Y \<equiv> \<lambda> \<tau>. \<lfloor>\<lfloor>Rep_Set_base X \<tau> = Rep_Set_base Y \<tau> \<rfloor>\<rfloor>"
subsection{* As a Motivation for the (infinite) Type Construction: Type-Extensions as Bags subsection\<open>As a Motivation for the (infinite) Type Construction: Type-Extensions as Bags
\label{sec:type-extensions}*} \label{sec:type-extensions}\<close>
text{* Our notion of typed bag goes beyond the usual notion of a finite executable bag and text\<open>Our notion of typed bag goes beyond the usual notion of a finite executable bag and
is powerful enough to capture \emph{the extension of a type} in UML and OCL. This means is powerful enough to capture \emph{the extension of a type} in UML and OCL. This means
we can have in Featherweight OCL Bags containing all possible elements of a type, not only we can have in Featherweight OCL Bags containing all possible elements of a type, not only
those (finite) ones representable in a state. This holds for base types as well as class types, those (finite) ones representable in a state. This holds for base types as well as class types,
@ -77,9 +77,9 @@ In a world with @{term invalid} and @{term null}, there are two notions extensio
\item the bag of all \emph{valid} values of a type @{term T}, so including @{term null} \item the bag of all \emph{valid} values of a type @{term T}, so including @{term null}
(for which we will introduce the constant @{term T\<^sub>n\<^sub>u\<^sub>l\<^sub>l}). (for which we will introduce the constant @{term T\<^sub>n\<^sub>u\<^sub>l\<^sub>l}).
\end{enumerate} \end{enumerate}
*} \<close>
text{* We define the bag extensions for the base type @{term Integer} as follows: *} text\<open>We define the bag extensions for the base type @{term Integer} as follows:\<close>
definition Integer :: "('\<AA>,Integer\<^sub>b\<^sub>a\<^sub>s\<^sub>e) Bag" definition Integer :: "('\<AA>,Integer\<^sub>b\<^sub>a\<^sub>s\<^sub>e) Bag"
where "Integer \<equiv> (\<lambda> \<tau>. (Abs_Bag\<^sub>b\<^sub>a\<^sub>s\<^sub>e o Some o Some) (\<lambda> None \<Rightarrow> 0 | Some None \<Rightarrow> 0 | _ \<Rightarrow> 1))" where "Integer \<equiv> (\<lambda> \<tau>. (Abs_Bag\<^sub>b\<^sub>a\<^sub>s\<^sub>e o Some o Some) (\<lambda> None \<Rightarrow> 0 | Some None \<Rightarrow> 0 | _ \<Rightarrow> 1))"
@ -96,20 +96,20 @@ apply(rule ext, auto simp: Integer\<^sub>n\<^sub>u\<^sub>l\<^sub>l_def defined_d
bot_fun_def null_fun_def null_option_def) bot_fun_def null_fun_def null_option_def)
by(simp_all add: Abs_Bag\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject bot_option_def bot_Bag\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_Bag\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_option_def) by(simp_all add: Abs_Bag\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject bot_option_def bot_Bag\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_Bag\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_option_def)
text{* This allows the theorems: text\<open>This allows the theorems:
@{text "\<tau> \<Turnstile> \<delta> x \<Longrightarrow> \<tau> \<Turnstile> (Integer->includes\<^sub>B\<^sub>a\<^sub>g(x))"} \<open>\<tau> \<Turnstile> \<delta> x \<Longrightarrow> \<tau> \<Turnstile> (Integer->includes\<^sub>B\<^sub>a\<^sub>g(x))\<close>
@{text "\<tau> \<Turnstile> \<delta> x \<Longrightarrow> \<tau> \<Turnstile> Integer \<triangleq> (Integer->including\<^sub>B\<^sub>a\<^sub>g(x))"} \<open>\<tau> \<Turnstile> \<delta> x \<Longrightarrow> \<tau> \<Turnstile> Integer \<triangleq> (Integer->including\<^sub>B\<^sub>a\<^sub>g(x))\<close>
and and
@{text "\<tau> \<Turnstile> \<upsilon> x \<Longrightarrow> \<tau> \<Turnstile> (Integer\<^sub>n\<^sub>u\<^sub>l\<^sub>l->includes\<^sub>B\<^sub>a\<^sub>g(x))"} \<open>\<tau> \<Turnstile> \<upsilon> x \<Longrightarrow> \<tau> \<Turnstile> (Integer\<^sub>n\<^sub>u\<^sub>l\<^sub>l->includes\<^sub>B\<^sub>a\<^sub>g(x))\<close>
@{text "\<tau> \<Turnstile> \<upsilon> x \<Longrightarrow> \<tau> \<Turnstile> Integer\<^sub>n\<^sub>u\<^sub>l\<^sub>l \<triangleq> (Integer\<^sub>n\<^sub>u\<^sub>l\<^sub>l->including\<^sub>B\<^sub>a\<^sub>g(x))"} \<open>\<tau> \<Turnstile> \<upsilon> x \<Longrightarrow> \<tau> \<Turnstile> Integer\<^sub>n\<^sub>u\<^sub>l\<^sub>l \<triangleq> (Integer\<^sub>n\<^sub>u\<^sub>l\<^sub>l->including\<^sub>B\<^sub>a\<^sub>g(x))\<close>
which characterize the infiniteness of these bags by a recursive property on these bags. which characterize the infiniteness of these bags by a recursive property on these bags.
*} \<close>
text{* In the same spirit, we proceed similarly for the remaining base types: *} text\<open>In the same spirit, we proceed similarly for the remaining base types:\<close>
definition Void\<^sub>n\<^sub>u\<^sub>l\<^sub>l :: "('\<AA>,Void\<^sub>b\<^sub>a\<^sub>s\<^sub>e) Bag" definition Void\<^sub>n\<^sub>u\<^sub>l\<^sub>l :: "('\<AA>,Void\<^sub>b\<^sub>a\<^sub>s\<^sub>e) Bag"
where "Void\<^sub>n\<^sub>u\<^sub>l\<^sub>l \<equiv> (\<lambda> \<tau>. (Abs_Bag\<^sub>b\<^sub>a\<^sub>s\<^sub>e o Some o Some) (\<lambda> x. if x = Abs_Void\<^sub>b\<^sub>a\<^sub>s\<^sub>e (Some None) then 1 else 0))" where "Void\<^sub>n\<^sub>u\<^sub>l\<^sub>l \<equiv> (\<lambda> \<tau>. (Abs_Bag\<^sub>b\<^sub>a\<^sub>s\<^sub>e o Some o Some) (\<lambda> x. if x = Abs_Void\<^sub>b\<^sub>a\<^sub>s\<^sub>e (Some None) then 1 else 0))"
@ -220,9 +220,9 @@ apply(rule ext, auto simp: Real\<^sub>n\<^sub>u\<^sub>l\<^sub>l_def defined_def
bot_fun_def null_fun_def null_option_def) bot_fun_def null_fun_def null_option_def)
by(simp_all add: Abs_Bag\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject bot_option_def bot_Bag\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_Bag\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_option_def) by(simp_all add: Abs_Bag\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject bot_option_def bot_Bag\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_Bag\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_option_def)
subsection{* Basic Properties of the Bag Type*} subsection\<open>Basic Properties of the Bag Type\<close>
text{* Every element in a defined bag is valid. *} text\<open>Every element in a defined bag is valid.\<close>
lemma Bag_inv_lemma: "\<tau> \<Turnstile> (\<delta> X) \<Longrightarrow> \<lceil>\<lceil>Rep_Bag\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \<tau>)\<rceil>\<rceil> bot = 0" lemma Bag_inv_lemma: "\<tau> \<Turnstile> (\<delta> X) \<Longrightarrow> \<lceil>\<lceil>Rep_Bag\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \<tau>)\<rceil>\<rceil> bot = 0"
apply(insert Rep_Bag\<^sub>b\<^sub>a\<^sub>s\<^sub>e [of "X \<tau>"], simp) apply(insert Rep_Bag\<^sub>b\<^sub>a\<^sub>s\<^sub>e [of "X \<tau>"], simp)
@ -271,18 +271,18 @@ apply(simp add: valid_def null_fun_def bot_fun_def bot_Bag\<^sub>b\<^sub>a\<^sub
apply(subst Abs_Bag\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject,simp_all add: null_option_def bot_option_def) apply(subst Abs_Bag\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject,simp_all add: null_option_def bot_option_def)
done done
text{* ... which means that we can have a type @{text "('\<AA>,('\<AA>,('\<AA>) Integer) Bag) Bag"} text\<open>... which means that we can have a type \<open>('\<AA>,('\<AA>,('\<AA>) Integer) Bag) Bag\<close>
corresponding exactly to Bag(Bag(Integer)) in OCL notation. Note that the parameter corresponding exactly to Bag(Bag(Integer)) in OCL notation. Note that the parameter
@{text "'\<AA>"} still refers to the object universe; making the OCL semantics entirely parametric \<open>'\<AA>\<close> still refers to the object universe; making the OCL semantics entirely parametric
in the object universe makes it possible to study (and prove) its properties in the object universe makes it possible to study (and prove) its properties
independently from a concrete class diagram. *} independently from a concrete class diagram.\<close>
subsection{* Definition: Strict Equality \label{sec:bag-strict-equality}*} subsection\<open>Definition: Strict Equality \label{sec:bag-strict-equality}\<close>
text{* After the part of foundational operations on bags, we detail here equality on bags. text\<open>After the part of foundational operations on bags, we detail here equality on bags.
Strong equality is inherited from the OCL core, but we have to consider Strong equality is inherited from the OCL core, but we have to consider
the case of the strict equality. We decide to overload strict equality in the the case of the strict equality. We decide to overload strict equality in the
same way we do for other value's in OCL:*} same way we do for other value's in OCL:\<close>
overloading StrictRefEq \<equiv> "StrictRefEq :: [('\<AA>,'\<alpha>::null)Bag,('\<AA>,'\<alpha>::null)Bag] \<Rightarrow> ('\<AA>)Boolean" overloading StrictRefEq \<equiv> "StrictRefEq :: [('\<AA>,'\<alpha>::null)Bag,('\<AA>,'\<alpha>::null)Bag] \<Rightarrow> ('\<AA>)Boolean"
begin begin
@ -292,20 +292,20 @@ begin
else invalid \<tau>" else invalid \<tau>"
end end
text{* One might object here that for the case of objects, this is an empty definition. text\<open>One might object here that for the case of objects, this is an empty definition.
The answer is no, we will restrain later on states and objects such that any object The answer is no, we will restrain later on states and objects such that any object
has its oid stored inside the object (so the ref, under which an object can be referenced has its oid stored inside the object (so the ref, under which an object can be referenced
in the store will represented in the object itself). For such well-formed stores that satisfy in the store will represented in the object itself). For such well-formed stores that satisfy
this invariant (the WFF-invariant), the referential equality and the this invariant (the WFF-invariant), the referential equality and the
strong equality---and therefore the strict equality on bags in the sense above---coincides.*} strong equality---and therefore the strict equality on bags in the sense above---coincides.\<close>
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"}*} text\<open>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"}\<close>
interpretation StrictRefEq\<^sub>B\<^sub>a\<^sub>g : profile_bin\<^sub>S\<^sub>t\<^sub>r\<^sub>o\<^sub>n\<^sub>g\<^sub>E\<^sub>q_\<^sub>v_\<^sub>v "\<lambda> x y. (x::('\<AA>,'\<alpha>::null)Bag) \<doteq> y" interpretation StrictRefEq\<^sub>B\<^sub>a\<^sub>g : profile_bin\<^sub>S\<^sub>t\<^sub>r\<^sub>o\<^sub>n\<^sub>g\<^sub>E\<^sub>q_\<^sub>v_\<^sub>v "\<lambda> x y. (x::('\<AA>,'\<alpha>::null)Bag) \<doteq> y"
by unfold_locales (auto simp: StrictRefEq\<^sub>B\<^sub>a\<^sub>g) by unfold_locales (auto simp: StrictRefEq\<^sub>B\<^sub>a\<^sub>g)
subsection{* Constants: mtBag *} subsection\<open>Constants: mtBag\<close>
definition mtBag::"('\<AA>,'\<alpha>::null) Bag" ("Bag{}") definition mtBag::"('\<AA>,'\<alpha>::null) Bag" ("Bag{}")
where "Bag{} \<equiv> (\<lambda> \<tau>. Abs_Bag\<^sub>b\<^sub>a\<^sub>s\<^sub>e \<lfloor>\<lfloor>\<lambda>_. 0::nat\<rfloor>\<rfloor> )" where "Bag{} \<equiv> (\<lambda> \<tau>. Abs_Bag\<^sub>b\<^sub>a\<^sub>s\<^sub>e \<lfloor>\<lfloor>\<lambda>_. 0::nat\<rfloor>\<rfloor> )"
@ -324,18 +324,18 @@ lemma mtBag_rep_bag: "\<lceil>\<lceil>Rep_Bag\<^sub>b\<^sub>a\<^sub>s\<^sub>e (B
apply(simp add: mtBag_def, subst Abs_Bag\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse) apply(simp add: mtBag_def, subst Abs_Bag\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse)
by(simp add: bot_option_def)+ by(simp add: bot_option_def)+
text_raw{* \isatagafp *} text_raw\<open>\isatagafp\<close>
lemma [simp,code_unfold]: "const Bag{}" lemma [simp,code_unfold]: "const Bag{}"
by(simp add: const_def mtBag_def) by(simp add: const_def mtBag_def)
text{* Note that the collection types in OCL allow for null to be included; text\<open>Note that the collection types in OCL allow for null to be included;
however, there is the null-collection into which inclusion yields invalid. *} however, there is the null-collection into which inclusion yields invalid.\<close>
text_raw{* \endisatagafp *} text_raw\<open>\endisatagafp\<close>
subsection{* Definition: Including *} subsection\<open>Definition: Including\<close>
definition OclIncluding :: "[('\<AA>,'\<alpha>::null) Bag,('\<AA>,'\<alpha>) val] \<Rightarrow> ('\<AA>,'\<alpha>) Bag" definition OclIncluding :: "[('\<AA>,'\<alpha>::null) Bag,('\<AA>,'\<alpha>) val] \<Rightarrow> ('\<AA>,'\<alpha>) Bag"
where "OclIncluding x y = (\<lambda> \<tau>. if (\<delta> x) \<tau> = true \<tau> \<and> (\<upsilon> y) \<tau> = true \<tau> where "OclIncluding x y = (\<lambda> \<tau>. if (\<delta> x) \<tau> = true \<tau> \<and> (\<upsilon> y) \<tau> = true \<tau>
@ -365,7 +365,7 @@ translations
"Bag{x}" == "CONST OclIncluding (Bag{}) x " "Bag{x}" == "CONST OclIncluding (Bag{}) x "
subsection{* Definition: Excluding *} subsection\<open>Definition: Excluding\<close>
definition OclExcluding :: "[('\<AA>,'\<alpha>::null) Bag,('\<AA>,'\<alpha>) val] \<Rightarrow> ('\<AA>,'\<alpha>) Bag" definition OclExcluding :: "[('\<AA>,'\<alpha>::null) Bag,('\<AA>,'\<alpha>) val] \<Rightarrow> ('\<AA>,'\<alpha>) Bag"
where "OclExcluding x y = (\<lambda> \<tau>. if (\<delta> x) \<tau> = true \<tau> \<and> (\<upsilon> y) \<tau> = true \<tau> where "OclExcluding x y = (\<lambda> \<tau>. if (\<delta> x) \<tau> = true \<tau> \<and> (\<upsilon> y) \<tau> = true \<tau>
@ -386,7 +386,7 @@ proof -
mem_Collect_eq null_option_def)+ mem_Collect_eq null_option_def)+
qed qed
subsection{* Definition: Includes *} subsection\<open>Definition: Includes\<close>
definition OclIncludes :: "[('\<AA>,'\<alpha>::null) Bag,('\<AA>,'\<alpha>) val] \<Rightarrow> '\<AA> Boolean" definition OclIncludes :: "[('\<AA>,'\<alpha>::null) Bag,('\<AA>,'\<alpha>) val] \<Rightarrow> '\<AA> Boolean"
where "OclIncludes x y = (\<lambda> \<tau>. if (\<delta> x) \<tau> = true \<tau> \<and> (\<upsilon> y) \<tau> = true \<tau> where "OclIncludes x y = (\<lambda> \<tau>. if (\<delta> x) \<tau> = true \<tau> \<and> (\<upsilon> y) \<tau> = true \<tau>
@ -397,21 +397,21 @@ notation OclIncludes ("_->includes\<^sub>B\<^sub>a\<^sub>g'(_')" (*[66,65]6
interpretation OclIncludes : profile_bin\<^sub>d_\<^sub>v OclIncludes "\<lambda>x y. \<lfloor>\<lfloor> \<lceil>\<lceil>Rep_Bag\<^sub>b\<^sub>a\<^sub>s\<^sub>e x\<rceil>\<rceil> y > 0 \<rfloor>\<rfloor>" interpretation OclIncludes : profile_bin\<^sub>d_\<^sub>v OclIncludes "\<lambda>x y. \<lfloor>\<lfloor> \<lceil>\<lceil>Rep_Bag\<^sub>b\<^sub>a\<^sub>s\<^sub>e x\<rceil>\<rceil> y > 0 \<rfloor>\<rfloor>"
by(unfold_locales, auto simp:OclIncludes_def bot_option_def null_option_def invalid_def) by(unfold_locales, auto simp:OclIncludes_def bot_option_def null_option_def invalid_def)
subsection{* Definition: Excludes *} subsection\<open>Definition: Excludes\<close>
definition OclExcludes :: "[('\<AA>,'\<alpha>::null) Bag,('\<AA>,'\<alpha>) val] \<Rightarrow> '\<AA> Boolean" definition OclExcludes :: "[('\<AA>,'\<alpha>::null) Bag,('\<AA>,'\<alpha>) val] \<Rightarrow> '\<AA> Boolean"
where "OclExcludes x y = (not(OclIncludes x y))" where "OclExcludes x y = (not(OclIncludes x y))"
notation OclExcludes ("_->excludes\<^sub>B\<^sub>a\<^sub>g'(_')" (*[66,65]65*)) notation OclExcludes ("_->excludes\<^sub>B\<^sub>a\<^sub>g'(_')" (*[66,65]65*))
text{* The case of the size definition is somewhat special, we admit text\<open>The case of the size definition is somewhat special, we admit
explicitly in Featherweight OCL the possibility of infinite bags. For explicitly in Featherweight OCL the possibility of infinite bags. For
the size definition, this requires an extra condition that assures the size definition, this requires an extra condition that assures
that the cardinality of the bag is actually a defined integer. *} that the cardinality of the bag is actually a defined integer.\<close>
interpretation OclExcludes : profile_bin\<^sub>d_\<^sub>v OclExcludes "\<lambda>x y. \<lfloor>\<lfloor> \<lceil>\<lceil>Rep_Bag\<^sub>b\<^sub>a\<^sub>s\<^sub>e x\<rceil>\<rceil> y \<le> 0 \<rfloor>\<rfloor>" interpretation OclExcludes : profile_bin\<^sub>d_\<^sub>v OclExcludes "\<lambda>x y. \<lfloor>\<lfloor> \<lceil>\<lceil>Rep_Bag\<^sub>b\<^sub>a\<^sub>s\<^sub>e x\<rceil>\<rceil> y \<le> 0 \<rfloor>\<rfloor>"
by(unfold_locales, auto simp:OclExcludes_def OclIncludes_def OclNot_def bot_option_def null_option_def invalid_def) by(unfold_locales, auto simp:OclExcludes_def OclIncludes_def OclNot_def bot_option_def null_option_def invalid_def)
subsection{* Definition: Size *} subsection\<open>Definition: Size\<close>
definition OclSize :: "('\<AA>,'\<alpha>::null)Bag \<Rightarrow> '\<AA> Integer" definition OclSize :: "('\<AA>,'\<alpha>::null)Bag \<Rightarrow> '\<AA> Integer"
where "OclSize x = (\<lambda> \<tau>. if (\<delta> x) \<tau> = true \<tau> \<and> finite (Rep_Bag_base x \<tau>) where "OclSize x = (\<lambda> \<tau>. if (\<delta> x) \<tau> = true \<tau> \<and> finite (Rep_Bag_base x \<tau>)
@ -420,16 +420,16 @@ where "OclSize x = (\<lambda> \<tau>. if (\<delta> x) \<tau> = true \<tau> \
notation (* standard ascii syntax *) notation (* standard ascii syntax *)
OclSize ("_->size\<^sub>B\<^sub>a\<^sub>g'(')" (*[66]*)) OclSize ("_->size\<^sub>B\<^sub>a\<^sub>g'(')" (*[66]*))
text{* The following definition follows the requirement of the text\<open>The following definition follows the requirement of the
standard to treat null as neutral element of bags. It is standard to treat null as neutral element of bags. It is
a well-documented exception from the general strictness a well-documented exception from the general strictness
rule and the rule that the distinguished argument self should rule and the rule that the distinguished argument self should
be non-null. *} be non-null.\<close>
(*TODO Locale - Equivalent*) (*TODO Locale - Equivalent*)
subsection{* Definition: IsEmpty *} subsection\<open>Definition: IsEmpty\<close>
definition OclIsEmpty :: "('\<AA>,'\<alpha>::null) Bag \<Rightarrow> '\<AA> Boolean" definition OclIsEmpty :: "('\<AA>,'\<alpha>::null) Bag \<Rightarrow> '\<AA> Boolean"
where "OclIsEmpty x = ((\<upsilon> x and not (\<delta> x)) or ((OclSize x) \<doteq> \<zero>))" where "OclIsEmpty x = ((\<upsilon> x and not (\<delta> x)) or ((OclSize x) \<doteq> \<zero>))"
@ -437,7 +437,7 @@ notation OclIsEmpty ("_->isEmpty\<^sub>B\<^sub>a\<^sub>g'(')" (*[66]*))
(*TODO Locale - Equivalent*) (*TODO Locale - Equivalent*)
subsection{* Definition: NotEmpty *} subsection\<open>Definition: NotEmpty\<close>
definition OclNotEmpty :: "('\<AA>,'\<alpha>::null) Bag \<Rightarrow> '\<AA> Boolean" definition OclNotEmpty :: "('\<AA>,'\<alpha>::null) Bag \<Rightarrow> '\<AA> Boolean"
where "OclNotEmpty x = not(OclIsEmpty x)" where "OclNotEmpty x = not(OclIsEmpty x)"
@ -445,7 +445,7 @@ notation OclNotEmpty ("_->notEmpty\<^sub>B\<^sub>a\<^sub>g'(')" (*[66]*))
(*TODO Locale - Equivalent*) (*TODO Locale - Equivalent*)
subsection{* Definition: Any *} subsection\<open>Definition: Any\<close>
(* Slight breach of naming convention in order to avoid naming conflict on constant.*) (* Slight breach of naming convention in order to avoid naming conflict on constant.*)
definition OclANY :: "[('\<AA>,'\<alpha>::null) Bag] \<Rightarrow> ('\<AA>,'\<alpha>) val" definition OclANY :: "[('\<AA>,'\<alpha>::null) Bag] \<Rightarrow> ('\<AA>,'\<alpha>) val"
@ -464,10 +464,10 @@ source->any\<^sub>B\<^sub>a\<^sub>g(iterator | body) =
source->select(iterator | body)->asSequence()->first(). Since we don't have sequences, source->select(iterator | body)->asSequence()->first(). Since we don't have sequences,
we have to go for a direct---restricted---definition. *) we have to go for a direct---restricted---definition. *)
subsection{* Definition: Forall *} subsection\<open>Definition: Forall\<close>
text{* The definition of OclForall mimics the one of @{term "OclAnd"}: text\<open>The definition of OclForall mimics the one of @{term "OclAnd"}:
OclForall is not a strict operation. *} OclForall is not a strict operation.\<close>
definition OclForall :: "[('\<AA>,'\<alpha>::null)Bag,('\<AA>,'\<alpha>)val\<Rightarrow>('\<AA>)Boolean] \<Rightarrow> '\<AA> Boolean" definition OclForall :: "[('\<AA>,'\<alpha>::null)Bag,('\<AA>,'\<alpha>)val\<Rightarrow>('\<AA>)Boolean] \<Rightarrow> '\<AA> Boolean"
where "OclForall S P = (\<lambda> \<tau>. if (\<delta> S) \<tau> = true \<tau> where "OclForall S P = (\<lambda> \<tau>. if (\<delta> S) \<tau> = true \<tau>
then if (\<exists>x\<in>Rep_Set_base S \<tau>. P (\<lambda>_. x) \<tau> = false \<tau>) then if (\<exists>x\<in>Rep_Set_base S \<tau>. P (\<lambda>_. x) \<tau> = false \<tau>)
@ -485,9 +485,9 @@ translations
(*TODO Locale - Equivalent*) (*TODO Locale - Equivalent*)
subsection{* Definition: Exists *} subsection\<open>Definition: Exists\<close>
text{* Like OclForall, OclExists is also not strict. *} text\<open>Like OclForall, OclExists is also not strict.\<close>
definition OclExists :: "[('\<AA>,'\<alpha>::null) Bag,('\<AA>,'\<alpha>)val\<Rightarrow>('\<AA>)Boolean] \<Rightarrow> '\<AA> Boolean" definition OclExists :: "[('\<AA>,'\<alpha>::null) Bag,('\<AA>,'\<alpha>)val\<Rightarrow>('\<AA>)Boolean] \<Rightarrow> '\<AA> Boolean"
where "OclExists S P = not(UML_Bag.OclForall S (\<lambda> X. not (P X)))" where "OclExists S P = not(UML_Bag.OclForall S (\<lambda> X. not (P X)))"
@ -498,7 +498,7 @@ translations
(*TODO Locale - Equivalent*) (*TODO Locale - Equivalent*)
subsection{* Definition: Iterate *} subsection\<open>Definition: Iterate\<close>
definition OclIterate :: "[('\<AA>,'\<alpha>::null) Bag,('\<AA>,'\<beta>::null)val, definition OclIterate :: "[('\<AA>,'\<alpha>::null) Bag,('\<AA>,'\<beta>::null)val,
('\<AA>,'\<alpha>)val\<Rightarrow>('\<AA>,'\<beta>)val\<Rightarrow>('\<AA>,'\<beta>)val] \<Rightarrow> ('\<AA>,'\<beta>)val" ('\<AA>,'\<alpha>)val\<Rightarrow>('\<AA>,'\<beta>)val\<Rightarrow>('\<AA>,'\<beta>)val] \<Rightarrow> ('\<AA>,'\<beta>)val"
@ -513,7 +513,7 @@ translations
(*TODO Locale - Equivalent*) (*TODO Locale - Equivalent*)
subsection{* Definition: Select *} subsection\<open>Definition: Select\<close>
definition OclSelect :: "[('\<AA>,'\<alpha>::null)Bag,('\<AA>,'\<alpha>)val\<Rightarrow>('\<AA>)Boolean] \<Rightarrow> ('\<AA>,'\<alpha>)Bag" definition OclSelect :: "[('\<AA>,'\<alpha>::null)Bag,('\<AA>,'\<alpha>)val\<Rightarrow>('\<AA>)Boolean] \<Rightarrow> ('\<AA>,'\<alpha>)Bag"
@ -534,7 +534,7 @@ translations
(*TODO Locale - Equivalent*) (*TODO Locale - Equivalent*)
subsection{* Definition: Reject *} subsection\<open>Definition: Reject\<close>
definition OclReject :: "[('\<AA>,'\<alpha>::null)Bag,('\<AA>,'\<alpha>)val\<Rightarrow>('\<AA>)Boolean] \<Rightarrow> ('\<AA>,'\<alpha>::null)Bag" definition OclReject :: "[('\<AA>,'\<alpha>::null)Bag,('\<AA>,'\<alpha>)val\<Rightarrow>('\<AA>)Boolean] \<Rightarrow> ('\<AA>,'\<alpha>::null)Bag"
where "OclReject S P = OclSelect S (not o P)" where "OclReject S P = OclSelect S (not o P)"
@ -545,7 +545,7 @@ translations
(*TODO Locale - Equivalent*) (*TODO Locale - Equivalent*)
subsection{* Definition: IncludesAll *} subsection\<open>Definition: IncludesAll\<close>
definition OclIncludesAll :: "[('\<AA>,'\<alpha>::null) Bag,('\<AA>,'\<alpha>) Bag] \<Rightarrow> '\<AA> Boolean" definition OclIncludesAll :: "[('\<AA>,'\<alpha>::null) Bag,('\<AA>,'\<alpha>) Bag] \<Rightarrow> '\<AA> Boolean"
where "OclIncludesAll x y = (\<lambda> \<tau>. if (\<delta> x) \<tau> = true \<tau> \<and> (\<delta> y) \<tau> = true \<tau> where "OclIncludesAll x y = (\<lambda> \<tau>. if (\<delta> x) \<tau> = true \<tau> \<and> (\<delta> y) \<tau> = true \<tau>
@ -557,7 +557,7 @@ interpretation OclIncludesAll : profile_bin\<^sub>d_\<^sub>d OclIncludesAll "\<l
by(unfold_locales, auto simp:OclIncludesAll_def bot_option_def null_option_def invalid_def by(unfold_locales, auto simp:OclIncludesAll_def bot_option_def null_option_def invalid_def
Rep_Bag_base_def Rep_Bag_base'_def) Rep_Bag_base_def Rep_Bag_base'_def)
subsection{* Definition: ExcludesAll *} subsection\<open>Definition: ExcludesAll\<close>
definition OclExcludesAll :: "[('\<AA>,'\<alpha>::null) Bag,('\<AA>,'\<alpha>) Bag] \<Rightarrow> '\<AA> Boolean" definition OclExcludesAll :: "[('\<AA>,'\<alpha>::null) Bag,('\<AA>,'\<alpha>) Bag] \<Rightarrow> '\<AA> Boolean"
where "OclExcludesAll x y = (\<lambda> \<tau>. if (\<delta> x) \<tau> = true \<tau> \<and> (\<delta> y) \<tau> = true \<tau> where "OclExcludesAll x y = (\<lambda> \<tau>. if (\<delta> x) \<tau> = true \<tau> \<and> (\<delta> y) \<tau> = true \<tau>
@ -569,7 +569,7 @@ interpretation OclExcludesAll : profile_bin\<^sub>d_\<^sub>d OclExcludesAll "\<l
by(unfold_locales, auto simp:OclExcludesAll_def bot_option_def null_option_def invalid_def by(unfold_locales, auto simp:OclExcludesAll_def bot_option_def null_option_def invalid_def
Rep_Bag_base_def Rep_Bag_base'_def) Rep_Bag_base_def Rep_Bag_base'_def)
subsection{* Definition: Union *} subsection\<open>Definition: Union\<close>
definition OclUnion :: "[('\<AA>,'\<alpha>::null) Bag,('\<AA>,'\<alpha>) Bag] \<Rightarrow> ('\<AA>,'\<alpha>) Bag" definition OclUnion :: "[('\<AA>,'\<alpha>::null) Bag,('\<AA>,'\<alpha>) Bag] \<Rightarrow> ('\<AA>,'\<alpha>) Bag"
where "OclUnion x y = (\<lambda> \<tau>. if (\<delta> x) \<tau> = true \<tau> \<and> (\<delta> y) \<tau> = true \<tau> where "OclUnion x y = (\<lambda> \<tau>. if (\<delta> x) \<tau> = true \<tau> \<and> (\<delta> y) \<tau> = true \<tau>
@ -592,7 +592,7 @@ proof -
null_option_def)+ null_option_def)+
qed qed
subsection{* Definition: Intersection *} subsection\<open>Definition: Intersection\<close>
definition OclIntersection :: "[('\<AA>,'\<alpha>::null) Bag,('\<AA>,'\<alpha>) Bag] \<Rightarrow> ('\<AA>,'\<alpha>) Bag" definition OclIntersection :: "[('\<AA>,'\<alpha>::null) Bag,('\<AA>,'\<alpha>) Bag] \<Rightarrow> ('\<AA>,'\<alpha>) Bag"
where "OclIntersection x y = (\<lambda> \<tau>. if (\<delta> x) \<tau> = true \<tau> \<and> (\<delta> y) \<tau> = true \<tau> where "OclIntersection x y = (\<lambda> \<tau>. if (\<delta> x) \<tau> = true \<tau> \<and> (\<delta> y) \<tau> = true \<tau>
@ -616,7 +616,7 @@ proof -
null_option_def)+ null_option_def)+
qed qed
subsection{* Definition: Count *} subsection\<open>Definition: Count\<close>
definition OclCount :: "[('\<AA>,'\<alpha>::null) Bag,('\<AA>,'\<alpha>) val] \<Rightarrow> ('\<AA>) Integer" definition OclCount :: "[('\<AA>,'\<alpha>::null) Bag,('\<AA>,'\<alpha>) val] \<Rightarrow> ('\<AA>) Integer"
where "OclCount x y = (\<lambda> \<tau>. if (\<delta> x) \<tau> = true \<tau> \<and> (\<delta> y) \<tau> = true \<tau> where "OclCount x y = (\<lambda> \<tau>. if (\<delta> x) \<tau> = true \<tau> \<and> (\<delta> y) \<tau> = true \<tau>
@ -627,16 +627,16 @@ notation OclCount ("_->count\<^sub>B\<^sub>a\<^sub>g'(_')" (*[66,65]65*))
interpretation OclCount : profile_bin\<^sub>d_\<^sub>d OclCount "\<lambda>x y. \<lfloor>\<lfloor>int(\<lceil>\<lceil>Rep_Bag\<^sub>b\<^sub>a\<^sub>s\<^sub>e x\<rceil>\<rceil> y)\<rfloor>\<rfloor>" interpretation OclCount : profile_bin\<^sub>d_\<^sub>d OclCount "\<lambda>x y. \<lfloor>\<lfloor>int(\<lceil>\<lceil>Rep_Bag\<^sub>b\<^sub>a\<^sub>s\<^sub>e x\<rceil>\<rceil> y)\<rfloor>\<rfloor>"
by(unfold_locales, auto simp:OclCount_def bot_option_def null_option_def) by(unfold_locales, auto simp:OclCount_def bot_option_def null_option_def)
subsection{* Definition (future operators) *} subsection\<open>Definition (future operators)\<close>
consts (* abstract bag collection operations *) consts (* abstract bag collection operations *)
OclSum :: " ('\<AA>,'\<alpha>::null) Bag \<Rightarrow> '\<AA> Integer" OclSum :: " ('\<AA>,'\<alpha>::null) Bag \<Rightarrow> '\<AA> Integer"
notation OclSum ("_->sum\<^sub>B\<^sub>a\<^sub>g'(')" (*[66]*)) notation OclSum ("_->sum\<^sub>B\<^sub>a\<^sub>g'(')" (*[66]*))
subsection{* Logical Properties *} subsection\<open>Logical Properties\<close>
text{* OclIncluding *} text\<open>OclIncluding\<close>
lemma OclIncluding_valid_args_valid: lemma OclIncluding_valid_args_valid:
"(\<tau> \<Turnstile> \<upsilon>(X->including\<^sub>B\<^sub>a\<^sub>g(x))) = ((\<tau> \<Turnstile>(\<delta> X)) \<and> (\<tau> \<Turnstile>(\<upsilon> x)))" "(\<tau> \<Turnstile> \<upsilon>(X->including\<^sub>B\<^sub>a\<^sub>g(x))) = ((\<tau> \<Turnstile>(\<delta> X)) \<and> (\<tau> \<Turnstile>(\<upsilon> x)))"
@ -646,10 +646,10 @@ lemma OclIncluding_valid_args_valid''[simp,code_unfold]:
"\<upsilon>(X->including\<^sub>B\<^sub>a\<^sub>g(x)) = ((\<delta> X) and (\<upsilon> x))" "\<upsilon>(X->including\<^sub>B\<^sub>a\<^sub>g(x)) = ((\<delta> X) and (\<upsilon> x))"
by (simp add: OclIncluding.def_valid_then_def) by (simp add: OclIncluding.def_valid_then_def)
text{* etc. etc. *} text\<open>etc. etc.\<close>
text_raw{* \isatagafp *} text_raw\<open>\isatagafp\<close>
text{* OclExcluding *} text\<open>OclExcluding\<close>
lemma OclExcluding_valid_args_valid: lemma OclExcluding_valid_args_valid:
"(\<tau> \<Turnstile> \<upsilon>(X->excluding\<^sub>B\<^sub>a\<^sub>g(x))) = ((\<tau> \<Turnstile>(\<delta> X)) \<and> (\<tau> \<Turnstile>(\<upsilon> x)))" "(\<tau> \<Turnstile> \<upsilon>(X->excluding\<^sub>B\<^sub>a\<^sub>g(x))) = ((\<tau> \<Turnstile>(\<delta> X)) \<and> (\<tau> \<Turnstile>(\<upsilon> x)))"
@ -659,7 +659,7 @@ lemma OclExcluding_valid_args_valid''[simp,code_unfold]:
"\<upsilon>(X->excluding\<^sub>B\<^sub>a\<^sub>g(x)) = ((\<delta> X) and (\<upsilon> x))" "\<upsilon>(X->excluding\<^sub>B\<^sub>a\<^sub>g(x)) = ((\<delta> X) and (\<upsilon> x))"
by (simp add: OclExcluding.def_valid_then_def) by (simp add: OclExcluding.def_valid_then_def)
text{* OclIncludes *} text\<open>OclIncludes\<close>
lemma OclIncludes_valid_args_valid: lemma OclIncludes_valid_args_valid:
"(\<tau> \<Turnstile> \<upsilon>(X->includes\<^sub>B\<^sub>a\<^sub>g(x))) = ((\<tau> \<Turnstile>(\<delta> X)) \<and> (\<tau> \<Turnstile>(\<upsilon> x)))" "(\<tau> \<Turnstile> \<upsilon>(X->includes\<^sub>B\<^sub>a\<^sub>g(x))) = ((\<tau> \<Turnstile>(\<delta> X)) \<and> (\<tau> \<Turnstile>(\<upsilon> x)))"
@ -669,7 +669,7 @@ lemma OclIncludes_valid_args_valid''[simp,code_unfold]:
"\<upsilon>(X->includes\<^sub>B\<^sub>a\<^sub>g(x)) = ((\<delta> X) and (\<upsilon> x))" "\<upsilon>(X->includes\<^sub>B\<^sub>a\<^sub>g(x)) = ((\<delta> X) and (\<upsilon> x))"
by (simp add: OclIncludes.def_valid_then_def) by (simp add: OclIncludes.def_valid_then_def)
text{* OclExcludes *} text\<open>OclExcludes\<close>
lemma OclExcludes_valid_args_valid: lemma OclExcludes_valid_args_valid:
"(\<tau> \<Turnstile> \<upsilon>(X->excludes\<^sub>B\<^sub>a\<^sub>g(x))) = ((\<tau> \<Turnstile>(\<delta> X)) \<and> (\<tau> \<Turnstile>(\<upsilon> x)))" "(\<tau> \<Turnstile> \<upsilon>(X->excludes\<^sub>B\<^sub>a\<^sub>g(x))) = ((\<tau> \<Turnstile>(\<delta> X)) \<and> (\<tau> \<Turnstile>(\<upsilon> x)))"
@ -679,7 +679,7 @@ lemma OclExcludes_valid_args_valid''[simp,code_unfold]:
"\<upsilon>(X->excludes\<^sub>B\<^sub>a\<^sub>g(x)) = ((\<delta> X) and (\<upsilon> x))" "\<upsilon>(X->excludes\<^sub>B\<^sub>a\<^sub>g(x)) = ((\<delta> X) and (\<upsilon> x))"
by (simp add: OclExcludes.def_valid_then_def) by (simp add: OclExcludes.def_valid_then_def)
text{* OclSize *} text\<open>OclSize\<close>
lemma OclSize_defined_args_valid: "\<tau> \<Turnstile> \<delta> (X->size\<^sub>B\<^sub>a\<^sub>g()) \<Longrightarrow> \<tau> \<Turnstile> \<delta> X" lemma OclSize_defined_args_valid: "\<tau> \<Turnstile> \<delta> (X->size\<^sub>B\<^sub>a\<^sub>g()) \<Longrightarrow> \<tau> \<Turnstile> \<delta> X"
by(auto simp: OclSize_def OclValid_def true_def valid_def false_def StrongEq_def by(auto simp: OclSize_def OclValid_def true_def valid_def false_def StrongEq_def
@ -712,7 +712,7 @@ lemma size_defined':
apply(simp add: defined_def bot_option_def bot_fun_def null_option_def null_fun_def X_finite) apply(simp add: defined_def bot_option_def bot_fun_def null_option_def null_fun_def X_finite)
done done
text{* OclIsEmpty *} text\<open>OclIsEmpty\<close>
lemma OclIsEmpty_defined_args_valid:"\<tau> \<Turnstile> \<delta> (X->isEmpty\<^sub>B\<^sub>a\<^sub>g()) \<Longrightarrow> \<tau> \<Turnstile> \<upsilon> X" lemma OclIsEmpty_defined_args_valid:"\<tau> \<Turnstile> \<delta> (X->isEmpty\<^sub>B\<^sub>a\<^sub>g()) \<Longrightarrow> \<tau> \<Turnstile> \<upsilon> X"
apply(auto simp: OclIsEmpty_def OclValid_def defined_def valid_def false_def true_def apply(auto simp: OclIsEmpty_def OclValid_def defined_def valid_def false_def true_def
@ -737,7 +737,7 @@ lemma OclIsEmpty_infinite: "\<tau> \<Turnstile> \<delta> X \<Longrightarrow> \<n
apply(case_tac x, simp add: null_option_def bot_option_def, simp) apply(case_tac x, simp add: null_option_def bot_option_def, simp)
by(simp add: OclSize_def StrictRefEq\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r valid_def bot_fun_def false_def true_def invalid_def) by(simp add: OclSize_def StrictRefEq\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r valid_def bot_fun_def false_def true_def invalid_def)
text{* OclNotEmpty *} text\<open>OclNotEmpty\<close>
lemma OclNotEmpty_defined_args_valid:"\<tau> \<Turnstile> \<delta> (X->notEmpty\<^sub>B\<^sub>a\<^sub>g()) \<Longrightarrow> \<tau> \<Turnstile> \<upsilon> X" lemma OclNotEmpty_defined_args_valid:"\<tau> \<Turnstile> \<delta> (X->notEmpty\<^sub>B\<^sub>a\<^sub>g()) \<Longrightarrow> \<tau> \<Turnstile> \<upsilon> X"
by (metis (hide_lams, no_types) OclNotEmpty_def OclNot_defargs OclNot_not foundation6 foundation9 by (metis (hide_lams, no_types) OclNotEmpty_def OclNot_defargs OclNot_not foundation6 foundation9
@ -779,7 +779,7 @@ lemma OclNotEmpty_has_elt' : "\<tau> \<Turnstile> \<delta> X \<Longrightarrow>
apply(drule OclNotEmpty_has_elt, simp) apply(drule OclNotEmpty_has_elt, simp)
by(simp add: Rep_Bag_base_def Rep_Set_base_def image_def) by(simp add: Rep_Bag_base_def Rep_Set_base_def image_def)
text{* OclANY *} text\<open>OclANY\<close>
lemma OclANY_defined_args_valid: "\<tau> \<Turnstile> \<delta> (X->any\<^sub>B\<^sub>a\<^sub>g()) \<Longrightarrow> \<tau> \<Turnstile> \<delta> X" lemma OclANY_defined_args_valid: "\<tau> \<Turnstile> \<delta> (X->any\<^sub>B\<^sub>a\<^sub>g()) \<Longrightarrow> \<tau> \<Turnstile> \<delta> X"
by(auto simp: OclANY_def OclValid_def true_def valid_def false_def StrongEq_def by(auto simp: OclANY_def OclValid_def true_def valid_def false_def StrongEq_def
@ -824,19 +824,19 @@ lemma OclANY_valid_args_valid''[simp,code_unfold]:
by(auto intro!: OclANY_valid_args_valid transform2_rev) by(auto intro!: OclANY_valid_args_valid transform2_rev)
(* and higher order ones : forall, exists, iterate, select, reject... *) (* and higher order ones : forall, exists, iterate, select, reject... *)
text_raw{* \endisatagafp *} text_raw\<open>\endisatagafp\<close>
subsection{* Execution Laws with Invalid or Null or Infinite Set as Argument *} subsection\<open>Execution Laws with Invalid or Null or Infinite Set as Argument\<close>
text{* OclIncluding *} (* properties already generated by the corresponding locale *) text\<open>OclIncluding\<close> (* properties already generated by the corresponding locale *)
text{* OclExcluding *} (* properties already generated by the corresponding locale *) text\<open>OclExcluding\<close> (* properties already generated by the corresponding locale *)
text{* OclIncludes *} (* properties already generated by the corresponding locale *) text\<open>OclIncludes\<close> (* properties already generated by the corresponding locale *)
text{* OclExcludes *} (* properties already generated by the corresponding locale *) text\<open>OclExcludes\<close> (* properties already generated by the corresponding locale *)
text{* OclSize *} text\<open>OclSize\<close>
lemma OclSize_invalid[simp,code_unfold]:"(invalid->size\<^sub>B\<^sub>a\<^sub>g()) = invalid" lemma OclSize_invalid[simp,code_unfold]:"(invalid->size\<^sub>B\<^sub>a\<^sub>g()) = invalid"
by(simp add: bot_fun_def OclSize_def invalid_def defined_def valid_def false_def true_def) by(simp add: bot_fun_def OclSize_def invalid_def defined_def valid_def false_def true_def)
@ -846,7 +846,7 @@ by(rule ext,
simp add: bot_fun_def null_fun_def null_is_valid OclSize_def simp add: bot_fun_def null_fun_def null_is_valid OclSize_def
invalid_def defined_def valid_def false_def true_def) invalid_def defined_def valid_def false_def true_def)
text{* OclIsEmpty *} text\<open>OclIsEmpty\<close>
lemma OclIsEmpty_invalid[simp,code_unfold]:"(invalid->isEmpty\<^sub>B\<^sub>a\<^sub>g()) = invalid" lemma OclIsEmpty_invalid[simp,code_unfold]:"(invalid->isEmpty\<^sub>B\<^sub>a\<^sub>g()) = invalid"
by(simp add: OclIsEmpty_def) by(simp add: OclIsEmpty_def)
@ -854,7 +854,7 @@ by(simp add: OclIsEmpty_def)
lemma OclIsEmpty_null[simp,code_unfold]:"(null->isEmpty\<^sub>B\<^sub>a\<^sub>g()) = true" lemma OclIsEmpty_null[simp,code_unfold]:"(null->isEmpty\<^sub>B\<^sub>a\<^sub>g()) = true"
by(simp add: OclIsEmpty_def) by(simp add: OclIsEmpty_def)
text{* OclNotEmpty *} text\<open>OclNotEmpty\<close>
lemma OclNotEmpty_invalid[simp,code_unfold]:"(invalid->notEmpty\<^sub>B\<^sub>a\<^sub>g()) = invalid" lemma OclNotEmpty_invalid[simp,code_unfold]:"(invalid->notEmpty\<^sub>B\<^sub>a\<^sub>g()) = invalid"
by(simp add: OclNotEmpty_def) by(simp add: OclNotEmpty_def)
@ -862,7 +862,7 @@ by(simp add: OclNotEmpty_def)
lemma OclNotEmpty_null[simp,code_unfold]:"(null->notEmpty\<^sub>B\<^sub>a\<^sub>g()) = false" lemma OclNotEmpty_null[simp,code_unfold]:"(null->notEmpty\<^sub>B\<^sub>a\<^sub>g()) = false"
by(simp add: OclNotEmpty_def) by(simp add: OclNotEmpty_def)
text{* OclANY *} text\<open>OclANY\<close>
lemma OclANY_invalid[simp,code_unfold]:"(invalid->any\<^sub>B\<^sub>a\<^sub>g()) = invalid" lemma OclANY_invalid[simp,code_unfold]:"(invalid->any\<^sub>B\<^sub>a\<^sub>g()) = invalid"
by(simp add: bot_fun_def OclANY_def invalid_def defined_def valid_def false_def true_def) by(simp add: bot_fun_def OclANY_def invalid_def defined_def valid_def false_def true_def)
@ -870,7 +870,7 @@ by(simp add: bot_fun_def OclANY_def invalid_def defined_def valid_def false_def
lemma OclANY_null[simp,code_unfold]:"(null->any\<^sub>B\<^sub>a\<^sub>g()) = null" lemma OclANY_null[simp,code_unfold]:"(null->any\<^sub>B\<^sub>a\<^sub>g()) = null"
by(simp add: OclANY_def false_def true_def) by(simp add: OclANY_def false_def true_def)
text{* OclForall *} text\<open>OclForall\<close>
lemma OclForall_invalid[simp,code_unfold]:"invalid->forAll\<^sub>B\<^sub>a\<^sub>g(a| P a) = invalid" lemma OclForall_invalid[simp,code_unfold]:"invalid->forAll\<^sub>B\<^sub>a\<^sub>g(a| P a) = invalid"
by(simp add: bot_fun_def invalid_def OclForall_def defined_def valid_def false_def true_def) by(simp add: bot_fun_def invalid_def OclForall_def defined_def valid_def false_def true_def)
@ -878,7 +878,7 @@ by(simp add: bot_fun_def invalid_def OclForall_def defined_def valid_def false_d
lemma OclForall_null[simp,code_unfold]:"null->forAll\<^sub>B\<^sub>a\<^sub>g(a | P a) = invalid" lemma OclForall_null[simp,code_unfold]:"null->forAll\<^sub>B\<^sub>a\<^sub>g(a | P a) = invalid"
by(simp add: bot_fun_def invalid_def OclForall_def defined_def valid_def false_def true_def) by(simp add: bot_fun_def invalid_def OclForall_def defined_def valid_def false_def true_def)
text{* OclExists *} text\<open>OclExists\<close>
lemma OclExists_invalid[simp,code_unfold]:"invalid->exists\<^sub>B\<^sub>a\<^sub>g(a| P a) = invalid" lemma OclExists_invalid[simp,code_unfold]:"invalid->exists\<^sub>B\<^sub>a\<^sub>g(a| P a) = invalid"
by(simp add: OclExists_def) by(simp add: OclExists_def)
@ -886,7 +886,7 @@ by(simp add: OclExists_def)
lemma OclExists_null[simp,code_unfold]:"null->exists\<^sub>B\<^sub>a\<^sub>g(a | P a) = invalid" lemma OclExists_null[simp,code_unfold]:"null->exists\<^sub>B\<^sub>a\<^sub>g(a | P a) = invalid"
by(simp add: OclExists_def) by(simp add: OclExists_def)
text{* OclIterate *} text\<open>OclIterate\<close>
lemma OclIterate_invalid[simp,code_unfold]:"invalid->iterate\<^sub>B\<^sub>a\<^sub>g(a; x = A | P a x) = invalid" lemma OclIterate_invalid[simp,code_unfold]:"invalid->iterate\<^sub>B\<^sub>a\<^sub>g(a; x = A | P a x) = invalid"
by(simp add: bot_fun_def invalid_def OclIterate_def defined_def valid_def false_def true_def) by(simp add: bot_fun_def invalid_def OclIterate_def defined_def valid_def false_def true_def)
@ -898,7 +898,7 @@ by(simp add: bot_fun_def invalid_def OclIterate_def defined_def valid_def false_
lemma OclIterate_invalid_args[simp,code_unfold]:"S->iterate\<^sub>B\<^sub>a\<^sub>g(a; x = invalid | P a x) = invalid" lemma OclIterate_invalid_args[simp,code_unfold]:"S->iterate\<^sub>B\<^sub>a\<^sub>g(a; x = invalid | P a x) = invalid"
by(simp add: bot_fun_def invalid_def OclIterate_def defined_def valid_def false_def true_def) by(simp add: bot_fun_def invalid_def OclIterate_def defined_def valid_def false_def true_def)
text{* An open question is this ... *} text\<open>An open question is this ...\<close>
lemma (*OclIterate_null_args[simp,code_unfold]:*) "S->iterate\<^sub>B\<^sub>a\<^sub>g(a; x = null | P a x) = invalid" lemma (*OclIterate_null_args[simp,code_unfold]:*) "S->iterate\<^sub>B\<^sub>a\<^sub>g(a; x = null | P a x) = invalid"
oops oops
(* In the definition above, this does not hold in general. (* In the definition above, this does not hold in general.
@ -911,7 +911,7 @@ apply(insert non_finite [THEN OclSize_infinite])
apply(subst (asm) foundation9, simp) apply(subst (asm) foundation9, simp)
by(metis OclIterate_def OclValid_def invalid_def) by(metis OclIterate_def OclValid_def invalid_def)
text{* OclSelect *} text\<open>OclSelect\<close>
lemma OclSelect_invalid[simp,code_unfold]:"invalid->select\<^sub>B\<^sub>a\<^sub>g(a | P a) = invalid" lemma OclSelect_invalid[simp,code_unfold]:"invalid->select\<^sub>B\<^sub>a\<^sub>g(a | P a) = invalid"
by(simp add: bot_fun_def invalid_def OclSelect_def defined_def valid_def false_def true_def) by(simp add: bot_fun_def invalid_def OclSelect_def defined_def valid_def false_def true_def)
@ -919,7 +919,7 @@ by(simp add: bot_fun_def invalid_def OclSelect_def defined_def valid_def false_d
lemma OclSelect_null[simp,code_unfold]:"null->select\<^sub>B\<^sub>a\<^sub>g(a | P a) = invalid" lemma OclSelect_null[simp,code_unfold]:"null->select\<^sub>B\<^sub>a\<^sub>g(a | P a) = invalid"
by(simp add: bot_fun_def invalid_def OclSelect_def defined_def valid_def false_def true_def) by(simp add: bot_fun_def invalid_def OclSelect_def defined_def valid_def false_def true_def)
text{* OclReject *} text\<open>OclReject\<close>
lemma OclReject_invalid[simp,code_unfold]:"invalid->reject\<^sub>B\<^sub>a\<^sub>g(a | P a) = invalid" lemma OclReject_invalid[simp,code_unfold]:"invalid->reject\<^sub>B\<^sub>a\<^sub>g(a | P a) = invalid"
by(simp add: OclReject_def) by(simp add: OclReject_def)
@ -927,9 +927,9 @@ by(simp add: OclReject_def)
lemma OclReject_null[simp,code_unfold]:"null->reject\<^sub>B\<^sub>a\<^sub>g(a | P a) = invalid" lemma OclReject_null[simp,code_unfold]:"null->reject\<^sub>B\<^sub>a\<^sub>g(a | P a) = invalid"
by(simp add: OclReject_def) by(simp add: OclReject_def)
text_raw{* \isatagafp *} text_raw\<open>\isatagafp\<close>
subsubsection{* Context Passing *} subsubsection\<open>Context Passing\<close>
lemma cp_OclIncludes1: lemma cp_OclIncludes1:
"(X->includes\<^sub>B\<^sub>a\<^sub>g(x)) \<tau> = (X->includes\<^sub>B\<^sub>a\<^sub>g(\<lambda> _. x \<tau>)) \<tau>" "(X->includes\<^sub>B\<^sub>a\<^sub>g(x)) \<tau> = (X->includes\<^sub>B\<^sub>a\<^sub>g(\<lambda> _. x \<tau>)) \<tau>"
@ -1012,7 +1012,7 @@ lemmas cp_intro''\<^sub>B\<^sub>a\<^sub>g[intro!,simp,code_unfold] =
cp_OclNotEmpty [THEN allI[THEN allI[THEN cpI1], of "OclNotEmpty"]] cp_OclNotEmpty [THEN allI[THEN allI[THEN cpI1], of "OclNotEmpty"]]
cp_OclANY [THEN allI[THEN allI[THEN cpI1], of "OclANY"]] cp_OclANY [THEN allI[THEN allI[THEN cpI1], of "OclANY"]]
subsubsection{* Const *} subsubsection\<open>Const\<close>
lemma const_OclIncluding[simp,code_unfold] : lemma const_OclIncluding[simp,code_unfold] :
assumes const_x : "const x" assumes const_x : "const x"
@ -1040,9 +1040,9 @@ lemma const_OclIncluding[simp,code_unfold] :
apply(subst (1 2) const_charn[OF const_S]) apply(subst (1 2) const_charn[OF const_S])
by simp by simp
qed qed
text_raw{* \endisatagafp *} text_raw\<open>\endisatagafp\<close>
subsection{* Test Statements *} subsection\<open>Test Statements\<close>
(*Assert "(\<tau> \<Turnstile> (Bag{\<lambda>_. \<lfloor>\<lfloor>x\<rfloor>\<rfloor>} \<doteq> Bag{\<lambda>_. \<lfloor>\<lfloor>x\<rfloor>\<rfloor>}))" (*Assert "(\<tau> \<Turnstile> (Bag{\<lambda>_. \<lfloor>\<lfloor>x\<rfloor>\<rfloor>} \<doteq> Bag{\<lambda>_. \<lfloor>\<lfloor>x\<rfloor>\<rfloor>}))"
Assert "(\<tau> \<Turnstile> (Bag{\<lambda>_. \<lfloor>x\<rfloor>} \<doteq> Bag{\<lambda>_. \<lfloor>x\<rfloor>}))"*) Assert "(\<tau> \<Turnstile> (Bag{\<lambda>_. \<lfloor>x\<rfloor>} \<doteq> Bag{\<lambda>_. \<lfloor>x\<rfloor>}))"*)

View File

@ -44,15 +44,15 @@ theory UML_Pair
imports "../UML_PropertyProfiles" imports "../UML_PropertyProfiles"
begin begin
section{* Collection Type Pairs: Operations \label{sec:collection_pairs} *} section\<open>Collection Type Pairs: Operations \label{sec:collection_pairs}\<close>
text{* The OCL standard provides the concept of \emph{Tuples}, \ie{} a family of record-types text\<open>The OCL standard provides the concept of \emph{Tuples}, \ie{} a family of record-types
with projection functions. In FeatherWeight OCL, only the theory of a special case is with projection functions. In FeatherWeight OCL, only the theory of a special case is
developped, namely the type of Pairs, which is, however, sufficient for all applications developped, namely the type of Pairs, which is, however, sufficient for all applications
since it can be used to mimick all tuples. In particular, it can be used to express operations since it can be used to mimick all tuples. In particular, it can be used to express operations
with multiple arguments, roles of n-ary associations, ... *} with multiple arguments, roles of n-ary associations, ...\<close>
subsection{* Semantic Properties of the Type Constructor *} subsection\<open>Semantic Properties of the Type Constructor\<close>
lemma A[simp]:"Rep_Pair\<^sub>b\<^sub>a\<^sub>s\<^sub>e x \<noteq> None \<Longrightarrow> Rep_Pair\<^sub>b\<^sub>a\<^sub>s\<^sub>e x \<noteq> null \<Longrightarrow> (fst \<lceil>\<lceil>Rep_Pair\<^sub>b\<^sub>a\<^sub>s\<^sub>e x\<rceil>\<rceil>) \<noteq> bot" lemma A[simp]:"Rep_Pair\<^sub>b\<^sub>a\<^sub>s\<^sub>e x \<noteq> None \<Longrightarrow> Rep_Pair\<^sub>b\<^sub>a\<^sub>s\<^sub>e x \<noteq> null \<Longrightarrow> (fst \<lceil>\<lceil>Rep_Pair\<^sub>b\<^sub>a\<^sub>s\<^sub>e x\<rceil>\<rceil>) \<noteq> bot"
by(insert Rep_Pair\<^sub>b\<^sub>a\<^sub>s\<^sub>e[of x],auto simp:null_option_def bot_option_def) by(insert Rep_Pair\<^sub>b\<^sub>a\<^sub>s\<^sub>e[of x],auto simp:null_option_def bot_option_def)
@ -82,12 +82,12 @@ apply(subst Rep_Pair\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject[symmetric], simp)
apply(subst Pair\<^sub>b\<^sub>a\<^sub>s\<^sub>e.Abs_Pair\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse, simp_all,simp add: null_option_def bot_option_def) apply(subst Pair\<^sub>b\<^sub>a\<^sub>s\<^sub>e.Abs_Pair\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse, simp_all,simp add: null_option_def bot_option_def)
done done
subsection{* Fundamental Properties of Strict Equality \label{sec:pair-strict-eq}*} subsection\<open>Fundamental Properties of Strict Equality \label{sec:pair-strict-eq}\<close>
text{* After the part of foundational operations on sets, we detail here equality on sets. text\<open>After the part of foundational operations on sets, we detail here equality on sets.
Strong equality is inherited from the OCL core, but we have to consider Strong equality is inherited from the OCL core, but we have to consider
the case of the strict equality. We decide to overload strict equality in the the case of the strict equality. We decide to overload strict equality in the
same way we do for other value's in OCL:*} same way we do for other value's in OCL:\<close>
overloading overloading
StrictRefEq \<equiv> "StrictRefEq :: [('\<AA>,'\<alpha>::null,'\<beta>::null)Pair,('\<AA>,'\<alpha>::null,'\<beta>::null)Pair] \<Rightarrow> ('\<AA>)Boolean" StrictRefEq \<equiv> "StrictRefEq :: [('\<AA>,'\<alpha>::null,'\<beta>::null)Pair,('\<AA>,'\<alpha>::null,'\<beta>::null)Pair] \<Rightarrow> ('\<AA>)Boolean"
@ -98,15 +98,15 @@ begin
else invalid \<tau>)" else invalid \<tau>)"
end 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"}*} text\<open>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"}\<close>
interpretation StrictRefEq\<^sub>P\<^sub>a\<^sub>i\<^sub>r : profile_bin\<^sub>S\<^sub>t\<^sub>r\<^sub>o\<^sub>n\<^sub>g\<^sub>E\<^sub>q_\<^sub>v_\<^sub>v "\<lambda> x y. (x::('\<AA>,'\<alpha>::null,'\<beta>::null)Pair) \<doteq> y" interpretation StrictRefEq\<^sub>P\<^sub>a\<^sub>i\<^sub>r : profile_bin\<^sub>S\<^sub>t\<^sub>r\<^sub>o\<^sub>n\<^sub>g\<^sub>E\<^sub>q_\<^sub>v_\<^sub>v "\<lambda> x y. (x::('\<AA>,'\<alpha>::null,'\<beta>::null)Pair) \<doteq> y"
by unfold_locales (auto simp: StrictRefEq\<^sub>P\<^sub>a\<^sub>i\<^sub>r) by unfold_locales (auto simp: StrictRefEq\<^sub>P\<^sub>a\<^sub>i\<^sub>r)
subsection{* Standard Operations Definitions *} subsection\<open>Standard Operations Definitions\<close>
text{* This part provides a collection of operators for the Pair type. *} text\<open>This part provides a collection of operators for the Pair type.\<close>
subsubsection{* Definition: Pair Constructor *} subsubsection\<open>Definition: Pair Constructor\<close>
definition OclPair::"('\<AA>, '\<alpha>) val \<Rightarrow> definition OclPair::"('\<AA>, '\<alpha>) val \<Rightarrow>
('\<AA>, '\<beta>) val \<Rightarrow> ('\<AA>, '\<beta>) val \<Rightarrow>
@ -121,7 +121,7 @@ interpretation OclPair : profile_bin\<^sub>v_\<^sub>v
by(auto simp: Abs_Pair\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject null_option_def bot_option_def) by(auto simp: Abs_Pair\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject null_option_def bot_option_def)
subsubsection{* Definition: First *} subsubsection\<open>Definition: First\<close>
definition OclFirst::" ('\<AA>,'\<alpha>::null,'\<beta>::null) Pair \<Rightarrow> ('\<AA>, '\<alpha>) val" (" _ .First'(')") definition OclFirst::" ('\<AA>,'\<alpha>::null,'\<beta>::null) Pair \<Rightarrow> ('\<AA>, '\<alpha>) val" (" _ .First'(')")
where "X .First() \<equiv> (\<lambda> \<tau>. if (\<delta> X) \<tau> = true \<tau> where "X .First() \<equiv> (\<lambda> \<tau>. if (\<delta> X) \<tau> = true \<tau>
@ -132,7 +132,7 @@ where "X .First() \<equiv> (\<lambda> \<tau>. if (\<delta> X) \<tau> = true
interpretation OclFirst : profile_mono\<^sub>d OclFirst "\<lambda>x. fst \<lceil>\<lceil>Rep_Pair\<^sub>b\<^sub>a\<^sub>s\<^sub>e (x)\<rceil>\<rceil>" interpretation OclFirst : profile_mono\<^sub>d OclFirst "\<lambda>x. fst \<lceil>\<lceil>Rep_Pair\<^sub>b\<^sub>a\<^sub>s\<^sub>e (x)\<rceil>\<rceil>"
by unfold_locales (auto simp: OclFirst_def) by unfold_locales (auto simp: OclFirst_def)
subsubsection{* Definition: Second *} subsubsection\<open>Definition: Second\<close>
definition OclSecond::" ('\<AA>,'\<alpha>::null,'\<beta>::null) Pair \<Rightarrow> ('\<AA>, '\<beta>) val" ("_ .Second'(')") definition OclSecond::" ('\<AA>,'\<alpha>::null,'\<beta>::null) Pair \<Rightarrow> ('\<AA>, '\<beta>) val" ("_ .Second'(')")
where "X .Second() \<equiv> (\<lambda> \<tau>. if (\<delta> X) \<tau> = true \<tau> where "X .Second() \<equiv> (\<lambda> \<tau>. if (\<delta> X) \<tau> = true \<tau>
@ -142,7 +142,7 @@ where "X .Second() \<equiv> (\<lambda> \<tau>. if (\<delta> X) \<tau> = true
interpretation OclSecond : profile_mono\<^sub>d OclSecond "\<lambda>x. snd \<lceil>\<lceil>Rep_Pair\<^sub>b\<^sub>a\<^sub>s\<^sub>e (x)\<rceil>\<rceil>" interpretation OclSecond : profile_mono\<^sub>d OclSecond "\<lambda>x. snd \<lceil>\<lceil>Rep_Pair\<^sub>b\<^sub>a\<^sub>s\<^sub>e (x)\<rceil>\<rceil>"
by unfold_locales (auto simp: OclSecond_def) by unfold_locales (auto simp: OclSecond_def)
subsection{* Logical Properties *} subsection\<open>Logical Properties\<close>
lemma 1 : "\<tau> \<Turnstile> \<upsilon> Y \<Longrightarrow> \<tau> \<Turnstile> Pair{X,Y} .First() \<triangleq> X" lemma 1 : "\<tau> \<Turnstile> \<upsilon> Y \<Longrightarrow> \<tau> \<Turnstile> Pair{X,Y} .First() \<triangleq> X"
apply(case_tac "\<not>(\<tau> \<Turnstile> \<upsilon> X)") apply(case_tac "\<not>(\<tau> \<Turnstile> \<upsilon> X)")
@ -162,7 +162,7 @@ apply(auto simp: OclValid_def valid_def defined_def StrongEq_def OclSecond_def O
apply(auto simp: Abs_Pair\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject null_option_def bot_option_def bot_Pair\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_Pair\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def) apply(auto simp: Abs_Pair\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject null_option_def bot_option_def bot_Pair\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_Pair\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def)
by(simp add: Abs_Pair\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse) by(simp add: Abs_Pair\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse)
subsection{* Algebraic Execution Properties *} subsection\<open>Algebraic Execution Properties\<close>
lemma proj1_exec [simp, code_unfold] : "Pair{X,Y} .First() = (if (\<upsilon> Y) then X else invalid endif)" lemma proj1_exec [simp, code_unfold] : "Pair{X,Y} .First() = (if (\<upsilon> Y) then X else invalid endif)"
apply(rule ext, rename_tac "\<tau>", simp add: foundation22[symmetric]) apply(rule ext, rename_tac "\<tau>", simp add: foundation22[symmetric])
@ -185,7 +185,7 @@ by(erule 2)
(* < *) (* < *)
subsection{* Test Statements*} subsection\<open>Test Statements\<close>
(* (*
Assert "(\<tau> \<Turnstile> (Pair{\<lambda>_. \<lfloor>\<lfloor>x\<rfloor>\<rfloor>,\<lambda>_. \<lfloor>\<lfloor>x\<rfloor>\<rfloor>} \<doteq> Pair{\<lambda>_. \<lfloor>\<lfloor>x\<rfloor>\<rfloor>,\<lambda>_. \<lfloor>\<lfloor>x\<rfloor>\<rfloor>}))" Assert "(\<tau> \<Turnstile> (Pair{\<lambda>_. \<lfloor>\<lfloor>x\<rfloor>\<rfloor>,\<lambda>_. \<lfloor>\<lfloor>x\<rfloor>\<rfloor>} \<doteq> Pair{\<lambda>_. \<lfloor>\<lfloor>x\<rfloor>\<rfloor>,\<lambda>_. \<lfloor>\<lfloor>x\<rfloor>\<rfloor>}))"
Assert "(\<tau> \<Turnstile> (Pair{\<lambda>_. \<lfloor>x\<rfloor>,\<lambda>_. \<lfloor>x\<rfloor>} \<doteq> Pair{\<lambda>_. \<lfloor>x\<rfloor>,\<lambda>_. \<lfloor>x\<rfloor>}))" Assert "(\<tau> \<Turnstile> (Pair{\<lambda>_. \<lfloor>x\<rfloor>,\<lambda>_. \<lfloor>x\<rfloor>} \<doteq> Pair{\<lambda>_. \<lfloor>x\<rfloor>,\<lambda>_. \<lfloor>x\<rfloor>}))"

View File

@ -47,11 +47,11 @@ imports "../basic_types/UML_Boolean"
begin begin
no_notation None ("\<bottom>") no_notation None ("\<bottom>")
section{* Collection Type Sequence: Operations *} section\<open>Collection Type Sequence: Operations\<close>
subsection{* Basic Properties of the Sequence Type *} subsection\<open>Basic Properties of the Sequence Type\<close>
text{* Every element in a defined sequence is valid. *} text\<open>Every element in a defined sequence is valid.\<close>
lemma Sequence_inv_lemma: "\<tau> \<Turnstile> (\<delta> X) \<Longrightarrow> \<forall>x\<in>set \<lceil>\<lceil>Rep_Sequence\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \<tau>)\<rceil>\<rceil>. x \<noteq> bot" lemma Sequence_inv_lemma: "\<tau> \<Turnstile> (\<delta> X) \<Longrightarrow> \<forall>x\<in>set \<lceil>\<lceil>Rep_Sequence\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \<tau>)\<rceil>\<rceil>. x \<noteq> bot"
apply(insert Rep_Sequence\<^sub>b\<^sub>a\<^sub>s\<^sub>e [of "X \<tau>"], simp) apply(insert Rep_Sequence\<^sub>b\<^sub>a\<^sub>s\<^sub>e [of "X \<tau>"], simp)
@ -66,12 +66,12 @@ apply(subst Abs_Sequence\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject[symmetric], rule
apply(simp add: Rep_Sequence\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse null_option_def) apply(simp add: Rep_Sequence\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse null_option_def)
by (simp add: bot_option_def) by (simp add: bot_option_def)
subsection{* Definition: Strict Equality \label{sec:seq-strict-equality}*} subsection\<open>Definition: Strict Equality \label{sec:seq-strict-equality}\<close>
text{* After the part of foundational operations on sets, we detail here equality on sets. text\<open>After the part of foundational operations on sets, we detail here equality on sets.
Strong equality is inherited from the OCL core, but we have to consider Strong equality is inherited from the OCL core, but we have to consider
the case of the strict equality. We decide to overload strict equality in the the case of the strict equality. We decide to overload strict equality in the
same way we do for other value's in OCL:*} same way we do for other value's in OCL:\<close>
overloading overloading
StrictRefEq \<equiv> "StrictRefEq :: [('\<AA>,'\<alpha>::null)Sequence,('\<AA>,'\<alpha>::null)Sequence] \<Rightarrow> ('\<AA>)Boolean" StrictRefEq \<equiv> "StrictRefEq :: [('\<AA>,'\<alpha>::null)Sequence,('\<AA>,'\<alpha>::null)Sequence] \<Rightarrow> ('\<AA>)Boolean"
@ -82,22 +82,22 @@ begin
else invalid \<tau>)" else invalid \<tau>)"
end end
text_raw{* \isatagafp *} text_raw\<open>\isatagafp\<close>
text{* One might object here that for the case of objects, this is an empty definition. text\<open>One might object here that for the case of objects, this is an empty definition.
The answer is no, we will restrain later on states and objects such that any object The answer is no, we will restrain later on states and objects such that any object
has its oid stored inside the object (so the ref, under which an object can be referenced has its oid stored inside the object (so the ref, under which an object can be referenced
in the store will represented in the object itself). For such well-formed stores that satisfy in the store will represented in the object itself). For such well-formed stores that satisfy
this invariant (the WFF-invariant), the referential equality and the this invariant (the WFF-invariant), the referential equality and the
strong equality---and therefore the strict equality on sequences in the sense above---coincides.*} strong equality---and therefore the strict equality on sequences in the sense above---coincides.\<close>
text_raw{* \endisatagafp *} text_raw\<open>\endisatagafp\<close>
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"}*} text\<open>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"}\<close>
interpretation StrictRefEq\<^sub>S\<^sub>e\<^sub>q : profile_bin\<^sub>S\<^sub>t\<^sub>r\<^sub>o\<^sub>n\<^sub>g\<^sub>E\<^sub>q_\<^sub>v_\<^sub>v "\<lambda> x y. (x::('\<AA>,'\<alpha>::null)Sequence) \<doteq> y" interpretation StrictRefEq\<^sub>S\<^sub>e\<^sub>q : profile_bin\<^sub>S\<^sub>t\<^sub>r\<^sub>o\<^sub>n\<^sub>g\<^sub>E\<^sub>q_\<^sub>v_\<^sub>v "\<lambda> x y. (x::('\<AA>,'\<alpha>::null)Sequence) \<doteq> y"
by unfold_locales (auto simp: StrictRefEq\<^sub>S\<^sub>e\<^sub>q) by unfold_locales (auto simp: StrictRefEq\<^sub>S\<^sub>e\<^sub>q)
subsection{* Constants: mtSequence *} subsection\<open>Constants: mtSequence\<close>
definition mtSequence ::"('\<AA>,'\<alpha>::null) Sequence" ("Sequence{}") definition mtSequence ::"('\<AA>,'\<alpha>::null) Sequence" ("Sequence{}")
where "Sequence{} \<equiv> (\<lambda> \<tau>. Abs_Sequence\<^sub>b\<^sub>a\<^sub>s\<^sub>e \<lfloor>\<lfloor>[]::'\<alpha> list\<rfloor>\<rfloor> )" where "Sequence{} \<equiv> (\<lambda> \<tau>. Abs_Sequence\<^sub>b\<^sub>a\<^sub>s\<^sub>e \<lfloor>\<lfloor>[]::'\<alpha> list\<rfloor>\<rfloor> )"
@ -116,18 +116,18 @@ lemma mtSequence_rep_set: "\<lceil>\<lceil>Rep_Sequence\<^sub>b\<^sub>a\<^sub>s\
apply(simp add: mtSequence_def, subst Abs_Sequence\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse) apply(simp add: mtSequence_def, subst Abs_Sequence\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse)
by(simp add: bot_option_def)+ by(simp add: bot_option_def)+
text_raw{* \isatagafp *} text_raw\<open>\isatagafp\<close>
lemma [simp,code_unfold]: "const Sequence{}" lemma [simp,code_unfold]: "const Sequence{}"
by(simp add: const_def mtSequence_def) by(simp add: const_def mtSequence_def)
text{* Note that the collection types in OCL allow for null to be included; text\<open>Note that the collection types in OCL allow for null to be included;
however, there is the null-collection into which inclusion yields invalid. *} however, there is the null-collection into which inclusion yields invalid.\<close>
text_raw{* \endisatagafp *} text_raw\<open>\endisatagafp\<close>
subsection{* Definition: Prepend *} subsection\<open>Definition: Prepend\<close>
definition OclPrepend :: "[('\<AA>,'\<alpha>::null) Sequence,('\<AA>,'\<alpha>) val] \<Rightarrow> ('\<AA>,'\<alpha>) Sequence" definition OclPrepend :: "[('\<AA>,'\<alpha>::null) Sequence,('\<AA>,'\<alpha>) val] \<Rightarrow> ('\<AA>,'\<alpha>) Sequence"
where "OclPrepend x y = (\<lambda> \<tau>. if (\<delta> x) \<tau> = true \<tau> \<and> (\<upsilon> y) \<tau> = true \<tau> where "OclPrepend x y = (\<lambda> \<tau>. if (\<delta> x) \<tau> = true \<tau> \<and> (\<upsilon> y) \<tau> = true \<tau>
then Abs_Sequence\<^sub>b\<^sub>a\<^sub>s\<^sub>e \<lfloor>\<lfloor> (y \<tau>)#\<lceil>\<lceil>Rep_Sequence\<^sub>b\<^sub>a\<^sub>s\<^sub>e (x \<tau>)\<rceil>\<rceil> \<rfloor>\<rfloor> then Abs_Sequence\<^sub>b\<^sub>a\<^sub>s\<^sub>e \<lfloor>\<lfloor> (y \<tau>)#\<lceil>\<lceil>Rep_Sequence\<^sub>b\<^sub>a\<^sub>s\<^sub>e (x \<tau>)\<rceil>\<rceil> \<rfloor>\<rfloor>
@ -163,7 +163,7 @@ translations
"Sequence{x, xs}" == "CONST OclPrepend (Sequence{xs}) x" "Sequence{x, xs}" == "CONST OclPrepend (Sequence{xs}) x"
"Sequence{x}" == "CONST OclPrepend (Sequence{}) x " "Sequence{x}" == "CONST OclPrepend (Sequence{}) x "
subsection{* Definition: Including *} subsection\<open>Definition: Including\<close>
definition OclIncluding :: "[('\<AA>,'\<alpha>::null) Sequence,('\<AA>,'\<alpha>) val] \<Rightarrow> ('\<AA>,'\<alpha>) Sequence" definition OclIncluding :: "[('\<AA>,'\<alpha>::null) Sequence,('\<AA>,'\<alpha>) val] \<Rightarrow> ('\<AA>,'\<alpha>) Sequence"
where "OclIncluding x y = (\<lambda> \<tau>. if (\<delta> x) \<tau> = true \<tau> \<and> (\<upsilon> y) \<tau> = true \<tau> where "OclIncluding x y = (\<lambda> \<tau>. if (\<delta> x) \<tau> = true \<tau> \<and> (\<upsilon> y) \<tau> = true \<tau>
@ -223,7 +223,7 @@ lemma [simp,code_unfold] : "((S->prepend\<^sub>S\<^sub>e\<^sub>q(a))->including\
by (metis OclValid_def foundation16 invalid_def)+ by (metis OclValid_def foundation16 invalid_def)+
qed qed
subsection{* Definition: Excluding *} subsection\<open>Definition: Excluding\<close>
definition OclExcluding :: "[('\<AA>,'\<alpha>::null) Sequence,('\<AA>,'\<alpha>) val] \<Rightarrow> ('\<AA>,'\<alpha>) Sequence" definition OclExcluding :: "[('\<AA>,'\<alpha>::null) Sequence,('\<AA>,'\<alpha>) val] \<Rightarrow> ('\<AA>,'\<alpha>) Sequence"
where "OclExcluding x y = (\<lambda> \<tau>. if (\<delta> x) \<tau> = true \<tau> \<and> (\<upsilon> y) \<tau> = true \<tau> where "OclExcluding x y = (\<lambda> \<tau>. if (\<delta> x) \<tau> = true \<tau> \<and> (\<upsilon> y) \<tau> = true \<tau>
then Abs_Sequence\<^sub>b\<^sub>a\<^sub>s\<^sub>e \<lfloor>\<lfloor> filter (\<lambda>x. x = y \<tau>) then Abs_Sequence\<^sub>b\<^sub>a\<^sub>s\<^sub>e \<lfloor>\<lfloor> filter (\<lambda>x. x = y \<tau>)
@ -243,8 +243,8 @@ proof -
done done
qed qed
subsection{* Definition: Append *} subsection\<open>Definition: Append\<close>
text{* Identical to OclIncluding. *} text\<open>Identical to OclIncluding.\<close>
definition OclAppend :: "[('\<AA>,'\<alpha>::null) Sequence,('\<AA>,'\<alpha>) val] \<Rightarrow> ('\<AA>,'\<alpha>) Sequence" definition OclAppend :: "[('\<AA>,'\<alpha>::null) Sequence,('\<AA>,'\<alpha>) val] \<Rightarrow> ('\<AA>,'\<alpha>) Sequence"
where "OclAppend = OclIncluding" where "OclAppend = OclIncluding"
notation OclAppend ("_->append\<^sub>S\<^sub>e\<^sub>q'(_')") notation OclAppend ("_->append\<^sub>S\<^sub>e\<^sub>q'(_')")
@ -255,7 +255,7 @@ interpretation OclAppend :
by(auto simp: OclAppend_def bin_def bin'_def by(auto simp: OclAppend_def bin_def bin'_def
OclIncluding.def_scheme OclIncluding.def_body) OclIncluding.def_scheme OclIncluding.def_body)
subsection{* Definition: Union *} subsection\<open>Definition: Union\<close>
definition OclUnion :: "[('\<AA>,'\<alpha>::null) Sequence,('\<AA>,'\<alpha>) Sequence] \<Rightarrow> ('\<AA>,'\<alpha>) Sequence" definition OclUnion :: "[('\<AA>,'\<alpha>::null) Sequence,('\<AA>,'\<alpha>) Sequence] \<Rightarrow> ('\<AA>,'\<alpha>) Sequence"
where "OclUnion x y = (\<lambda> \<tau>. if (\<delta> x) \<tau> = true \<tau> \<and> (\<delta> y) \<tau> = true \<tau> where "OclUnion x y = (\<lambda> \<tau>. if (\<delta> x) \<tau> = true \<tau> \<and> (\<delta> y) \<tau> = true \<tau>
then Abs_Sequence\<^sub>b\<^sub>a\<^sub>s\<^sub>e \<lfloor>\<lfloor> \<lceil>\<lceil>Rep_Sequence\<^sub>b\<^sub>a\<^sub>s\<^sub>e (x \<tau>)\<rceil>\<rceil> @ then Abs_Sequence\<^sub>b\<^sub>a\<^sub>s\<^sub>e \<lfloor>\<lfloor> \<lceil>\<lceil>Rep_Sequence\<^sub>b\<^sub>a\<^sub>s\<^sub>e (x \<tau>)\<rceil>\<rceil> @
@ -277,7 +277,7 @@ proof -
simp_all add: bot_option_def null_option_def Set.ball_Un A null_Sequence\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def bot_Sequence\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def)+ simp_all add: bot_option_def null_option_def Set.ball_Un A null_Sequence\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def bot_Sequence\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def)+
qed qed
subsection{* Definition: At *} subsection\<open>Definition: At\<close>
definition OclAt :: "[('\<AA>,'\<alpha>::null) Sequence,('\<AA>) Integer] \<Rightarrow> ('\<AA>,'\<alpha>) val" definition OclAt :: "[('\<AA>,'\<alpha>::null) Sequence,('\<AA>) Integer] \<Rightarrow> ('\<AA>,'\<alpha>) val"
where "OclAt x y = (\<lambda> \<tau>. if (\<delta> x) \<tau> = true \<tau> \<and> (\<delta> y) \<tau> = true \<tau> where "OclAt x y = (\<lambda> \<tau>. if (\<delta> x) \<tau> = true \<tau> \<and> (\<delta> y) \<tau> = true \<tau>
then if 1 \<le> \<lceil>\<lceil>y \<tau>\<rceil>\<rceil> \<and> \<lceil>\<lceil>y \<tau>\<rceil>\<rceil> \<le> length\<lceil>\<lceil>Rep_Sequence\<^sub>b\<^sub>a\<^sub>s\<^sub>e (x \<tau>)\<rceil>\<rceil> then if 1 \<le> \<lceil>\<lceil>y \<tau>\<rceil>\<rceil> \<and> \<lceil>\<lceil>y \<tau>\<rceil>\<rceil> \<le> length\<lceil>\<lceil>Rep_Sequence\<^sub>b\<^sub>a\<^sub>s\<^sub>e (x \<tau>)\<rceil>\<rceil>
@ -288,7 +288,7 @@ notation OclAt ("_->at\<^sub>S\<^sub>e\<^sub>q'(_')")
(*TODO Locale - Equivalent*) (*TODO Locale - Equivalent*)
subsection{* Definition: First *} subsection\<open>Definition: First\<close>
definition OclFirst :: "[('\<AA>,'\<alpha>::null) Sequence] \<Rightarrow> ('\<AA>,'\<alpha>) val" definition OclFirst :: "[('\<AA>,'\<alpha>::null) Sequence] \<Rightarrow> ('\<AA>,'\<alpha>) val"
where "OclFirst x = (\<lambda> \<tau>. if (\<delta> x) \<tau> = true \<tau> then where "OclFirst x = (\<lambda> \<tau>. if (\<delta> x) \<tau> = true \<tau> then
case \<lceil>\<lceil>Rep_Sequence\<^sub>b\<^sub>a\<^sub>s\<^sub>e (x \<tau>)\<rceil>\<rceil> of [] \<Rightarrow> invalid \<tau> case \<lceil>\<lceil>Rep_Sequence\<^sub>b\<^sub>a\<^sub>s\<^sub>e (x \<tau>)\<rceil>\<rceil> of [] \<Rightarrow> invalid \<tau>
@ -298,7 +298,7 @@ notation OclFirst ("_->first\<^sub>S\<^sub>e\<^sub>q'(_')")
(*TODO Locale - Equivalent*) (*TODO Locale - Equivalent*)
subsection{* Definition: Last *} subsection\<open>Definition: Last\<close>
definition OclLast :: "[('\<AA>,'\<alpha>::null) Sequence] \<Rightarrow> ('\<AA>,'\<alpha>) val" definition OclLast :: "[('\<AA>,'\<alpha>::null) Sequence] \<Rightarrow> ('\<AA>,'\<alpha>) val"
where "OclLast x = (\<lambda> \<tau>. if (\<delta> x) \<tau> = true \<tau> then where "OclLast x = (\<lambda> \<tau>. if (\<delta> x) \<tau> = true \<tau> then
if \<lceil>\<lceil>Rep_Sequence\<^sub>b\<^sub>a\<^sub>s\<^sub>e (x \<tau>)\<rceil>\<rceil> = [] then if \<lceil>\<lceil>Rep_Sequence\<^sub>b\<^sub>a\<^sub>s\<^sub>e (x \<tau>)\<rceil>\<rceil> = [] then
@ -309,7 +309,7 @@ where "OclLast x = (\<lambda> \<tau>. if (\<delta> x) \<tau> = true \<tau> t
notation OclLast ("_->last\<^sub>S\<^sub>e\<^sub>q'(_')") notation OclLast ("_->last\<^sub>S\<^sub>e\<^sub>q'(_')")
(*TODO Locale - Equivalent*) (*TODO Locale - Equivalent*)
subsection{* Definition: Iterate *} subsection\<open>Definition: Iterate\<close>
definition OclIterate :: "[('\<AA>,'\<alpha>::null) Sequence,('\<AA>,'\<beta>::null)val, definition OclIterate :: "[('\<AA>,'\<alpha>::null) Sequence,('\<AA>,'\<beta>::null)val,
('\<AA>,'\<alpha>)val\<Rightarrow>('\<AA>,'\<beta>)val\<Rightarrow>('\<AA>,'\<beta>)val] \<Rightarrow> ('\<AA>,'\<beta>)val" ('\<AA>,'\<alpha>)val\<Rightarrow>('\<AA>,'\<beta>)val\<Rightarrow>('\<AA>,'\<beta>)val] \<Rightarrow> ('\<AA>,'\<beta>)val"
@ -326,7 +326,7 @@ translations
subsection{* Definition: Forall *} subsection\<open>Definition: Forall\<close>
definition OclForall :: "[('\<AA>,'\<alpha>::null) Sequence,('\<AA>,'\<alpha>)val\<Rightarrow>('\<AA>)Boolean] \<Rightarrow> '\<AA> Boolean" definition OclForall :: "[('\<AA>,'\<alpha>::null) Sequence,('\<AA>,'\<alpha>)val\<Rightarrow>('\<AA>)Boolean] \<Rightarrow> '\<AA> Boolean"
where "OclForall S P = (S->iterate\<^sub>S\<^sub>e\<^sub>q(b; x = true | x and (P b)))" where "OclForall S P = (S->iterate\<^sub>S\<^sub>e\<^sub>q(b; x = true | x and (P b)))"
@ -337,7 +337,7 @@ translations
(*TODO Locale - Equivalent*) (*TODO Locale - Equivalent*)
subsection{* Definition: Exists *} subsection\<open>Definition: Exists\<close>
definition OclExists :: "[('\<AA>,'\<alpha>::null) Sequence,('\<AA>,'\<alpha>)val\<Rightarrow>('\<AA>)Boolean] \<Rightarrow> '\<AA> Boolean" definition OclExists :: "[('\<AA>,'\<alpha>::null) Sequence,('\<AA>,'\<alpha>)val\<Rightarrow>('\<AA>)Boolean] \<Rightarrow> '\<AA> Boolean"
where "OclExists S P = (S->iterate\<^sub>S\<^sub>e\<^sub>q(b; x = false | x or (P b)))" where "OclExists S P = (S->iterate\<^sub>S\<^sub>e\<^sub>q(b; x = false | x or (P b)))"
@ -348,7 +348,7 @@ translations
(*TODO Locale - Equivalent*) (*TODO Locale - Equivalent*)
subsection{* Definition: Collect *} subsection\<open>Definition: Collect\<close>
definition OclCollect :: "[('\<AA>,'\<alpha>::null)Sequence,('\<AA>,'\<alpha>)val\<Rightarrow>('\<AA>,'\<beta>)val]\<Rightarrow>('\<AA>,'\<beta>::null)Sequence" definition OclCollect :: "[('\<AA>,'\<alpha>::null)Sequence,('\<AA>,'\<alpha>)val\<Rightarrow>('\<AA>,'\<beta>)val]\<Rightarrow>('\<AA>,'\<beta>::null)Sequence"
where "OclCollect S P = (S->iterate\<^sub>S\<^sub>e\<^sub>q(b; x = Sequence{} | x->prepend\<^sub>S\<^sub>e\<^sub>q(P b)))" where "OclCollect S P = (S->iterate\<^sub>S\<^sub>e\<^sub>q(b; x = Sequence{} | x->prepend\<^sub>S\<^sub>e\<^sub>q(P b)))"
@ -359,7 +359,7 @@ translations
(*TODO Locale - Equivalent*) (*TODO Locale - Equivalent*)
subsection{* Definition: Select *} subsection\<open>Definition: Select\<close>
definition OclSelect :: "[('\<AA>,'\<alpha>::null)Sequence,('\<AA>,'\<alpha>)val\<Rightarrow>('\<AA>)Boolean]\<Rightarrow>('\<AA>,'\<alpha>::null)Sequence" definition OclSelect :: "[('\<AA>,'\<alpha>::null)Sequence,('\<AA>,'\<alpha>)val\<Rightarrow>('\<AA>)Boolean]\<Rightarrow>('\<AA>,'\<alpha>::null)Sequence"
where "OclSelect S P = where "OclSelect S P =
(S->iterate\<^sub>S\<^sub>e\<^sub>q(b; x = Sequence{} | if P b then x->prepend\<^sub>S\<^sub>e\<^sub>q(b) else x endif))" (S->iterate\<^sub>S\<^sub>e\<^sub>q(b; x = Sequence{} | if P b then x->prepend\<^sub>S\<^sub>e\<^sub>q(b) else x endif))"
@ -371,20 +371,20 @@ translations
(*TODO Locale - Equivalent*) (*TODO Locale - Equivalent*)
subsection{* Definition: Size *} subsection\<open>Definition: Size\<close>
definition OclSize :: "[('\<AA>,'\<alpha>::null)Sequence]\<Rightarrow>('\<AA>)Integer" ("(_)->size\<^sub>S\<^sub>e\<^sub>q'(')") definition OclSize :: "[('\<AA>,'\<alpha>::null)Sequence]\<Rightarrow>('\<AA>)Integer" ("(_)->size\<^sub>S\<^sub>e\<^sub>q'(')")
where "OclSize S = (S->iterate\<^sub>S\<^sub>e\<^sub>q(b; x = \<zero> | x +\<^sub>i\<^sub>n\<^sub>t \<one> ))" where "OclSize S = (S->iterate\<^sub>S\<^sub>e\<^sub>q(b; x = \<zero> | x +\<^sub>i\<^sub>n\<^sub>t \<one> ))"
(*TODO Locale - Equivalent*) (*TODO Locale - Equivalent*)
subsection{* Definition: IsEmpty *} subsection\<open>Definition: IsEmpty\<close>
definition OclIsEmpty :: "('\<AA>,'\<alpha>::null) Sequence \<Rightarrow> '\<AA> Boolean" definition OclIsEmpty :: "('\<AA>,'\<alpha>::null) Sequence \<Rightarrow> '\<AA> Boolean"
where "OclIsEmpty x = ((\<upsilon> x and not (\<delta> x)) or ((OclSize x) \<doteq> \<zero>))" where "OclIsEmpty x = ((\<upsilon> x and not (\<delta> x)) or ((OclSize x) \<doteq> \<zero>))"
notation OclIsEmpty ("_->isEmpty\<^sub>S\<^sub>e\<^sub>q'(')" (*[66]*)) notation OclIsEmpty ("_->isEmpty\<^sub>S\<^sub>e\<^sub>q'(')" (*[66]*))
(*TODO Locale - Equivalent*) (*TODO Locale - Equivalent*)
subsection{* Definition: NotEmpty *} subsection\<open>Definition: NotEmpty\<close>
definition OclNotEmpty :: "('\<AA>,'\<alpha>::null) Sequence \<Rightarrow> '\<AA> Boolean" definition OclNotEmpty :: "('\<AA>,'\<alpha>::null) Sequence \<Rightarrow> '\<AA> Boolean"
where "OclNotEmpty x = not(OclIsEmpty x)" where "OclNotEmpty x = not(OclIsEmpty x)"
@ -392,7 +392,7 @@ notation OclNotEmpty ("_->notEmpty\<^sub>S\<^sub>e\<^sub>q'(')" (*[66]*))
(*TODO Locale - Equivalent*) (*TODO Locale - Equivalent*)
subsection{* Definition: Any *} subsection\<open>Definition: Any\<close>
definition "OclANY x = (\<lambda> \<tau>. definition "OclANY x = (\<lambda> \<tau>.
if x \<tau> = invalid \<tau> then if x \<tau> = invalid \<tau> then
@ -404,7 +404,7 @@ notation OclANY ("_->any\<^sub>S\<^sub>e\<^sub>q'(')")
(*TODO Locale - Equivalent*) (*TODO Locale - Equivalent*)
subsection{* Definition (future operators) *} subsection\<open>Definition (future operators)\<close>
consts (* abstract set collection operations *) consts (* abstract set collection operations *)
OclCount :: "[('\<AA>,'\<alpha>::null) Sequence,('\<AA>,'\<alpha>) Sequence] \<Rightarrow> '\<AA> Integer" OclCount :: "[('\<AA>,'\<alpha>::null) Sequence,('\<AA>,'\<alpha>) Sequence] \<Rightarrow> '\<AA> Integer"
@ -418,11 +418,11 @@ consts (* abstract set collection operations *)
notation OclCount ("_->count\<^sub>S\<^sub>e\<^sub>q'(_')" (*[66,65]65*)) notation OclCount ("_->count\<^sub>S\<^sub>e\<^sub>q'(_')" (*[66,65]65*))
notation OclSum ("_->sum\<^sub>S\<^sub>e\<^sub>q'(')" (*[66]*)) notation OclSum ("_->sum\<^sub>S\<^sub>e\<^sub>q'(')" (*[66]*))
subsection{* Logical Properties *} subsection\<open>Logical Properties\<close>
subsection{* Execution Laws with Invalid or Null as Argument *} subsection\<open>Execution Laws with Invalid or Null as Argument\<close>
text{* OclIterate *} text\<open>OclIterate\<close>
lemma OclIterate_invalid[simp,code_unfold]:"invalid->iterate\<^sub>S\<^sub>e\<^sub>q(a; x = A | P a x) = invalid" lemma OclIterate_invalid[simp,code_unfold]:"invalid->iterate\<^sub>S\<^sub>e\<^sub>q(a; x = A | P a x) = invalid"
by(simp add: OclIterate_def false_def true_def, simp add: invalid_def) by(simp add: OclIterate_def false_def true_def, simp add: invalid_def)
@ -433,9 +433,9 @@ by(simp add: OclIterate_def false_def true_def, simp add: invalid_def)
lemma OclIterate_invalid_args[simp,code_unfold]:"S->iterate\<^sub>S\<^sub>e\<^sub>q(a; x = invalid | P a x) = invalid" lemma OclIterate_invalid_args[simp,code_unfold]:"S->iterate\<^sub>S\<^sub>e\<^sub>q(a; x = invalid | P a x) = invalid"
by(simp add: bot_fun_def invalid_def OclIterate_def defined_def valid_def false_def true_def) by(simp add: bot_fun_def invalid_def OclIterate_def defined_def valid_def false_def true_def)
text_raw{* \isatagafp *} text_raw\<open>\isatagafp\<close>
subsubsection{* Context Passing *} subsubsection\<open>Context Passing\<close>
lemma cp_OclIncluding: lemma cp_OclIncluding:
"(X->including\<^sub>S\<^sub>e\<^sub>q(x)) \<tau> = ((\<lambda> _. X \<tau>)->including\<^sub>S\<^sub>e\<^sub>q(\<lambda> _. x \<tau>)) \<tau>" "(X->including\<^sub>S\<^sub>e\<^sub>q(x)) \<tau> = ((\<lambda> _. X \<tau>)->including\<^sub>S\<^sub>e\<^sub>q(\<lambda> _. x \<tau>)) \<tau>"
@ -450,12 +450,12 @@ by(simp add: OclIterate_def cp_defined[symmetric])
lemmas cp_intro''\<^sub>S\<^sub>e\<^sub>q[intro!,simp,code_unfold] = lemmas cp_intro''\<^sub>S\<^sub>e\<^sub>q[intro!,simp,code_unfold] =
cp_OclIncluding [THEN allI[THEN allI[THEN allI[THEN cpI2]], of "OclIncluding"]] cp_OclIncluding [THEN allI[THEN allI[THEN allI[THEN cpI2]], of "OclIncluding"]]
subsubsection{* Const *} subsubsection\<open>Const\<close>
text_raw{* \endisatagafp *} text_raw\<open>\endisatagafp\<close>
subsection{* General Algebraic Execution Rules *} subsection\<open>General Algebraic Execution Rules\<close>
subsubsection{* Execution Rules on Iterate *} subsubsection\<open>Execution Rules on Iterate\<close>
lemma OclIterate_empty[simp,code_unfold]:"Sequence{}->iterate\<^sub>S\<^sub>e\<^sub>q(a; x = A | P a x) = A" lemma OclIterate_empty[simp,code_unfold]:"Sequence{}->iterate\<^sub>S\<^sub>e\<^sub>q(a; x = A | P a x) = A"
apply(simp add: OclIterate_def foundation22[symmetric] foundation13, apply(simp add: OclIterate_def foundation22[symmetric] foundation13,
@ -464,7 +464,7 @@ apply(case_tac "\<tau> \<Turnstile> \<upsilon> A", simp_all add: foundation18')
apply(simp add: mtSequence_def) apply(simp add: mtSequence_def)
apply(subst Abs_Sequence\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse) by auto apply(subst Abs_Sequence\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse) by auto
text{* In particular, this does hold for A = null. *} text\<open>In particular, this does hold for A = null.\<close>
lemma OclIterate_including[simp,code_unfold]: lemma OclIterate_including[simp,code_unfold]:
assumes strict1 : "\<And>X. P invalid X = invalid" assumes strict1 : "\<And>X. P invalid X = invalid"
@ -540,7 +540,7 @@ qed
(* < *) (* < *)
subsection{* Test Statements *} subsection\<open>Test Statements\<close>
(* (*
Assert "(\<tau> \<Turnstile> (Sequence{\<lambda>_. \<lfloor>\<lfloor>x\<rfloor>\<rfloor>} \<doteq> Sequence{\<lambda>_. \<lfloor>\<lfloor>x\<rfloor>\<rfloor>}))" Assert "(\<tau> \<Turnstile> (Sequence{\<lambda>_. \<lfloor>\<lfloor>x\<rfloor>\<rfloor>} \<doteq> Sequence{\<lambda>_. \<lfloor>\<lfloor>x\<rfloor>\<rfloor>}))"
Assert "(\<tau> \<Turnstile> (Sequence{\<lambda>_. \<lfloor>x\<rfloor>} \<doteq> Sequence{\<lambda>_. \<lfloor>x\<rfloor>}))" Assert "(\<tau> \<Turnstile> (Sequence{\<lambda>_. \<lfloor>x\<rfloor>} \<doteq> Sequence{\<lambda>_. \<lfloor>x\<rfloor>}))"

View File

@ -50,12 +50,12 @@ imports "../basic_types/UML_Void"
begin begin
no_notation None ("\<bottom>") no_notation None ("\<bottom>")
section{* Collection Type Set: Operations \label{formal-set}*} section\<open>Collection Type Set: Operations \label{formal-set}\<close>
subsection{* As a Motivation for the (infinite) Type Construction: Type-Extensions as Sets subsection\<open>As a Motivation for the (infinite) Type Construction: Type-Extensions as Sets
\label{sec:type-extensions}*} \label{sec:type-extensions}\<close>
text{* Our notion of typed set goes beyond the usual notion of a finite executable set and text\<open>Our notion of typed set goes beyond the usual notion of a finite executable set and
is powerful enough to capture \emph{the extension of a type} in UML and OCL. This means is powerful enough to capture \emph{the extension of a type} in UML and OCL. This means
we can have in Featherweight OCL Sets containing all possible elements of a type, not only we can have in Featherweight OCL Sets containing all possible elements of a type, not only
those (finite) ones representable in a state. This holds for base types as well as class types, those (finite) ones representable in a state. This holds for base types as well as class types,
@ -69,9 +69,9 @@ In a world with @{term invalid} and @{term null}, there are two notions extensio
\item the set of all \emph{valid} values of a type @{term T}, so including @{term null} \item the set of all \emph{valid} values of a type @{term T}, so including @{term null}
(for which we will introduce the constant @{term T\<^sub>n\<^sub>u\<^sub>l\<^sub>l}). (for which we will introduce the constant @{term T\<^sub>n\<^sub>u\<^sub>l\<^sub>l}).
\end{enumerate} \end{enumerate}
*} \<close>
text{* We define the set extensions for the base type @{term Integer} as follows: *} text\<open>We define the set extensions for the base type @{term Integer} as follows:\<close>
definition Integer :: "('\<AA>,Integer\<^sub>b\<^sub>a\<^sub>s\<^sub>e) Set" definition Integer :: "('\<AA>,Integer\<^sub>b\<^sub>a\<^sub>s\<^sub>e) Set"
where "Integer \<equiv> (\<lambda> \<tau>. (Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e o Some o Some) ((Some o Some) ` (UNIV::int set)))" where "Integer \<equiv> (\<lambda> \<tau>. (Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e o Some o Some) ((Some o Some) ` (UNIV::int set)))"
@ -88,20 +88,20 @@ apply(rule ext, auto simp: Integer\<^sub>n\<^sub>u\<^sub>l\<^sub>l_def defined_d
bot_fun_def null_fun_def null_option_def) bot_fun_def null_fun_def null_option_def)
by(simp_all add: Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject bot_option_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_option_def) by(simp_all add: Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject bot_option_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_option_def)
text{* This allows the theorems: text\<open>This allows the theorems:
@{text "\<tau> \<Turnstile> \<delta> x \<Longrightarrow> \<tau> \<Turnstile> (Integer->includes\<^sub>S\<^sub>e\<^sub>t(x))"} \<open>\<tau> \<Turnstile> \<delta> x \<Longrightarrow> \<tau> \<Turnstile> (Integer->includes\<^sub>S\<^sub>e\<^sub>t(x))\<close>
@{text "\<tau> \<Turnstile> \<delta> x \<Longrightarrow> \<tau> \<Turnstile> Integer \<triangleq> (Integer->including\<^sub>S\<^sub>e\<^sub>t(x))"} \<open>\<tau> \<Turnstile> \<delta> x \<Longrightarrow> \<tau> \<Turnstile> Integer \<triangleq> (Integer->including\<^sub>S\<^sub>e\<^sub>t(x))\<close>
and and
@{text "\<tau> \<Turnstile> \<upsilon> x \<Longrightarrow> \<tau> \<Turnstile> (Integer\<^sub>n\<^sub>u\<^sub>l\<^sub>l->includes\<^sub>S\<^sub>e\<^sub>t(x))"} \<open>\<tau> \<Turnstile> \<upsilon> x \<Longrightarrow> \<tau> \<Turnstile> (Integer\<^sub>n\<^sub>u\<^sub>l\<^sub>l->includes\<^sub>S\<^sub>e\<^sub>t(x))\<close>
@{text "\<tau> \<Turnstile> \<upsilon> x \<Longrightarrow> \<tau> \<Turnstile> Integer\<^sub>n\<^sub>u\<^sub>l\<^sub>l \<triangleq> (Integer\<^sub>n\<^sub>u\<^sub>l\<^sub>l->including\<^sub>S\<^sub>e\<^sub>t(x))"} \<open>\<tau> \<Turnstile> \<upsilon> x \<Longrightarrow> \<tau> \<Turnstile> Integer\<^sub>n\<^sub>u\<^sub>l\<^sub>l \<triangleq> (Integer\<^sub>n\<^sub>u\<^sub>l\<^sub>l->including\<^sub>S\<^sub>e\<^sub>t(x))\<close>
which characterize the infiniteness of these sets by a recursive property on these sets. which characterize the infiniteness of these sets by a recursive property on these sets.
*} \<close>
text{* In the same spirit, we proceed similarly for the remaining base types: *} text\<open>In the same spirit, we proceed similarly for the remaining base types:\<close>
definition Void\<^sub>n\<^sub>u\<^sub>l\<^sub>l :: "('\<AA>,Void\<^sub>b\<^sub>a\<^sub>s\<^sub>e) Set" definition Void\<^sub>n\<^sub>u\<^sub>l\<^sub>l :: "('\<AA>,Void\<^sub>b\<^sub>a\<^sub>s\<^sub>e) Set"
where "Void\<^sub>n\<^sub>u\<^sub>l\<^sub>l \<equiv> (\<lambda> \<tau>. (Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e o Some o Some) {Abs_Void\<^sub>b\<^sub>a\<^sub>s\<^sub>e (Some None)})" where "Void\<^sub>n\<^sub>u\<^sub>l\<^sub>l \<equiv> (\<lambda> \<tau>. (Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e o Some o Some) {Abs_Void\<^sub>b\<^sub>a\<^sub>s\<^sub>e (Some None)})"
@ -213,9 +213,9 @@ apply(rule ext, auto simp: Real\<^sub>n\<^sub>u\<^sub>l\<^sub>l_def defined_def
bot_fun_def null_fun_def null_option_def) bot_fun_def null_fun_def null_option_def)
by(simp_all add: Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject bot_option_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_option_def) by(simp_all add: Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject bot_option_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_option_def)
subsection{* Basic Properties of the Set Type*} subsection\<open>Basic Properties of the Set Type\<close>
text{* Every element in a defined set is valid. *} text\<open>Every element in a defined set is valid.\<close>
lemma Set_inv_lemma: "\<tau> \<Turnstile> (\<delta> X) \<Longrightarrow> \<forall>x\<in>\<lceil>\<lceil>Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \<tau>)\<rceil>\<rceil>. x \<noteq> bot" lemma Set_inv_lemma: "\<tau> \<Turnstile> (\<delta> X) \<Longrightarrow> \<forall>x\<in>\<lceil>\<lceil>Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \<tau>)\<rceil>\<rceil>. x \<noteq> bot"
apply(insert Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e [of "X \<tau>"], simp) apply(insert Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e [of "X \<tau>"], simp)
@ -276,18 +276,18 @@ apply(simp add: valid_def null_fun_def bot_fun_def bot_Set\<^sub>b\<^sub>a\<^sub
apply(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject,simp_all add: null_option_def bot_option_def) apply(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject,simp_all add: null_option_def bot_option_def)
done done
text{* ... which means that we can have a type @{text "('\<AA>,('\<AA>,('\<AA>) Integer) Set) Set"} text\<open>... which means that we can have a type \<open>('\<AA>,('\<AA>,('\<AA>) Integer) Set) Set\<close>
corresponding exactly to Set(Set(Integer)) in OCL notation. Note that the parameter corresponding exactly to Set(Set(Integer)) in OCL notation. Note that the parameter
@{text "'\<AA>"} still refers to the object universe; making the OCL semantics entirely parametric \<open>'\<AA>\<close> still refers to the object universe; making the OCL semantics entirely parametric
in the object universe makes it possible to study (and prove) its properties in the object universe makes it possible to study (and prove) its properties
independently from a concrete class diagram. *} independently from a concrete class diagram.\<close>
subsection{* Definition: Strict Equality \label{sec:set-strict-equality}*} subsection\<open>Definition: Strict Equality \label{sec:set-strict-equality}\<close>
text{* After the part of foundational operations on sets, we detail here equality on sets. text\<open>After the part of foundational operations on sets, we detail here equality on sets.
Strong equality is inherited from the OCL core, but we have to consider Strong equality is inherited from the OCL core, but we have to consider
the case of the strict equality. We decide to overload strict equality in the the case of the strict equality. We decide to overload strict equality in the
same way we do for other value's in OCL:*} same way we do for other value's in OCL:\<close>
overloading overloading
StrictRefEq \<equiv> "StrictRefEq :: [('\<AA>,'\<alpha>::null)Set,('\<AA>,'\<alpha>::null)Set] \<Rightarrow> ('\<AA>)Boolean" StrictRefEq \<equiv> "StrictRefEq :: [('\<AA>,'\<alpha>::null)Set,('\<AA>,'\<alpha>::null)Set] \<Rightarrow> ('\<AA>)Boolean"
@ -298,20 +298,20 @@ begin
else invalid \<tau>" else invalid \<tau>"
end end
text{* One might object here that for the case of objects, this is an empty definition. text\<open>One might object here that for the case of objects, this is an empty definition.
The answer is no, we will restrain later on states and objects such that any object The answer is no, we will restrain later on states and objects such that any object
has its oid stored inside the object (so the ref, under which an object can be referenced has its oid stored inside the object (so the ref, under which an object can be referenced
in the store will represented in the object itself). For such well-formed stores that satisfy in the store will represented in the object itself). For such well-formed stores that satisfy
this invariant (the WFF-invariant), the referential equality and the this invariant (the WFF-invariant), the referential equality and the
strong equality---and therefore the strict equality on sets in the sense above---coincides.*} strong equality---and therefore the strict equality on sets in the sense above---coincides.\<close>
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"}*} text\<open>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"}\<close>
interpretation StrictRefEq\<^sub>S\<^sub>e\<^sub>t : profile_bin\<^sub>S\<^sub>t\<^sub>r\<^sub>o\<^sub>n\<^sub>g\<^sub>E\<^sub>q_\<^sub>v_\<^sub>v "\<lambda> x y. (x::('\<AA>,'\<alpha>::null)Set) \<doteq> y" interpretation StrictRefEq\<^sub>S\<^sub>e\<^sub>t : profile_bin\<^sub>S\<^sub>t\<^sub>r\<^sub>o\<^sub>n\<^sub>g\<^sub>E\<^sub>q_\<^sub>v_\<^sub>v "\<lambda> x y. (x::('\<AA>,'\<alpha>::null)Set) \<doteq> y"
by unfold_locales (auto simp: StrictRefEq\<^sub>S\<^sub>e\<^sub>t) by unfold_locales (auto simp: StrictRefEq\<^sub>S\<^sub>e\<^sub>t)
subsection{* Constants: mtSet *} subsection\<open>Constants: mtSet\<close>
definition mtSet::"('\<AA>,'\<alpha>::null) Set" ("Set{}") definition mtSet::"('\<AA>,'\<alpha>::null) Set" ("Set{}")
where "Set{} \<equiv> (\<lambda> \<tau>. Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \<lfloor>\<lfloor>{}::'\<alpha> set\<rfloor>\<rfloor> )" where "Set{} \<equiv> (\<lambda> \<tau>. Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \<lfloor>\<lfloor>{}::'\<alpha> set\<rfloor>\<rfloor> )"
@ -334,10 +334,10 @@ lemma [simp,code_unfold]: "const Set{}"
by(simp add: const_def mtSet_def) by(simp add: const_def mtSet_def)
text{* Note that the collection types in OCL allow for null to be included; text\<open>Note that the collection types in OCL allow for null to be included;
however, there is the null-collection into which inclusion yields invalid. *} however, there is the null-collection into which inclusion yields invalid.\<close>
subsection{* Definition: Including *} subsection\<open>Definition: Including\<close>
definition OclIncluding :: "[('\<AA>,'\<alpha>::null) Set,('\<AA>,'\<alpha>) val] \<Rightarrow> ('\<AA>,'\<alpha>) Set" definition OclIncluding :: "[('\<AA>,'\<alpha>::null) Set,('\<AA>,'\<alpha>) val] \<Rightarrow> ('\<AA>,'\<alpha>) Set"
where "OclIncluding x y = (\<lambda> \<tau>. if (\<delta> x) \<tau> = true \<tau> \<and> (\<upsilon> y) \<tau> = true \<tau> where "OclIncluding x y = (\<lambda> \<tau>. if (\<delta> x) \<tau> = true \<tau> \<and> (\<upsilon> y) \<tau> = true \<tau>
@ -373,7 +373,7 @@ translations
"Set{x}" == "CONST OclIncluding (Set{}) x " "Set{x}" == "CONST OclIncluding (Set{}) x "
subsection{* Definition: Excluding *} subsection\<open>Definition: Excluding\<close>
definition OclExcluding :: "[('\<AA>,'\<alpha>::null) Set,('\<AA>,'\<alpha>) val] \<Rightarrow> ('\<AA>,'\<alpha>) Set" definition OclExcluding :: "[('\<AA>,'\<alpha>::null) Set,('\<AA>,'\<alpha>) val] \<Rightarrow> ('\<AA>,'\<alpha>) Set"
where "OclExcluding x y = (\<lambda> \<tau>. if (\<delta> x) \<tau> = true \<tau> \<and> (\<upsilon> y) \<tau> = true \<tau> where "OclExcluding x y = (\<lambda> \<tau>. if (\<delta> x) \<tau> = true \<tau> \<and> (\<upsilon> y) \<tau> = true \<tau>
@ -409,7 +409,7 @@ proof -
qed qed
subsection{* Definition: Includes *} subsection\<open>Definition: Includes\<close>
definition OclIncludes :: "[('\<AA>,'\<alpha>::null) Set,('\<AA>,'\<alpha>) val] \<Rightarrow> '\<AA> Boolean" definition OclIncludes :: "[('\<AA>,'\<alpha>::null) Set,('\<AA>,'\<alpha>) val] \<Rightarrow> '\<AA> Boolean"
where "OclIncludes x y = (\<lambda> \<tau>. if (\<delta> x) \<tau> = true \<tau> \<and> (\<upsilon> y) \<tau> = true \<tau> where "OclIncludes x y = (\<lambda> \<tau>. if (\<delta> x) \<tau> = true \<tau> \<and> (\<upsilon> y) \<tau> = true \<tau>
@ -421,21 +421,21 @@ interpretation OclIncludes : profile_bin\<^sub>d_\<^sub>v OclIncludes "\<lambda>
by(unfold_locales, auto simp:OclIncludes_def bot_option_def null_option_def invalid_def) by(unfold_locales, auto simp:OclIncludes_def bot_option_def null_option_def invalid_def)
subsection{* Definition: Excludes *} subsection\<open>Definition: Excludes\<close>
definition OclExcludes :: "[('\<AA>,'\<alpha>::null) Set,('\<AA>,'\<alpha>) val] \<Rightarrow> '\<AA> Boolean" definition OclExcludes :: "[('\<AA>,'\<alpha>::null) Set,('\<AA>,'\<alpha>) val] \<Rightarrow> '\<AA> Boolean"
where "OclExcludes x y = (not(OclIncludes x y))" where "OclExcludes x y = (not(OclIncludes x y))"
notation OclExcludes ("_->excludes\<^sub>S\<^sub>e\<^sub>t'(_')" (*[66,65]65*)) notation OclExcludes ("_->excludes\<^sub>S\<^sub>e\<^sub>t'(_')" (*[66,65]65*))
text{* The case of the size definition is somewhat special, we admit text\<open>The case of the size definition is somewhat special, we admit
explicitly in Featherweight OCL the possibility of infinite sets. For explicitly in Featherweight OCL the possibility of infinite sets. For
the size definition, this requires an extra condition that assures the size definition, this requires an extra condition that assures
that the cardinality of the set is actually a defined integer. *} that the cardinality of the set is actually a defined integer.\<close>
interpretation OclExcludes : profile_bin\<^sub>d_\<^sub>v OclExcludes "\<lambda>x y. \<lfloor>\<lfloor>y \<notin> \<lceil>\<lceil>Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e x\<rceil>\<rceil>\<rfloor>\<rfloor>" interpretation OclExcludes : profile_bin\<^sub>d_\<^sub>v OclExcludes "\<lambda>x y. \<lfloor>\<lfloor>y \<notin> \<lceil>\<lceil>Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e x\<rceil>\<rceil>\<rfloor>\<rfloor>"
by(unfold_locales, auto simp:OclExcludes_def OclIncludes_def OclNot_def bot_option_def null_option_def invalid_def) by(unfold_locales, auto simp:OclExcludes_def OclIncludes_def OclNot_def bot_option_def null_option_def invalid_def)
subsection{* Definition: Size *} subsection\<open>Definition: Size\<close>
definition OclSize :: "('\<AA>,'\<alpha>::null)Set \<Rightarrow> '\<AA> Integer" definition OclSize :: "('\<AA>,'\<alpha>::null)Set \<Rightarrow> '\<AA> Integer"
where "OclSize x = (\<lambda> \<tau>. if (\<delta> x) \<tau> = true \<tau> \<and> finite(\<lceil>\<lceil>Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (x \<tau>)\<rceil>\<rceil>) where "OclSize x = (\<lambda> \<tau>. if (\<delta> x) \<tau> = true \<tau> \<and> finite(\<lceil>\<lceil>Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (x \<tau>)\<rceil>\<rceil>)
@ -444,16 +444,16 @@ where "OclSize x = (\<lambda> \<tau>. if (\<delta> x) \<tau> = true \<tau> \
notation (* standard ascii syntax *) notation (* standard ascii syntax *)
OclSize ("_->size\<^sub>S\<^sub>e\<^sub>t'(')" (*[66]*)) OclSize ("_->size\<^sub>S\<^sub>e\<^sub>t'(')" (*[66]*))
text{* The following definition follows the requirement of the text\<open>The following definition follows the requirement of the
standard to treat null as neutral element of sets. It is standard to treat null as neutral element of sets. It is
a well-documented exception from the general strictness a well-documented exception from the general strictness
rule and the rule that the distinguished argument self should rule and the rule that the distinguished argument self should
be non-null. *} be non-null.\<close>
(*TODO Locale - Equivalent*) (*TODO Locale - Equivalent*)
subsection{* Definition: IsEmpty *} subsection\<open>Definition: IsEmpty\<close>
definition OclIsEmpty :: "('\<AA>,'\<alpha>::null) Set \<Rightarrow> '\<AA> Boolean" definition OclIsEmpty :: "('\<AA>,'\<alpha>::null) Set \<Rightarrow> '\<AA> Boolean"
where "OclIsEmpty x = ((\<upsilon> x and not (\<delta> x)) or ((OclSize x) \<doteq> \<zero>))" where "OclIsEmpty x = ((\<upsilon> x and not (\<delta> x)) or ((OclSize x) \<doteq> \<zero>))"
@ -462,7 +462,7 @@ notation OclIsEmpty ("_->isEmpty\<^sub>S\<^sub>e\<^sub>t'(')" (*[66]*))
(*TODO Locale - Equivalent*) (*TODO Locale - Equivalent*)
subsection{* Definition: NotEmpty *} subsection\<open>Definition: NotEmpty\<close>
definition OclNotEmpty :: "('\<AA>,'\<alpha>::null) Set \<Rightarrow> '\<AA> Boolean" definition OclNotEmpty :: "('\<AA>,'\<alpha>::null) Set \<Rightarrow> '\<AA> Boolean"
where "OclNotEmpty x = not(OclIsEmpty x)" where "OclNotEmpty x = not(OclIsEmpty x)"
@ -470,7 +470,7 @@ notation OclNotEmpty ("_->notEmpty\<^sub>S\<^sub>e\<^sub>t'(')" (*[66]*))
(*TODO Locale - Equivalent*) (*TODO Locale - Equivalent*)
subsection{* Definition: Any *} subsection\<open>Definition: Any\<close>
(* Slight breach of naming convention in order to avoid naming conflict on constant.*) (* Slight breach of naming convention in order to avoid naming conflict on constant.*)
definition OclANY :: "[('\<AA>,'\<alpha>::null) Set] \<Rightarrow> ('\<AA>,'\<alpha>) val" definition OclANY :: "[('\<AA>,'\<alpha>::null) Set] \<Rightarrow> ('\<AA>,'\<alpha>) val"
@ -491,10 +491,10 @@ we have to go for a direct---restricted---definition. *)
subsection{* Definition: Forall *} subsection\<open>Definition: Forall\<close>
text{* The definition of OclForall mimics the one of @{term "OclAnd"}: text\<open>The definition of OclForall mimics the one of @{term "OclAnd"}:
OclForall is not a strict operation. *} OclForall is not a strict operation.\<close>
definition OclForall :: "[('\<AA>,'\<alpha>::null)Set,('\<AA>,'\<alpha>)val\<Rightarrow>('\<AA>)Boolean] \<Rightarrow> '\<AA> Boolean" definition OclForall :: "[('\<AA>,'\<alpha>::null)Set,('\<AA>,'\<alpha>)val\<Rightarrow>('\<AA>)Boolean] \<Rightarrow> '\<AA> Boolean"
where "OclForall S P = (\<lambda> \<tau>. if (\<delta> S) \<tau> = true \<tau> where "OclForall S P = (\<lambda> \<tau>. if (\<delta> S) \<tau> = true \<tau>
then if (\<exists>x\<in>\<lceil>\<lceil>Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \<tau>)\<rceil>\<rceil>. P(\<lambda> _. x) \<tau> = false \<tau>) then if (\<exists>x\<in>\<lceil>\<lceil>Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \<tau>)\<rceil>\<rceil>. P(\<lambda> _. x) \<tau> = false \<tau>)
@ -512,9 +512,9 @@ translations
(*TODO Locale - Equivalent*) (*TODO Locale - Equivalent*)
subsection{* Definition: Exists *} subsection\<open>Definition: Exists\<close>
text{* Like OclForall, OclExists is also not strict. *} text\<open>Like OclForall, OclExists is also not strict.\<close>
definition OclExists :: "[('\<AA>,'\<alpha>::null) Set,('\<AA>,'\<alpha>)val\<Rightarrow>('\<AA>)Boolean] \<Rightarrow> '\<AA> Boolean" definition OclExists :: "[('\<AA>,'\<alpha>::null) Set,('\<AA>,'\<alpha>)val\<Rightarrow>('\<AA>)Boolean] \<Rightarrow> '\<AA> Boolean"
where "OclExists S P = not(UML_Set.OclForall S (\<lambda> X. not (P X)))" where "OclExists S P = not(UML_Set.OclForall S (\<lambda> X. not (P X)))"
@ -525,7 +525,7 @@ translations
(*TODO Locale - Equivalent*) (*TODO Locale - Equivalent*)
subsection{* Definition: Iterate *} subsection\<open>Definition: Iterate\<close>
definition OclIterate :: "[('\<AA>,'\<alpha>::null) Set,('\<AA>,'\<beta>::null)val, definition OclIterate :: "[('\<AA>,'\<alpha>::null) Set,('\<AA>,'\<beta>::null)val,
('\<AA>,'\<alpha>)val\<Rightarrow>('\<AA>,'\<beta>)val\<Rightarrow>('\<AA>,'\<beta>)val] \<Rightarrow> ('\<AA>,'\<beta>)val" ('\<AA>,'\<alpha>)val\<Rightarrow>('\<AA>,'\<beta>)val\<Rightarrow>('\<AA>,'\<beta>)val] \<Rightarrow> ('\<AA>,'\<beta>)val"
@ -540,7 +540,7 @@ translations
(*TODO Locale - Equivalent*) (*TODO Locale - Equivalent*)
subsection{* Definition: Select *} subsection\<open>Definition: Select\<close>
definition OclSelect :: "[('\<AA>,'\<alpha>::null)Set,('\<AA>,'\<alpha>)val\<Rightarrow>('\<AA>)Boolean] \<Rightarrow> ('\<AA>,'\<alpha>)Set" definition OclSelect :: "[('\<AA>,'\<alpha>::null)Set,('\<AA>,'\<alpha>)val\<Rightarrow>('\<AA>)Boolean] \<Rightarrow> ('\<AA>,'\<alpha>)Set"
where "OclSelect S P = (\<lambda>\<tau>. if (\<delta> S) \<tau> = true \<tau> where "OclSelect S P = (\<lambda>\<tau>. if (\<delta> S) \<tau> = true \<tau>
@ -555,7 +555,7 @@ translations
(*TODO Locale - Equivalent*) (*TODO Locale - Equivalent*)
subsection{* Definition: Reject *} subsection\<open>Definition: Reject\<close>
definition OclReject :: "[('\<AA>,'\<alpha>::null)Set,('\<AA>,'\<alpha>)val\<Rightarrow>('\<AA>)Boolean] \<Rightarrow> ('\<AA>,'\<alpha>::null)Set" definition OclReject :: "[('\<AA>,'\<alpha>::null)Set,('\<AA>,'\<alpha>)val\<Rightarrow>('\<AA>)Boolean] \<Rightarrow> ('\<AA>,'\<alpha>::null)Set"
where "OclReject S P = OclSelect S (not o P)" where "OclReject S P = OclSelect S (not o P)"
@ -566,7 +566,7 @@ translations
(*TODO Locale - Equivalent*) (*TODO Locale - Equivalent*)
subsection{* Definition: IncludesAll *} subsection\<open>Definition: IncludesAll\<close>
definition OclIncludesAll :: "[('\<AA>,'\<alpha>::null) Set,('\<AA>,'\<alpha>) Set] \<Rightarrow> '\<AA> Boolean" definition OclIncludesAll :: "[('\<AA>,'\<alpha>::null) Set,('\<AA>,'\<alpha>) Set] \<Rightarrow> '\<AA> Boolean"
where "OclIncludesAll x y = (\<lambda> \<tau>. if (\<delta> x) \<tau> = true \<tau> \<and> (\<delta> y) \<tau> = true \<tau> where "OclIncludesAll x y = (\<lambda> \<tau>. if (\<delta> x) \<tau> = true \<tau> \<and> (\<delta> y) \<tau> = true \<tau>
@ -577,7 +577,7 @@ notation OclIncludesAll ("_->includesAll\<^sub>S\<^sub>e\<^sub>t'(_')" (*[66,6
interpretation OclIncludesAll : profile_bin\<^sub>d_\<^sub>d OclIncludesAll "\<lambda>x y. \<lfloor>\<lfloor>\<lceil>\<lceil>Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e y\<rceil>\<rceil> \<subseteq> \<lceil>\<lceil>Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e x\<rceil>\<rceil>\<rfloor>\<rfloor>" interpretation OclIncludesAll : profile_bin\<^sub>d_\<^sub>d OclIncludesAll "\<lambda>x y. \<lfloor>\<lfloor>\<lceil>\<lceil>Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e y\<rceil>\<rceil> \<subseteq> \<lceil>\<lceil>Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e x\<rceil>\<rceil>\<rfloor>\<rfloor>"
by(unfold_locales, auto simp:OclIncludesAll_def bot_option_def null_option_def invalid_def) by(unfold_locales, auto simp:OclIncludesAll_def bot_option_def null_option_def invalid_def)
subsection{* Definition: ExcludesAll *} subsection\<open>Definition: ExcludesAll\<close>
definition OclExcludesAll :: "[('\<AA>,'\<alpha>::null) Set,('\<AA>,'\<alpha>) Set] \<Rightarrow> '\<AA> Boolean" definition OclExcludesAll :: "[('\<AA>,'\<alpha>::null) Set,('\<AA>,'\<alpha>) Set] \<Rightarrow> '\<AA> Boolean"
where "OclExcludesAll x y = (\<lambda> \<tau>. if (\<delta> x) \<tau> = true \<tau> \<and> (\<delta> y) \<tau> = true \<tau> where "OclExcludesAll x y = (\<lambda> \<tau>. if (\<delta> x) \<tau> = true \<tau> \<and> (\<delta> y) \<tau> = true \<tau>
@ -588,7 +588,7 @@ notation OclExcludesAll ("_->excludesAll\<^sub>S\<^sub>e\<^sub>t'(_')" (*[66,65
interpretation OclExcludesAll : profile_bin\<^sub>d_\<^sub>d OclExcludesAll "\<lambda>x y. \<lfloor>\<lfloor>\<lceil>\<lceil>Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e y\<rceil>\<rceil> \<inter> \<lceil>\<lceil>Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e x\<rceil>\<rceil> = {}\<rfloor>\<rfloor>" interpretation OclExcludesAll : profile_bin\<^sub>d_\<^sub>d OclExcludesAll "\<lambda>x y. \<lfloor>\<lfloor>\<lceil>\<lceil>Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e y\<rceil>\<rceil> \<inter> \<lceil>\<lceil>Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e x\<rceil>\<rceil> = {}\<rfloor>\<rfloor>"
by(unfold_locales, auto simp:OclExcludesAll_def bot_option_def null_option_def invalid_def) by(unfold_locales, auto simp:OclExcludesAll_def bot_option_def null_option_def invalid_def)
subsection{* Definition: Union *} subsection\<open>Definition: Union\<close>
definition OclUnion :: "[('\<AA>,'\<alpha>::null) Set,('\<AA>,'\<alpha>) Set] \<Rightarrow> ('\<AA>,'\<alpha>) Set" definition OclUnion :: "[('\<AA>,'\<alpha>::null) Set,('\<AA>,'\<alpha>) Set] \<Rightarrow> ('\<AA>,'\<alpha>) Set"
where "OclUnion x y = (\<lambda> \<tau>. if (\<delta> x) \<tau> = true \<tau> \<and> (\<delta> y) \<tau> = true \<tau> where "OclUnion x y = (\<lambda> \<tau>. if (\<delta> x) \<tau> = true \<tau> \<and> (\<delta> y) \<tau> = true \<tau>
@ -627,7 +627,7 @@ proof -
done done
qed qed
subsection{* Definition: Intersection *} subsection\<open>Definition: Intersection\<close>
definition OclIntersection :: "[('\<AA>,'\<alpha>::null) Set,('\<AA>,'\<alpha>) Set] \<Rightarrow> ('\<AA>,'\<alpha>) Set" definition OclIntersection :: "[('\<AA>,'\<alpha>::null) Set,('\<AA>,'\<alpha>) Set] \<Rightarrow> ('\<AA>,'\<alpha>) Set"
where "OclIntersection x y = (\<lambda> \<tau>. if (\<delta> x) \<tau> = true \<tau> \<and> (\<delta> y) \<tau> = true \<tau> where "OclIntersection x y = (\<lambda> \<tau>. if (\<delta> x) \<tau> = true \<tau> \<and> (\<delta> y) \<tau> = true \<tau>
@ -667,7 +667,7 @@ proof -
done done
qed qed
subsection{* Definition (future operators) *} subsection\<open>Definition (future operators)\<close>
consts (* abstract set collection operations *) consts (* abstract set collection operations *)
OclCount :: "[('\<AA>,'\<alpha>::null) Set,('\<AA>,'\<alpha>) Set] \<Rightarrow> '\<AA> Integer" OclCount :: "[('\<AA>,'\<alpha>::null) Set,('\<AA>,'\<alpha>) Set] \<Rightarrow> '\<AA> Integer"
@ -676,9 +676,9 @@ consts (* abstract set collection operations *)
notation OclCount ("_->count\<^sub>S\<^sub>e\<^sub>t'(_')" (*[66,65]65*)) notation OclCount ("_->count\<^sub>S\<^sub>e\<^sub>t'(_')" (*[66,65]65*))
notation OclSum ("_->sum\<^sub>S\<^sub>e\<^sub>t'(')" (*[66]*)) notation OclSum ("_->sum\<^sub>S\<^sub>e\<^sub>t'(')" (*[66]*))
subsection{* Logical Properties *} subsection\<open>Logical Properties\<close>
text{* OclIncluding *} text\<open>OclIncluding\<close>
lemma OclIncluding_valid_args_valid: lemma OclIncluding_valid_args_valid:
"(\<tau> \<Turnstile> \<upsilon>(X->including\<^sub>S\<^sub>e\<^sub>t(x))) = ((\<tau> \<Turnstile>(\<delta> X)) \<and> (\<tau> \<Turnstile>(\<upsilon> x)))" "(\<tau> \<Turnstile> \<upsilon>(X->including\<^sub>S\<^sub>e\<^sub>t(x))) = ((\<tau> \<Turnstile>(\<delta> X)) \<and> (\<tau> \<Turnstile>(\<upsilon> x)))"
@ -688,10 +688,10 @@ lemma OclIncluding_valid_args_valid''[simp,code_unfold]:
"\<upsilon>(X->including\<^sub>S\<^sub>e\<^sub>t(x)) = ((\<delta> X) and (\<upsilon> x))" "\<upsilon>(X->including\<^sub>S\<^sub>e\<^sub>t(x)) = ((\<delta> X) and (\<upsilon> x))"
by (simp add: OclIncluding.def_valid_then_def) by (simp add: OclIncluding.def_valid_then_def)
text{* etc. etc. *} text\<open>etc. etc.\<close>
text_raw{* \isatagafp *} text_raw\<open>\isatagafp\<close>
text{* OclExcluding *} text\<open>OclExcluding\<close>
lemma OclExcluding_valid_args_valid: lemma OclExcluding_valid_args_valid:
"(\<tau> \<Turnstile> \<upsilon>(X->excluding\<^sub>S\<^sub>e\<^sub>t(x))) = ((\<tau> \<Turnstile>(\<delta> X)) \<and> (\<tau> \<Turnstile>(\<upsilon> x)))" "(\<tau> \<Turnstile> \<upsilon>(X->excluding\<^sub>S\<^sub>e\<^sub>t(x))) = ((\<tau> \<Turnstile>(\<delta> X)) \<and> (\<tau> \<Turnstile>(\<upsilon> x)))"
@ -701,7 +701,7 @@ lemma OclExcluding_valid_args_valid''[simp,code_unfold]:
"\<upsilon>(X->excluding\<^sub>S\<^sub>e\<^sub>t(x)) = ((\<delta> X) and (\<upsilon> x))" "\<upsilon>(X->excluding\<^sub>S\<^sub>e\<^sub>t(x)) = ((\<delta> X) and (\<upsilon> x))"
by (simp add: OclExcluding.def_valid_then_def) by (simp add: OclExcluding.def_valid_then_def)
text{* OclIncludes *} text\<open>OclIncludes\<close>
lemma OclIncludes_valid_args_valid: lemma OclIncludes_valid_args_valid:
"(\<tau> \<Turnstile> \<upsilon>(X->includes\<^sub>S\<^sub>e\<^sub>t(x))) = ((\<tau> \<Turnstile>(\<delta> X)) \<and> (\<tau> \<Turnstile>(\<upsilon> x)))" "(\<tau> \<Turnstile> \<upsilon>(X->includes\<^sub>S\<^sub>e\<^sub>t(x))) = ((\<tau> \<Turnstile>(\<delta> X)) \<and> (\<tau> \<Turnstile>(\<upsilon> x)))"
@ -711,7 +711,7 @@ lemma OclIncludes_valid_args_valid''[simp,code_unfold]:
"\<upsilon>(X->includes\<^sub>S\<^sub>e\<^sub>t(x)) = ((\<delta> X) and (\<upsilon> x))" "\<upsilon>(X->includes\<^sub>S\<^sub>e\<^sub>t(x)) = ((\<delta> X) and (\<upsilon> x))"
by (simp add: OclIncludes.def_valid_then_def) by (simp add: OclIncludes.def_valid_then_def)
text{* OclExcludes *} text\<open>OclExcludes\<close>
lemma OclExcludes_valid_args_valid: lemma OclExcludes_valid_args_valid:
"(\<tau> \<Turnstile> \<upsilon>(X->excludes\<^sub>S\<^sub>e\<^sub>t(x))) = ((\<tau> \<Turnstile>(\<delta> X)) \<and> (\<tau> \<Turnstile>(\<upsilon> x)))" "(\<tau> \<Turnstile> \<upsilon>(X->excludes\<^sub>S\<^sub>e\<^sub>t(x))) = ((\<tau> \<Turnstile>(\<delta> X)) \<and> (\<tau> \<Turnstile>(\<upsilon> x)))"
@ -721,7 +721,7 @@ lemma OclExcludes_valid_args_valid''[simp,code_unfold]:
"\<upsilon>(X->excludes\<^sub>S\<^sub>e\<^sub>t(x)) = ((\<delta> X) and (\<upsilon> x))" "\<upsilon>(X->excludes\<^sub>S\<^sub>e\<^sub>t(x)) = ((\<delta> X) and (\<upsilon> x))"
by (simp add: OclExcludes.def_valid_then_def) by (simp add: OclExcludes.def_valid_then_def)
text{* OclSize *} text\<open>OclSize\<close>
lemma OclSize_defined_args_valid: "\<tau> \<Turnstile> \<delta> (X->size\<^sub>S\<^sub>e\<^sub>t()) \<Longrightarrow> \<tau> \<Turnstile> \<delta> X" lemma OclSize_defined_args_valid: "\<tau> \<Turnstile> \<delta> (X->size\<^sub>S\<^sub>e\<^sub>t()) \<Longrightarrow> \<tau> \<Turnstile> \<delta> X"
by(auto simp: OclSize_def OclValid_def true_def valid_def false_def StrongEq_def by(auto simp: OclSize_def OclValid_def true_def valid_def false_def StrongEq_def
@ -755,7 +755,7 @@ lemma size_defined':
apply(simp add: defined_def bot_option_def bot_fun_def null_option_def null_fun_def X_finite) apply(simp add: defined_def bot_option_def bot_fun_def null_option_def null_fun_def X_finite)
done done
text{* OclIsEmpty *} text\<open>OclIsEmpty\<close>
lemma OclIsEmpty_defined_args_valid:"\<tau> \<Turnstile> \<delta> (X->isEmpty\<^sub>S\<^sub>e\<^sub>t()) \<Longrightarrow> \<tau> \<Turnstile> \<upsilon> X" lemma OclIsEmpty_defined_args_valid:"\<tau> \<Turnstile> \<delta> (X->isEmpty\<^sub>S\<^sub>e\<^sub>t()) \<Longrightarrow> \<tau> \<Turnstile> \<upsilon> X"
apply(auto simp: OclIsEmpty_def OclValid_def defined_def valid_def false_def true_def apply(auto simp: OclIsEmpty_def OclValid_def defined_def valid_def false_def true_def
@ -780,7 +780,7 @@ lemma OclIsEmpty_infinite: "\<tau> \<Turnstile> \<delta> X \<Longrightarrow> \<n
apply(case_tac x, simp add: null_option_def bot_option_def, simp) apply(case_tac x, simp add: null_option_def bot_option_def, simp)
by(simp add: OclSize_def StrictRefEq\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r valid_def bot_fun_def false_def true_def invalid_def) by(simp add: OclSize_def StrictRefEq\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r valid_def bot_fun_def false_def true_def invalid_def)
text{* OclNotEmpty *} text\<open>OclNotEmpty\<close>
lemma OclNotEmpty_defined_args_valid:"\<tau> \<Turnstile> \<delta> (X->notEmpty\<^sub>S\<^sub>e\<^sub>t()) \<Longrightarrow> \<tau> \<Turnstile> \<upsilon> X" lemma OclNotEmpty_defined_args_valid:"\<tau> \<Turnstile> \<delta> (X->notEmpty\<^sub>S\<^sub>e\<^sub>t()) \<Longrightarrow> \<tau> \<Turnstile> \<upsilon> X"
by (metis (hide_lams, no_types) OclNotEmpty_def OclNot_defargs OclNot_not foundation6 foundation9 by (metis (hide_lams, no_types) OclNotEmpty_def OclNot_defargs OclNot_not foundation6 foundation9
@ -808,7 +808,7 @@ lemma OclNotEmpty_has_elt : "\<tau> \<Turnstile> \<delta> X \<Longrightarrow>
simp_all add: false_def true_def bot_option_def bot_fun_def OclInt0_def) simp_all add: false_def true_def bot_option_def bot_fun_def OclInt0_def)
by (metis equals0I) by (metis equals0I)
text{* OclANY *} text\<open>OclANY\<close>
lemma OclANY_defined_args_valid: "\<tau> \<Turnstile> \<delta> (X->any\<^sub>S\<^sub>e\<^sub>t()) \<Longrightarrow> \<tau> \<Turnstile> \<delta> X" lemma OclANY_defined_args_valid: "\<tau> \<Turnstile> \<delta> (X->any\<^sub>S\<^sub>e\<^sub>t()) \<Longrightarrow> \<tau> \<Turnstile> \<delta> X"
by(auto simp: OclANY_def OclValid_def true_def valid_def false_def StrongEq_def by(auto simp: OclANY_def OclValid_def true_def valid_def false_def StrongEq_def
@ -852,19 +852,19 @@ lemma OclANY_valid_args_valid''[simp,code_unfold]:
by(auto intro!: OclANY_valid_args_valid transform2_rev) by(auto intro!: OclANY_valid_args_valid transform2_rev)
(* and higher order ones : forall, exists, iterate, select, reject... *) (* and higher order ones : forall, exists, iterate, select, reject... *)
text_raw{* \endisatagafp *} text_raw\<open>\endisatagafp\<close>
subsection{* Execution Laws with Invalid or Null or Infinite Set as Argument *} subsection\<open>Execution Laws with Invalid or Null or Infinite Set as Argument\<close>
text{* OclIncluding *} (* properties already generated by the corresponding locale *) text\<open>OclIncluding\<close> (* properties already generated by the corresponding locale *)
text{* OclExcluding *} (* properties already generated by the corresponding locale *) text\<open>OclExcluding\<close> (* properties already generated by the corresponding locale *)
text{* OclIncludes *} (* properties already generated by the corresponding locale *) text\<open>OclIncludes\<close> (* properties already generated by the corresponding locale *)
text{* OclExcludes *} (* properties already generated by the corresponding locale *) text\<open>OclExcludes\<close> (* properties already generated by the corresponding locale *)
text{* OclSize *} text\<open>OclSize\<close>
lemma OclSize_invalid[simp,code_unfold]:"(invalid->size\<^sub>S\<^sub>e\<^sub>t()) = invalid" lemma OclSize_invalid[simp,code_unfold]:"(invalid->size\<^sub>S\<^sub>e\<^sub>t()) = invalid"
by(simp add: bot_fun_def OclSize_def invalid_def defined_def valid_def false_def true_def) by(simp add: bot_fun_def OclSize_def invalid_def defined_def valid_def false_def true_def)
@ -874,7 +874,7 @@ by(rule ext,
simp add: bot_fun_def null_fun_def null_is_valid OclSize_def simp add: bot_fun_def null_fun_def null_is_valid OclSize_def
invalid_def defined_def valid_def false_def true_def) invalid_def defined_def valid_def false_def true_def)
text{* OclIsEmpty *} text\<open>OclIsEmpty\<close>
lemma OclIsEmpty_invalid[simp,code_unfold]:"(invalid->isEmpty\<^sub>S\<^sub>e\<^sub>t()) = invalid" lemma OclIsEmpty_invalid[simp,code_unfold]:"(invalid->isEmpty\<^sub>S\<^sub>e\<^sub>t()) = invalid"
by(simp add: OclIsEmpty_def) by(simp add: OclIsEmpty_def)
@ -882,7 +882,7 @@ by(simp add: OclIsEmpty_def)
lemma OclIsEmpty_null[simp,code_unfold]:"(null->isEmpty\<^sub>S\<^sub>e\<^sub>t()) = true" lemma OclIsEmpty_null[simp,code_unfold]:"(null->isEmpty\<^sub>S\<^sub>e\<^sub>t()) = true"
by(simp add: OclIsEmpty_def) by(simp add: OclIsEmpty_def)
text{* OclNotEmpty *} text\<open>OclNotEmpty\<close>
lemma OclNotEmpty_invalid[simp,code_unfold]:"(invalid->notEmpty\<^sub>S\<^sub>e\<^sub>t()) = invalid" lemma OclNotEmpty_invalid[simp,code_unfold]:"(invalid->notEmpty\<^sub>S\<^sub>e\<^sub>t()) = invalid"
by(simp add: OclNotEmpty_def) by(simp add: OclNotEmpty_def)
@ -890,7 +890,7 @@ by(simp add: OclNotEmpty_def)
lemma OclNotEmpty_null[simp,code_unfold]:"(null->notEmpty\<^sub>S\<^sub>e\<^sub>t()) = false" lemma OclNotEmpty_null[simp,code_unfold]:"(null->notEmpty\<^sub>S\<^sub>e\<^sub>t()) = false"
by(simp add: OclNotEmpty_def) by(simp add: OclNotEmpty_def)
text{* OclANY *} text\<open>OclANY\<close>
lemma OclANY_invalid[simp,code_unfold]:"(invalid->any\<^sub>S\<^sub>e\<^sub>t()) = invalid" lemma OclANY_invalid[simp,code_unfold]:"(invalid->any\<^sub>S\<^sub>e\<^sub>t()) = invalid"
by(simp add: bot_fun_def OclANY_def invalid_def defined_def valid_def false_def true_def) by(simp add: bot_fun_def OclANY_def invalid_def defined_def valid_def false_def true_def)
@ -898,7 +898,7 @@ by(simp add: bot_fun_def OclANY_def invalid_def defined_def valid_def false_def
lemma OclANY_null[simp,code_unfold]:"(null->any\<^sub>S\<^sub>e\<^sub>t()) = null" lemma OclANY_null[simp,code_unfold]:"(null->any\<^sub>S\<^sub>e\<^sub>t()) = null"
by(simp add: OclANY_def false_def true_def) by(simp add: OclANY_def false_def true_def)
text{* OclForall *} text\<open>OclForall\<close>
lemma OclForall_invalid[simp,code_unfold]:"invalid->forAll\<^sub>S\<^sub>e\<^sub>t(a| P a) = invalid" lemma OclForall_invalid[simp,code_unfold]:"invalid->forAll\<^sub>S\<^sub>e\<^sub>t(a| P a) = invalid"
by(simp add: bot_fun_def invalid_def OclForall_def defined_def valid_def false_def true_def) by(simp add: bot_fun_def invalid_def OclForall_def defined_def valid_def false_def true_def)
@ -906,7 +906,7 @@ by(simp add: bot_fun_def invalid_def OclForall_def defined_def valid_def false_d
lemma OclForall_null[simp,code_unfold]:"null->forAll\<^sub>S\<^sub>e\<^sub>t(a | P a) = invalid" lemma OclForall_null[simp,code_unfold]:"null->forAll\<^sub>S\<^sub>e\<^sub>t(a | P a) = invalid"
by(simp add: bot_fun_def invalid_def OclForall_def defined_def valid_def false_def true_def) by(simp add: bot_fun_def invalid_def OclForall_def defined_def valid_def false_def true_def)
text{* OclExists *} text\<open>OclExists\<close>
lemma OclExists_invalid[simp,code_unfold]:"invalid->exists\<^sub>S\<^sub>e\<^sub>t(a| P a) = invalid" lemma OclExists_invalid[simp,code_unfold]:"invalid->exists\<^sub>S\<^sub>e\<^sub>t(a| P a) = invalid"
by(simp add: OclExists_def) by(simp add: OclExists_def)
@ -914,7 +914,7 @@ by(simp add: OclExists_def)
lemma OclExists_null[simp,code_unfold]:"null->exists\<^sub>S\<^sub>e\<^sub>t(a | P a) = invalid" lemma OclExists_null[simp,code_unfold]:"null->exists\<^sub>S\<^sub>e\<^sub>t(a | P a) = invalid"
by(simp add: OclExists_def) by(simp add: OclExists_def)
text{* OclIterate *} text\<open>OclIterate\<close>
lemma OclIterate_invalid[simp,code_unfold]:"invalid->iterate\<^sub>S\<^sub>e\<^sub>t(a; x = A | P a x) = invalid" lemma OclIterate_invalid[simp,code_unfold]:"invalid->iterate\<^sub>S\<^sub>e\<^sub>t(a; x = A | P a x) = invalid"
by(simp add: bot_fun_def invalid_def OclIterate_def defined_def valid_def false_def true_def) by(simp add: bot_fun_def invalid_def OclIterate_def defined_def valid_def false_def true_def)
@ -926,7 +926,7 @@ by(simp add: bot_fun_def invalid_def OclIterate_def defined_def valid_def false_
lemma OclIterate_invalid_args[simp,code_unfold]:"S->iterate\<^sub>S\<^sub>e\<^sub>t(a; x = invalid | P a x) = invalid" lemma OclIterate_invalid_args[simp,code_unfold]:"S->iterate\<^sub>S\<^sub>e\<^sub>t(a; x = invalid | P a x) = invalid"
by(simp add: bot_fun_def invalid_def OclIterate_def defined_def valid_def false_def true_def) by(simp add: bot_fun_def invalid_def OclIterate_def defined_def valid_def false_def true_def)
text{* An open question is this ... *} text\<open>An open question is this ...\<close>
lemma (*OclIterate_null_args[simp,code_unfold]:*) "S->iterate\<^sub>S\<^sub>e\<^sub>t(a; x = null | P a x) = invalid" lemma (*OclIterate_null_args[simp,code_unfold]:*) "S->iterate\<^sub>S\<^sub>e\<^sub>t(a; x = null | P a x) = invalid"
oops oops
(* In the definition above, this does not hold in general. (* In the definition above, this does not hold in general.
@ -939,7 +939,7 @@ apply(insert non_finite [THEN OclSize_infinite])
apply(subst (asm) foundation9, simp) apply(subst (asm) foundation9, simp)
by(metis OclIterate_def OclValid_def invalid_def) by(metis OclIterate_def OclValid_def invalid_def)
text{* OclSelect *} text\<open>OclSelect\<close>
lemma OclSelect_invalid[simp,code_unfold]:"invalid->select\<^sub>S\<^sub>e\<^sub>t(a | P a) = invalid" lemma OclSelect_invalid[simp,code_unfold]:"invalid->select\<^sub>S\<^sub>e\<^sub>t(a | P a) = invalid"
by(simp add: bot_fun_def invalid_def OclSelect_def defined_def valid_def false_def true_def) by(simp add: bot_fun_def invalid_def OclSelect_def defined_def valid_def false_def true_def)
@ -947,7 +947,7 @@ by(simp add: bot_fun_def invalid_def OclSelect_def defined_def valid_def false_d
lemma OclSelect_null[simp,code_unfold]:"null->select\<^sub>S\<^sub>e\<^sub>t(a | P a) = invalid" lemma OclSelect_null[simp,code_unfold]:"null->select\<^sub>S\<^sub>e\<^sub>t(a | P a) = invalid"
by(simp add: bot_fun_def invalid_def OclSelect_def defined_def valid_def false_def true_def) by(simp add: bot_fun_def invalid_def OclSelect_def defined_def valid_def false_def true_def)
text{* OclReject *} text\<open>OclReject\<close>
lemma OclReject_invalid[simp,code_unfold]:"invalid->reject\<^sub>S\<^sub>e\<^sub>t(a | P a) = invalid" lemma OclReject_invalid[simp,code_unfold]:"invalid->reject\<^sub>S\<^sub>e\<^sub>t(a | P a) = invalid"
by(simp add: OclReject_def) by(simp add: OclReject_def)
@ -955,9 +955,9 @@ by(simp add: OclReject_def)
lemma OclReject_null[simp,code_unfold]:"null->reject\<^sub>S\<^sub>e\<^sub>t(a | P a) = invalid" lemma OclReject_null[simp,code_unfold]:"null->reject\<^sub>S\<^sub>e\<^sub>t(a | P a) = invalid"
by(simp add: OclReject_def) by(simp add: OclReject_def)
text_raw{* \isatagafp *} text_raw\<open>\isatagafp\<close>
subsubsection{* Context Passing *} subsubsection\<open>Context Passing\<close>
lemma cp_OclIncludes1: lemma cp_OclIncludes1:
"(X->includes\<^sub>S\<^sub>e\<^sub>t(x)) \<tau> = (X->includes\<^sub>S\<^sub>e\<^sub>t(\<lambda> _. x \<tau>)) \<tau>" "(X->includes\<^sub>S\<^sub>e\<^sub>t(x)) \<tau> = (X->includes\<^sub>S\<^sub>e\<^sub>t(\<lambda> _. x \<tau>)) \<tau>"
@ -1040,7 +1040,7 @@ lemmas cp_intro''\<^sub>S\<^sub>e\<^sub>t[intro!,simp,code_unfold] =
cp_OclNotEmpty [THEN allI[THEN allI[THEN cpI1], of "OclNotEmpty"]] cp_OclNotEmpty [THEN allI[THEN allI[THEN cpI1], of "OclNotEmpty"]]
cp_OclANY [THEN allI[THEN allI[THEN cpI1], of "OclANY"]] cp_OclANY [THEN allI[THEN allI[THEN cpI1], of "OclANY"]]
subsubsection{* Const *} subsubsection\<open>Const\<close>
lemma const_OclIncluding[simp,code_unfold] : lemma const_OclIncluding[simp,code_unfold] :
assumes const_x : "const x" assumes const_x : "const x"
@ -1068,10 +1068,10 @@ lemma const_OclIncluding[simp,code_unfold] :
apply(subst const_charn[OF const_S]) apply(subst const_charn[OF const_S])
by simp by simp
qed qed
text_raw{* \endisatagafp *} text_raw\<open>\endisatagafp\<close>
subsection{* General Algebraic Execution Rules *} subsection\<open>General Algebraic Execution Rules\<close>
subsubsection{* Execution Rules on Including *} subsubsection\<open>Execution Rules on Including\<close>
lemma OclIncluding_finite_rep_set : lemma OclIncluding_finite_rep_set :
assumes X_def : "\<tau> \<Turnstile> \<delta> X" assumes X_def : "\<tau> \<Turnstile> \<delta> X"
@ -1221,7 +1221,7 @@ proof -
qed qed
subsubsection{* Execution Rules on Excluding *} subsubsection\<open>Execution Rules on Excluding\<close>
lemma OclExcluding_finite_rep_set : lemma OclExcluding_finite_rep_set :
assumes X_def : "\<tau> \<Turnstile> \<delta> X" assumes X_def : "\<tau> \<Turnstile> \<delta> X"
@ -1508,7 +1508,7 @@ proof -
qed qed
text{* One would like a generic theorem of the form: text\<open>One would like a generic theorem of the form:
\begin{isar}[mathescape] \begin{isar}[mathescape]
lemma OclExcluding_charn_exec: lemma OclExcluding_charn_exec:
"(X->including$_{Set}$(x::('$\mathfrak{A}$,'a::null)val)->excluding$_{Set}$(y)) = "(X->including$_{Set}$(x::('$\mathfrak{A}$,'a::null)val)->excluding$_{Set}$(y)) =
@ -1522,14 +1522,14 @@ Unfortunately, this does not hold in general, since referential equality is
an overloaded concept and has to be defined for each type individually. an overloaded concept and has to be defined for each type individually.
Consequently, it is only valid for concrete type instances for Boolean, Consequently, it is only valid for concrete type instances for Boolean,
Integer, and Sets thereof... Integer, and Sets thereof...
*} \<close>
text{* The computational law \emph{OclExcluding-charn-exec} becomes generic since it text\<open>The computational law \emph{OclExcluding-charn-exec} becomes generic since it
uses strict equality which in itself is generic. It is possible to prove uses strict equality which in itself is generic. It is possible to prove
the following generic theorem and instantiate it later (using properties the following generic theorem and instantiate it later (using properties
that link the polymorphic logical strong equality with the concrete instance that link the polymorphic logical strong equality with the concrete instance
of strict quality).*} of strict quality).\<close>
lemma OclExcluding_charn_exec: lemma OclExcluding_charn_exec:
assumes strict1: "(invalid \<doteq> y) = invalid" assumes strict1: "(invalid \<doteq> y) = invalid"
and strict2: "(x \<doteq> invalid) = invalid" and strict2: "(x \<doteq> invalid) = invalid"
@ -1632,7 +1632,7 @@ by(rule OclExcluding_charn_exec[OF StrictRefEq\<^sub>S\<^sub>e\<^sub>t.strict1 S
StrictRefEq\<^sub>S\<^sub>e\<^sub>t.cp0 StrictRefEq\<^sub>S\<^sub>e\<^sub>t.StrictRefEq_vs_StrongEq], simp_all) StrictRefEq\<^sub>S\<^sub>e\<^sub>t.cp0 StrictRefEq\<^sub>S\<^sub>e\<^sub>t.StrictRefEq_vs_StrongEq], simp_all)
subsubsection{* Execution Rules on Includes *} subsubsection\<open>Execution Rules on Includes\<close>
lemma OclIncludes_charn0[simp]: lemma OclIncludes_charn0[simp]:
assumes val_x:"\<tau> \<Turnstile> (\<upsilon> x)" assumes val_x:"\<tau> \<Turnstile> (\<upsilon> x)"
@ -1698,7 +1698,7 @@ proof -
by(metis foundation22 foundation6 foundation9 neq) by(metis foundation22 foundation6 foundation9 neq)
qed qed
text{* Here is again a generic theorem similar as above. *} text\<open>Here is again a generic theorem similar as above.\<close>
lemma OclIncludes_execute_generic: lemma OclIncludes_execute_generic:
assumes strict1: "(invalid \<doteq> y) = invalid" assumes strict1: "(invalid \<doteq> y) = invalid"
@ -1813,7 +1813,7 @@ qed
lemmas OclIncludes_including\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r = lemmas OclIncludes_including\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r =
OclIncludes_including_generic[OF OclIncludes_execute\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r StrictRefEq\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r.def_homo] OclIncludes_including_generic[OF OclIncludes_execute\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r StrictRefEq\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r.def_homo]
subsubsection{* Execution Rules on Excludes *} subsubsection\<open>Execution Rules on Excludes\<close>
lemma OclExcludes_charn1: lemma OclExcludes_charn1:
assumes def_X:"\<tau> \<Turnstile> (\<delta> X)" assumes def_X:"\<tau> \<Turnstile> (\<delta> X)"
@ -1834,7 +1834,7 @@ proof -
by(simp add: OclAnd_def def_X[simplified OclValid_def] val_x[simplified OclValid_def] true_def) by(simp add: OclAnd_def def_X[simplified OclValid_def] val_x[simplified OclValid_def] true_def)
qed qed
subsubsection{* Execution Rules on Size *} subsubsection\<open>Execution Rules on Size\<close>
lemma [simp,code_unfold]: "Set{} ->size\<^sub>S\<^sub>e\<^sub>t() = \<zero>" lemma [simp,code_unfold]: "Set{} ->size\<^sub>S\<^sub>e\<^sub>t() = \<zero>"
apply(rule ext) apply(rule ext)
@ -1922,7 +1922,7 @@ proof -
qed qed
qed qed
subsubsection{* Execution Rules on IsEmpty *} subsubsection\<open>Execution Rules on IsEmpty\<close>
lemma [simp,code_unfold]: "Set{}->isEmpty\<^sub>S\<^sub>e\<^sub>t() = true" lemma [simp,code_unfold]: "Set{}->isEmpty\<^sub>S\<^sub>e\<^sub>t() = true"
by(simp add: OclIsEmpty_def) by(simp add: OclIsEmpty_def)
@ -1960,7 +1960,7 @@ proof -
OclIncluding_notempty_rep_set[OF X_def a_val]) OclIncluding_notempty_rep_set[OF X_def a_val])
qed qed
subsubsection{* Execution Rules on NotEmpty *} subsubsection\<open>Execution Rules on NotEmpty\<close>
lemma [simp,code_unfold]: "Set{}->notEmpty\<^sub>S\<^sub>e\<^sub>t() = false" lemma [simp,code_unfold]: "Set{}->notEmpty\<^sub>S\<^sub>e\<^sub>t() = false"
by(simp add: OclNotEmpty_def) by(simp add: OclNotEmpty_def)
@ -1974,7 +1974,7 @@ shows "X->including\<^sub>S\<^sub>e\<^sub>t(a)->notEmpty\<^sub>S\<^sub>e\<^sub>t
apply(subst cp_OclNot, subst OclIsEmpty_including, simp_all add: assms) apply(subst cp_OclNot, subst OclIsEmpty_including, simp_all add: assms)
by (metis OclNot4 cp_OclNot) by (metis OclNot4 cp_OclNot)
subsubsection{* Execution Rules on Any *} subsubsection\<open>Execution Rules on Any\<close>
lemma [simp,code_unfold]: "Set{}->any\<^sub>S\<^sub>e\<^sub>t() = null" lemma [simp,code_unfold]: "Set{}->any\<^sub>S\<^sub>e\<^sub>t() = null"
by(rule ext, simp add: OclANY_def, simp add: false_def true_def) by(rule ext, simp add: OclANY_def, simp add: false_def true_def)
@ -2005,7 +2005,7 @@ lemma OclANY_singleton_exec[simp,code_unfold]:
simp add: mtSet_defined[simplified mtSet_def] valid_def bot_fun_def) simp add: mtSet_defined[simplified mtSet_def] valid_def bot_fun_def)
by(simp add: cp_OclAnd[symmetric], rule impI, simp add: false_def true_def) by(simp add: cp_OclAnd[symmetric], rule impI, simp add: false_def true_def)
subsubsection{* Execution Rules on Forall *} subsubsection\<open>Execution Rules on Forall\<close>
lemma OclForall_mtSet_exec[simp,code_unfold] :"((Set{})->forAll\<^sub>S\<^sub>e\<^sub>t(z| P(z))) = true" lemma OclForall_mtSet_exec[simp,code_unfold] :"((Set{})->forAll\<^sub>S\<^sub>e\<^sub>t(z| P(z))) = true"
apply(simp add: OclForall_def) apply(simp add: OclForall_def)
@ -2014,7 +2014,7 @@ apply(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse, simp_all add: true_
done done
text{* The following rule is a main theorem of our approach: From a denotational definition text\<open>The following rule is a main theorem of our approach: From a denotational definition
that assures consistency, but may be --- as in the case of the @{term "X->forAll\<^sub>S\<^sub>e\<^sub>t(x | P x)"} --- that assures consistency, but may be --- as in the case of the @{term "X->forAll\<^sub>S\<^sub>e\<^sub>t(x | P x)"} ---
dauntingly complex, we derive operational rules that can serve as a gold-standard for operational dauntingly complex, we derive operational rules that can serve as a gold-standard for operational
execution, since they may be evaluated in whatever situation and according to whatever strategy. execution, since they may be evaluated in whatever situation and according to whatever strategy.
@ -2029,7 +2029,7 @@ or even:
@{term "Integer ->forAll\<^sub>S\<^sub>e\<^sub>t(x | (Integer ->forAll\<^sub>S\<^sub>e\<^sub>t(y | x +\<^sub>i\<^sub>n\<^sub>t y \<doteq> y +\<^sub>i\<^sub>n\<^sub>t x)))"} @{term "Integer ->forAll\<^sub>S\<^sub>e\<^sub>t(x | (Integer ->forAll\<^sub>S\<^sub>e\<^sub>t(y | x +\<^sub>i\<^sub>n\<^sub>t y \<doteq> y +\<^sub>i\<^sub>n\<^sub>t x)))"}
are valid OCL statements in any context $\tau$. are valid OCL statements in any context $\tau$.
*} \<close>
theorem OclForall_including_exec[simp,code_unfold] : theorem OclForall_including_exec[simp,code_unfold] :
assumes cp0 : "cp P" assumes cp0 : "cp P"
@ -2215,7 +2215,7 @@ qed
subsubsection{* Execution Rules on Exists *} subsubsection\<open>Execution Rules on Exists\<close>
lemma OclExists_mtSet_exec[simp,code_unfold] : lemma OclExists_mtSet_exec[simp,code_unfold] :
"((Set{})->exists\<^sub>S\<^sub>e\<^sub>t(z | P(z))) = false" "((Set{})->exists\<^sub>S\<^sub>e\<^sub>t(z | P(z))) = false"
@ -2230,7 +2230,7 @@ lemma OclExists_including_exec[simp,code_unfold] :
by(simp add: OclExists_def OclOr_def cp OclNot_inject) by(simp add: OclExists_def OclOr_def cp OclNot_inject)
subsubsection{* Execution Rules on Iterate *} subsubsection\<open>Execution Rules on Iterate\<close>
lemma OclIterate_empty[simp,code_unfold]: "((Set{})->iterate\<^sub>S\<^sub>e\<^sub>t(a; x = A | P a x)) = A" lemma OclIterate_empty[simp,code_unfold]: "((Set{})->iterate\<^sub>S\<^sub>e\<^sub>t(a; x = A | P a x)) = A"
proof - proof -
@ -2244,7 +2244,7 @@ proof -
done done
qed qed
text{* In particular, this does hold for A = null. *} text\<open>In particular, this does hold for A = null.\<close>
lemma OclIterate_including: lemma OclIterate_including:
assumes S_finite: "\<tau> \<Turnstile> \<delta>(S->size\<^sub>S\<^sub>e\<^sub>t())" assumes S_finite: "\<tau> \<Turnstile> \<delta>(S->size\<^sub>S\<^sub>e\<^sub>t())"
@ -2314,7 +2314,7 @@ proof -
by(subst Finite_Set.comp_fun_commute.fold_insert_remove[OF F_commute], simp+) by(subst Finite_Set.comp_fun_commute.fold_insert_remove[OF F_commute], simp+)
qed qed
subsubsection{* Execution Rules on Select *} subsubsection\<open>Execution Rules on Select\<close>
lemma OclSelect_mtSet_exec[simp,code_unfold]: "OclSelect mtSet P = mtSet" lemma OclSelect_mtSet_exec[simp,code_unfold]: "OclSelect mtSet P = mtSet"
apply(rule ext, rename_tac \<tau>) apply(rule ext, rename_tac \<tau>)
@ -2560,7 +2560,7 @@ proof -
done done
qed qed
subsubsection{* Execution Rules on Reject *} subsubsection\<open>Execution Rules on Reject\<close>
lemma OclReject_mtSet_exec[simp,code_unfold]: "OclReject mtSet P = mtSet" lemma OclReject_mtSet_exec[simp,code_unfold]: "OclReject mtSet P = mtSet"
by(simp add: OclReject_def) by(simp add: OclReject_def)
@ -2571,9 +2571,9 @@ lemma OclReject_including_exec[simp,code_unfold]:
apply(simp add: OclReject_def comp_def, rule OclSelect_including_exec) apply(simp add: OclReject_def comp_def, rule OclSelect_including_exec)
by (metis assms cp_intro'(5)) by (metis assms cp_intro'(5))
subsubsection{* Execution Rules Combining Previous Operators *} subsubsection\<open>Execution Rules Combining Previous Operators\<close>
text{* OclIncluding *} text\<open>OclIncluding\<close>
(* logical level : *) (* logical level : *)
lemma OclIncluding_idem0 : lemma OclIncluding_idem0 :
@ -2615,7 +2615,7 @@ proof -
done done
qed qed
text{* OclExcluding *} text\<open>OclExcluding\<close>
(* logical level : *) (* logical level : *)
lemma OclExcluding_idem0 : lemma OclExcluding_idem0 :
@ -2657,7 +2657,7 @@ proof -
done done
qed qed
text{* OclIncludes *} text\<open>OclIncludes\<close>
lemma OclIncludes_any[simp,code_unfold]: lemma OclIncludes_any[simp,code_unfold]:
@ -2748,7 +2748,7 @@ proof -
simp add: false_def true_def OclIf_false[simplified false_def] invalid_def) simp add: false_def true_def OclIf_false[simplified false_def] invalid_def)
qed qed
text{* OclSize *} text\<open>OclSize\<close>
lemma [simp,code_unfold]: "\<delta> (Set{} ->size\<^sub>S\<^sub>e\<^sub>t()) = true" lemma [simp,code_unfold]: "\<delta> (Set{} ->size\<^sub>S\<^sub>e\<^sub>t()) = true"
by simp by simp
@ -2869,7 +2869,7 @@ lemma [simp]:
by(simp add: size_defined[OF X_finite] del: OclSize_including_exec) by(simp add: size_defined[OF X_finite] del: OclSize_including_exec)
text{* OclForall *} text\<open>OclForall\<close>
lemma OclForall_rep_set_false: lemma OclForall_rep_set_false:
assumes "\<tau> \<Turnstile> \<delta> X" assumes "\<tau> \<Turnstile> \<delta> X"
@ -2923,7 +2923,7 @@ lemma OclForall_not_includes :
apply(simp add: OclForall_rep_set_false[OF x_def], apply(simp add: OclForall_rep_set_false[OF x_def],
simp add: OclIncludes_def OclValid_def y_def[simplified OclValid_def]) simp add: OclIncludes_def OclValid_def y_def[simplified OclValid_def])
apply(insert Set_inv_lemma[OF x_def], simp add: valid_def false_def true_def bot_fun_def) apply(insert Set_inv_lemma[OF x_def], simp add: valid_def false_def true_def bot_fun_def)
by(rule iffI, metis set_rev_mp, metis subsetI) by(rule iffI, metis rev_subsetD, metis subsetI)
lemma OclForall_iterate: lemma OclForall_iterate:
assumes S_finite: "finite \<lceil>\<lceil>Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \<tau>)\<rceil>\<rceil>" assumes S_finite: "finite \<lceil>\<lceil>Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \<tau>)\<rceil>\<rceil>"
@ -3012,7 +3012,7 @@ proof -
by (simp add: assms) by (simp add: assms)
qed qed
text{* Strict Equality *} text\<open>Strict Equality\<close>
lemma StrictRefEq\<^sub>S\<^sub>e\<^sub>t_defined : lemma StrictRefEq\<^sub>S\<^sub>e\<^sub>t_defined :
assumes x_def: "\<tau> \<Turnstile> \<delta> x" assumes x_def: "\<tau> \<Turnstile> \<delta> x"
@ -3139,7 +3139,7 @@ lemma const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including :
apply(rule StrictRefEq\<^sub>S\<^sub>e\<^sub>t.const, assumption) apply(rule StrictRefEq\<^sub>S\<^sub>e\<^sub>t.const, assumption)
by(rule const_OclIncluding) by(rule const_OclIncluding)
subsection{* Test Statements *} subsection\<open>Test Statements\<close>
Assert "(\<tau> \<Turnstile> (Set{\<lambda>_. \<lfloor>\<lfloor>x\<rfloor>\<rfloor>} \<doteq> Set{\<lambda>_. \<lfloor>\<lfloor>x\<rfloor>\<rfloor>}))" Assert "(\<tau> \<Turnstile> (Set{\<lambda>_. \<lfloor>\<lfloor>x\<rfloor>\<rfloor>} \<doteq> Set{\<lambda>_. \<lfloor>\<lfloor>x\<rfloor>\<rfloor>}))"
Assert "(\<tau> \<Turnstile> (Set{\<lambda>_. \<lfloor>x\<rfloor>} \<doteq> Set{\<lambda>_. \<lfloor>x\<rfloor>}))" Assert "(\<tau> \<Turnstile> (Set{\<lambda>_. \<lfloor>x\<rfloor>} \<doteq> Set{\<lambda>_. \<lfloor>x\<rfloor>}))"

View File

@ -9,7 +9,7 @@
# {\providecommand{\isbn}{\textsc{isbn}} } # {\providecommand{\isbn}{\textsc{isbn}} }
# {\providecommand{\Cpp}{C++} } # {\providecommand{\Cpp}{C++} }
# {\providecommand{\Specsharp}{Spec\#} } # {\providecommand{\Specsharp}{Spec\#} }
# {\providecommand{\doi}[1]{\href{http://dx.doi.org/#1}{doi: # {\providecommand{\doi}[1]{\href{https://doi.org/#1}{doi:
{\urlstyle{rm}\nolinkurl{#1}}}}} } {\urlstyle{rm}\nolinkurl{#1}}}}} }
@STRING{conf-tphols="{TPHOLs}" } @STRING{conf-tphols="{TPHOLs}" }
@STRING{iso = {International Organization for Standardization} } @STRING{iso = {International Organization for Standardization} }
@ -1425,7 +1425,7 @@
year = 2006, year = 2006,
pages = {160--174}, pages = {160--174},
crossref = {altenkirch.ea:types:2007}, crossref = {altenkirch.ea:types:2007},
url = {http://dx.doi.org/10.1007/978-3-540-74464-1_11}, url = {https://doi.org/10.1007/978-3-540-74464-1_11},
doi = {10.1007/978-3-540-74464-1_11}, doi = {10.1007/978-3-540-74464-1_11},
timestamp = {Thu, 04 Sep 2014 22:14:34 +0200}, timestamp = {Thu, 04 Sep 2014 22:14:34 +0200},
biburl = {http://dblp.uni-trier.de/rec/bib/conf/types/HaftmannW06} biburl = {http://dblp.uni-trier.de/rec/bib/conf/types/HaftmannW06}

View File

@ -120,7 +120,7 @@
\newcommand{\ie}{i.\,e.\xspace} \newcommand{\ie}{i.\,e.\xspace}
\newcommand{\eg}{e.\,g.\xspace} \newcommand{\eg}{e.\,g.\xspace}
\newenvironment{isamarkuplazy_text}{\par \isacommand{lazy{\isacharunderscore}text}\isamarkupfalse\isacharverbatimopen\isastyletext\begin{isapar}}{\end{isapar}\isacharverbatimclose} \newenvironment{isamarkuplazy_text}{\par \isacommand{lazy{\isacharunderscore}text}\isamarkupfalse\isacartoucheopen\isastyletext\begin{isapar}}{\end{isapar}\isacartoucheclose}
\renewcommand{\isasymguillemotleft}{\isatext{\textquotedblleft}} \renewcommand{\isasymguillemotleft}{\isatext{\textquotedblleft}}
\renewcommand{\isasymguillemotright}{\isatext{\textquotedblright}} \renewcommand{\isasymguillemotright}{\isatext{\textquotedblright}}
\begin{document} \begin{document}

View File

@ -45,10 +45,10 @@ theory
imports imports
Analysis_UML Analysis_UML
begin begin
text {* \label{ex:employee-analysis:ocl} *} text \<open>\label{ex:employee-analysis:ocl}\<close>
section{* OCL Part: Invariant *} section\<open>OCL Part: Invariant\<close>
text{* These recursive predicates can be defined conservatively text\<open>These recursive predicates can be defined conservatively
by greatest fix-point by greatest fix-point
constructions---automatically. See~\cite{brucker.ea:hol-ocl-book:2006,brucker:interactive:2007} constructions---automatically. See~\cite{brucker.ea:hol-ocl-book:2006,brucker:interactive:2007}
for details. For the purpose of this example, we state them as axioms for details. For the purpose of this example, we state them as axioms
@ -58,7 +58,7 @@ here.
context Person context Person
inv label : self .boss <> null implies (self .salary \<le> ((self .boss) .salary)) inv label : self .boss <> null implies (self .salary \<le> ((self .boss) .salary))
\end{ocl} \end{ocl}
*} \<close>
definition Person_label\<^sub>i\<^sub>n\<^sub>v :: "Person \<Rightarrow> Boolean" definition Person_label\<^sub>i\<^sub>n\<^sub>v :: "Person \<Rightarrow> Boolean"
where "Person_label\<^sub>i\<^sub>n\<^sub>v (self) \<equiv> where "Person_label\<^sub>i\<^sub>n\<^sub>v (self) \<equiv>
@ -87,7 +87,7 @@ lemma REC_pre : "\<tau> \<Turnstile> Person_label\<^sub>g\<^sub>l\<^sub>o\<^sub>
oops (* Attempt to allegiate the burden of he following axiomatizations: could be oops (* Attempt to allegiate the burden of he following axiomatizations: could be
a witness for a constant specification ...*) a witness for a constant specification ...*)
text{* This allows to state a predicate: *} text\<open>This allows to state a predicate:\<close>
axiomatization inv\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\<^sub>_\<^sub>l\<^sub>a\<^sub>b\<^sub>e\<^sub>l :: "Person \<Rightarrow> Boolean" axiomatization inv\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\<^sub>_\<^sub>l\<^sub>a\<^sub>b\<^sub>e\<^sub>l :: "Person \<Rightarrow> Boolean"
where inv\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\<^sub>_\<^sub>l\<^sub>a\<^sub>b\<^sub>e\<^sub>l_def: where inv\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\<^sub>_\<^sub>l\<^sub>a\<^sub>b\<^sub>e\<^sub>l_def:
@ -121,8 +121,8 @@ lemma inv_2 :
(\<tau> \<Turnstile> (inv\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\<^sub>_\<^sub>l\<^sub>a\<^sub>b\<^sub>e\<^sub>l\<^sub>A\<^sub>T\<^sub>p\<^sub>r\<^sub>e(self .boss@pre)))))" (\<tau> \<Turnstile> (inv\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\<^sub>_\<^sub>l\<^sub>a\<^sub>b\<^sub>e\<^sub>l\<^sub>A\<^sub>T\<^sub>p\<^sub>r\<^sub>e(self .boss@pre)))))"
oops (* Let's hope that this holds ... *) oops (* Let's hope that this holds ... *)
text{* A very first attempt to characterize the axiomatization by an inductive text\<open>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!) *} definition - this can not be the last word since too weak (should be equality!)\<close>
coinductive inv :: "Person \<Rightarrow> (\<AA>)st \<Rightarrow> bool" where coinductive inv :: "Person \<Rightarrow> (\<AA>)st \<Rightarrow> bool" where
"(\<tau> \<Turnstile> (\<delta> self)) \<Longrightarrow> ((\<tau> \<Turnstile> (self .boss \<doteq> null)) \<or> "(\<tau> \<Turnstile> (\<delta> self)) \<Longrightarrow> ((\<tau> \<Turnstile> (self .boss \<doteq> null)) \<or>
(\<tau> \<Turnstile> (self .boss <> null) \<and> (\<tau> \<Turnstile> (self .boss .salary \<le>\<^sub>i\<^sub>n\<^sub>t self .salary)) \<and> (\<tau> \<Turnstile> (self .boss <> null) \<and> (\<tau> \<Turnstile> (self .boss .salary \<le>\<^sub>i\<^sub>n\<^sub>t self .salary)) \<and>
@ -130,8 +130,8 @@ coinductive inv :: "Person \<Rightarrow> (\<AA>)st \<Rightarrow> bool" where
\<Longrightarrow> ( inv self \<tau>)" \<Longrightarrow> ( inv self \<tau>)"
section{* OCL Part: The Contract of a Recursive Query *} section\<open>OCL Part: The Contract of a Recursive Query\<close>
text{* The original specification of a recursive query : text\<open>The original specification of a recursive query :
\begin{ocl} \begin{ocl}
context Person::contents():Set(Integer) context Person::contents():Set(Integer)
pre: true pre: true
@ -139,11 +139,11 @@ post: result = if self.boss = null
then Set{i} then Set{i}
else self.boss.contents()->including(i) else self.boss.contents()->including(i)
endif endif
\end{ocl} *} \end{ocl}\<close>
text{* For the case of recursive queries, we use at present just axiomatizations: *} text\<open>For the case of recursive queries, we use at present just axiomatizations:\<close>
axiomatization contents :: "Person \<Rightarrow> Set_Integer" ("(1(_).contents'('))" 50) axiomatization contents :: "Person \<Rightarrow> Set_Integer" ("(1(_).contents'('))" 50)
where contents_def: where contents_def:
@ -206,8 +206,8 @@ interpretation contents : contract0 "contents" "\<lambda> self. true"
qed qed
text{* Specializing @{thm contents.unfold2}, one gets the following more practical rewrite text\<open>Specializing @{thm contents.unfold2}, one gets the following more practical rewrite
rule that is amenable to symbolic evaluation: *} rule that is amenable to symbolic evaluation:\<close>
theorem unfold_contents : theorem unfold_contents :
assumes "cp E" assumes "cp E"
and "\<tau> \<Turnstile> \<delta> self" and "\<tau> \<Turnstile> \<delta> self"
@ -218,8 +218,8 @@ theorem unfold_contents :
by(rule contents.unfold2[of _ _ _ "\<lambda> X. true"], simp_all add: assms) by(rule contents.unfold2[of _ _ _ "\<lambda> X. true"], simp_all add: assms)
text{* Since we have only one interpretation function, we need the corresponding text\<open>Since we have only one interpretation function, we need the corresponding
operation on the pre-state: *} operation on the pre-state:\<close>
consts contentsATpre :: "Person \<Rightarrow> Set_Integer" ("(1(_).contents@pre'('))" 50) consts contentsATpre :: "Person \<Rightarrow> Set_Integer" ("(1(_).contents@pre'('))" 50)
@ -281,8 +281,8 @@ interpretation contentsATpre : contract0 "contentsATpre" "\<lambda> self. true"
by(simp add: A B C) by(simp add: A B C)
qed qed
text{* Again, we derive via @{thm [source] contents.unfold2} a Knaster-Tarski like Fixpoint rule text\<open>Again, we derive via @{thm [source] contents.unfold2} a Knaster-Tarski like Fixpoint rule
that is amenable to symbolic evaluation: *} that is amenable to symbolic evaluation:\<close>
theorem unfold_contentsATpre : theorem unfold_contentsATpre :
assumes "cp E" assumes "cp E"
and "\<tau> \<Turnstile> \<delta> self" and "\<tau> \<Turnstile> \<delta> self"
@ -293,11 +293,11 @@ theorem unfold_contentsATpre :
by(rule contentsATpre.unfold2[of _ _ _ "\<lambda> X. true"], simp_all add: assms) by(rule contentsATpre.unfold2[of _ _ _ "\<lambda> X. true"], simp_all add: assms)
text{* Note that these \inlineocl{@pre} variants on methods are only available on queries, \ie, text\<open>Note that these \inlineocl{@pre} variants on methods are only available on queries, \ie,
operations without side-effect. *} operations without side-effect.\<close>
section{* OCL Part: The Contract of a User-defined Method *} section\<open>OCL Part: The Contract of a User-defined Method\<close>
text{* text\<open>
The example specification in high-level OCL input syntax reads as follows: The example specification in high-level OCL input syntax reads as follows:
\begin{ocl} \begin{ocl}
context Person::insert(x:Integer) context Person::insert(x:Integer)
@ -307,7 +307,7 @@ contents() = contents@pre()->including(x)
\end{ocl} \end{ocl}
This boils down to: This boils down to:
*} \<close>
definition insert :: "Person \<Rightarrow>Integer \<Rightarrow> Void" ("(1(_).insert'(_'))" 50) definition insert :: "Person \<Rightarrow>Integer \<Rightarrow> Void" ("(1(_).insert'(_'))" 50)
where "self .insert(x) \<equiv> where "self .insert(x) \<equiv>
@ -317,7 +317,7 @@ where "self .insert(x) \<equiv>
(\<tau> \<Turnstile> ((self).contents() \<triangleq> (self).contents@pre()->including\<^sub>S\<^sub>e\<^sub>t(x)))) (\<tau> \<Turnstile> ((self).contents() \<triangleq> (self).contents@pre()->including\<^sub>S\<^sub>e\<^sub>t(x))))
else \<tau> \<Turnstile> res \<triangleq> invalid)" else \<tau> \<Turnstile> res \<triangleq> invalid)"
text{* The semantic consequences of this definition were computed inside this locale interpretation:*} text\<open>The semantic consequences of this definition were computed inside this locale interpretation:\<close>
interpretation insert : contract1 "insert" "\<lambda> self x. true" interpretation insert : contract1 "insert" "\<lambda> self x. true"
"\<lambda> self x res. ((self .contents()) \<triangleq> "\<lambda> self x res. ((self .contents()) \<triangleq>
(self .contents@pre()->including\<^sub>S\<^sub>e\<^sub>t(x)))" (self .contents@pre()->including\<^sub>S\<^sub>e\<^sub>t(x)))"
@ -330,7 +330,7 @@ interpretation insert : contract1 "insert" "\<lambda> self x. true"
by(simp) (* an extremely hacky proof that cries for reformulation and automation - bu *) by(simp) (* an extremely hacky proof that cries for reformulation and automation - bu *)
text{* The result of this locale interpretation for our @{term insert} contract is the following text\<open>The result of this locale interpretation for our @{term insert} contract is the following
set of properties, which serves as basis for automated deduction on them: set of properties, which serves as basis for automated deduction on them:
\begin{table}[htbp] \begin{table}[htbp]
@ -357,6 +357,6 @@ set of properties, which serves as basis for automated deduction on them:
\label{tab:sem_operation_contract} \label{tab:sem_operation_contract}
\end{table} \end{table}
*} \<close>
end end

View File

@ -40,7 +40,7 @@
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE. * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
******************************************************************************) ******************************************************************************)
chapter{* Example: The Employee Analysis Model *} (* UML part *) chapter\<open>Example: The Employee Analysis Model\<close> (* UML part *)
theory theory
Analysis_UML Analysis_UML
@ -48,10 +48,10 @@ imports
"../../../UML_Main" "../../../UML_Main"
begin begin
text {* \label{ex:employee-analysis:uml} *} text \<open>\label{ex:employee-analysis:uml}\<close>
section{* Introduction *} section\<open>Introduction\<close>
text{* text\<open>
For certain concepts like classes and class-types, only a generic For certain concepts like classes and class-types, only a generic
definition for its resulting semantics can be given. Generic means, definition for its resulting semantics can be given. Generic means,
there is a function outside HOL that ``compiles'' a concrete, there is a function outside HOL that ``compiles'' a concrete,
@ -59,19 +59,19 @@ text{*
consisting of a bunch of definitions for classes, accessors, method, consisting of a bunch of definitions for classes, accessors, method,
casts, and tests for actual types, as well as proofs for the casts, and tests for actual types, as well as proofs for the
fundamental properties of these operations in this concrete data fundamental properties of these operations in this concrete data
model. *} model.\<close>
text{* Such generic function or ``compiler'' can be implemented in text\<open>Such generic function or ``compiler'' can be implemented in
Isabelle on the ML level. This has been done, for a semantics Isabelle on the ML level. This has been done, for a semantics
following the open-world assumption, for UML 2.0 following the open-world assumption, for UML 2.0
in~\cite{brucker.ea:extensible:2008-b, brucker:interactive:2007}. In in~\cite{brucker.ea:extensible:2008-b, brucker:interactive:2007}. In
this paper, we follow another approach for UML 2.4: we define the this paper, we follow another approach for UML 2.4: we define the
concepts of the compilation informally, and present a concrete concepts of the compilation informally, and present a concrete
example which is verified in Isabelle/HOL. *} example which is verified in Isabelle/HOL.\<close>
subsection{* Outlining the Example *} subsection\<open>Outlining the Example\<close>
text{* We are presenting here an ``analysis-model'' of the (slightly text\<open>We are presenting here an ``analysis-model'' of the (slightly
modified) example Figure 7.3, page 20 of modified) example Figure 7.3, page 20 of
the OCL standard~\cite{omg:ocl:2012}. the OCL standard~\cite{omg:ocl:2012}.
Here, analysis model means that associations Here, analysis model means that associations
@ -85,28 +85,28 @@ done in our ``design model''
(see \url{http://isa-afp.org/entries/Featherweight_OCL.shtml}). (see \url{http://isa-afp.org/entries/Featherweight_OCL.shtml}).
\endisatagannexa \endisatagannexa
To be precise, this theory contains the formalization of the data-part To be precise, this theory contains the formalization of the data-part
covered by the UML class model (see \autoref{fig:person-ana}):*} covered by the UML class model (see \autoref{fig:person-ana}):\<close>
text{* text\<open>
\begin{figure} \begin{figure}
\centering\scalebox{.3}{\includegraphics{figures/person.png}}% \centering\scalebox{.3}{\includegraphics{figures/person.png}}%
\caption{A simple UML class model drawn from Figure 7.3, \caption{A simple UML class model drawn from Figure 7.3,
page 20 of~\cite{omg:ocl:2012}. \label{fig:person-ana}} page 20 of~\cite{omg:ocl:2012}. \label{fig:person-ana}}
\end{figure} \end{figure}
*} \<close>
text{* This means that the association (attached to the association class text\<open>This means that the association (attached to the association class
\inlineocl{EmployeeRanking}) with the association ends \inlineocl+boss+ and \inlineocl+employees+ is implemented \inlineocl{EmployeeRanking}) with the association ends \inlineocl+boss+ and \inlineocl+employees+ is implemented
by the attribute \inlineocl+boss+ and the operation \inlineocl+employees+ (to be discussed in the OCL part by the attribute \inlineocl+boss+ and the operation \inlineocl+employees+ (to be discussed in the OCL part
captured by the subsequent theory). captured by the subsequent theory).
*} \<close>
section{* Example Data-Universe and its Infrastructure *} section\<open>Example Data-Universe and its Infrastructure\<close>
text{* Ideally, the following is generated automatically from a UML class model. *} text\<open>Ideally, the following is generated automatically from a UML class model.\<close>
text{* Our data universe consists in the concrete class diagram just of node's, text\<open>Our data universe consists in the concrete class diagram just of node's,
and implicitly of the class object. Each class implies the existence of a class and implicitly of the class object. Each class implies the existence of a class
type defined for the corresponding object representations as follows: *} type defined for the corresponding object representations as follows:\<close>
datatype type\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n = mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n oid (* the oid to the person itself *) datatype type\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n = mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n oid (* the oid to the person itself *)
"int option" (* the attribute "salary" or null *) "int option" (* the attribute "salary" or null *)
@ -119,15 +119,15 @@ datatype type\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y = mk\<^sub>O\<^sub
in case of existence of several subclasses in case of existence of several subclasses
of oclany, sums of extensions have to be provided. *) of oclany, sums of extensions have to be provided. *)
text{* Now, we construct a concrete ``universe of OclAny types'' by injection into a text\<open>Now, we construct a concrete ``universe of OclAny types'' by injection into a
sum type containing the class types. This type of OclAny will be used as instance sum type containing the class types. This type of OclAny will be used as instance
for all respective type-variables. *} for all respective type-variables.\<close>
datatype \<AA> = in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n type\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n | in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y type\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y datatype \<AA> = in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n type\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n | in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y type\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y
text{* Having fixed the object universe, we can introduce type synonyms that exactly correspond text\<open>Having fixed the object universe, we can introduce type synonyms that exactly correspond
to OCL types. Again, we exploit that our representation of OCL is a ``shallow embedding'' with a to OCL types. Again, we exploit that our representation of OCL is a ``shallow embedding'' with a
one-to-one correspondance of OCL-types to types of the meta-language HOL. *} one-to-one correspondance of OCL-types to types of the meta-language HOL.\<close>
type_synonym Boolean = " \<AA> Boolean" type_synonym Boolean = " \<AA> Boolean"
type_synonym Integer = " \<AA> Integer" type_synonym Integer = " \<AA> Integer"
type_synonym Void = " \<AA> Void" type_synonym Void = " \<AA> Void"
@ -136,12 +136,12 @@ type_synonym Person = "(\<AA>, type\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o
type_synonym Set_Integer = "(\<AA>, int option option) Set" type_synonym Set_Integer = "(\<AA>, int option option) Set"
type_synonym Set_Person = "(\<AA>, type\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n option option) Set" type_synonym Set_Person = "(\<AA>, type\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n option option) Set"
text{* Just a little check: *} text\<open>Just a little check:\<close>
typ "Boolean" typ "Boolean"
text{* To reuse key-elements of the library like referential equality, we have text\<open>To reuse key-elements of the library like referential equality, we have
to show that the object universe belongs to the type class ``oclany,'' \ie, to show that the object universe belongs to the type class ``oclany,'' \ie,
each class type has to provide a function @{term oid_of} yielding the object id (oid) of the object. *} each class type has to provide a function @{term oid_of} yielding the object id (oid) of the object.\<close>
instantiation type\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n :: object instantiation type\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n :: object
begin begin
definition oid_of_type\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_def: "oid_of x = (case x of mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n oid _ \<Rightarrow> oid)" definition oid_of_type\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_def: "oid_of x = (case x of mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n oid _ \<Rightarrow> oid)"
@ -165,9 +165,9 @@ end
section{* Instantiation of the Generic Strict Equality *} section\<open>Instantiation of the Generic Strict Equality\<close>
text{* We instantiate the referential equality text\<open>We instantiate the referential equality
on @{text "Person"} and @{text "OclAny"} *} on \<open>Person\<close> and \<open>OclAny\<close>\<close>
overloading StrictRefEq \<equiv> "StrictRefEq :: [Person,Person] \<Rightarrow> Boolean" overloading StrictRefEq \<equiv> "StrictRefEq :: [Person,Person] \<Rightarrow> Boolean"
begin begin
@ -195,17 +195,17 @@ lemmas cps23 =
[of "x::Person", [of "x::Person",
simplified StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n[symmetric]] simplified StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n[symmetric]]
for x y \<tau> P Q for x y \<tau> P Q
text{* For each Class \emph{C}, we will have a casting operation \inlineocl{.oclAsType($C$)}, text\<open>For each Class \emph{C}, we will have a casting operation \inlineocl{.oclAsType($C$)},
a test on the actual type \inlineocl{.oclIsTypeOf($C$)} as well as its relaxed form a test on the actual type \inlineocl{.oclIsTypeOf($C$)} as well as its relaxed form
\inlineocl{.oclIsKindOf($C$)} (corresponding exactly to Java's \verb+instanceof+-operator. \inlineocl{.oclIsKindOf($C$)} (corresponding exactly to Java's \verb+instanceof+-operator.
*} \<close>
text{* Thus, since we have two class-types in our concrete class hierarchy, we have text\<open>Thus, since we have two class-types in our concrete class hierarchy, we have
two operations to declare and to provide two overloading definitions for the two static types. two operations to declare and to provide two overloading definitions for the two static types.
*} \<close>
section{* OclAsType *} section\<open>OclAsType\<close>
subsection{* Definition *} subsection\<open>Definition\<close>
consts OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y :: "'\<alpha> \<Rightarrow> OclAny" ("(_) .oclAsType'(OclAny')") consts OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y :: "'\<alpha> \<Rightarrow> OclAny" ("(_) .oclAsType'(OclAny')")
consts OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n :: "'\<alpha> \<Rightarrow> Person" ("(_) .oclAsType'(Person')") consts OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n :: "'\<alpha> \<Rightarrow> Person" ("(_) .oclAsType'(Person')")
@ -254,12 +254,12 @@ begin
"(X::Person) .oclAsType(Person) \<equiv> X " (* to avoid identity for null ? *) "(X::Person) .oclAsType(Person) \<equiv> X " (* to avoid identity for null ? *)
end end
text_raw{* \isatagafp *} text_raw\<open>\isatagafp\<close>
lemmas [simp] = lemmas [simp] =
OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny
OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person
subsection{* Context Passing *} subsection\<open>Context Passing\<close>
lemma cp_OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_Person: "cp P \<Longrightarrow> cp(\<lambda>X. (P (X::Person)::Person) .oclAsType(OclAny))" lemma cp_OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_Person: "cp P \<Longrightarrow> cp(\<lambda>X. (P (X::Person)::Person) .oclAsType(OclAny))"
by(rule cpI1, simp_all add: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person) by(rule cpI1, simp_all add: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person)
@ -290,9 +290,9 @@ lemmas [simp] =
cp_OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_OclAny cp_OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_OclAny
cp_OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_Person cp_OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_Person
text_raw{* \endisatagafp*} text_raw\<open>\endisatagafp\<close>
subsection{* Execution with Invalid or Null as Argument *} subsection\<open>Execution with Invalid or Null as Argument\<close>
lemma OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_strict : "(invalid::OclAny) .oclAsType(OclAny) = invalid" by(simp) lemma OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_strict : "(invalid::OclAny) .oclAsType(OclAny) = invalid" by(simp)
lemma OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_nullstrict : "(null::OclAny) .oclAsType(OclAny) = null" by(simp) lemma OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_nullstrict : "(null::OclAny) .oclAsType(OclAny) = null" by(simp)
@ -307,9 +307,9 @@ lemma OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_nullstric
lemma OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_strict : "(invalid::Person) .oclAsType(Person) = invalid" by(simp) lemma OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_strict : "(invalid::Person) .oclAsType(Person) = invalid" by(simp)
lemma OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_nullstrict : "(null::Person) .oclAsType(Person) = null" by(simp) lemma OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_nullstrict : "(null::Person) .oclAsType(Person) = null" by(simp)
section{* OclIsTypeOf *} section\<open>OclIsTypeOf\<close>
subsection{* Definition *} subsection\<open>Definition\<close>
consts OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y :: "'\<alpha> \<Rightarrow> Boolean" ("(_).oclIsTypeOf'(OclAny')") consts OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y :: "'\<alpha> \<Rightarrow> Boolean" ("(_).oclIsTypeOf'(OclAny')")
consts OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n :: "'\<alpha> \<Rightarrow> Boolean" ("(_).oclIsTypeOf'(Person')") consts OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n :: "'\<alpha> \<Rightarrow> Boolean" ("(_).oclIsTypeOf'(Person')")
@ -375,8 +375,8 @@ begin
\<bottom> \<Rightarrow> invalid \<tau> \<bottom> \<Rightarrow> invalid \<tau>
| _ \<Rightarrow> true \<tau>)" (* for (* \<lfloor>\<lfloor> _ \<rfloor>\<rfloor> \<Rightarrow> true \<tau> *) : must have actual type Node otherwise *) | _ \<Rightarrow> true \<tau>)" (* for (* \<lfloor>\<lfloor> _ \<rfloor>\<rfloor> \<Rightarrow> true \<tau> *) : must have actual type Node otherwise *)
end end
text_raw{* \isatagafp *} text_raw\<open>\isatagafp\<close>
subsection{* Context Passing *} subsection\<open>Context Passing\<close>
lemma cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_Person: "cp P \<Longrightarrow> cp(\<lambda>X.(P(X::Person)::Person).oclIsTypeOf(OclAny))" lemma cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_Person: "cp P \<Longrightarrow> cp(\<lambda>X.(P(X::Person)::Person).oclIsTypeOf(OclAny))"
by(rule cpI1, simp_all add: OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person) by(rule cpI1, simp_all add: OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person)
@ -407,9 +407,9 @@ lemmas [simp] =
cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_Person cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_Person
cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_OclAny cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_OclAny
cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_Person cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_Person
text_raw{* \endisatagafp *} text_raw\<open>\endisatagafp\<close>
subsection{* Execution with Invalid or Null as Argument *} subsection\<open>Execution with Invalid or Null as Argument\<close>
lemma OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_strict1[simp]: lemma OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_strict1[simp]:
"(invalid::OclAny) .oclIsTypeOf(OclAny) = invalid" "(invalid::OclAny) .oclIsTypeOf(OclAny) = invalid"
@ -444,7 +444,7 @@ lemma OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_strict2
by(rule ext, simp add: null_fun_def null_option_def bot_option_def null_def invalid_def by(rule ext, simp add: null_fun_def null_option_def bot_option_def null_def invalid_def
OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person) OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person)
subsection{* Up Down Casting *} subsection\<open>Up Down Casting\<close>
lemma actualType_larger_staticType: lemma actualType_larger_staticType:
assumes isdef: "\<tau> \<Turnstile> (\<delta> X)" assumes isdef: "\<tau> \<Turnstile> (\<delta> X)"
@ -503,8 +503,8 @@ shows "\<tau> \<Turnstile> (X .oclIsTypeOf(Person) implies (X .oclAsType(OclAn
by simp by simp
section{* OclIsKindOf *} section\<open>OclIsKindOf\<close>
subsection{* Definition *} subsection\<open>Definition\<close>
consts OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y :: "'\<alpha> \<Rightarrow> Boolean" ("(_).oclIsKindOf'(OclAny')") consts OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y :: "'\<alpha> \<Rightarrow> Boolean" ("(_).oclIsKindOf'(OclAny')")
consts OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n :: "'\<alpha> \<Rightarrow> Boolean" ("(_).oclIsKindOf'(Person')") consts OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n :: "'\<alpha> \<Rightarrow> Boolean" ("(_).oclIsKindOf'(Person')")
@ -547,8 +547,8 @@ begin
\<bottom> \<Rightarrow> invalid \<tau> \<bottom> \<Rightarrow> invalid \<tau>
| _ \<Rightarrow> true \<tau>)" | _ \<Rightarrow> true \<tau>)"
end end
text_raw{* \isatagafp *} text_raw\<open>\isatagafp\<close>
subsection{* Context Passing *} subsection\<open>Context Passing\<close>
lemma cp_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_Person: "cp P \<Longrightarrow> cp(\<lambda>X.(P(X::Person)::Person).oclIsKindOf(OclAny))" lemma cp_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_Person: "cp P \<Longrightarrow> cp(\<lambda>X.(P(X::Person)::Person).oclIsKindOf(OclAny))"
by(rule cpI1, simp_all add: OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person) by(rule cpI1, simp_all add: OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person)
@ -578,8 +578,8 @@ lemmas [simp] =
cp_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_Person cp_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_Person
cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_OclAny cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_OclAny
cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_Person cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_Person
text_raw{* \endisatagafp *} text_raw\<open>\endisatagafp\<close>
subsection{* Execution with Invalid or Null as Argument *} subsection\<open>Execution with Invalid or Null as Argument\<close>
lemma OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_strict1[simp] : "(invalid::OclAny) .oclIsKindOf(OclAny) = invalid" lemma OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_strict1[simp] : "(invalid::OclAny) .oclIsKindOf(OclAny) = invalid"
by(rule ext, simp add: invalid_def bot_option_def by(rule ext, simp add: invalid_def bot_option_def
@ -606,7 +606,7 @@ lemma OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_strict2
by(rule ext, simp add: null_fun_def null_option_def bot_option_def null_def invalid_def by(rule ext, simp add: null_fun_def null_option_def bot_option_def null_def invalid_def
OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person) OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person)
subsection{* Up Down Casting *} subsection\<open>Up Down Casting\<close>
lemma actualKind_larger_staticKind: lemma actualKind_larger_staticKind:
assumes isdef: "\<tau> \<Turnstile> (\<delta> X)" assumes isdef: "\<tau> \<Turnstile> (\<delta> X)"
@ -625,11 +625,11 @@ apply(auto simp : bot_fun_def null_fun_def null_option_def bot_option_def null_d
split: option.split type\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y.split type\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n.split) split: option.split type\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y.split type\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n.split)
by(simp add: OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny OclValid_def false_def true_def) by(simp add: OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny OclValid_def false_def true_def)
section{* OclAllInstances *} section\<open>OclAllInstances\<close>
text{* To denote OCL-types occurring in OCL expressions syntactically---as, for example, as text\<open>To denote OCL-types occurring in OCL expressions syntactically---as, for example, as
``argument'' of \inlineisar{oclAllInstances()}---we use the inverses of the injection ``argument'' of \inlineisar{oclAllInstances()}---we use the inverses of the injection
functions into the object universes; we show that this is sufficient ``characterization.'' *} functions into the object universes; we show that this is sufficient ``characterization.''\<close>
definition "Person \<equiv> OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_\<AA>" definition "Person \<equiv> OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_\<AA>"
definition "OclAny \<equiv> OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\<AA>" definition "OclAny \<equiv> OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\<AA>"
@ -656,7 +656,7 @@ lemma OclAllInstances_at_pre\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_exe
unfolding OclAllInstances_at_pre_def unfolding OclAllInstances_at_pre_def
by(rule OclAllInstances_generic\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_exec) by(rule OclAllInstances_generic\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_exec)
subsection{* OclIsTypeOf *} subsection\<open>OclIsTypeOf\<close>
lemma OclAny_allInstances_generic_oclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y1: lemma OclAny_allInstances_generic_oclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y1:
assumes [simp]: "\<And>x. pre_post (x, x) = x" assumes [simp]: "\<And>x. pre_post (x, x) = x"
@ -720,7 +720,7 @@ lemma Person_allInstances_at_pre_oclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^s
unfolding OclAllInstances_at_pre_def unfolding OclAllInstances_at_pre_def
by(rule Person_allInstances_generic_oclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n) by(rule Person_allInstances_generic_oclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)
subsection{* OclIsKindOf *} subsection\<open>OclIsKindOf\<close>
lemma OclAny_allInstances_generic_oclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y: lemma OclAny_allInstances_generic_oclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y:
"\<tau> \<Turnstile> ((OclAllInstances_generic pre_post OclAny)->forAll\<^sub>S\<^sub>e\<^sub>t(X|X .oclIsKindOf(OclAny)))" "\<tau> \<Turnstile> ((OclAllInstances_generic pre_post OclAny)->forAll\<^sub>S\<^sub>e\<^sub>t(X|X .oclIsKindOf(OclAny)))"
apply(simp add: OclValid_def del: OclAllInstances_generic_def) apply(simp add: OclValid_def del: OclAllInstances_generic_def)
@ -778,23 +778,23 @@ lemma Person_allInstances_at_pre_oclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^s
unfolding OclAllInstances_at_pre_def unfolding OclAllInstances_at_pre_def
by(rule Person_allInstances_generic_oclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n) by(rule Person_allInstances_generic_oclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)
section{* The Accessors (any, boss, salary) *} section\<open>The Accessors (any, boss, salary)\<close>
text{*\label{sec:eam-accessors}*} text\<open>\label{sec:eam-accessors}\<close>
text{* Should be generated entirely from a class-diagram. *} text\<open>Should be generated entirely from a class-diagram.\<close>
subsection{* Definition (of the association Employee-Boss) *} subsection\<open>Definition (of the association Employee-Boss)\<close>
text{* We start with a oid for the association; this oid can be used text\<open>We start with a oid for the association; this oid can be used
in presence of association classes to represent the association inside an object, in presence of association classes to represent the association inside an object,
pretty much similar to the \inlineisar+Design_UML+, where we stored pretty much similar to the \inlineisar+Design_UML+, where we stored
an \verb+oid+ inside the class as ``pointer.'' *} an \verb+oid+ inside the class as ``pointer.''\<close>
definition oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\<B>\<O>\<S>\<S> ::"oid" where "oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\<B>\<O>\<S>\<S> = 10" definition oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\<B>\<O>\<S>\<S> ::"oid" where "oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\<B>\<O>\<S>\<S> = 10"
text{* From there on, we can already define an empty state which must contain text\<open>From there on, we can already define an empty state which must contain
for $\mathit{oid}_{Person}\mathcal{BOSS}$ the empty relation (encoded as association list, since there are for $\mathit{oid}_{Person}\mathcal{BOSS}$ the empty relation (encoded as association list, since there are
associations with a Sequence-like structure).*} associations with a Sequence-like structure).\<close>
definition eval_extract :: "('\<AA>,('a::object) option option) val definition eval_extract :: "('\<AA>,('a::object) option option) val
@ -823,9 +823,9 @@ where "deref_assocs\<^sub>2 pre_post to_from assoc_oid f oid =
| _ \<Rightarrow> invalid \<tau>)" | _ \<Rightarrow> invalid \<tau>)"
text{* The @{text pre_post}-parameter is configured with @{text fst} or text\<open>The \<open>pre_post\<close>-parameter is configured with \<open>fst\<close> or
@{text snd}, the @{text to_from}-parameter either with the identity @{term id} or \<open>snd\<close>, the \<open>to_from\<close>-parameter either with the identity @{term id} or
the following combinator @{text switch}: *} the following combinator \<open>switch\<close>:\<close>
definition "switch\<^sub>2_1 = (\<lambda>[x,y]\<Rightarrow> (x,y))" definition "switch\<^sub>2_1 = (\<lambda>[x,y]\<Rightarrow> (x,y))"
definition "switch\<^sub>2_2 = (\<lambda>[x,y]\<Rightarrow> (y,x))" definition "switch\<^sub>2_2 = (\<lambda>[x,y]\<Rightarrow> (y,x))"
definition "switch\<^sub>3_1 = (\<lambda>[x,y,z]\<Rightarrow> (x,y))" definition "switch\<^sub>3_1 = (\<lambda>[x,y,z]\<Rightarrow> (x,y))"
@ -853,7 +853,7 @@ where "deref_oid\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y fst_snd f oid =
\<lfloor> in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y obj \<rfloor> \<Rightarrow> f obj \<tau> \<lfloor> in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y obj \<rfloor> \<Rightarrow> f obj \<tau>
| _ \<Rightarrow> invalid \<tau>)" | _ \<Rightarrow> invalid \<tau>)"
text{* pointer undefined in state or not referencing a type conform object representation *} text\<open>pointer undefined in state or not referencing a type conform object representation\<close>
definition "select\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y\<A>\<N>\<Y> f = (\<lambda> X. case X of definition "select\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y\<A>\<N>\<Y> f = (\<lambda> X. case X of
@ -923,7 +923,7 @@ lemmas dot_accessor =
dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\<B>\<O>\<S>\<S>_at_pre_def dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\<B>\<O>\<S>\<S>_at_pre_def
dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\<S>\<A>\<L>\<A>\<R>\<Y>_at_pre_def dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\<S>\<A>\<L>\<A>\<R>\<Y>_at_pre_def
subsection{* Context Passing *} subsection\<open>Context Passing\<close>
lemmas [simp] = eval_extract_def lemmas [simp] = eval_extract_def
@ -956,7 +956,7 @@ lemmas cp_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\<S>\<A>\<L>\<A>\<R
cp_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\<S>\<A>\<L>\<A>\<R>\<Y>_at_pre[THEN allI[THEN allI], cp_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\<S>\<A>\<L>\<A>\<R>\<Y>_at_pre[THEN allI[THEN allI],
of "\<lambda> X _. X" "\<lambda> _ \<tau>. \<tau>", THEN cpI1] of "\<lambda> X _. X" "\<lambda> _ \<tau>. \<tau>", THEN cpI1]
subsection{* Execution with Invalid or Null as Argument *} subsection\<open>Execution with Invalid or Null as Argument\<close>
lemma dot\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y\<A>\<N>\<Y>_nullstrict [simp]: "(null).any = invalid" lemma dot\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y\<A>\<N>\<Y>_nullstrict [simp]: "(null).any = invalid"
by(rule ext, simp add: dot_accessor null_fun_def null_option_def bot_option_def null_def invalid_def) by(rule ext, simp add: dot_accessor null_fun_def null_option_def bot_option_def null_def invalid_def)
@ -987,7 +987,7 @@ by(rule ext, simp add: dot_accessor null_fun_def null_option_def bot_option_def
lemma dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\<S>\<A>\<L>\<A>\<R>\<Y>_at_pre_strict [simp] : "(invalid).salary@pre = invalid" lemma dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\<S>\<A>\<L>\<A>\<R>\<Y>_at_pre_strict [simp] : "(invalid).salary@pre = invalid"
by(rule ext, simp add: dot_accessor null_fun_def null_option_def bot_option_def null_def invalid_def) by(rule ext, simp add: dot_accessor null_fun_def null_option_def bot_option_def null_def invalid_def)
subsection{* Representation in States *} subsection\<open>Representation in States\<close>
lemma dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\<B>\<O>\<S>\<S>_def_mono:"\<tau> \<Turnstile> \<delta>(X .boss) \<Longrightarrow> \<tau> \<Turnstile> \<delta>(X)" lemma dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\<B>\<O>\<S>\<S>_def_mono:"\<tau> \<Turnstile> \<delta>(X .boss) \<Longrightarrow> \<tau> \<Turnstile> \<delta>(X)"
apply(case_tac "\<tau> \<Turnstile> (X \<triangleq> invalid)", insert StrongEq_L_subst2[where P = "(\<lambda>x. (\<delta> (x .boss)))" and \<tau> = "\<tau>" and x = "X" and y = "invalid"], simp add: foundation16') apply(case_tac "\<tau> \<Turnstile> (X \<triangleq> invalid)", insert StrongEq_L_subst2[where P = "(\<lambda>x. (\<delta> (x .boss)))" and \<tau> = "\<tau>" and x = "X" and y = "invalid"], simp add: foundation16')
@ -1008,9 +1008,9 @@ assumes A: "\<tau> \<Turnstile> \<delta>(x .boss)"
shows "\<tau> \<Turnstile> ((Person .allInstances()) ->includes\<^sub>S\<^sub>e\<^sub>t(x .boss))" shows "\<tau> \<Turnstile> ((Person .allInstances()) ->includes\<^sub>S\<^sub>e\<^sub>t(x .boss))"
oops oops
section{* A Little Infra-structure on Example States *} section\<open>A Little Infra-structure on Example States\<close>
text{* text\<open>
The example we are defining in this section comes from the figure~\ref{fig:eam1_system-states}. The example we are defining in this section comes from the figure~\ref{fig:eam1_system-states}.
\begin{figure} \begin{figure}
\includegraphics[width=\textwidth]{figures/pre-post.pdf} \includegraphics[width=\textwidth]{figures/pre-post.pdf}
@ -1018,9 +1018,9 @@ The example we are defining in this section comes from the figure~\ref{fig:eam1_
(b) post-state $\sigma_1'$.} (b) post-state $\sigma_1'$.}
\label{fig:eam1_system-states} \label{fig:eam1_system-states}
\end{figure} \end{figure}
*} \<close>
text_raw{* \isatagafp*} text_raw\<open>\isatagafp\<close>
definition OclInt1000 ("\<one>\<zero>\<zero>\<zero>") where "OclInt1000 = (\<lambda> _ . \<lfloor>\<lfloor>1000\<rfloor>\<rfloor>)" definition OclInt1000 ("\<one>\<zero>\<zero>\<zero>") where "OclInt1000 = (\<lambda> _ . \<lfloor>\<lfloor>1000\<rfloor>\<rfloor>)"
definition OclInt1200 ("\<one>\<two>\<zero>\<zero>") where "OclInt1200 = (\<lambda> _ . \<lfloor>\<lfloor>1200\<rfloor>\<rfloor>)" definition OclInt1200 ("\<one>\<two>\<zero>\<zero>") where "OclInt1200 = (\<lambda> _ . \<lfloor>\<lfloor>1200\<rfloor>\<rfloor>)"
@ -1051,7 +1051,7 @@ definition "person7 \<equiv> mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y
definition "person8 \<equiv> mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y oid7 None" definition "person8 \<equiv> mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y oid7 None"
definition "person9 \<equiv> mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n oid8 \<lfloor>0\<rfloor>" definition "person9 \<equiv> mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n oid8 \<lfloor>0\<rfloor>"
text_raw{* \endisatagafp*} text_raw\<open>\endisatagafp\<close>
definition definition
"\<sigma>\<^sub>1 \<equiv> \<lparr> heap = Map.empty(oid0 \<mapsto> in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n oid0 \<lfloor>1000\<rfloor>)) "\<sigma>\<^sub>1 \<equiv> \<lparr> heap = Map.empty(oid0 \<mapsto> in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n oid0 \<lfloor>1000\<rfloor>))
@ -1093,7 +1093,7 @@ by(auto simp: \<sigma>\<^sub>1_def)
lemma [simp,code_unfold]: "dom (heap \<sigma>\<^sub>1') = {oid0,oid1,oid2,oid3\<^cancel>\<open>,oid4\<close>,oid5,oid6,oid7,oid8}" lemma [simp,code_unfold]: "dom (heap \<sigma>\<^sub>1') = {oid0,oid1,oid2,oid3\<^cancel>\<open>,oid4\<close>,oid5,oid6,oid7,oid8}"
by(auto simp: \<sigma>\<^sub>1'_def) by(auto simp: \<sigma>\<^sub>1'_def)
text_raw{* \isatagafp *} text_raw\<open>\isatagafp\<close>
definition "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 :: Person \<equiv> \<lambda> _ .\<lfloor>\<lfloor> person1 \<rfloor>\<rfloor>" definition "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 :: Person \<equiv> \<lambda> _ .\<lfloor>\<lfloor> person1 \<rfloor>\<rfloor>"
definition "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2 :: Person \<equiv> \<lambda> _ .\<lfloor>\<lfloor> person2 \<rfloor>\<rfloor>" definition "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2 :: Person \<equiv> \<lambda> _ .\<lfloor>\<lfloor> person2 \<rfloor>\<rfloor>"
@ -1123,7 +1123,7 @@ lemmas [simp,code_unfold] =
OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person
OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny
OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person
text_raw{* \endisatagafp *} text_raw\<open>\endisatagafp\<close>
Assert "\<And>s\<^sub>p\<^sub>r\<^sub>e . (s\<^sub>p\<^sub>r\<^sub>e,\<sigma>\<^sub>1') \<Turnstile> (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 .salary <> \<one>\<zero>\<zero>\<zero>)" Assert "\<And>s\<^sub>p\<^sub>r\<^sub>e . (s\<^sub>p\<^sub>r\<^sub>e,\<sigma>\<^sub>1') \<Turnstile> (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 .salary <> \<one>\<zero>\<zero>\<zero>)"
Assert "\<And>s\<^sub>p\<^sub>r\<^sub>e . (s\<^sub>p\<^sub>r\<^sub>e,\<sigma>\<^sub>1') \<Turnstile> (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 .salary \<doteq> \<one>\<three>\<zero>\<zero>)" Assert "\<And>s\<^sub>p\<^sub>r\<^sub>e . (s\<^sub>p\<^sub>r\<^sub>e,\<sigma>\<^sub>1') \<Turnstile> (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 .salary \<doteq> \<one>\<three>\<zero>\<zero>)"

View File

@ -45,10 +45,10 @@ theory
imports imports
Design_UML Design_UML
begin begin
text {* \label{ex:employee-design:ocl} *} text \<open>\label{ex:employee-design:ocl}\<close>
section{* OCL Part: Invariant *} section\<open>OCL Part: Invariant\<close>
text{* These recursive predicates can be defined conservatively text\<open>These recursive predicates can be defined conservatively
by greatest fix-point by greatest fix-point
constructions---automatically. See~\cite{brucker.ea:hol-ocl-book:2006,brucker:interactive:2007} constructions---automatically. See~\cite{brucker.ea:hol-ocl-book:2006,brucker:interactive:2007}
for details. For the purpose of this example, we state them as axioms for details. For the purpose of this example, we state them as axioms
@ -58,7 +58,7 @@ here.
context Person context Person
inv label : self .boss <> null implies (self .salary \<le> ((self .boss) .salary)) inv label : self .boss <> null implies (self .salary \<le> ((self .boss) .salary))
\end{ocl} \end{ocl}
*} \<close>
definition Person_label\<^sub>i\<^sub>n\<^sub>v :: "Person \<Rightarrow> Boolean" definition Person_label\<^sub>i\<^sub>n\<^sub>v :: "Person \<Rightarrow> Boolean"
where "Person_label\<^sub>i\<^sub>n\<^sub>v (self) \<equiv> where "Person_label\<^sub>i\<^sub>n\<^sub>v (self) \<equiv>
@ -87,7 +87,7 @@ lemma REC_pre : "\<tau> \<Turnstile> Person_label\<^sub>g\<^sub>l\<^sub>o\<^sub>
oops (* Attempt to allegiate the burden of he following axiomatizations: could be oops (* Attempt to allegiate the burden of he following axiomatizations: could be
a witness for a constant specification ...*) a witness for a constant specification ...*)
text{* This allows to state a predicate: *} text\<open>This allows to state a predicate:\<close>
axiomatization inv\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\<^sub>_\<^sub>l\<^sub>a\<^sub>b\<^sub>e\<^sub>l :: "Person \<Rightarrow> Boolean" axiomatization inv\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\<^sub>_\<^sub>l\<^sub>a\<^sub>b\<^sub>e\<^sub>l :: "Person \<Rightarrow> Boolean"
where inv\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\<^sub>_\<^sub>l\<^sub>a\<^sub>b\<^sub>e\<^sub>l_def: where inv\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\<^sub>_\<^sub>l\<^sub>a\<^sub>b\<^sub>e\<^sub>l_def:
@ -121,8 +121,8 @@ lemma inv_2 :
(\<tau> \<Turnstile> (inv\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\<^sub>_\<^sub>l\<^sub>a\<^sub>b\<^sub>e\<^sub>l\<^sub>A\<^sub>T\<^sub>p\<^sub>r\<^sub>e(self .boss@pre)))))" (\<tau> \<Turnstile> (inv\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\<^sub>_\<^sub>l\<^sub>a\<^sub>b\<^sub>e\<^sub>l\<^sub>A\<^sub>T\<^sub>p\<^sub>r\<^sub>e(self .boss@pre)))))"
oops (* Let's hope that this holds ... *) oops (* Let's hope that this holds ... *)
text{* A very first attempt to characterize the axiomatization by an inductive text\<open>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!) *} definition - this can not be the last word since too weak (should be equality!)\<close>
coinductive inv :: "Person \<Rightarrow> (\<AA>)st \<Rightarrow> bool" where coinductive inv :: "Person \<Rightarrow> (\<AA>)st \<Rightarrow> bool" where
"(\<tau> \<Turnstile> (\<delta> self)) \<Longrightarrow> ((\<tau> \<Turnstile> (self .boss \<doteq> null)) \<or> "(\<tau> \<Turnstile> (\<delta> self)) \<Longrightarrow> ((\<tau> \<Turnstile> (self .boss \<doteq> null)) \<or>
(\<tau> \<Turnstile> (self .boss <> null) \<and> (\<tau> \<Turnstile> (self .boss .salary \<le>\<^sub>i\<^sub>n\<^sub>t self .salary)) \<and> (\<tau> \<Turnstile> (self .boss <> null) \<and> (\<tau> \<Turnstile> (self .boss .salary \<le>\<^sub>i\<^sub>n\<^sub>t self .salary)) \<and>
@ -130,8 +130,8 @@ coinductive inv :: "Person \<Rightarrow> (\<AA>)st \<Rightarrow> bool" where
\<Longrightarrow> ( inv self \<tau>)" \<Longrightarrow> ( inv self \<tau>)"
section{* OCL Part: The Contract of a Recursive Query *} section\<open>OCL Part: The Contract of a Recursive Query\<close>
text{* This part is analogous to the Analysis Model and skipped here. *} text\<open>This part is analogous to the Analysis Model and skipped here.\<close>
end end

View File

@ -40,7 +40,7 @@
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE. * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
******************************************************************************) ******************************************************************************)
chapter{* Example: The Employee Design Model *} (* UML part *) chapter\<open>Example: The Employee Design Model\<close> (* UML part *)
theory theory
Design_UML Design_UML
@ -48,10 +48,10 @@ imports
"../../../UML_Main" "../../../UML_Main"
begin begin
text {* \label{ex:employee-design:uml} *} text \<open>\label{ex:employee-design:uml}\<close>
section{* Introduction *} section\<open>Introduction\<close>
text{* text\<open>
For certain concepts like classes and class-types, only a generic For certain concepts like classes and class-types, only a generic
definition for its resulting semantics can be given. Generic means, definition for its resulting semantics can be given. Generic means,
there is a function outside HOL that ``compiles'' a concrete, there is a function outside HOL that ``compiles'' a concrete,
@ -59,43 +59,43 @@ text{*
consisting of a bunch of definitions for classes, accessors, method, consisting of a bunch of definitions for classes, accessors, method,
casts, and tests for actual types, as well as proofs for the casts, and tests for actual types, as well as proofs for the
fundamental properties of these operations in this concrete data fundamental properties of these operations in this concrete data
model. *} model.\<close>
text{* Such generic function or ``compiler'' can be implemented in text\<open>Such generic function or ``compiler'' can be implemented in
Isabelle on the ML level. This has been done, for a semantics Isabelle on the ML level. This has been done, for a semantics
following the open-world assumption, for UML 2.0 following the open-world assumption, for UML 2.0
in~\cite{brucker.ea:extensible:2008-b, brucker:interactive:2007}. In in~\cite{brucker.ea:extensible:2008-b, brucker:interactive:2007}. In
this paper, we follow another approach for UML 2.4: we define the this paper, we follow another approach for UML 2.4: we define the
concepts of the compilation informally, and present a concrete concepts of the compilation informally, and present a concrete
example which is verified in Isabelle/HOL. *} example which is verified in Isabelle/HOL.\<close>
subsection{* Outlining the Example *} subsection\<open>Outlining the Example\<close>
text{* We are presenting here a ``design-model'' of the (slightly text\<open>We are presenting here a ``design-model'' of the (slightly
modified) example Figure 7.3, page 20 of modified) example Figure 7.3, page 20 of
the OCL standard~\cite{omg:ocl:2012}. To be precise, this theory contains the formalization of the OCL standard~\cite{omg:ocl:2012}. To be precise, this theory contains the formalization of
the data-part covered by the UML class model (see \autoref{fig:person}):*} the data-part covered by the UML class model (see \autoref{fig:person}):\<close>
text{* text\<open>
\begin{figure} \begin{figure}
\centering\scalebox{.3}{\includegraphics{figures/person.png}}% \centering\scalebox{.3}{\includegraphics{figures/person.png}}%
\caption{A simple UML class model drawn from Figure 7.3, \caption{A simple UML class model drawn from Figure 7.3,
page 20 of~\cite{omg:ocl:2012}. \label{fig:person}} page 20 of~\cite{omg:ocl:2012}. \label{fig:person}}
\end{figure} \end{figure}
*} \<close>
text{* This means that the association (attached to the association class text\<open>This means that the association (attached to the association class
\inlineocl{EmployeeRanking}) with the association ends \inlineocl+boss+ and \inlineocl+employees+ is implemented \inlineocl{EmployeeRanking}) with the association ends \inlineocl+boss+ and \inlineocl+employees+ is implemented
by the attribute \inlineocl+boss+ and the operation \inlineocl+employees+ (to be discussed in the OCL part by the attribute \inlineocl+boss+ and the operation \inlineocl+employees+ (to be discussed in the OCL part
captured by the subsequent theory). captured by the subsequent theory).
*} \<close>
section{* Example Data-Universe and its Infrastructure *} section\<open>Example Data-Universe and its Infrastructure\<close>
text{* Ideally, the following is generated automatically from a UML class model. *} text\<open>Ideally, the following is generated automatically from a UML class model.\<close>
text{* Our data universe consists in the concrete class diagram just of node's, text\<open>Our data universe consists in the concrete class diagram just of node's,
and implicitly of the class object. Each class implies the existence of a class and implicitly of the class object. Each class implies the existence of a class
type defined for the corresponding object representations as follows: *} type defined for the corresponding object representations as follows:\<close>
datatype type\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n = mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n oid (* the oid to the person itself *) datatype type\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n = mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n oid (* the oid to the person itself *)
"int option" (* the attribute "salary" or null *) "int option" (* the attribute "salary" or null *)
@ -109,15 +109,15 @@ datatype type\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y = mk\<^sub>O\<^sub
in case of existence of several subclasses in case of existence of several subclasses
of oclany, sums of extensions have to be provided. *) of oclany, sums of extensions have to be provided. *)
text{* Now, we construct a concrete ``universe of OclAny types'' by injection into a text\<open>Now, we construct a concrete ``universe of OclAny types'' by injection into a
sum type containing the class types. This type of OclAny will be used as instance sum type containing the class types. This type of OclAny will be used as instance
for all respective type-variables. *} for all respective type-variables.\<close>
datatype \<AA> = in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n type\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n | in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y type\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y datatype \<AA> = in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n type\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n | in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y type\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y
text{* Having fixed the object universe, we can introduce type synonyms that exactly correspond text\<open>Having fixed the object universe, we can introduce type synonyms that exactly correspond
to OCL types. Again, we exploit that our representation of OCL is a ``shallow embedding'' with a to OCL types. Again, we exploit that our representation of OCL is a ``shallow embedding'' with a
one-to-one correspondance of OCL-types to types of the meta-language HOL. *} one-to-one correspondance of OCL-types to types of the meta-language HOL.\<close>
type_synonym Boolean = " \<AA> Boolean" type_synonym Boolean = " \<AA> Boolean"
type_synonym Integer = " \<AA> Integer" type_synonym Integer = " \<AA> Integer"
type_synonym Void = " \<AA> Void" type_synonym Void = " \<AA> Void"
@ -126,12 +126,12 @@ type_synonym Person = "(\<AA>, type\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o
type_synonym Set_Integer = "(\<AA>, int option option) Set" type_synonym Set_Integer = "(\<AA>, int option option) Set"
type_synonym Set_Person = "(\<AA>, type\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n option option) Set" type_synonym Set_Person = "(\<AA>, type\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n option option) Set"
text{* Just a little check: *} text\<open>Just a little check:\<close>
typ "Boolean" typ "Boolean"
text{* To reuse key-elements of the library like referential equality, we have text\<open>To reuse key-elements of the library like referential equality, we have
to show that the object universe belongs to the type class ``oclany,'' \ie, to show that the object universe belongs to the type class ``oclany,'' \ie,
each class type has to provide a function @{term oid_of} yielding the object id (oid) of the object. *} each class type has to provide a function @{term oid_of} yielding the object id (oid) of the object.\<close>
instantiation type\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n :: object instantiation type\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n :: object
begin begin
definition oid_of_type\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_def: "oid_of x = (case x of mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n oid _ _ \<Rightarrow> oid)" definition oid_of_type\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_def: "oid_of x = (case x of mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n oid _ _ \<Rightarrow> oid)"
@ -155,9 +155,9 @@ end
section{* Instantiation of the Generic Strict Equality *} section\<open>Instantiation of the Generic Strict Equality\<close>
text{* We instantiate the referential equality text\<open>We instantiate the referential equality
on @{text "Person"} and @{text "OclAny"} *} on \<open>Person\<close> and \<open>OclAny\<close>\<close>
overloading StrictRefEq \<equiv> "StrictRefEq :: [Person,Person] \<Rightarrow> Boolean" overloading StrictRefEq \<equiv> "StrictRefEq :: [Person,Person] \<Rightarrow> Boolean"
begin begin
@ -185,17 +185,17 @@ lemmas cps23 =
[of "x::Person", [of "x::Person",
simplified StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n[symmetric]] simplified StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n[symmetric]]
for x y \<tau> P Q for x y \<tau> P Q
text{* For each Class \emph{C}, we will have a casting operation \inlineocl{.oclAsType($C$)}, text\<open>For each Class \emph{C}, we will have a casting operation \inlineocl{.oclAsType($C$)},
a test on the actual type \inlineocl{.oclIsTypeOf($C$)} as well as its relaxed form a test on the actual type \inlineocl{.oclIsTypeOf($C$)} as well as its relaxed form
\inlineocl{.oclIsKindOf($C$)} (corresponding exactly to Java's \verb+instanceof+-operator. \inlineocl{.oclIsKindOf($C$)} (corresponding exactly to Java's \verb+instanceof+-operator.
*} \<close>
text{* Thus, since we have two class-types in our concrete class hierarchy, we have text\<open>Thus, since we have two class-types in our concrete class hierarchy, we have
two operations to declare and to provide two overloading definitions for the two static types. two operations to declare and to provide two overloading definitions for the two static types.
*} \<close>
section{* OclAsType *} section\<open>OclAsType\<close>
subsection{* Definition *} subsection\<open>Definition\<close>
consts OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y :: "'\<alpha> \<Rightarrow> OclAny" ("(_) .oclAsType'(OclAny')") consts OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y :: "'\<alpha> \<Rightarrow> OclAny" ("(_) .oclAsType'(OclAny')")
consts OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n :: "'\<alpha> \<Rightarrow> Person" ("(_) .oclAsType'(Person')") consts OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n :: "'\<alpha> \<Rightarrow> Person" ("(_) .oclAsType'(Person')")
@ -243,12 +243,12 @@ begin
definition OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person: definition OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person:
"(X::Person) .oclAsType(Person) \<equiv> X " (* to avoid identity for null ? *) "(X::Person) .oclAsType(Person) \<equiv> X " (* to avoid identity for null ? *)
end end
text_raw{* \isatagafp *} text_raw\<open>\isatagafp\<close>
lemmas [simp] = lemmas [simp] =
OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny
OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person
subsection{* Context Passing *} subsection\<open>Context Passing\<close>
lemma cp_OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_Person: "cp P \<Longrightarrow> cp(\<lambda>X. (P (X::Person)::Person) .oclAsType(OclAny))" lemma cp_OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_Person: "cp P \<Longrightarrow> cp(\<lambda>X. (P (X::Person)::Person) .oclAsType(OclAny))"
by(rule cpI1, simp_all add: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person) by(rule cpI1, simp_all add: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person)
@ -279,9 +279,9 @@ lemmas [simp] =
cp_OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_OclAny cp_OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_OclAny
cp_OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_Person cp_OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_Person
text_raw{* \endisatagafp*} text_raw\<open>\endisatagafp\<close>
subsection{* Execution with Invalid or Null as Argument *} subsection\<open>Execution with Invalid or Null as Argument\<close>
lemma OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_strict : "(invalid::OclAny) .oclAsType(OclAny) = invalid" by(simp) lemma OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_strict : "(invalid::OclAny) .oclAsType(OclAny) = invalid" by(simp)
lemma OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_nullstrict : "(null::OclAny) .oclAsType(OclAny) = null" by(simp) lemma OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_nullstrict : "(null::OclAny) .oclAsType(OclAny) = null" by(simp)
@ -296,9 +296,9 @@ lemma OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_nullstric
lemma OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_strict : "(invalid::Person) .oclAsType(Person) = invalid" by(simp) lemma OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_strict : "(invalid::Person) .oclAsType(Person) = invalid" by(simp)
lemma OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_nullstrict : "(null::Person) .oclAsType(Person) = null" by(simp) lemma OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_nullstrict : "(null::Person) .oclAsType(Person) = null" by(simp)
section{* OclIsTypeOf *} section\<open>OclIsTypeOf\<close>
subsection{* Definition *} subsection\<open>Definition\<close>
consts OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y :: "'\<alpha> \<Rightarrow> Boolean" ("(_).oclIsTypeOf'(OclAny')") consts OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y :: "'\<alpha> \<Rightarrow> Boolean" ("(_).oclIsTypeOf'(OclAny')")
consts OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n :: "'\<alpha> \<Rightarrow> Boolean" ("(_).oclIsTypeOf'(Person')") consts OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n :: "'\<alpha> \<Rightarrow> Boolean" ("(_).oclIsTypeOf'(Person')")
@ -364,8 +364,8 @@ begin
\<bottom> \<Rightarrow> invalid \<tau> \<bottom> \<Rightarrow> invalid \<tau>
| _ \<Rightarrow> true \<tau>)" (* for (* \<lfloor>\<lfloor> _ \<rfloor>\<rfloor> \<Rightarrow> true \<tau> *) : must have actual type Node otherwise *) | _ \<Rightarrow> true \<tau>)" (* for (* \<lfloor>\<lfloor> _ \<rfloor>\<rfloor> \<Rightarrow> true \<tau> *) : must have actual type Node otherwise *)
end end
text_raw{* \isatagafp *} text_raw\<open>\isatagafp\<close>
subsection{* Context Passing *} subsection\<open>Context Passing\<close>
lemma cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_Person: "cp P \<Longrightarrow> cp(\<lambda>X.(P(X::Person)::Person).oclIsTypeOf(OclAny))" lemma cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_Person: "cp P \<Longrightarrow> cp(\<lambda>X.(P(X::Person)::Person).oclIsTypeOf(OclAny))"
by(rule cpI1, simp_all add: OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person) by(rule cpI1, simp_all add: OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person)
@ -396,9 +396,9 @@ lemmas [simp] =
cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_Person cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_Person
cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_OclAny cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_OclAny
cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_Person cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_Person
text_raw{* \endisatagafp *} text_raw\<open>\endisatagafp\<close>
subsection{* Execution with Invalid or Null as Argument *} subsection\<open>Execution with Invalid or Null as Argument\<close>
lemma OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_strict1[simp]: lemma OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_strict1[simp]:
"(invalid::OclAny) .oclIsTypeOf(OclAny) = invalid" "(invalid::OclAny) .oclIsTypeOf(OclAny) = invalid"
@ -433,7 +433,7 @@ lemma OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_strict2
by(rule ext, simp add: null_fun_def null_option_def bot_option_def null_def invalid_def by(rule ext, simp add: null_fun_def null_option_def bot_option_def null_def invalid_def
OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person) OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person)
subsection{* Up Down Casting *} subsection\<open>Up Down Casting\<close>
lemma actualType_larger_staticType: lemma actualType_larger_staticType:
assumes isdef: "\<tau> \<Turnstile> (\<delta> X)" assumes isdef: "\<tau> \<Turnstile> (\<delta> X)"
@ -492,8 +492,8 @@ shows "\<tau> \<Turnstile> (X .oclIsTypeOf(Person) implies (X .oclAsType(OclAn
by simp by simp
section{* OclIsKindOf *} section\<open>OclIsKindOf\<close>
subsection{* Definition *} subsection\<open>Definition\<close>
consts OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y :: "'\<alpha> \<Rightarrow> Boolean" ("(_).oclIsKindOf'(OclAny')") consts OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y :: "'\<alpha> \<Rightarrow> Boolean" ("(_).oclIsKindOf'(OclAny')")
consts OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n :: "'\<alpha> \<Rightarrow> Boolean" ("(_).oclIsKindOf'(Person')") consts OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n :: "'\<alpha> \<Rightarrow> Boolean" ("(_).oclIsKindOf'(Person')")
@ -536,8 +536,8 @@ begin
\<bottom> \<Rightarrow> invalid \<tau> \<bottom> \<Rightarrow> invalid \<tau>
| _ \<Rightarrow> true \<tau>)" | _ \<Rightarrow> true \<tau>)"
end end
text_raw{* \isatagafp *} text_raw\<open>\isatagafp\<close>
subsection{* Context Passing *} subsection\<open>Context Passing\<close>
lemma cp_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_Person: "cp P \<Longrightarrow> cp(\<lambda>X.(P(X::Person)::Person).oclIsKindOf(OclAny))" lemma cp_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_Person: "cp P \<Longrightarrow> cp(\<lambda>X.(P(X::Person)::Person).oclIsKindOf(OclAny))"
by(rule cpI1, simp_all add: OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person) by(rule cpI1, simp_all add: OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person)
@ -567,8 +567,8 @@ lemmas [simp] =
cp_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_Person cp_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_Person
cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_OclAny cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_OclAny
cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_Person cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_Person
text_raw{* \endisatagafp *} text_raw\<open>\endisatagafp\<close>
subsection{* Execution with Invalid or Null as Argument *} subsection\<open>Execution with Invalid or Null as Argument\<close>
lemma OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_strict1[simp] : "(invalid::OclAny) .oclIsKindOf(OclAny) = invalid" lemma OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_strict1[simp] : "(invalid::OclAny) .oclIsKindOf(OclAny) = invalid"
by(rule ext, simp add: invalid_def bot_option_def by(rule ext, simp add: invalid_def bot_option_def
@ -595,7 +595,7 @@ lemma OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_strict2
by(rule ext, simp add: null_fun_def null_option_def bot_option_def null_def invalid_def by(rule ext, simp add: null_fun_def null_option_def bot_option_def null_def invalid_def
OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person) OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person)
subsection{* Up Down Casting *} subsection\<open>Up Down Casting\<close>
lemma actualKind_larger_staticKind: lemma actualKind_larger_staticKind:
assumes isdef: "\<tau> \<Turnstile> (\<delta> X)" assumes isdef: "\<tau> \<Turnstile> (\<delta> X)"
@ -614,11 +614,11 @@ apply(auto simp : bot_fun_def null_fun_def null_option_def bot_option_def null_d
split: option.split type\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y.split type\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n.split) split: option.split type\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y.split type\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n.split)
by(simp add: OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny OclValid_def false_def true_def) by(simp add: OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny OclValid_def false_def true_def)
section{* OclAllInstances *} section\<open>OclAllInstances\<close>
text{* To denote OCL-types occurring in OCL expressions syntactically---as, for example, as text\<open>To denote OCL-types occurring in OCL expressions syntactically---as, for example, as
``argument'' of \inlineisar{oclAllInstances()}---we use the inverses of the injection ``argument'' of \inlineisar{oclAllInstances()}---we use the inverses of the injection
functions into the object universes; we show that this is sufficient ``characterization.'' *} functions into the object universes; we show that this is sufficient ``characterization.''\<close>
definition "Person \<equiv> OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_\<AA>" definition "Person \<equiv> OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_\<AA>"
definition "OclAny \<equiv> OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\<AA>" definition "OclAny \<equiv> OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\<AA>"
@ -645,7 +645,7 @@ lemma OclAllInstances_at_pre\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_exe
unfolding OclAllInstances_at_pre_def unfolding OclAllInstances_at_pre_def
by(rule OclAllInstances_generic\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_exec) by(rule OclAllInstances_generic\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_exec)
subsection{* OclIsTypeOf *} subsection\<open>OclIsTypeOf\<close>
lemma OclAny_allInstances_generic_oclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y1: lemma OclAny_allInstances_generic_oclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y1:
assumes [simp]: "\<And>x. pre_post (x, x) = x" assumes [simp]: "\<And>x. pre_post (x, x) = x"
@ -709,7 +709,7 @@ lemma Person_allInstances_at_pre_oclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^s
unfolding OclAllInstances_at_pre_def unfolding OclAllInstances_at_pre_def
by(rule Person_allInstances_generic_oclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n) by(rule Person_allInstances_generic_oclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)
subsection{* OclIsKindOf *} subsection\<open>OclIsKindOf\<close>
lemma OclAny_allInstances_generic_oclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y: lemma OclAny_allInstances_generic_oclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y:
"\<tau> \<Turnstile> ((OclAllInstances_generic pre_post OclAny)->forAll\<^sub>S\<^sub>e\<^sub>t(X|X .oclIsKindOf(OclAny)))" "\<tau> \<Turnstile> ((OclAllInstances_generic pre_post OclAny)->forAll\<^sub>S\<^sub>e\<^sub>t(X|X .oclIsKindOf(OclAny)))"
apply(simp add: OclValid_def del: OclAllInstances_generic_def) apply(simp add: OclValid_def del: OclAllInstances_generic_def)
@ -767,12 +767,12 @@ lemma Person_allInstances_at_pre_oclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^s
unfolding OclAllInstances_at_pre_def unfolding OclAllInstances_at_pre_def
by(rule Person_allInstances_generic_oclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n) by(rule Person_allInstances_generic_oclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)
section{* The Accessors (any, boss, salary) *} section\<open>The Accessors (any, boss, salary)\<close>
text{*\label{sec:edm-accessors}*} text\<open>\label{sec:edm-accessors}\<close>
text{* Should be generated entirely from a class-diagram. *} text\<open>Should be generated entirely from a class-diagram.\<close>
subsection{* Definition *} subsection\<open>Definition\<close>
definition eval_extract :: "('\<AA>,('a::object) option option) val definition eval_extract :: "('\<AA>,('a::object) option option) val
\<Rightarrow> (oid \<Rightarrow> ('\<AA>,'c::null) val) \<Rightarrow> (oid \<Rightarrow> ('\<AA>,'c::null) val)
@ -801,7 +801,7 @@ where "deref_oid\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y fst_snd f oid =
\<lfloor> in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y obj \<rfloor> \<Rightarrow> f obj \<tau> \<lfloor> in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y obj \<rfloor> \<Rightarrow> f obj \<tau>
| _ \<Rightarrow> invalid \<tau>)" | _ \<Rightarrow> invalid \<tau>)"
text{* pointer undefined in state or not referencing a type conform object representation *} text\<open>pointer undefined in state or not referencing a type conform object representation\<close>
definition "select\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y\<A>\<N>\<Y> f = (\<lambda> X. case X of definition "select\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y\<A>\<N>\<Y> f = (\<lambda> X. case X of
@ -868,7 +868,7 @@ lemmas dot_accessor =
dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\<B>\<O>\<S>\<S>_at_pre_def dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\<B>\<O>\<S>\<S>_at_pre_def
dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\<S>\<A>\<L>\<A>\<R>\<Y>_at_pre_def dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\<S>\<A>\<L>\<A>\<R>\<Y>_at_pre_def
subsection{* Context Passing *} subsection\<open>Context Passing\<close>
lemmas [simp] = eval_extract_def lemmas [simp] = eval_extract_def
@ -901,7 +901,7 @@ lemmas cp_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\<S>\<A>\<L>\<A>\<R
cp_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\<S>\<A>\<L>\<A>\<R>\<Y>_at_pre[THEN allI[THEN allI], cp_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\<S>\<A>\<L>\<A>\<R>\<Y>_at_pre[THEN allI[THEN allI],
of "\<lambda> X _. X" "\<lambda> _ \<tau>. \<tau>", THEN cpI1] of "\<lambda> X _. X" "\<lambda> _ \<tau>. \<tau>", THEN cpI1]
subsection{* Execution with Invalid or Null as Argument *} subsection\<open>Execution with Invalid or Null as Argument\<close>
lemma dot\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y\<A>\<N>\<Y>_nullstrict [simp]: "(null).any = invalid" lemma dot\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y\<A>\<N>\<Y>_nullstrict [simp]: "(null).any = invalid"
by(rule ext, simp add: dot_accessor null_fun_def null_option_def bot_option_def null_def invalid_def) by(rule ext, simp add: dot_accessor null_fun_def null_option_def bot_option_def null_def invalid_def)
@ -932,7 +932,7 @@ by(rule ext, simp add: dot_accessor null_fun_def null_option_def bot_option_def
lemma dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\<S>\<A>\<L>\<A>\<R>\<Y>_at_pre_strict [simp] : "(invalid).salary@pre = invalid" lemma dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\<S>\<A>\<L>\<A>\<R>\<Y>_at_pre_strict [simp] : "(invalid).salary@pre = invalid"
by(rule ext, simp add: dot_accessor null_fun_def null_option_def bot_option_def null_def invalid_def) by(rule ext, simp add: dot_accessor null_fun_def null_option_def bot_option_def null_def invalid_def)
subsection{* Representation in States *} subsection\<open>Representation in States\<close>
lemma dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\<B>\<O>\<S>\<S>_def_mono:"\<tau> \<Turnstile> \<delta>(X .boss) \<Longrightarrow> \<tau> \<Turnstile> \<delta>(X)" lemma dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\<B>\<O>\<S>\<S>_def_mono:"\<tau> \<Turnstile> \<delta>(X .boss) \<Longrightarrow> \<tau> \<Turnstile> \<delta>(X)"
apply(case_tac "\<tau> \<Turnstile> (X \<triangleq> invalid)", insert StrongEq_L_subst2[where P = "(\<lambda>x. (\<delta> (x .boss)))" and \<tau> = "\<tau>" and x = "X" and y = "invalid"], simp add: foundation16') apply(case_tac "\<tau> \<Turnstile> (X \<triangleq> invalid)", insert StrongEq_L_subst2[where P = "(\<lambda>x. (\<delta> (x .boss)))" and \<tau> = "\<tau>" and x = "X" and y = "invalid"], simp add: foundation16')
@ -968,9 +968,9 @@ proof -
by(simp add: image_comp B true_def) by(simp add: image_comp B true_def)
qed qed
section{* A Little Infra-structure on Example States *} section\<open>A Little Infra-structure on Example States\<close>
text{* text\<open>
The example we are defining in this section comes from the figure~\ref{fig:edm1_system-states}. The example we are defining in this section comes from the figure~\ref{fig:edm1_system-states}.
\begin{figure} \begin{figure}
\includegraphics[width=\textwidth]{figures/pre-post.pdf} \includegraphics[width=\textwidth]{figures/pre-post.pdf}
@ -978,9 +978,9 @@ The example we are defining in this section comes from the figure~\ref{fig:edm1_
(b) post-state $\sigma_1'$.} (b) post-state $\sigma_1'$.}
\label{fig:edm1_system-states} \label{fig:edm1_system-states}
\end{figure} \end{figure}
*} \<close>
text_raw{* \isatagafp*} text_raw\<open>\isatagafp\<close>
definition OclInt1000 ("\<one>\<zero>\<zero>\<zero>") where "OclInt1000 = (\<lambda> _ . \<lfloor>\<lfloor>1000\<rfloor>\<rfloor>)" definition OclInt1000 ("\<one>\<zero>\<zero>\<zero>") where "OclInt1000 = (\<lambda> _ . \<lfloor>\<lfloor>1000\<rfloor>\<rfloor>)"
definition OclInt1200 ("\<one>\<two>\<zero>\<zero>") where "OclInt1200 = (\<lambda> _ . \<lfloor>\<lfloor>1200\<rfloor>\<rfloor>)" definition OclInt1200 ("\<one>\<two>\<zero>\<zero>") where "OclInt1200 = (\<lambda> _ . \<lfloor>\<lfloor>1200\<rfloor>\<rfloor>)"
@ -1051,7 +1051,7 @@ by(auto simp: \<sigma>\<^sub>1_def)
lemma [simp,code_unfold]: "dom (heap \<sigma>\<^sub>1') = {oid0,oid1,oid2,oid3\<^cancel>\<open>,oid4\<close>,oid5,oid6,oid7,oid8}" lemma [simp,code_unfold]: "dom (heap \<sigma>\<^sub>1') = {oid0,oid1,oid2,oid3\<^cancel>\<open>,oid4\<close>,oid5,oid6,oid7,oid8}"
by(auto simp: \<sigma>\<^sub>1'_def) by(auto simp: \<sigma>\<^sub>1'_def)
text_raw{* \isatagafp *} text_raw\<open>\isatagafp\<close>
definition "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 :: Person \<equiv> \<lambda> _ .\<lfloor>\<lfloor> person1 \<rfloor>\<rfloor>" definition "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 :: Person \<equiv> \<lambda> _ .\<lfloor>\<lfloor> person1 \<rfloor>\<rfloor>"
definition "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2 :: Person \<equiv> \<lambda> _ .\<lfloor>\<lfloor> person2 \<rfloor>\<rfloor>" definition "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2 :: Person \<equiv> \<lambda> _ .\<lfloor>\<lfloor> person2 \<rfloor>\<rfloor>"
@ -1081,7 +1081,7 @@ lemmas [simp,code_unfold] =
OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person
OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny
OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person
text_raw{* \endisatagafp *} text_raw\<open>\endisatagafp\<close>
Assert "\<And>s\<^sub>p\<^sub>r\<^sub>e . (s\<^sub>p\<^sub>r\<^sub>e,\<sigma>\<^sub>1') \<Turnstile> (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 .salary <> \<one>\<zero>\<zero>\<zero>)" Assert "\<And>s\<^sub>p\<^sub>r\<^sub>e . (s\<^sub>p\<^sub>r\<^sub>e,\<sigma>\<^sub>1') \<Turnstile> (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 .salary <> \<one>\<zero>\<zero>\<zero>)"
Assert "\<And>s\<^sub>p\<^sub>r\<^sub>e . (s\<^sub>p\<^sub>r\<^sub>e,\<sigma>\<^sub>1') \<Turnstile> (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 .salary \<doteq> \<one>\<three>\<zero>\<zero>)" Assert "\<And>s\<^sub>p\<^sub>r\<^sub>e . (s\<^sub>p\<^sub>r\<^sub>e,\<sigma>\<^sub>1') \<Turnstile> (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 .salary \<doteq> \<one>\<three>\<zero>\<zero>)"