(***************************************************************************** * Featherweight-OCL --- A Formal Semantics for UML-OCL Version OCL 2.5 * for the OMG Standard. * http://www.brucker.ch/projects/hol-testgen/ * * UML_Contracts.thy --- * This file is part of HOL-TestGen. * * Copyright (c) 2013-2015 Université Paris-Saclay, Univ. Paris-Sud, France * 2013-2015 IRT SystemX, France * * All rights reserved. * * Redistribution and use in source and binary forms, with or without * modification, are permitted provided that the following conditions are * met: * * * Redistributions of source code must retain the above copyright * notice, this list of conditions and the following disclaimer. * * * Redistributions in binary form must reproduce the above * copyright notice, this list of conditions and the following * disclaimer in the documentation and/or other materials provided * with the distribution. * * * Neither the name of the copyright holders nor the names of its * contributors may be used to endorse or promote products derived * from this software without specific prior written permission. * * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS * "AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT * LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR * A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT * OWNER OR CONTRIBUTORS BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL, * SPECIAL, EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT * LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE, * DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY * THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT * (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE. ******************************************************************************) theory UML_Contracts imports UML_State begin text\Modeling of an operation contract for an operation with 2 arguments, (so depending on three parameters if one takes "self" into account).\ locale contract_scheme = fixes f_\ fixes f_lam fixes f :: "('\,'\0::null)val \ 'b \ ('\,'res::null)val" fixes PRE fixes POST assumes def_scheme': "f self x \ (\ \. SOME res. let res = \ _. res in if (\ \ (\ self)) \ f_\ x \ then (\ \ PRE self x) \ (\ \ POST self x res) else \ \ res \ invalid)" assumes all_post': "\ \ \' \''. ((\,\') \ PRE self x) = ((\,\'') \ PRE self x)" (* PRE is really a pre-condition semantically, i.e. it does not depend on the post-state. ... *) assumes cp\<^sub>P\<^sub>R\<^sub>E': "PRE (self) x \ = PRE (\ _. self \) (f_lam x \) \ " (* this interface is preferable than : assumes "cp self' \ cp a1' \ cp a2' \ cp (\X. PRE (self' X) (a1' X) (a2' X) )" which is too polymorphic. *) assumes cp\<^sub>P\<^sub>O\<^sub>S\<^sub>T':"POST (self) x (res) \ = POST (\ _. self \) (f_lam x \) (\ _. res \) \" assumes f_\_val: "\a1. f_\ (f_lam a1 \) \ = f_\ a1 \" begin lemma strict0 [simp]: "f invalid X = invalid" by(rule ext, rename_tac "\", simp add: def_scheme' StrongEq_def OclValid_def false_def true_def) lemma nullstrict0[simp]: "f null X = invalid" by(rule ext, rename_tac "\", simp add: def_scheme' StrongEq_def OclValid_def false_def true_def) lemma cp0 : "f self a1 \ = f (\ _. self \) (f_lam a1 \) \" proof - have A: "(\ \ \ (\_. self \)) = (\ \ \ self)" by(simp add: OclValid_def cp_defined[symmetric]) have B: "f_\ (f_lam a1 \) \ = f_\ a1 \" by (rule f_\_val) have D: "(\ \ PRE (\_. self \) (f_lam a1 \)) = ( \ \ PRE self a1 )" by(simp add: OclValid_def cp\<^sub>P\<^sub>R\<^sub>E'[symmetric]) show ?thesis apply(auto simp: def_scheme' A B D) apply(simp add: OclValid_def) by(subst cp\<^sub>P\<^sub>O\<^sub>S\<^sub>T', simp) qed theorem unfold' : assumes context_ok: "cp E" and args_def_or_valid: "(\ \ \ self) \ f_\ a1 \" and pre_satisfied: "\ \ PRE self a1" and post_satisfiable: " \res. (\ \ POST self a1 (\ _. res))" and sat_for_sols_post: "(\res. \ \ POST self a1 (\ _. res) \ \ \ E (\ _. res))" shows "\ \ E(f self a1)" proof - have cp0: "\ X \. E X \ = E (\_. X \) \" by(insert context_ok[simplified cp_def], auto) show ?thesis apply(simp add: OclValid_def, subst cp0, fold OclValid_def) apply(simp add:def_scheme' args_def_or_valid pre_satisfied) apply(insert post_satisfiable, elim exE) apply(rule Hilbert_Choice.someI2, assumption) by(rule sat_for_sols_post, simp) qed lemma unfold2' : assumes context_ok: "cp E" and args_def_or_valid: "(\ \ \ self) \ (f_\ a1 \)" and pre_satisfied: "\ \ PRE self a1" and postsplit_satisfied: "\ \ POST' self a1" (* split constraint holds on post-state *) and post_decomposable : "\ res. (POST self a1 res) = ((POST' self a1) and (res \ (BODY self a1)))" shows "(\ \ E(f self a1)) = (\ \ E(BODY self a1))" proof - have cp0: "\ X \. E X \ = E (\_. X \) \" by(insert context_ok[simplified cp_def], auto) show ?thesis apply(simp add: OclValid_def, subst cp0, fold OclValid_def) apply(simp add:def_scheme' args_def_or_valid pre_satisfied post_decomposable postsplit_satisfied foundation10') apply(subst some_equality) apply(simp add: OclValid_def StrongEq_def true_def)+ by(subst (2) cp0, rule refl) qed end locale contract0 = fixes f :: "('\,'\0::null)val \ ('\,'res::null)val" fixes PRE fixes POST assumes def_scheme: "f self \ (\ \. SOME res. let res = \ _. res in if (\ \ (\ self)) then (\ \ PRE self) \ (\ \ POST self res) else \ \ res \ invalid)" assumes all_post: "\ \ \' \''. ((\,\') \ PRE self) = ((\,\'') \ PRE self)" (* PRE is really a pre-condition semantically, i.e. it does not depend on the post-state. ... *) assumes cp\<^sub>P\<^sub>R\<^sub>E: "PRE (self) \ = PRE (\ _. self \) \ " (* this interface is preferable than : assumes "cp self' \ cp a1' \ cp a2' \ cp (\X. PRE (self' X) (a1' X) (a2' X) )" which is too polymorphic. *) assumes cp\<^sub>P\<^sub>O\<^sub>S\<^sub>T:"POST (self) (res) \ = POST (\ _. self \) (\ _. res \) \" sublocale contract0 < contract_scheme "\_ _. True" "\x _. x" "\x _. f x" "\x _. PRE x" "\x _. POST x" apply(unfold_locales) apply(simp add: def_scheme, rule all_post, rule cp\<^sub>P\<^sub>R\<^sub>E, rule cp\<^sub>P\<^sub>O\<^sub>S\<^sub>T) by simp context contract0 begin lemma cp_pre: "cp self' \ cp (\X. PRE (self' X) )" by(rule_tac f=PRE in cpI1, auto intro: cp\<^sub>P\<^sub>R\<^sub>E) lemma cp_post: "cp self' \ cp res' \ cp (\X. POST (self' X) (res' X))" by(rule_tac f=POST in cpI2, auto intro: cp\<^sub>P\<^sub>O\<^sub>S\<^sub>T) lemma cp [simp]: "cp self' \ cp res' \ cp (\X. f (self' X) )" by(rule_tac f=f in cpI1, auto intro:cp0) lemmas unfold = unfold'[simplified] lemma unfold2 : assumes "cp E" and "(\ \ \ self)" and "\ \ PRE self" and "\ \ POST' self" (* split constraint holds on post-state *) and "\ res. (POST self res) = ((POST' self) and (res \ (BODY self)))" shows "(\ \ E(f self)) = (\ \ E(BODY self))" apply(rule unfold2'[simplified]) by((rule assms)+) end locale contract1 = fixes f :: "('\,'\0::null)val \ ('\,'\1::null)val \ ('\,'res::null)val" fixes PRE fixes POST assumes def_scheme: "f self a1 \ (\ \. SOME res. let res = \ _. res in if (\ \ (\ self)) \ (\ \ \ a1) then (\ \ PRE self a1) \ (\ \ POST self a1 res) else \ \ res \ invalid) " assumes all_post: "\ \ \' \''. ((\,\') \ PRE self a1) = ((\,\'') \ PRE self a1)" (* PRE is really a pre-condition semantically, i.e. it does not depend on the post-state. ... *) assumes cp\<^sub>P\<^sub>R\<^sub>E: "PRE (self) (a1) \ = PRE (\ _. self \) (\ _. a1 \) \ " (* this interface is preferable than : assumes "cp self' \ cp a1' \ cp a2' \ cp (\X. PRE (self' X) (a1' X) (a2' X) )" which is too polymorphic. *) assumes cp\<^sub>P\<^sub>O\<^sub>S\<^sub>T:"POST (self) (a1) (res) \ = POST (\ _. self \)(\ _. a1 \) (\ _. res \) \" sublocale contract1 < contract_scheme "\a1 \. (\ \ \ a1)" "\a1 \. (\ _. a1 \)" apply(unfold_locales) apply(rule def_scheme, rule all_post, rule cp\<^sub>P\<^sub>R\<^sub>E, rule cp\<^sub>P\<^sub>O\<^sub>S\<^sub>T) by(simp add: OclValid_def cp_valid[symmetric]) context contract1 begin lemma strict1[simp]: "f self invalid = invalid" by(rule ext, rename_tac "\", simp add: def_scheme StrongEq_def OclValid_def false_def true_def) lemma defined_mono : "\ \\(f Y Z) \ (\ \\ Y) \ (\ \\ Z)" by(auto simp: valid_def bot_fun_def invalid_def def_scheme StrongEq_def OclValid_def false_def true_def split: if_split_asm) lemma cp_pre: "cp self' \ cp a1' \ cp (\X. PRE (self' X) (a1' X) )" by(rule_tac f=PRE in cpI2, auto intro: cp\<^sub>P\<^sub>R\<^sub>E) lemma cp_post: "cp self' \ cp a1' \ cp res' \ cp (\X. POST (self' X) (a1' X) (res' X))" by(rule_tac f=POST in cpI3, auto intro: cp\<^sub>P\<^sub>O\<^sub>S\<^sub>T) lemma cp [simp]: "cp self' \ cp a1' \ cp res' \ cp (\X. f (self' X) (a1' X))" by(rule_tac f=f in cpI2, auto intro:cp0) lemmas unfold = unfold' lemmas unfold2 = unfold2' end locale contract2 = fixes f :: "('\,'\0::null)val \ ('\,'\1::null)val \ ('\,'\2::null)val \ ('\,'res::null)val" fixes PRE fixes POST assumes def_scheme: "f self a1 a2 \ (\ \. SOME res. let res = \ _. res in if (\ \ (\ self)) \ (\ \ \ a1) \ (\ \ \ a2) then (\ \ PRE self a1 a2) \ (\ \ POST self a1 a2 res) else \ \ res \ invalid) " assumes all_post: "\ \ \' \''. ((\,\') \ PRE self a1 a2) = ((\,\'') \ PRE self a1 a2)" (* PRE is really a pre-condition semantically, i.e. it does not depend on the post-state. ... *) assumes cp\<^sub>P\<^sub>R\<^sub>E: "PRE (self) (a1) (a2) \ = PRE (\ _. self \) (\ _. a1 \) (\ _. a2 \) \ " (* this interface is preferable than : assumes "cp self' \ cp a1' \ cp a2' \ cp (\X. PRE (self' X) (a1' X) (a2' X) )" which is too polymorphic. *) assumes cp\<^sub>P\<^sub>O\<^sub>S\<^sub>T:"\res. POST (self) (a1) (a2) (res) \ = POST (\ _. self \)(\ _. a1 \)(\ _. a2 \) (\ _. res \) \" sublocale contract2 < contract_scheme "\(a1,a2) \. (\ \ \ a1) \ (\ \ \ a2)" "\(a1,a2) \. (\ _.a1 \, \ _.a2 \)" "(\x (a,b). f x a b)" "(\x (a,b). PRE x a b)" "(\x (a,b). POST x a b)" apply(unfold_locales) apply(auto simp add: def_scheme) apply (metis all_post, metis all_post) apply(subst cp\<^sub>P\<^sub>R\<^sub>E, simp) apply(subst cp\<^sub>P\<^sub>O\<^sub>S\<^sub>T, simp) by(simp_all add: OclValid_def cp_valid[symmetric]) context contract2 begin lemma strict0'[simp] : "f invalid X Y = invalid" by(insert strict0[of "(X,Y)"], simp) lemma nullstrict0'[simp]: "f null X Y = invalid" by(insert nullstrict0[of "(X,Y)"], simp) lemma strict1[simp]: "f self invalid Y = invalid" by(rule ext, rename_tac "\", simp add: def_scheme StrongEq_def OclValid_def false_def true_def) lemma strict2[simp]: "f self X invalid = invalid" by(rule ext, rename_tac "\", simp add: def_scheme StrongEq_def OclValid_def false_def true_def) lemma defined_mono : "\ \\(f X Y Z) \ (\ \\ X) \ (\ \\ Y) \ (\ \\ Z)" by(auto simp: valid_def bot_fun_def invalid_def def_scheme StrongEq_def OclValid_def false_def true_def split: if_split_asm) lemma cp_pre: "cp self' \ cp a1' \ cp a2' \ cp (\X. PRE (self' X) (a1' X) (a2' X) )" by(rule_tac f=PRE in cpI3, auto intro: cp\<^sub>P\<^sub>R\<^sub>E) lemma cp_post: "cp self' \ cp a1' \ cp a2' \ cp res' \ cp (\X. POST (self' X) (a1' X) (a2' X) (res' X))" by(rule_tac f=POST in cpI4, auto intro: cp\<^sub>P\<^sub>O\<^sub>S\<^sub>T) lemma cp0' : "f self a1 a2 \ = f (\ _. self \) (\ _. a1 \) (\ _. a2 \) \" by (rule cp0[of _ "(a1,a2)", simplified]) lemma cp [simp]: "cp self' \ cp a1' \ cp a2' \ cp res' \ cp (\X. f (self' X) (a1' X) (a2' X))" by(rule_tac f=f in cpI3, auto intro:cp0') theorem unfold : assumes "cp E" and "(\ \ \ self) \ (\ \ \ a1) \ (\ \ \ a2)" and "\ \ PRE self a1 a2" and " \res. (\ \ POST self a1 a2 (\ _. res))" and "(\res. \ \ POST self a1 a2 (\ _. res) \ \ \ E (\ _. res))" shows "\ \ E(f self a1 a2)" apply(rule unfold'[of _ _ _ "(a1, a2)", simplified]) by((rule assms)+) lemma unfold2 : assumes "cp E" and "(\ \ \ self) \ (\ \ \ a1) \ (\ \ \ a2)" and "\ \ PRE self a1 a2" and "\ \ POST' self a1 a2" (* split constraint holds on post-state *) and "\ res. (POST self a1 a2 res) = ((POST' self a1 a2) and (res \ (BODY self a1 a2)))" shows "(\ \ E(f self a1 a2)) = (\ \ E(BODY self a1 a2))" apply(rule unfold2'[of _ _ _ "(a1, a2)", simplified]) by((rule assms)+) end locale contract3 = fixes f :: "('\,'\0::null)val \ ('\,'\1::null)val \ ('\,'\2::null)val \ ('\,'\3::null)val \ ('\,'res::null)val" fixes PRE fixes POST assumes def_scheme: "f self a1 a2 a3 \ (\ \. SOME res. let res = \ _. res in if (\ \ (\ self)) \ (\ \ \ a1) \ (\ \ \ a2) \ (\ \ \ a3) then (\ \ PRE self a1 a2 a3) \ (\ \ POST self a1 a2 a3 res) else \ \ res \ invalid) " assumes all_post: "\ \ \' \''. ((\,\') \ PRE self a1 a2 a3) = ((\,\'') \ PRE self a1 a2 a3)" (* PRE is really a pre-condition semantically, i.e. it does not depend on the post-state. ... *) assumes cp\<^sub>P\<^sub>R\<^sub>E: "PRE (self) (a1) (a2) (a3) \ = PRE (\ _. self \) (\ _. a1 \) (\ _. a2 \) (\ _. a3 \) \ " (* this interface is preferable than : assumes "cp self' \ cp a1' \ cp a2' \ cp a3' \ cp (\X. PRE (self' X) (a1' X) (a2' X) (a3' X) )" which is too polymorphic. *) assumes cp\<^sub>P\<^sub>O\<^sub>S\<^sub>T:"\res. POST (self) (a1) (a2) (a3) (res) \ = POST (\ _. self \)(\ _. a1 \)(\ _. a2 \)(\ _. a3 \) (\ _. res \) \" sublocale contract3 < contract_scheme "\(a1,a2,a3) \. (\ \ \ a1) \ (\ \ \ a2)\ (\ \ \ a3)" "\(a1,a2,a3) \. (\ _.a1 \, \ _.a2 \, \ _.a3 \)" "(\x (a,b,c). f x a b c)" "(\x (a,b,c). PRE x a b c)" "(\x (a,b,c). POST x a b c)" apply(unfold_locales) apply(auto simp add: def_scheme) apply (metis all_post, metis all_post) apply(subst cp\<^sub>P\<^sub>R\<^sub>E, simp) apply(subst cp\<^sub>P\<^sub>O\<^sub>S\<^sub>T, simp) by(simp_all add: OclValid_def cp_valid[symmetric]) context contract3 begin lemma strict0'[simp] : "f invalid X Y Z = invalid" by(rule ext, rename_tac "\", simp add: def_scheme StrongEq_def OclValid_def false_def true_def) lemma nullstrict0'[simp]: "f null X Y Z = invalid" by(rule ext, rename_tac "\", simp add: def_scheme StrongEq_def OclValid_def false_def true_def) lemma strict1[simp]: "f self invalid Y Z = invalid" by(rule ext, rename_tac "\", simp add: def_scheme StrongEq_def OclValid_def false_def true_def) lemma strict2[simp]: "f self X invalid Z = invalid" by(rule ext, rename_tac "\", simp add: def_scheme StrongEq_def OclValid_def false_def true_def) lemma defined_mono : "\ \\(f W X Y Z) \ (\ \\ W) \ (\ \\ X) \ (\ \\ Y) \ (\ \\ Z)" by(auto simp: valid_def bot_fun_def invalid_def def_scheme StrongEq_def OclValid_def false_def true_def split: if_split_asm) lemma cp_pre: "cp self' \ cp a1' \ cp a2'\ cp a3' \ cp (\X. PRE (self' X) (a1' X) (a2' X) (a3' X) )" by(rule_tac f=PRE in cpI4, auto intro: cp\<^sub>P\<^sub>R\<^sub>E) lemma cp_post: "cp self' \ cp a1' \ cp a2' \ cp a3' \ cp res' \ cp (\X. POST (self' X) (a1' X) (a2' X) (a3' X) (res' X))" by(rule_tac f=POST in cpI5, auto intro: cp\<^sub>P\<^sub>O\<^sub>S\<^sub>T) lemma cp0' : "f self a1 a2 a3 \ = f (\ _. self \) (\ _. a1 \) (\ _. a2 \) (\ _. a3 \) \" by (rule cp0[of _ "(a1,a2,a3)", simplified]) lemma cp [simp]: "cp self' \ cp a1' \ cp a2' \ cp a3' \ cp res' \ cp (\X. f (self' X) (a1' X) (a2' X) (a3' X))" by(rule_tac f=f in cpI4, auto intro:cp0') theorem unfold : assumes "cp E" and "(\ \ \ self) \ (\ \ \ a1) \ (\ \ \ a2) \ (\ \ \ a3)" and "\ \ PRE self a1 a2 a3" and " \res. (\ \ POST self a1 a2 a3 (\ _. res))" and "(\res. \ \ POST self a1 a2 a3 (\ _. res) \ \ \ E (\ _. res))" shows "\ \ E(f self a1 a2 a3)" apply(rule unfold'[of _ _ _ "(a1, a2, a3)", simplified]) by((rule assms)+) lemma unfold2 : assumes "cp E" and "(\ \ \ self) \ (\ \ \ a1) \ (\ \ \ a2) \ (\ \ \ a3)" and "\ \ PRE self a1 a2 a3" and "\ \ POST' self a1 a2 a3" (* split constraint holds on post-state *) and "\ res. (POST self a1 a2 a3 res) = ((POST' self a1 a2 a3) and (res \ (BODY self a1 a2 a3)))" shows "(\ \ E(f self a1 a2 a3)) = (\ \ E(BODY self a1 a2 a3))" apply(rule unfold2'[of _ _ _ "(a1, a2, a3)", simplified]) by((rule assms)+) end end