Up to index of Isabelle/HOL/HOL-Nominal/CCS
theory Weak_Sim_Pres(*
Title: The Calculus of Communicating Systems
Author/Maintainer: Jesper Bengtson (jebe@itu.dk), 2012
*)
theory Weak_Sim_Pres
imports Weak_Sim
begin
lemma actPres:
fixes P :: ccs
and Q :: ccs
and Rel :: "(ccs × ccs) set"
and a :: name
and Rel' :: "(ccs × ccs) set"
assumes "(P, Q) ∈ Rel"
shows "α.(P) \<leadsto>⇧^<Rel> α.(Q)"
using assms
by(fastsimp simp add: weakSimulation_def elim: actCases intro: weakAction)
lemma sumPres:
fixes P :: ccs
and Q :: ccs
and Rel :: "(ccs × ccs) set"
assumes "P \<leadsto>⇧^<Rel> Q"
and "Rel ⊆ Rel'"
and "Id ⊆ Rel'"
and C1: "!!S T U. (S, T) ∈ Rel ==> (S ⊕ U, T) ∈ Rel'"
shows "P ⊕ R \<leadsto>⇧^<Rel'> Q ⊕ R"
proof(induct rule: weakSimI)
case(Sim α QR)
from `Q ⊕ R \<longmapsto>α \<prec> QR` show ?case
proof(induct rule: sumCases)
case(cSum1 Q')
from `P \<leadsto>⇧^<Rel> Q` `Q \<longmapsto>α \<prec> Q'`
obtain P' where "P ==>⇧^α \<prec> P'" and "(P', Q') ∈ Rel"
by(blast dest: weakSimE)
thus ?case
proof(induct rule: weakTransCases)
case Base
have "P ⊕ R ==>⇧^τ \<prec> P ⊕ R" by simp
moreover from `(P, Q') ∈ Rel` have "(P ⊕ R, Q') ∈ Rel'" by(rule C1)
ultimately show ?case by blast
next
case Step
from `P ==>α \<prec> P'` have "P ⊕ R ==>α \<prec> P'" by(rule weakCongSum1)
hence "P ⊕ R ==>⇧^α \<prec> P'" by(simp add: weakTrans_def)
thus ?case using `(P', Q') ∈ Rel` `Rel ⊆ Rel'` by blast
qed
next
case(cSum2 R')
from `R \<longmapsto>α \<prec> R'` have "R ==>α \<prec> R'" by(rule transitionWeakCongTransition)
hence "P ⊕ R ==>α \<prec> R'" by(rule weakCongSum2)
hence "P ⊕ R ==>⇧^α \<prec> R'" by(simp add: weakTrans_def)
thus ?case using `Id ⊆ Rel'` by blast
qed
qed
lemma parPresAux:
fixes P :: ccs
and Q :: ccs
and R :: ccs
and T :: ccs
and Rel :: "(ccs × ccs) set"
and Rel' :: "(ccs × ccs) set"
and Rel'' :: "(ccs × ccs) set"
assumes "P \<leadsto>⇧^<Rel> Q"
and "(P, Q) ∈ Rel"
and "R \<leadsto>⇧^<Rel'> T"
and "(R, T) ∈ Rel'"
and C1: "!!P' Q' R' T'. [|(P', Q') ∈ Rel; (R', T') ∈ Rel'|] ==> (P' \<parallel> R', Q' \<parallel> T') ∈ Rel''"
shows "P \<parallel> R \<leadsto>⇧^<Rel''> Q \<parallel> T"
proof(induct rule: weakSimI)
case(Sim α QT)
from `Q \<parallel> T \<longmapsto>α \<prec> QT`
show ?case
proof(induct rule: parCases)
case(cPar1 Q')
from `P \<leadsto>⇧^<Rel> Q` `Q \<longmapsto>α \<prec> Q'` obtain P' where "P ==>⇧^α \<prec> P'" and "(P', Q') ∈ Rel"
by(rule weakSimE)
from `P ==>⇧^α \<prec> P'` have "P \<parallel> R ==>⇧^α \<prec> P' \<parallel> R" by(rule weakPar1)
moreover from `(P', Q') ∈ Rel` `(R, T) ∈ Rel'` have "(P' \<parallel> R, Q' \<parallel> T) ∈ Rel''" by(rule C1)
ultimately show ?case by blast
next
case(cPar2 T')
from `R \<leadsto>⇧^<Rel'> T` `T \<longmapsto>α \<prec> T'` obtain R' where "R ==>⇧^α \<prec> R'" and "(R', T') ∈ Rel'"
by(rule weakSimE)
from `R ==>⇧^α \<prec> R'` have "P \<parallel> R ==>⇧^α \<prec> P \<parallel> R'" by(rule weakPar2)
moreover from `(P, Q) ∈ Rel` `(R', T') ∈ Rel'` have "(P \<parallel> R', Q \<parallel> T') ∈ Rel''" by(rule C1)
ultimately show ?case by blast
next
case(cComm Q' T' α)
from `P \<leadsto>⇧^<Rel> Q` `Q \<longmapsto>α \<prec> Q'` obtain P' where "P ==>⇧^α \<prec> P'" and "(P', Q') ∈ Rel"
by(rule weakSimE)
from `R \<leadsto>⇧^<Rel'> T` `T \<longmapsto>(coAction α) \<prec> T'` obtain R' where "R ==>⇧^(coAction α) \<prec> R'" and "(R', T') ∈ Rel'"
by(rule weakSimE)
from `P ==>⇧^α \<prec> P'` `R ==>⇧^(coAction α) \<prec> R'` `α ≠ τ` have "P \<parallel> R ==>τ \<prec> P' \<parallel> R'"
by(auto intro: weakCongSync simp add: weakTrans_def)
hence "P \<parallel> R ==>⇧^τ \<prec> P' \<parallel> R'" by(simp add: weakTrans_def)
moreover from `(P', Q') ∈ Rel` `(R', T') ∈ Rel'` have "(P' \<parallel> R', Q' \<parallel> T') ∈ Rel''" by(rule C1)
ultimately show ?case by blast
qed
qed
lemma parPres:
fixes P :: ccs
and Q :: ccs
and R :: ccs
and Rel :: "(ccs × ccs) set"
and Rel' :: "(ccs × ccs) set"
assumes "P \<leadsto>⇧^<Rel> Q"
and "(P, Q) ∈ Rel"
and C1: "!!S T U. (S, T) ∈ Rel ==> (S \<parallel> U, T \<parallel> U) ∈ Rel'"
shows "P \<parallel> R \<leadsto>⇧^<Rel'> Q \<parallel> R"
using assms
by(rule_tac parPresAux[where Rel'=Id and Rel''=Rel']) (auto intro: reflexive)
lemma resPres:
fixes P :: ccs
and Rel :: "(ccs × ccs) set"
and Q :: ccs
and x :: name
assumes "P \<leadsto>⇧^<Rel> Q"
and "!!R S y. (R, S) ∈ Rel ==> ((|νy|)),R, (|νy|)),S) ∈ Rel'"
shows "(|νx|)),P \<leadsto>⇧^<Rel'> (|νx|)),Q"
using assms
by(fastsimp simp add: weakSimulation_def elim: resCases intro: weakRes)
lemma bangPres:
fixes P :: ccs
and Rel :: "(ccs × ccs) set"
and Q :: ccs
assumes "(P, Q) ∈ Rel"
and C1: "!!R S. (R, S) ∈ Rel ==> R \<leadsto>⇧^<Rel> S"
and Par: "!!R S T U. [|(R, S) ∈ Rel; (T, U) ∈ Rel'|] ==> (R \<parallel> T, S \<parallel> U) ∈ Rel'"
and C2: "bangRel Rel ⊆ Rel'"
and C3: "!!R S. (R \<parallel> !R, S) ∈ Rel' ==> (!R, S) ∈ Rel'"
shows "!P \<leadsto>⇧^<Rel'> !Q"
proof(induct rule: weakSimI)
case(Sim α Q')
{
fix Pa α Q'
assume "!Q \<longmapsto>α \<prec> Q'" and "(Pa, !Q) ∈ bangRel Rel"
hence "∃P'. Pa ==>⇧^α \<prec> P' ∧ (P', Q') ∈ Rel'"
proof(nominal_induct arbitrary: Pa rule: bangInduct)
case(cPar1 α Q')
from `(Pa, Q \<parallel> !Q) ∈ bangRel Rel`
show ?case
proof(induct rule: BRParCases)
case(BRPar P R)
from `(P, Q) ∈ Rel` have "P \<leadsto>⇧^<Rel> Q" by(rule C1)
with `Q \<longmapsto>α \<prec> Q'` obtain P' where "P ==>⇧^α \<prec> P'" and "(P', Q') ∈ Rel"
by(blast dest: weakSimE)
from `P ==>⇧^α \<prec> P'` have "P \<parallel> R ==>⇧^α \<prec> P' \<parallel> R" by(rule weakPar1)
moreover from `(P', Q') ∈ Rel` `(R, !Q) ∈ bangRel Rel` C2 have "(P' \<parallel> R, Q' \<parallel> !Q) ∈ Rel'"
by(blast intro: Par)
ultimately show ?case by blast
qed
next
case(cPar2 α Q')
from `(Pa, Q \<parallel> !Q) ∈ bangRel Rel`
show ?case
proof(induct rule: BRParCases)
case(BRPar P R)
from `(R, !Q) ∈ bangRel Rel` obtain R' where "R ==>⇧^α \<prec> R'" and "(R', Q') ∈ Rel'" using cPar2
by blast
from `R ==>⇧^α \<prec> R'` have "P \<parallel> R ==>⇧^α \<prec> P \<parallel> R'" by(rule weakPar2)
moreover from `(P, Q) ∈ Rel` `(R', Q') ∈ Rel'` have "(P \<parallel> R', Q \<parallel> Q') ∈ Rel'" by(rule Par)
ultimately show ?case by blast
qed
next
case(cComm a Q' Q'' Pa)
from `(Pa, Q \<parallel> !Q) ∈ bangRel Rel`
show ?case
proof(induct rule: BRParCases)
case(BRPar P R)
from `(P, Q) ∈ Rel` have "P \<leadsto>⇧^<Rel> Q" by(rule C1)
with `Q \<longmapsto>a \<prec> Q'` obtain P' where "P ==>⇧^a \<prec> P'" and "(P', Q') ∈ Rel"
by(blast dest: weakSimE)
from `(R, !Q) ∈ bangRel Rel` obtain R' where "R ==>⇧^(coAction a) \<prec> R'" and "(R', Q'') ∈ Rel'" using cComm
by blast
from `P ==>⇧^a \<prec> P'` `R ==>⇧^(coAction a) \<prec> R'` `a ≠ τ` have "P \<parallel> R ==>⇧^τ \<prec> P' \<parallel> R'"
by(auto intro: weakCongSync simp add: weakTrans_def)
moreover from `(P', Q') ∈ Rel` `(R', Q'') ∈ Rel'` have "(P' \<parallel> R', Q' \<parallel> Q'') ∈ Rel'" by(rule Par)
ultimately show ?case by blast
qed
next
case(cBang α Q' Pa)
from `(Pa, !Q) ∈ bangRel Rel`
show ?case
proof(induct rule: BRBangCases)
case(BRBang P)
from `(P, Q) ∈ Rel` have "(!P, !Q) ∈ bangRel Rel" by(rule bangRel.BRBang)
with `(P, Q) ∈ Rel` have "(P \<parallel> !P, Q \<parallel> !Q) ∈ bangRel Rel" by(rule bangRel.BRPar)
then obtain P' where "P \<parallel> !P ==>⇧^α \<prec> P'" and "(P', Q') ∈ Rel'" using cBang
by blast
from `P \<parallel> !P ==>⇧^α \<prec> P'`
show ?case
proof(induct rule: weakTransCases)
case Base
have "!P ==>⇧^τ \<prec> !P" by simp
moreover from `(P', Q') ∈ Rel'` `P \<parallel> !P = P'` have "(!P, Q') ∈ Rel'" by(blast intro: C3)
ultimately show ?case by blast
next
case Step
from `P \<parallel> !P ==>α \<prec> P'` have "!P ==>α \<prec> P'" by(rule weakCongRepl)
hence "!P ==>⇧^α \<prec> P'" by(simp add: weakTrans_def)
with `(P', Q') ∈ Rel'` show ?case by blast
qed
qed
qed
}
moreover from `(P, Q) ∈ Rel` have "(!P, !Q) ∈ bangRel Rel" by(rule BRBang)
ultimately show ?case using `!Q \<longmapsto> α \<prec> Q'` by blast
qed
end