• Aucun résultat trouvé

Denition (of the association Employee-Boss)

Dans le document FEATHERWEIGHT OCL (Page 167-178)

1.37. Example I : The Employee Analysis Model (UML)

1.45.1. Denition (of the association Employee-Boss)

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, pretty much similar to the Design_UML, where we stored an oid inside the class as pointer.

denition oidP er sonBOSS ::oid where oidP er sonBOSS=10

From there on, we can already dene an empty state which must contain for oidP ersonBOSS the empty relation (encoded as association list, since there are associations with a Sequence-like structure).

denition eval-extract :: (0A,(0a::object)option option)val

⇒(oid ⇒(0A,0c::null) val)

⇒(0A,0c::null) val where eval-extract X f = (λ τ .case X τ of

⊥ ⇒invalid τ (∗exception propagation ∗)

|x ⊥y⇒invalid τ (∗dereferencing null pointer ∗)

|xxobjyy⇒f (oid-of obj)τ)

denition choose2-1 =fst denition choose2-2 =snd

denition List-atten = (λl.(foldl ((λacc.(λl.(foldl ((λacc.(λl.(Cons(l) (acc))))) (acc) ((rev(l))))))) (Nil) ((rev (l)))))

denition deref-assocs2 :: (0Astate × 0Astate ⇒ 0Astate)

⇒(oid list list ⇒oid list ×oid list)

⇒oid

⇒(oid list ⇒(0A,0f)val)

⇒oid

⇒(0A, 0f::null)val

where deref-assocs2 pre-post to-from assoc-oid f oid = (λτ .case (assocs (pre-post τ))assoc-oid of

xSy⇒f (List-atten (map (choose2-2 ◦to-from)

(lter (λp.List.member (choose2-1 (to-from p))oid) S))) τ

|- ⇒invalid τ)

The pre-post-parameter is congured with fst or snd, the to-from-parameter either with the identity id or the following combinator switch:

denition switch2-1 = (λ[x,y]⇒(x,y)) denition switch2-2 = (λ[x,y]⇒(y,x)) denition switch3-1 = (λ[x,y,z]⇒(x,y)) denition switch3-2 = (λ[x,y,z]⇒(x,z)) denition switch3-3 = (λ[x,y,z]⇒(y,x)) denition switch3-4 = (λ[x,y,z]⇒(y,z)) denition switch3-5 = (λ[x,y,z]⇒(z,x)) denition switch3-6 = (λ[x,y,z]⇒(z,y))

denition deref-oidP er son:: (Astate ×Astate ⇒Astate)

⇒(typeP er son⇒(A, 0c::null)val)

⇒oid

⇒(A, 0c::null)val

where deref-oidP er sonfst-snd f oid = (λτ .case (heap (fst-snd τ))oid of xinP er sonobjy⇒f obj τ

|- ⇒invalid τ)

denition deref-oidO clAny :: (Astate ×Astate ⇒Astate)

⇒(typeO clAny ⇒(A, 0c::null)val)

⇒oid

⇒(A, 0c::null)val

where deref-oidO clAnyfst-snd f oid = (λτ .case(heap(fst-snd τ))oid of xinO clAny objy⇒f obj τ

|- ⇒invalid τ)

pointer undened in state or not referencing a type conform object representation denition selectO clAnyAN Y f = (λX.case X of

(mkO clAny -⊥)⇒null

|(mkO clAny -xanyy) ⇒f (λx -.xxxyy)any)

denition selectP er sonBOSSf =select-object mtSet UML-Set.OclIncluding UML-Set.OclANY (f (λx -.xxxyy))

denition selectP er sonSALARY f = (λX.case X of (mkP er son-⊥)⇒null

|(mkP er son-xsalaryy)⇒f (λx -.xxxyy) salary)

denition deref-assocs2BOSS fst-snd f = (λmkP er sonoid -⇒ deref-assocs2 fst-snd switch2-1 oidP er sonBOSS f oid) denition in-pre-state =fst

denition in-post-state =snd

denition reconst-basetype= (λconvert x.convert x) denition dotO clAnyAN Y ::OclAny ⇒- ((1(-).any)50)

where(X).any =eval-extract X

(deref-oidO clAny in-post-state (selectO clAnyAN Y

reconst-basetype))

denition dotP er sonBOSS::Person ⇒Person ((1(-).boss)50) where(X).boss =eval-extract X

(deref-oidP er sonin-post-state (deref-assocs2BOSSin-post-state

(selectP er sonBOSS

(deref-oidP er son in-post-state))))

denition dotP er sonSALARY ::Person ⇒Integer ((1(-).salary)50) where(X).salary =eval-extract X

(deref-oidP er son in-post-state (selectP er sonSALARY

reconst-basetype))

denition dotO clAnyAN Y-at-pre ::OclAny ⇒- ((1(-).any@pre) 50) where(X).any@pre=eval-extract X

(deref-oidO clAny in-pre-state (selectO clAnyAN Y

reconst-basetype))

denition dotP er sonBOSS-at-pre::Person ⇒Person ((1(-).boss@pre)50) where(X).boss@pre=eval-extract X

(deref-oidP er son in-pre-state (deref-assocs2BOSSin-pre-state

(selectP er sonBOSS

(deref-oidP er sonin-pre-state))))

denition dotP er sonSALARY-at-pre::Person ⇒Integer ((1(-).salary@pre) 50) where(X).salary@pre=eval-extract X

(deref-oidP er son in-pre-state (selectP er sonSALARY

reconst-basetype))

lemmas dot-accessor = dotO clAnyAN Y-def dotP er sonBOSS-def dotP er sonSALARY-def dotO clAnyAN Y-at-pre-def dotP er sonBOSS-at-pre-def dotP er sonSALARY-at-pre-def

1.45.2. Context Passing

lemmas [simp] =eval-extract-def

lemma cp-dotO clAnyAN Y: ((X).any)τ = ((λ-.X τ).any)τ hproofi lemma cp-dotP er sonBOSS: ((X).boss)τ = ((λ-.X τ).boss)τ hproofi lemma cp-dotP er sonSALARY: ((X).salary)τ = ((λ-.X τ).salary) τ hproofi lemma cp-dotO clAnyAN Y-at-pre: ((X).any@pre) τ = ((λ-.X τ).any@pre)τ hproofi lemma cp-dotP er sonBOSS-at-pre: ((X).boss@pre)τ = ((λ-.X τ).boss@pre)τ hproofi lemma cp-dotP er sonSALARY-at-pre: ((X).salary@pre) τ = ((λ-.X τ).salary@pre)τ hproofi lemmas cp-dotO clAnyAN Y-I [simp,intro!]=

cp-dotO clAnyAN Y[THEN allI[THEN allI], of λX -.X λ-τ . τ ,THEN cpI1] lemmas cp-dotO clAnyAN Y-at-pre-I [simp,intro!]=

cp-dotO clAnyAN Y-at-pre[THEN allI[THEN allI], of λX -.X λ-τ . τ ,THEN cpI1] lemmas cp-dotP er sonBOSS-I [simp,intro!]=

cp-dotP er sonBOSS[THEN allI[THEN allI], of λX -.X λ-τ . τ ,THEN cpI1] lemmas cp-dotP er sonBOSS-at-pre-I [simp,intro!]=

cp-dotP er sonBOSS-at-pre[THEN allI[THEN allI], of λX -.X λ-τ . τ ,THEN cpI1] lemmas cp-dotP er sonSALARY-I [simp,intro!]=

cp-dotP er sonSALARY[THEN allI[THEN allI], of λX -.X λ-τ . τ ,THEN cpI1] lemmas cp-dotP er sonSALARY-at-pre-I [simp,intro!]=

cp-dotP er sonSALARY-at-pre[THEN allI[THEN allI], of λX -.X λ-τ . τ ,THEN cpI1]

1.45.3. Execution with Invalid or Null as Argument

lemma dotO clAnyAN Y-nullstrict [simp]: (null).any=invalid hproofi

lemma dotO clAnyAN Y-at-pre-nullstrict [simp] : (null).any@pre =invalid hproofi

lemma dotO clAnyAN Y-strict [simp] : (invalid).any =invalid hproofi

lemma dotO clAnyAN Y-at-pre-strict [simp] : (invalid).any@pre =invalid hproofi

lemma dotP er sonBOSS-nullstrict [simp]: (null).boss =invalid hproofi

lemma dotP er sonBOSS-at-pre-nullstrict [simp] : (null).boss@pre=invalid

p1:Person

lemma dotP er sonBOSS-strict [simp] : (invalid).boss =invalid hproofi

lemma dotP er sonBOSS-at-pre-strict [simp] : (invalid).boss@pre =invalid hproofi

lemma dotP er sonSALARY-nullstrict [simp]: (null).salary =invalid hproofi

lemma dotP er sonSALARY-at-pre-nullstrict [simp] : (null).salary@pre=invalid hproofi

lemma dotP er sonSALARY-strict [simp] : (invalid).salary =invalid hproofi

lemma dotP er sonSALARY-at-pre-strict [simp] : (invalid).salary@pre=invalid hproofi

shows is-represented-in-state in-post-state(x .boss)Person τ hproofi

lemma repr-bossX : assumes A:τ |=δ(x .boss)

showsτ |= ((Person .allInstances())−>includesS et(x .boss)) hproofi

1.46. A Little Infra-structure on Example States

The example we are dening in this section comes from the gure 1.6.

denition OclInt1000 (1000)where OclInt1000 = (λ-.xx1000yy) denition OclInt1200 (1200)where OclInt1200 = (λ-.xx1200yy) denition OclInt1300 (1300)where OclInt1300 = (λ-.xx1300yy) denition OclInt1800 (1800)where OclInt1800 = (λ-.xx1800yy) denition OclInt2600 (2600)where OclInt2600 = (λ-.xx2600yy) denition OclInt2900 (2900)where OclInt2900 = (λ-.xx2900yy) denition OclInt3200 (3200)where OclInt3200 = (λ-.xx3200yy) denition OclInt3500 (3500)where OclInt3500 = (λ-.xx3500yy)

denition oid0 ≡0 denition oid1 ≡1 denition oid2 ≡2 denition oid3 ≡3 denition oid4 ≡4 denition oid5 ≡5 denition oid6 ≡6 denition oid7 ≡7 denition oid8 ≡8

denition person1 ≡mkP er sonoid0 x1300y denition person2 ≡mkP er sonoid1 x1800y denition person3 ≡mkP er sonoid2 None denition person4 ≡mkP er sonoid3 x2900y denition person5 ≡mkP er sonoid4 x3500y denition person6 ≡mkP er sonoid5 x2500y denition person7 ≡mkO clAnyoid6 xx3200yy denition person8 ≡mkO clAnyoid7 None

denition person9 ≡mkP er sonoid8 x0ydenition

σ1 ≡(|heap =empty(oid0 7→inP er son(mkP er sonoid0 x1000y)) (oid1 7→inP er son(mkP er sonoid1 x1200y)) (∗oid2∗)

(oid3 7→inP er son(mkP er sonoid3 x2600y)) (oid4 7→inP er sonperson5)

(oid5 7→inP er son(mkP er sonoid5 x2300y)) (∗oid6∗)

(∗oid7∗)

(oid8 7→inP er sonperson9),

assocs =empty(oidP er sonBOSS 7→[[[oid0],[oid1]],[[oid3],[oid4]],[[oid5],[oid3]]])|) denition

σ10≡(|heap =empty(oid0 7→inP er sonperson1) (oid1 7→inP er sonperson2) (oid2 7→inP er sonperson3) (oid3 7→inP er sonperson4) (∗oid4∗)

(oid5 7→inP er sonperson6) (oid6 7→inO clAny person7) (oid7 7→inO clAny person8) (oid8 7→inP er sonperson9),

assocs =empty(oidP er sonBOSS 7→[[[oid0],[oid1]],[[oid1],[oid1]],[[oid5],[oid6]],[[oid6],[oid6]]])|) denitionσ0 ≡(|heap =empty,assocs =empty |)

lemma basic-τ-w:WFF(σ110

) hproofi

lemma[simp,code-unfold]:dom (heap σ1) ={oid0,oid1,(∗,oid2∗)oid3,oid4,oid5(∗,oid6,oid7∗),oid8} hproofi

lemma[simp,code-unfold]:dom (heap σ10

) ={oid0,oid1,oid2,oid3,(∗,oid4∗)oid5,oid6,oid7,oid8} hproofidenition XP er son1 ::Person ≡λ-.xxperson1yy

denition XP er son2 ::Person ≡λ-.xxperson2yy denition XP er son3 ::Person ≡λ-.xxperson3yy denition XP er son4 ::Person ≡λ-.xxperson4yy

denition XP er son5 ::Person ≡λ-.xxperson5yy

AssertV spost. (σ1,spost)|= (XP er son5 .salary@pre .

=3500) lemma (σ110

)|= (XP er son5 .oclIsDeleted()) hproofi

lemma (σ110

)|= (XP er son6 .oclIsMaintained()) hproofi

AssertVspr e spost. (spr e,spost) |= υ(XP er son7 .oclAsType(Person))

lemmaVspr espost. (spr e,spost)|= ((XP er son7 .oclAsType(Person).oclAsType(OclAny) .oclAsType(Person))

= (. XP er son7 .oclAsType(Person))) hproofi

lemma (σ110

)|= (XP er son7 .oclIsNew()) hproofi

AssertVspr e spost. (spr e,spost) |= (XP er son8 <>XP er son7) AssertVspr e spost. (spr e,spost) |=not(υ(XP er son8 .oclAsType(Person))) AssertVspr e spost. (spr e,spost) |= (XP er son8 .oclIsTypeOf(OclAny)) AssertVspr e spost. (spr e,spost) |= not(XP er son8 .oclIsTypeOf(Person)) AssertVspr e spost. (spr e,spost) |= not(XP er son8 .oclIsKindOf(Person)) AssertVspr e spost. (spr e,spost) |= (XP er son8 .oclIsKindOf(OclAny)) lemmaσ-modiedonly: (σ110

)|= (Set{XP er son1 .oclAsType(OclAny) ,XP er son2 .oclAsType(OclAny)

(∗,XP er son3 .oclAsType(OclAny)∗) ,XP er son4 .oclAsType(OclAny) (∗,XP er son5 .oclAsType(OclAny)∗)

,XP er son6 .oclAsType(OclAny) (∗,XP er son7 .oclAsType(OclAny)∗) (∗,XP er son8 .oclAsType(OclAny)∗)

(∗,XP er son9 .oclAsType(OclAny)∗)}−>oclIsModiedOnly()) hproofi

lemma(σ110

)|= ((XP er son9 @pre (λx.xOclAsTypeP er son-Axy)) ,XP er son9) hproofi

lemma(σ110

)|= ((XP er son9 @post (λx.xOclAsTypeP er son-Axy)) , XP er son9) hproofi

lemma(σ110

)|= (((XP er son9 .oclAsType(OclAny)) @pre (λx.xOclAsTypeO clAny-Axy)), ((XP er son9 .oclAsType(OclAny)) @post(λx.xOclAsTypeO clAny-Axy))) hproofi

lemma perm-σ10

10

= (|heap=empty (oid8 7→inP er sonperson9) (oid7 7→inO clAny person8) (oid6 7→inO clAny person7) (oid5 7→inP er sonperson6) (∗oid4∗)

(oid3 7→inP er sonperson4)

(oid2 7→inP er sonperson3) (oid1 7→inP er sonperson2) (oid0 7→inP er sonperson1) ,assocs =assocsσ10|) hproofi

declare const-ss [simp] lemmaV

σ1. (σ110

) |= (Person .allInstances() .

= Set{ XP er son1, XP er son2, XP er son3, XP er son4(∗, XP er son5∗), XP er son6,

XP er son7 .oclAsType(Person)(∗,XP er son8∗),XP er son9 }) hproofi

lemmaV σ1. (σ110

) |= (OclAny .allInstances() .

=Set{XP er son1 .oclAsType(OclAny),XP er son2 .oclAsType(OclAny), XP er son3 .oclAsType(OclAny),XP er son4 .oclAsType(OclAny)

(∗,XP er son5∗),XP er son6 .oclAsType(OclAny),

XP er son7,XP er son8,XP er son9 .oclAsType(OclAny)}) hproofi

end

theory Analysis-OCL imports

Analysis-UML begin

1.47. OCL Part: Standard State Infrastructure

Ideally, these denitions are automatically generated from the class model.

1.48. Invariant

These recursive predicates can be dened conservatively by greatest x-point constructions automatically. See [4, 6] for details. For the purpose of this example, we state them as axioms here.

context Person

inv label : self . boss <> null implies ( self . salary \<le >

(( self . boss ) . salary ))

denition Person-labelinv ::Person ⇒Boolean where Person-labelinv (self) ≡

(self .boss <>null implies (self .salary ≤int ((self .boss).salary)))

denition Person-labelinv AT pr e::Person ⇒Boolean where Person-labelinv AT pr e (self)≡

(self .boss@pre <>null implies (self .salary@pre ≤int ((self .boss@pre).salary@pre))) denition Person-labelg lobalinv ::Boolean

where Person-labelg lobalinv ≡(Person .allInstances()−>forAllS et(x |Person-labelinv (x))and (Person .allInstances@pre()−>forAllS et(x |Person-labelinv AT pr e (x))))

lemmaτ |=δ (X .boss) =⇒τ |=Person .allInstances()−>includesS et(X .boss) ∧ τ |=Person .allInstances()−>includesS et(X)

hproofi

lemma REC-pre :τ |=Person-labelg lobalinv

=⇒τ |=Person .allInstances()−>includesS et(X) (∗X represented object in state ∗)

=⇒ ∃ REC. τ |=REC(X) , (Person-labelinv (X)and (X .boss <>null implies REC(X .boss))) hproofi

This allows to state a predicate:

axiomatization invP er son-label ::Person ⇒Boolean where invP er son-label-def:

(τ |=Person .allInstances()−>includesS et(self)) =⇒ (τ |= (invP er son-label(self), (self .boss <>null implies

(self .salary ≤int ((self .boss).salary))and invP er son-label(self .boss))))

axiomatization invP er son-labelAT pr e ::Person ⇒Boolean where invP er son-labelAT pr e-def:

(τ |=Person .allInstances@pre()−>includesS et(self)) =⇒

(τ |= (invP er son-labelAT pr e(self), (self .boss@pre <>null implies

(self .salary@pre ≤int ((self .boss@pre).salary@pre))and invP er son-labelAT pr e(self .boss@pre))))

lemma inv-1 :

(τ |=Person .allInstances()−>includesS et(self)) =⇒ (τ |=invP er son-label(self) = ((τ |= (self .boss .

=null))∨ (τ |= (self .boss <>null)∧

τ |= ((self .salary) ≤int (self .boss .salary)) ∧ τ |= (invP er son-label(self .boss)))))

hproofi

lemma inv-2 :

(τ |=Person .allInstances@pre()−>includesS et(self)) =⇒ (τ |=invP er son-labelAT pr e(self)) = ((τ |= (self .boss@pre .

=null))∨ (τ |= (self .boss@pre<>null)∧

(τ |= (self .boss@pre.salary@pre≤int self .salary@pre)) ∧ (τ |= (invP er son-labelAT pr e(self .boss@pre)))))

hproofi

A very rst attempt to characterize the axiomatization by an inductive denition - this can not be the last word since too weak (should be equality!)

coinductive inv ::Person ⇒(A)st ⇒bool where (τ |= (δself)) =⇒((τ |= (self .boss .

=null))∨

(τ |= (self .boss <>null)∧(τ |= (self .boss .salary ≤intself .salary)) ∧ ( (inv(self .boss))τ )))

=⇒( inv self τ)

1.49. The Contract of a Recursive Query

The original specication of a recursive query : context Person :: contents (): Set ( Integer ) pre : true

post : result = if self . boss = null then Set {i}

else self . boss . contents ()-> including (i) endif

For the case of recursive queries, we use at present just axiomatizations:

axiomatization contents::Person ⇒Set-Integer ((1(-).contents0(0))50) where contents-def: interpretation contents :contract0 contentsλself.true

λself res. res ,if (self .boss . practical rewrite rule that is amenable to symbolic evaluation:

theorem unfold-contents :

else self .boss .contents()−>includingS et(self .salary)endif)) hproofi

Since we have only one interpretation function, we need the corresponding operation on the pre-state:

consts contentsATpre::Person ⇒Set-Integer ((1(-).contents@pre0(0))50) axiomatization where contentsATpre-def:

else (self).boss@pre.contents@pre()

−>includingS et(self .salary@pre) endif)))

else τ |=res ,invalid)

and cp0-contents-at-pre:(X .contents@pre())τ = ((λ-.X τ).contents@pre()) τ interpretation contentsATpre :contract0 contentsATpre λself.true

λself res. res ,if (self .boss@pre .

=null)

then(Set{self .salary@pre}) else(self .boss@pre.contents@pre()

−>includingS et(self .salary@pre)) endif

hproofi

Again, we derive via contents.unfold2 a Knaster-Tarski like Fixpoint rule that is amenable to symbolic evaluation:

theorem unfold-contentsATpre : assumes cp E

and τ |=δ self

shows (τ |=E (self .contents@pre())) = (τ |=E (if self .boss@pre .

=null then Set{self .salary@pre}

else self .boss@pre .contents@pre()−>includingS et(self .salary@pre)endif)) hproofi

Note that these @pre variants on methods are only available on queries, i. e., operations without side-eect.

Dans le document FEATHERWEIGHT OCL (Page 167-178)