Theory Weak_Cong_Sim

Up to index of Isabelle/HOL/HOL-Nominal/CCS

theory Weak_Cong_Sim
imports Weak_Sim
(* 
Title: The Calculus of Communicating Systems
Author/Maintainer: Jesper Bengtson (jebe@itu.dk), 2012
*)

theory Weak_Cong_Sim
imports Weak_Cong_Semantics Weak_Sim Strong_Sim
begin

definition weakCongSimulation :: "ccs => (ccs × ccs) set => ccs => bool" ("_ \<leadsto><_> _" [80, 80, 80] 80)
where
"P \<leadsto><Rel> Q ≡ ∀a Q'. Q \<longmapsto>a \<prec> Q' --> (∃P'. P ==>a \<prec> P' ∧ (P', Q') ∈ Rel)"

lemma weakSimI[case_names Sim]:
fixes P :: ccs
and Rel :: "(ccs × ccs) set"
and Q :: ccs

assumes "!!α Q'. Q \<longmapsto>α \<prec> Q' ==> ∃P'. P ==>α \<prec> P' ∧ (P', Q') ∈ Rel"

shows "P \<leadsto><Rel> Q"
using assms
by(auto simp add: weakCongSimulation_def)

lemma weakSimE:
fixes P :: ccs
and Rel :: "(ccs × ccs) set"
and Q :: ccs
and α :: act
and Q' :: ccs

assumes "P \<leadsto><Rel> Q"
and "Q \<longmapsto>α \<prec> Q'"

obtains P' where "P ==>α \<prec> P'" and "(P', Q') ∈ Rel"
using assms
by(auto simp add: weakCongSimulation_def)

lemma simWeakSim:
fixes P :: ccs
and Rel :: "(ccs × ccs) set"
and Q :: ccs

assumes "P \<leadsto>[Rel] Q"

shows "P \<leadsto><Rel> Q"
using assms
by(rule_tac weakSimI, auto)
(blast dest: simE transitionWeakCongTransition)

lemma weakCongSimWeakSim:
fixes P :: ccs
and Rel :: "(ccs × ccs) set"
and Q :: ccs

assumes "P \<leadsto><Rel> Q"

shows "P \<leadsto>^<Rel> Q"
using assms
by(rule_tac Weak_Sim.weakSimI, auto)
(blast dest: weakSimE weakCongTransitionWeakTransition)

lemma test:
assumes "P ==>τ P'"

shows "P = P' ∨ (∃P''. P \<longmapsto>τ \<prec> P'' ∧ P'' ==>τ P')"
using assms
by(induct rule: tauChainInduct) auto

lemma tauChainCasesSym[consumes 1, case_names cTauNil cTauStep]:
assumes "P ==>τ P'"
and "Prop P"
and "!!P''. [|P \<longmapsto>τ \<prec> P''; P'' ==>τ P'|] ==> Prop P'"

shows "Prop P'"
using assms
by(blast dest: test)

lemma simE2:
fixes P :: ccs
and Rel :: "(ccs × ccs) set"
and Q :: ccs
and α :: act
and Q' :: ccs

assumes "P \<leadsto><Rel> Q"
and "Q ==>α \<prec> Q'"
and Sim: "!!R S. (R, S) ∈ Rel ==> R \<leadsto>^<Rel> S"

obtains P' where "P ==>α \<prec> P'" and "(P', Q') ∈ Rel"
proof -
assume Goal: "!!P'. [|P ==>α \<prec> P'; (P', Q') ∈ Rel|] ==> thesis"
from `Q ==>α \<prec> Q'` obtain Q''' Q''
where QChain: "Q ==>τ Q'''" and Q'''Trans: "Q''' \<longmapsto>α \<prec> Q''" and Q''Chain: "Q'' ==>τ Q'"
by(rule weakCongTransE)
from QChain Q'''Trans show ?thesis
proof(induct rule: tauChainCasesSym)
case cTauNil
from `P \<leadsto><Rel> Q` `Q \<longmapsto>α \<prec> Q''` obtain P''' where PTrans: "P ==>α \<prec> P'''" and "(P''', Q'') ∈ Rel"
by(blast dest: weakSimE)
moreover from Q''Chain `(P''', Q'') ∈ Rel` Sim obtain P' where P''Chain: "P''' ==>τ P'" and "(P', Q') ∈ Rel"
by(rule simTauChain)
with PTrans P''Chain show ?thesis
by(force intro: Goal simp add: weakCongTrans_def weakTrans_def)
next
case(cTauStep Q'''')
from `P \<leadsto><Rel> Q` `Q \<longmapsto>τ \<prec> Q''''` obtain P'''' where PChain: "P ==>τ \<prec> P''''" and "(P'''', Q'''') ∈ Rel"
by(drule_tac weakSimE) auto
from `Q'''' ==>τ Q'''` `(P'''', Q'''') ∈ Rel` Sim obtain P''' where P''''Chain: "P'''' ==>τ P'''" and "(P''', Q''') ∈ Rel"
by(rule simTauChain)
from `(P''', Q''') ∈ Rel` have "P''' \<leadsto>^<Rel> Q'''" by(rule Sim)
then obtain P'' where P'''Trans: "P''' ==>^α \<prec> P''" and "(P'', Q'') ∈ Rel" using Q'''Trans by(rule Weak_Sim.weakSimE)
from Q''Chain `(P'', Q'') ∈ Rel` Sim obtain P' where P''Chain: "P'' ==>τ P'" and "(P', Q') ∈ Rel"
by(rule simTauChain)
from PChain P''''Chain P'''Trans P''Chain
have "P ==>α \<prec> P'"
apply(auto simp add: weakCongTrans_def weakTrans_def)
apply(rule_tac x=P''aa in exI)
apply auto
defer
apply blast
by(auto simp add: tauChain_def)

with `(P', Q') ∈ Rel` show ?thesis
by(force intro: Goal simp add: weakCongTrans_def weakTrans_def)
qed
qed

lemma reflexive:
fixes P :: ccs
and Rel :: "(ccs × ccs) set"

assumes "Id ⊆ Rel"

shows "P \<leadsto><Rel> P"
using assms
by(auto simp add: weakCongSimulation_def intro: transitionWeakCongTransition)

lemma transitive:
fixes P :: ccs
and Rel :: "(ccs × ccs) set"
and Q :: ccs
and Rel' :: "(ccs × ccs) set"
and R :: ccs
and Rel'' :: "(ccs × ccs) set"

assumes "P \<leadsto><Rel> Q"
and "Q \<leadsto><Rel'> R"
and "Rel O Rel' ⊆ Rel''"
and "!!S T. (S, T) ∈ Rel ==> S \<leadsto>^<Rel> T"

shows "P \<leadsto><Rel''> R"
proof(induct rule: weakSimI)
case(Sim α R')
thus ?case using assms
apply(drule_tac Q=R in weakSimE, auto)
by(drule_tac Q=Q in simE2) auto
qed

lemma weakMonotonic:
fixes P :: ccs
and A :: "(ccs × ccs) set"
and Q :: ccs
and B :: "(ccs × ccs) set"

assumes "P \<leadsto><A> Q"
and "A ⊆ B"

shows "P \<leadsto><B> Q"
using assms
by(fastsimp simp add: weakCongSimulation_def)

end