Theory Conclusion
chapter‹ Conclusion ›
theory Conclusion
imports DeadlockResults "HOL-Library.LaTeXsugar"
begin
text ‹In this session, we defined five architectural operators: \<^const>‹MultiDet›, \<^const>‹MultiNdet›
and \<^const>‹GlobalNdet›, \<^const>‹MultiSync›, and \<^const>‹MultiSeq› as respective generalizations
of \<^term>‹P □ Q›, \<^term>‹P ⊓ Q›, \<^term>‹P ⟦S⟧ Q›, and \<^term>‹P ❙; Q›.›
text ‹We did this in a fully-abstract way, that is:
▪ \<^const>‹Det› is commutative, idempotent and admits \<^const>‹STOP› as a neutral element so
we defined \<^const>‹MultiDet› on a \<^const>‹finite› \<^typ>‹'α set› \<^term>‹A› by making it equal
to \<^const>‹STOP› when \<^term>‹A = {}›.
▪ \<^const>‹Ndet› is also commutative and idempotent so we defined \<^const>‹MultiNdet›
on a \<^const>‹finite› \<^typ>‹'α set› \<^term>‹A› by making it equal
to \<^const>‹STOP› when \<^term>‹A = {}›. Beware of the fact that \<^const>‹STOP› is not the
neutral element for \<^const>‹Ndet› (this operator does not admit a neutral element) so
we ❙‹do not have› the equality
@{term [display, eta_contract = false] ‹(⨅p ∈ {a}. P p) = P a ⊓ (⨅p ∈ {}. P p)›}
while this holds for \<^const>‹Det› and \<^const>‹MultiDet›).
As its failures and divergences can easily be generalized to the infinite case,
we also defined \<^const>‹GlobalNdet› verifying
@{thm [display, eta_contract = false] finite_GlobalNdet_is_MultiNdet[no_vars]}
▪ \<^const>‹Sync› is commutative but is not idempotent so we defined \<^const>‹MultiSync›
on a \<^typ>‹'α multiset› \<^term>‹M› to keep the multiplicity of the processes.
We made it equal to \<^const>‹STOP› when \<^term>‹M = {#}› but like \<^const>‹Ndet›,
\<^const>‹Sync› does not admit a neutral element so beware of the fact that in general
@{term [display, eta_contract = false] ‹(❙⟦S❙⟧p ∈# {#a#}. P p) ≠ P a ⟦S⟧ (❙⟦S❙⟧p ∈# {#}. P p)›}
▪ \<^const>‹Seq› is neither commutative nor idempotent, so we defined \<^const>‹MultiSeq›
on a \<^typ>‹'α list› \<^term>‹L› to keep the multiplicity and the order of the processes.
Since \<^const>‹SKIP› is the neutral element for \<^const>‹Seq›, we have
@{term [display, eta_contract = false] ‹(SEQ p ∈@ [a]. P p) = (SEQ p ∈@ []. P p) ❙; P a›}
@{term [display, eta_contract = false] ‹(SEQ p ∈@ [a]. P p) = P a ❙; (SEQ p ∈@ []. P p)›}›
text ‹On our architectural operators we proved continuity (under weakest liberal assumptions),
wrote refinements rules and obtained results about the behaviour with other
operators inherited from the binary rules.›
text ‹We presented two examples: Dining Philosophers, and POTS.
In both, we underlined the usefulness of the architectural operators
for modeling complex systems.›
text ‹Finally we provided powerful results on \<^const>‹events_of› and \<^const>‹deadlock_free›
among which the most important is undoubtedly :
@{thm [display, eta_contract = false] MultiInter_deadlock_free[no_vars]}
This theorem allows, for example, to establish:
@{term [display, eta_contract = false]
‹0 < n ⟹ deadlock_free (❙|❙|❙| i ∈# mset [0..<n]. P i)›}
under the assumption that a family of processes parameterized by
\<^term>‹i› @{text "::"} \<^typ>‹nat› verifies \<^term>‹∀i < n::nat. deadlock_free (P i)›.
›
end