Theory Strong_Sim_Pres

(* 
   Title: The Calculus of Communicating Systems   
   Author/Maintainer: Jesper Bengtson (jebe@itu.dk), 2012
*)
theory Strong_Sim_Pres
  imports Strong_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) ↝[Rel] α.(Q)"
using assms
by(fastforce simp add: simulation_def elim: actCases intro: Action)

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

  assumes "P ↝[Rel] Q"
  and     "Rel  Rel'"
  and     "Id  Rel'"

  shows "P  R ↝[Rel'] Q  R"
using assms
by(force simp add: simulation_def elim: sumCases intro: Sum1 Sum2)

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

  assumes "P ↝[Rel] Q"
  and     "(P, Q)  Rel"
  and     "R ↝[Rel'] T"
  and     "(R, T)  Rel'"
  and     C1: "P' Q' R' T'. (P', Q')  Rel; (R', T')  Rel'  (P'  R', Q'  T')  Rel''"

  shows "P  R ↝[Rel''] Q  T"
proof(induct rule: simI)
  case(Sim a QT)
  from Q  T a  QT
  show ?case
  proof(induct rule: parCases)
    case(cPar1 Q')
    from P ↝[Rel] Q Q a  Q' obtain P' where "P a  P'" and "(P', Q')  Rel" 
      by(rule simE)
    from P a  P' have "P  R a  P'  R" by(rule Par1)
    moreover from (P', Q')  Rel (R, T)  Rel' have "(P'  R, Q'  T)  Rel''" by(rule C1)
    ultimately show ?case by blast
  next
    case(cPar2 T')
    from R ↝[Rel'] T T a  T' obtain R' where "R a  R'" and "(R', T')  Rel'" 
      by(rule simE)
    from R a  R' have "P  R a  P  R'" by(rule Par2)
    moreover from (P, Q)  Rel (R', T')  Rel' have "(P  R', Q  T')  Rel''" by(rule C1)
    ultimately show ?case by blast
  next
    case(cComm Q' T' a)
    from P ↝[Rel] Q Q a  Q' obtain P' where "P a  P'" and "(P', Q')  Rel" 
      by(rule simE)
    from R ↝[Rel'] T T (coAction a)  T' obtain R' where "R (coAction a)  R'" and "(R', T')  Rel'" 
      by(rule simE)
    from P a  P' R (coAction a)  R' a  τ have "P  R τ  P'  R'" by(rule Comm)
    moreover from (P', Q')  Rel (R', T')  Rel' have "(P'  R', Q'  T')  Rel''" by(rule C1)
    ultimately show ?case by blast
  qed
qed

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

  assumes "P ↝[Rel] Q"
  and     "(P, Q)  Rel"
  and     C1: "S T U. (S, T)  Rel  (S  U, T  U)  Rel'"

  shows "P  R ↝[Rel'] Q  R"
using assms
by(rule_tac parPresAux[where Rel''=Rel' and Rel'=Id]) (auto intro: reflexive)

lemma resPres:
  fixes P   :: ccs
  and   Rel :: "(ccs × ccs) set"
  and   Q   :: ccs
  and   x   :: name

  assumes "P ↝[Rel] Q"
  and     "R S y. (R, S)  Rel  (⦇νyR, ⦇νyS)  Rel'"

  shows "⦇νxP ↝[Rel'] ⦇νxQ"
using assms
by(fastforce simp add: simulation_def elim: resCases intro: Res)

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 ↝[Rel] S"

  shows "!P ↝[bangRel Rel] !Q"
proof(induct rule: simI)
  case(Sim α Q')
  {
    fix Pa α Q'
    assume "!Q α  Q'" and "(Pa, !Q)  bangRel Rel"
    hence "P'. Pa α  P'  (P', Q')  bangRel Rel"
    proof(nominal_induct arbitrary: Pa rule: bangInduct)
      case(cPar1 α Q')
      from (Pa, Q  !Q)  bangRel Rel 
      show ?case
      proof(induct rule: BRParCases)
        case(BRPar P R)
        from (P, Q)  Rel have "P ↝[Rel] Q" by(rule C1)
        with Q α  Q' obtain P' where "P α  P'" and "(P', Q')  Rel"
          by(blast dest: simE)
        from P α  P' have "P  R α  P'  R" by(rule Par1)
        moreover from (P', Q')  Rel (R, !Q)  bangRel Rel have "(P'  R, Q'  !Q)  bangRel Rel"
          by(rule bangRel.BRPar)
        ultimately show ?case by blast
      qed
    next
      case(cPar2 α Q')
      from (Pa, Q  !Q)  bangRel Rel
      show ?case
      proof(induct rule: BRParCases)
        case(BRPar P R)
        from (R, !Q)  bangRel Rel obtain R' where "R α  R'" and "(R', Q')  bangRel Rel" using cPar2
          by blast
        from R α  R' have "P  R α  P  R'" by(rule Par2)
        moreover from (P, Q)  Rel (R', Q')  bangRel Rel have "(P  R', Q  Q')  bangRel Rel" by(rule bangRel.BRPar)
        ultimately show ?case by blast
      qed
    next
      case(cComm a Q' Q'' Pa)
      from (Pa, Q  !Q)  bangRel Rel
      show ?case
      proof(induct rule: BRParCases)
        case(BRPar P R)
        from (P, Q)  Rel have "P ↝[Rel] Q" by(rule C1)
        with Q a  Q' obtain P' where "P a  P'" and "(P', Q')  Rel"
          by(blast dest: simE)
        from (R, !Q)  bangRel Rel obtain R' where "R (coAction a)  R'" and "(R', Q'')  bangRel Rel" using cComm
          by blast
        from P a  P' R (coAction a)  R' a  τ have "P  R τ  P'  R'" by(rule Comm)
        moreover from (P', Q')  Rel (R', Q'')  bangRel Rel have "(P'  R', Q'  Q'')  bangRel Rel" by(rule bangRel.BRPar)
        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  !P, Q  !Q)  bangRel Rel" by(rule bangRel.BRPar)
        then obtain P' where "P  !P α  P'" and "(P', Q')  bangRel Rel" using cBang
          by blast
        from P  !P α  P' have "!P α  P'" by(rule Bang)
        thus ?case using (P', Q')  bangRel Rel by blast
      qed
    qed
  }

  moreover from (P, Q)  Rel have "(!P, !Q)  bangRel Rel" by(rule BRBang) 
  ultimately show ?case using !Q  α  Q' by blast
qed

end