Theory Sim_Struct_Cong
theory Sim_Struct_Cong
imports Simulation "HOL-Library.Multiset"
begin
text ‹This file is a (heavily modified) variant of the theory {\it Psi\_Calculi.Sim\_Struct\_Cong}
from~\cite{DBLP:journals/afp/Bengtson12}.›
lemma partitionListLeft:
assumes "xs@ys=xs'@y#ys'"
and "y ∈ set xs"
and "distinct(xs@ys)"
obtains zs where "xs = xs'@y#zs" and "ys'=zs@ys"
using assms
by(force simp add: append_eq_append_conv2 append_eq_Cons_conv)
lemma partitionListRight:
assumes "xs@ys=xs'@y#ys'"
and "y ∈ set ys"
and "distinct(xs@ys)"
obtains zs where "xs' = xs@zs" and "ys=zs@y#ys'"
using assms
by(force simp add: append_eq_append_conv2 append_eq_Cons_conv)
context env begin
lemma resOutputCases''''[consumes 8, case_names cOpen cRes]:
fixes Ψ :: 'b
and x :: name
and zvec :: "name list"
and P :: "('a, 'b, 'c) psi"
and α :: "'a action"
and P' :: "('a, 'b, 'c) psi"
and C :: "'f::fs_name"
assumes Trans: "Ψ ⊳ ⦇νx⦈P ⟼α ≺ P'"
and 1: "x ♯ Ψ"
and 2: "x ♯ α"
and 3: "x ♯ P'"
and 4: "bn α ♯* Ψ"
and 5: "bn α ♯* P"
and 6: "bn α ♯* subject α"
and "α = M⦇ν*zvec⦈⟨N⟩"
and rOpen: "⋀M xvec yvec y N P'. ⟦Ψ ⊳ ([(x, y)] ∙ P) ⟼M⦇ν*(xvec@yvec)⦈⟨N⟩ ≺ P'; y ∈ supp N;
x ♯ N; x ♯ P'; x ≠ y; y ♯ xvec; y ♯ yvec; y ♯ M; distinct xvec; distinct yvec;
xvec ♯* Ψ; y ♯ Ψ; yvec ♯* Ψ; xvec ♯* P; y ♯ P; yvec ♯* P; xvec ♯* M; y ♯ M;
yvec ♯* M; xvec ♯* yvec⟧ ⟹
Prop (M⦇ν*(xvec@y#yvec)⦈⟨N⟩) P'"
and rScope: "⋀P'. ⟦Ψ ⊳ P ⟼α ≺ P'⟧ ⟹ Prop α (⦇νx⦈P')"
shows "Prop α P'"
proof -
from Trans have "distinct (bn α)" by(auto dest: boundOutputDistinct)
show ?thesis using Trans 1 2 3 4 5 6 ‹α=M⦇ν*zvec⦈⟨N⟩› rOpen rScope
proof(induct rule: resCases'[where C="(zvec, C)"])
case cBrOpen
then show ?case
by(auto simp add: residualInject boundOutputApp)
next
case cRes
then show ?case
by(auto simp add: residualInject boundOutputApp)
next
case cBrClose
then show ?case
by(auto simp add: residualInject boundOutputApp)
next
case(cOpen M' xvec yvec y N' P')
show ?case
using ‹Ψ ⊳ [(x, y)] ∙ P ⟼ M'⦇ν*(xvec @ yvec)⦈⟨N'⟩ ≺ P'› ‹y ∈ supp N'› ‹x ♯ N'› ‹x ♯ P'›
‹x ≠ y› ‹y ♯ xvec› ‹y ♯ yvec› ‹y ♯ M'› ‹distinct xvec› ‹distinct yvec› ‹xvec ♯* Ψ›
‹y ♯ Ψ› ‹yvec ♯* Ψ› ‹xvec ♯* P› ‹y ♯ P› ‹yvec ♯* P› ‹xvec ♯* M'› ‹y ♯ M'›
‹yvec ♯* M'› ‹xvec ♯* yvec›
by(rule cOpen(22))
qed
qed
lemma resOutputCases'''''[consumes 7, case_names cOpen cRes]:
fixes Ψ :: 'b
and x :: name
and zvec :: "name list"
and P :: "('a, 'b, 'c) psi"
and P' :: "('a, 'b, 'c) psi"
and C :: "'f::fs_name"
assumes Trans: "Ψ ⊳ ⦇νx⦈P ⟼M⦇ν*zvec⦈⟨N⟩ ≺ P'"
and 1: "x ♯ Ψ"
and "x ♯ M⦇ν*zvec⦈⟨N⟩"
and 3: "x ♯ P'"
and "zvec ♯* Ψ"
and "zvec ♯* P"
and "zvec ♯* M"
and rOpen: "⋀M' xvec yvec y N' P'. ⟦Ψ ⊳ ([(x, y)] ∙ P) ⟼M'⦇ν*(xvec@yvec)⦈⟨N'⟩ ≺ P'; y ∈ supp N';
x ♯ N'; x ♯ P'; x ≠ y; y ♯ xvec; y ♯ yvec; y ♯ M'; distinct xvec; distinct yvec;
xvec ♯* Ψ; y ♯ Ψ; yvec ♯* Ψ; xvec ♯* P; y ♯ P; yvec ♯* P; xvec ♯* M'; y ♯ M';
yvec ♯* M'; xvec ♯* yvec; M'⦇ν*(xvec@y#yvec)⦈⟨N'⟩ = M⦇ν*zvec⦈⟨N⟩⟧ ⟹
Prop P'"
and rScope: "⋀P'. ⟦Ψ ⊳ P ⟼M⦇ν*zvec⦈⟨N⟩ ≺ P'⟧ ⟹ Prop (⦇νx⦈P')"
shows "Prop P'"
proof -
from Trans have "distinct zvec" by(auto dest: boundOutputDistinct)
obtain al where "al=M⦇ν*zvec⦈⟨N⟩" by simp
from ‹al = M⦇ν*zvec⦈⟨N⟩› Trans ‹zvec ♯* Ψ› ‹zvec ♯* P› ‹zvec ♯* M›
have αTrans: "Ψ ⊳ ⦇νx⦈P ⟼al ≺ P'" and 4: "bn al ♯* Ψ" and 5: "bn al ♯* P" and 6: "bn al ♯* subject al"
by simp+
from ‹x ♯ M⦇ν*zvec⦈⟨N⟩› ‹al=M⦇ν*zvec⦈⟨N⟩› have 2: "x ♯ al" by simp
show ?thesis using αTrans 1 2 3 4 5 6 ‹al=M⦇ν*zvec⦈⟨N⟩› rOpen rScope
proof(induct rule: resCases'[where C="(zvec, C)"])
case cBrOpen
then show ?case
by(auto simp add: residualInject boundOutputApp)
next
case cBrClose
then show ?case
by(auto simp add: residualInject boundOutputApp)
next
case(cOpen M' xvec yvec y N' P')
show ?case
using ‹Ψ ⊳ [(x, y)] ∙ P ⟼ M'⦇ν*(xvec @ yvec)⦈⟨N'⟩ ≺ P'› ‹y ∈ supp N'› ‹x ♯ N'› ‹x ♯ P'›
‹x ≠ y› ‹y ♯ xvec› ‹y ♯ yvec› ‹y ♯ M'› ‹distinct xvec› ‹distinct yvec› ‹xvec ♯* Ψ›
‹y ♯ Ψ› ‹yvec ♯* Ψ› ‹xvec ♯* P› ‹y ♯ P› ‹yvec ♯* P› ‹xvec ♯* M'› ‹y ♯ M'›
‹yvec ♯* M'› ‹xvec ♯* yvec› ‹M'⦇ν*(xvec @ y # yvec)⦈⟨N'⟩ = M⦇ν*zvec⦈⟨N⟩›
by(rule cOpen(22))
next
case (cRes P')
from ‹Ψ ⊳ P ⟼ al ≺ P'› ‹al = M⦇ν*zvec⦈⟨N⟩›
show ?case
by (simp add: cRes(4))
qed
qed
lemma resBrOutputCases'[consumes 7, case_names cBrOpen cRes]:
fixes Ψ :: 'b
and x :: name
and zvec :: "name list"
and P :: "('a, 'b, 'c) psi"
and P' :: "('a, 'b, 'c) psi"
and C :: "'f::fs_name"
assumes Trans: "Ψ ⊳ ⦇νx⦈P ⟼¡M⦇ν*zvec⦈⟨N⟩ ≺ P'"
and 1: "x ♯ Ψ"
and "x ♯ ¡M⦇ν*zvec⦈⟨N⟩"
and 3: "x ♯ P'"
and "zvec ♯* Ψ"
and "zvec ♯* P"
and "zvec ♯* M"
and rBrOpen: "⋀M' xvec yvec y N' P'. ⟦Ψ ⊳ ([(x, y)] ∙ P) ⟼¡M'⦇ν*(xvec@yvec)⦈⟨N'⟩ ≺ P'; y ∈ supp N';
x ♯ N'; x ♯ P'; x ≠ y; y ♯ xvec; y ♯ yvec; y ♯ M'; distinct xvec; distinct yvec;
xvec ♯* Ψ; y ♯ Ψ; yvec ♯* Ψ; xvec ♯* P; y ♯ P; yvec ♯* P; xvec ♯* M'; y ♯ M';
yvec ♯* M'; xvec ♯* yvec; ¡M'⦇ν*(xvec@y#yvec)⦈⟨N'⟩ = ¡M⦇ν*zvec⦈⟨N⟩⟧ ⟹
Prop P'"
and rScope: "⋀P'. ⟦Ψ ⊳ P ⟼¡M⦇ν*zvec⦈⟨N⟩ ≺ P'⟧ ⟹ Prop (⦇νx⦈P')"
shows "Prop P'"
proof -
from Trans have "distinct zvec" by(auto dest: boundOutputDistinct)
obtain al where "al=¡M⦇ν*zvec⦈⟨N⟩" by simp
from ‹al = ¡M⦇ν*zvec⦈⟨N⟩› Trans ‹zvec ♯* Ψ› ‹zvec ♯* P› ‹zvec ♯* M›
have αTrans: "Ψ ⊳ ⦇νx⦈P ⟼al ≺ P'" and 4: "bn al ♯* Ψ" and 5: "bn al ♯* P" and 6: "bn al ♯* subject al"
by simp+
from ‹x ♯ ¡M⦇ν*zvec⦈⟨N⟩› ‹al=¡M⦇ν*zvec⦈⟨N⟩› have 2: "x ♯ al" by simp
show ?thesis using αTrans 1 2 3 4 5 6 ‹al=¡M⦇ν*zvec⦈⟨N⟩› rBrOpen rScope
proof(induct rule: resCases'[where C="(zvec, C)"])
case cBrOpen
then show ?case
by(auto simp add: residualInject boundOutputApp)
next
case cBrClose
then show ?case
by(auto simp add: residualInject boundOutputApp)
next
case(cOpen M' xvec yvec y N' P')
then show ?case
by(auto simp add: residualInject boundOutputApp)
next
case (cRes P')
from ‹Ψ ⊳ P ⟼ al ≺ P'› ‹al = ¡M⦇ν*zvec⦈⟨N⟩›
show ?case
by (simp add: cRes(4))
qed
qed
lemma brOutputFreshSubject:
fixes x::name
assumes "Ψ ⊳ P ⟼ ¡M⦇ν*xvec⦈⟨N⟩ ≺ P'"
and "xvec ♯* M"
and "x ♯ P"
shows "x ♯ M"
using assms
proof(nominal_induct avoiding: x rule: brOutputInduct')
case(cAlpha Ψ P M xvec N P' p)
then show ?case by simp
next
case(cBrOutput Ψ M K N P)
then show ?case
by(auto simp add: fresh_def psi.supp dest: chanOutConSupp)
next
case(cCase Ψ P M xvec N P' φ Cs) then show ?case
by(induct Cs) auto
next
case(cPar1 Ψ Ψ⇩Q P M xvec N P' A⇩Q Q) then show ?case
by simp
next
case cPar2 then show ?case by simp
next
case cBrComm1 then show ?case by simp
next
case cBrComm2 then show ?case by simp
next
case cBrOpen then show ?case by(simp add: fresh_abs_fun_iff[OF pt_name_inst, OF at_name_inst, OF fin_supp])
next
case cScope then show ?case by(simp add: fresh_abs_fun_iff[OF pt_name_inst, OF at_name_inst, OF fin_supp])
next
case cBang then show ?case by simp
qed
lemma brInputFreshSubject:
fixes x::name
assumes "Ψ ⊳ P ⟼ ¿M⦇N⦈ ≺ P'"
and "x ♯ P"
shows "x ♯ M"
using assms
proof(nominal_induct avoiding: x rule: brInputInduct)
case(cBrInput Ψ K M xvec N Tvec P y)
then show ?case
by(auto simp add: fresh_def psi.supp dest: chanInConSupp)
next
case(cCase Ψ P M N P' φ Cs y) then show ?case
by(induct Cs) auto
next
case(cPar1 Ψ Ψ⇩Q P M N P' A⇩Q Q y) then show ?case
by simp
next
case cPar2 then show ?case by simp
next
case cBrMerge then show ?case by simp
next
case cScope then show ?case by(simp add: fresh_abs_fun_iff[OF pt_name_inst, OF at_name_inst, OF fin_supp])
next
case cBang then show ?case by simp
qed
lemma resComm:
fixes Ψ :: 'b
and x :: name
and y :: name
and Rel :: "('b × ('a, 'b, 'c) psi × ('a, 'b, 'c) psi) set"
and P :: "('a, 'b, 'c) psi"
assumes "x ♯ Ψ"
and "y ♯ Ψ"
and "eqvt Rel"
and R1: "⋀Ψ' Q. (Ψ', Q, Q) ∈ Rel"
and R2: "⋀Ψ' a b Q. ⟦a ♯ Ψ'; b ♯ Ψ'⟧ ⟹ (Ψ', ⦇νa⦈(⦇νb⦈Q), ⦇νb⦈(⦇νa⦈Q)) ∈ Rel"
and R3: "⋀Ψ' xvec yvec Q. ⟦xvec ♯* Ψ'; mset xvec = mset yvec⟧ ⟹ (Ψ', ⦇ν*xvec⦈Q, ⦇ν*yvec⦈Q) ∈ Rel"
shows "Ψ ⊳ ⦇νx⦈(⦇νy⦈P) ↝[Rel] ⦇νy⦈(⦇νx⦈P)"
proof(cases "x=y")
assume "x = y"
then show ?thesis using R1
by(force intro: reflexive)
next
assume "x ≠ y"
note ‹eqvt Rel›
moreover from ‹x ♯ Ψ› ‹y ♯ Ψ› have "[x, y] ♯* Ψ" by(simp add: fresh_star_def)
moreover have "[x, y] ♯* ⦇νx⦈(⦇νy⦈P)" by(simp add: abs_fresh)
moreover have "[x, y] ♯* ⦇νy⦈(⦇νx⦈P)" by(simp add: abs_fresh)
ultimately show ?thesis
proof(induct rule: simIChainFresh[where C="(x, y)"])
case(cSim α P')
from ‹bn α ♯* (x, y)› ‹bn α ♯* (⦇νx⦈(⦇νy⦈P))› have "x ♯ bn α" and "y ♯ bn α" and "bn α ♯* P" by simp+
from ‹[x, y] ♯* α› have "x ♯ α" and "y ♯ α" by simp+
from ‹[x, y] ♯* P'› have "x ♯ P'" and "y ♯ P'" by simp+
from ‹bn α ♯* P› ‹x ♯ α› have "bn α ♯* ⦇νx⦈P" by(simp add: abs_fresh)
with ‹Ψ ⊳ ⦇νy⦈(⦇νx⦈P) ⟼α ≺ P'› ‹y ♯ Ψ› ‹y ♯ α› ‹y ♯ P'› ‹bn α ♯* Ψ›
show ?case using ‹bn α ♯* subject α› ‹x ♯ α› ‹x ♯ P'› ‹bn α ♯* Ψ› ‹bn α ♯* P› ‹bn α ♯* subject α› ‹y ♯ α›
proof(induct rule: resCases'[where C=x])
case(cOpen M yvec1 yvec2 y' N P')
from ‹yvec1 ♯* yvec2› ‹distinct yvec1› ‹distinct yvec2› have "distinct(yvec1@yvec2)" by auto
from ‹x ♯ M⦇ν*(yvec1 @ y' # yvec2)⦈⟨N⟩› have "x ♯ M" and "x ♯ yvec1" and "x ≠ y'" and "x ♯ yvec2" and "x ♯ N"
by simp+
from ‹y ♯ M⦇ν*(yvec1 @ y' # yvec2)⦈⟨N⟩› have "y ♯ M" and "y ♯ yvec1" and "y ♯ yvec2"
by simp+
from ‹Ψ ⊳ ([(y, y')] ∙ ⦇νx⦈P) ⟼M⦇ν*(yvec1@yvec2)⦈⟨N⟩ ≺ P'› ‹x ≠ y› ‹x ≠ y'›
have "Ψ ⊳ ⦇νx⦈([(y, y')] ∙ P) ⟼M⦇ν*(yvec1@yvec2)⦈⟨N⟩ ≺ P'" by(simp add: eqvts)
moreover note ‹x ♯ Ψ›
moreover from ‹x ♯ N› ‹x ♯ yvec1› ‹x ♯ yvec2› ‹x ♯ M› have "x ♯ M⦇ν*(yvec1@yvec2)⦈⟨N⟩" by simp
moreover note ‹x ♯ P'›
moreover from ‹yvec1 ♯* Ψ› ‹yvec2 ♯* Ψ› have "(yvec1@yvec2) ♯* Ψ" by simp
moreover from ‹yvec1 ♯* ⦇νx⦈P› ‹yvec2 ♯* ⦇νx⦈P› ‹y ♯ yvec1› ‹y' ♯ yvec1› ‹y ♯ yvec2› ‹y' ♯ yvec2› ‹x ♯ yvec1› ‹x ♯ yvec2›
have "(yvec1@yvec2) ♯* ([(y, y')] ∙ P)" by simp
moreover from ‹yvec1 ♯* M› ‹yvec2 ♯* M› have "(yvec1 @ yvec2) ♯* M"
by simp
ultimately show ?case
proof(induct rule: resOutputCases''''')
case(cOpen M' xvec1 xvec2 x' N' P')
from ‹M'⦇ν*(xvec1 @ x' # xvec2)⦈⟨N'⟩ = M⦇ν*(yvec1 @ yvec2)⦈⟨N⟩› have "yvec1@yvec2 = xvec1@x'#xvec2" and "M = M'" and "N = N'" by (simp add: action.inject)+
from ‹x ♯ yvec1› ‹x ♯ yvec2› ‹y' ♯ yvec1› ‹y' ♯ yvec2› ‹y ♯ yvec1› ‹y ♯ yvec2›
have "x ♯ (yvec1@yvec2)" and "y ♯ (yvec1@yvec2)" and "y' ♯ (yvec1@yvec2)" by simp+
with ‹yvec1@yvec2 = xvec1@x'#xvec2›
have "x ♯ xvec1" and "x ≠ x'" and "x ♯ xvec2" and "y ♯ xvec1" and "y ≠ x'" and "y ♯ xvec2"
and "y' ♯ xvec1" and "x' ≠ y'" and "y' ♯ xvec2"
by auto
show ?case
proof(cases "x' ∈ set yvec1")
assume "x' ∈ set yvec1"
with ‹yvec1@yvec2 = xvec1@x'#xvec2› ‹distinct (yvec1@yvec2)›
obtain xvec2' where Eq1: "yvec1=xvec1@x'#xvec2'"
and Eq: "xvec2=xvec2'@yvec2"
by(metis partitionListLeft)
from ‹Ψ ⊳ ([(x, x')] ∙ [(y, y')] ∙ P) ⟼M'⦇ν*(xvec1@xvec2)⦈⟨N'⟩ ≺ P'› ‹y' ∈ supp N› ‹y' ♯ Ψ› ‹y' ♯ M› ‹y' ♯ xvec1› ‹y' ♯ xvec2› Eq ‹M=M'› ‹N = N'›
have "Ψ ⊳ ⦇νy'⦈([(x, x')] ∙ [(y, y')] ∙ P) ⟼M'⦇ν*((xvec1@xvec2')@y'#yvec2)⦈⟨N'⟩ ≺ P'"
by(intro Open) auto
then have "Ψ ⊳ ⦇νx'⦈(⦇νy'⦈([(x, x')] ∙ [(y, y')] ∙ P)) ⟼M⦇ν*(xvec1@x'#xvec2'@y'#yvec2)⦈⟨N⟩ ≺ P'"
using ‹x' ∈ supp N'› ‹x' ♯ Ψ› ‹x' ♯ M'› ‹x' ♯ xvec1› ‹x' ♯ xvec2› ‹x' ≠ y'› Eq ‹M=M'› ‹N=N'›
by(intro Open) auto
with ‹x' ≠ y'› ‹x ≠ y'› ‹x' ♯ [(y, y')] ∙ P›
have "Ψ ⊳ ⦇νx⦈(⦇νy'⦈([(y, y')] ∙ P)) ⟼M⦇ν*(xvec1@x'#xvec2'@y'#yvec2)⦈⟨N⟩ ≺ P'"
by(subst alphaRes[where y=x']) (simp add: calc_atm eqvts abs_fresh)+
with Eq1 ‹y' ♯ ⦇νx⦈P› ‹x ≠ y'› R1 show ?case
by(auto simp add: alphaRes abs_fresh)
next
assume "¬x' ∈ set yvec1"
then have "x' ♯ yvec1" by(simp add: fresh_def)
from ‹¬x' ∈ set yvec1› ‹yvec1@yvec2 = xvec1@x'#xvec2›
have "x' ∈ set yvec2"
by(auto simp add: append_eq_append_conv2 append_eq_Cons_conv)
with ‹yvec1@yvec2 = xvec1@x'#xvec2› ‹distinct (yvec1@yvec2)›
obtain xvec2' where Eq: "xvec1=yvec1@xvec2'"
and Eq1: "yvec2=xvec2'@x'#xvec2"
by(metis partitionListRight)
from ‹Ψ ⊳ ([(x, x')] ∙ [(y, y')] ∙ P) ⟼M'⦇ν*(xvec1@xvec2)⦈⟨N'⟩ ≺ P'› ‹y' ∈ supp N› ‹y' ♯ Ψ› ‹y' ♯ M› ‹y' ♯ xvec1› ‹y' ♯ xvec2› Eq ‹M=M'› ‹N = N'›
have "Ψ ⊳ ⦇νy'⦈([(x, x')] ∙ [(y, y')] ∙ P) ⟼M'⦇ν*(yvec1@y'#xvec2'@xvec2)⦈⟨N'⟩ ≺ P'"
by(intro Open) (assumption | simp)+
then have "Ψ ⊳ ⦇νx'⦈(⦇νy'⦈([(x, x')] ∙ [(y, y')] ∙ P)) ⟼M⦇ν*((yvec1@y'#xvec2')@x'#xvec2)⦈⟨N⟩ ≺ P'"
using ‹x' ∈ supp N'› ‹x' ♯ Ψ› ‹x' ♯ M'› ‹x' ♯ xvec1› ‹x' ♯ xvec2› ‹x' ≠ y'› Eq ‹M=M'› ‹N=N'›
by(intro Open) auto
with ‹x' ≠ y'› ‹x ≠ y'› ‹x' ♯ [(y, y')] ∙ P›
have "Ψ ⊳ ⦇νx⦈(⦇νy'⦈([(y, y')] ∙ P)) ⟼M⦇ν*((yvec1@y'#xvec2')@x'#xvec2)⦈⟨N⟩ ≺ P'"
by(subst alphaRes[where y=x']) (simp add: calc_atm eqvts abs_fresh)+
with Eq1 ‹y' ♯ ⦇νx⦈P› ‹x ≠ y'› R1 show ?case
by(auto simp add: alphaRes abs_fresh)
qed
next
case(cRes P')
from ‹Ψ ⊳ ([(y, y')] ∙ P) ⟼M⦇ν*(yvec1@yvec2)⦈⟨N⟩ ≺ P'› ‹y' ∈ supp N› ‹y' ♯ Ψ› ‹y' ♯ M› ‹y' ♯ yvec1› ‹y' ♯ yvec2›
have "Ψ ⊳ ⦇νy'⦈([(y, y')] ∙ P) ⟼M⦇ν*(yvec1@y'#yvec2)⦈⟨N⟩ ≺ P'" by(rule Open)
with ‹y' ♯ ⦇νx⦈P› ‹x ≠ y'›have "Ψ ⊳ ⦇νy⦈P ⟼M⦇ν*(yvec1@y'#yvec2)⦈⟨N⟩ ≺ P'" by(simp add: alphaRes abs_fresh)
then have "Ψ ⊳ ⦇νx⦈(⦇νy⦈P) ⟼M⦇ν*(yvec1@y'#yvec2)⦈⟨N⟩ ≺ ⦇νx⦈P'" using ‹x ♯ Ψ› ‹x ♯ M› ‹x ♯ yvec1› ‹x ≠ y'› ‹x ♯ yvec2› ‹x ♯ N›
by(intro Scope) auto
moreover have "(Ψ, ⦇νx⦈P', ⦇νx⦈P') ∈ Rel" by(rule R1)
ultimately show ?case by blast
qed
next
case(cBrOpen M yvec1 yvec2 y' N P')
from ‹yvec1 ♯* yvec2› ‹distinct yvec1› ‹distinct yvec2› have "distinct(yvec1@yvec2)" by auto
from ‹x ♯ ¡M⦇ν*(yvec1 @ y' # yvec2)⦈⟨N⟩› have "x ♯ M" and "x ♯ yvec1" and "x ≠ y'" and "x ♯ yvec2" and "x ♯ N"
by simp+
from ‹y ♯ ¡M⦇ν*(yvec1 @ y' # yvec2)⦈⟨N⟩› have "y ♯ M" and "y ♯ yvec1" and "y ♯ yvec2"
by simp+
from ‹Ψ ⊳ ([(y, y')] ∙ ⦇νx⦈P) ⟼¡M⦇ν*(yvec1@yvec2)⦈⟨N⟩ ≺ P'› ‹x ≠ y› ‹x ≠ y'›
have "Ψ ⊳ ⦇νx⦈([(y, y')] ∙ P) ⟼¡M⦇ν*(yvec1@yvec2)⦈⟨N⟩ ≺ P'" by(simp add: eqvts)
moreover note ‹x ♯ Ψ›
moreover from ‹x ♯ N› ‹x ♯ yvec1› ‹x ♯ yvec2› ‹x ♯ M› have "x ♯ ¡M⦇ν*(yvec1@yvec2)⦈⟨N⟩" by simp
moreover note ‹x ♯ P'›
moreover from ‹yvec1 ♯* Ψ› ‹yvec2 ♯* Ψ› have "(yvec1@yvec2) ♯* Ψ" by simp
moreover from ‹yvec1 ♯* ⦇νx⦈P› ‹yvec2 ♯* ⦇νx⦈P› ‹y ♯ yvec1› ‹y' ♯ yvec1› ‹y ♯ yvec2› ‹y' ♯ yvec2› ‹x ♯ yvec1› ‹x ♯ yvec2›
have "(yvec1@yvec2) ♯* ([(y, y')] ∙ P)" by simp
moreover from ‹yvec1 ♯* M› ‹yvec2 ♯* M› have "(yvec1 @ yvec2) ♯* M"
by simp
ultimately show ?case
proof(induct rule: resBrOutputCases')
case(cBrOpen M' xvec1 xvec2 x' N' P')
from ‹¡M'⦇ν*(xvec1 @ x' # xvec2)⦈⟨N'⟩ = ¡M⦇ν*(yvec1 @ yvec2)⦈⟨N⟩› have "yvec1@yvec2 = xvec1@x'#xvec2" and "M = M'" and "N = N'" by (simp add: action.inject)+
from ‹x ♯ yvec1› ‹x ♯ yvec2› ‹y' ♯ yvec1› ‹y' ♯ yvec2› ‹y ♯ yvec1› ‹y ♯ yvec2›
have "x ♯ (yvec1@yvec2)" and "y ♯ (yvec1@yvec2)" and "y' ♯ (yvec1@yvec2)" by simp+
with ‹yvec1@yvec2 = xvec1@x'#xvec2›
have "x ♯ xvec1" and "x ≠ x'" and "x ♯ xvec2" and "y ♯ xvec1" and "y ≠ x'" and "y ♯ xvec2"
and "y' ♯ xvec1" and "x' ≠ y'" and "y' ♯ xvec2"
by auto
show ?case
proof(cases "x' ∈ set yvec1")
assume "x' ∈ set yvec1"
with ‹yvec1@yvec2 = xvec1@x'#xvec2› ‹distinct (yvec1@yvec2)›
obtain xvec2' where Eq1: "yvec1=xvec1@x'#xvec2'"
and Eq: "xvec2=xvec2'@yvec2"
by(metis partitionListLeft)
from ‹Ψ ⊳ ([(x, x')] ∙ [(y, y')] ∙ P) ⟼¡M'⦇ν*(xvec1@xvec2)⦈⟨N'⟩ ≺ P'› ‹y' ∈ supp N› ‹y' ♯ Ψ› ‹y' ♯ M› ‹y' ♯ xvec1› ‹y' ♯ xvec2› Eq ‹M=M'› ‹N = N'›
have "Ψ ⊳ ⦇νy'⦈([(x, x')] ∙ [(y, y')] ∙ P) ⟼¡M'⦇ν*((xvec1@xvec2')@y'#yvec2)⦈⟨N'⟩ ≺ P'"
by(intro BrOpen) auto
then have "Ψ ⊳ ⦇νx'⦈(⦇νy'⦈([(x, x')] ∙ [(y, y')] ∙ P)) ⟼¡M⦇ν*(xvec1@x'#xvec2'@y'#yvec2)⦈⟨N⟩ ≺ P'"
using ‹x' ∈ supp N'› ‹x' ♯ Ψ› ‹x' ♯ M'› ‹x' ♯ xvec1› ‹x' ♯ xvec2› ‹x' ≠ y'› Eq ‹M=M'› ‹N=N'›
by(intro BrOpen) auto
with ‹x' ≠ y'› ‹x ≠ y'› ‹x' ♯ [(y, y')] ∙ P›
have "Ψ ⊳ ⦇νx⦈(⦇νy'⦈([(y, y')] ∙ P)) ⟼¡M⦇ν*(xvec1@x'#xvec2'@y'#yvec2)⦈⟨N⟩ ≺ P'"
by(subst alphaRes[where y=x']) (simp add: calc_atm eqvts abs_fresh)+
with Eq1 ‹y' ♯ ⦇νx⦈P› ‹x ≠ y'› R1 show ?case
by(auto simp add: alphaRes abs_fresh)
next
assume "¬x' ∈ set yvec1"
then have "x' ♯ yvec1" by(simp add: fresh_def)
from ‹¬x' ∈ set yvec1› ‹yvec1@yvec2 = xvec1@x'#xvec2›
have "x' ∈ set yvec2"
by(auto simp add: append_eq_append_conv2 append_eq_Cons_conv)
with ‹yvec1@yvec2 = xvec1@x'#xvec2› ‹distinct (yvec1@yvec2)›
obtain xvec2' where Eq: "xvec1=yvec1@xvec2'"
and Eq1: "yvec2=xvec2'@x'#xvec2"
by(metis partitionListRight)
from ‹Ψ ⊳ ([(x, x')] ∙ [(y, y')] ∙ P) ⟼¡M'⦇ν*(xvec1@xvec2)⦈⟨N'⟩ ≺ P'› ‹y' ∈ supp N› ‹y' ♯ Ψ› ‹y' ♯ M› ‹y' ♯ xvec1› ‹y' ♯ xvec2› Eq ‹M=M'› ‹N = N'›
have "Ψ ⊳ ⦇νy'⦈([(x, x')] ∙ [(y, y')] ∙ P) ⟼¡M'⦇ν*(yvec1@y'#xvec2'@xvec2)⦈⟨N'⟩ ≺ P'"
by(intro BrOpen) (assumption | simp)+
then have "Ψ ⊳ ⦇νx'⦈(⦇νy'⦈([(x, x')] ∙ [(y, y')] ∙ P)) ⟼¡M⦇ν*((yvec1@y'#xvec2')@x'#xvec2)⦈⟨N⟩ ≺ P'"
using ‹x' ∈ supp N'› ‹x' ♯ Ψ› ‹x' ♯ M'› ‹x' ♯ xvec1› ‹x' ♯ xvec2› ‹x' ≠ y'› Eq ‹M=M'› ‹N=N'›
by(intro BrOpen) auto
with ‹x' ≠ y'› ‹x ≠ y'› ‹x' ♯ [(y, y')] ∙ P›
have "Ψ ⊳ ⦇νx⦈(⦇νy'⦈([(y, y')] ∙ P)) ⟼¡M⦇ν*((yvec1@y'#xvec2')@x'#xvec2)⦈⟨N⟩ ≺ P'"
by(subst alphaRes[where y=x']) (simp add: calc_atm eqvts abs_fresh)+
with Eq1 ‹y' ♯ ⦇νx⦈P› ‹x ≠ y'› R1 show ?case
by(auto simp add: alphaRes abs_fresh)
qed
next
case(cRes P')
from ‹Ψ ⊳ ([(y, y')] ∙ P) ⟼¡M⦇ν*(yvec1@yvec2)⦈⟨N⟩ ≺ P'› ‹y' ∈ supp N› ‹y' ♯ Ψ› ‹y' ♯ M› ‹y' ♯ yvec1› ‹y' ♯ yvec2›
have "Ψ ⊳ ⦇νy'⦈([(y, y')] ∙ P) ⟼¡M⦇ν*(yvec1@y'#yvec2)⦈⟨N⟩ ≺ P'" by(rule BrOpen)
with ‹y' ♯ ⦇νx⦈P› ‹x ≠ y'›have "Ψ ⊳ ⦇νy⦈P ⟼¡M⦇ν*(yvec1@y'#yvec2)⦈⟨N⟩ ≺ P'" by(simp add: alphaRes abs_fresh)
then have "Ψ ⊳ ⦇νx⦈(⦇νy⦈P) ⟼¡M⦇ν*(yvec1@y'#yvec2)⦈⟨N⟩ ≺ ⦇νx⦈P'" using ‹x ♯ Ψ› ‹x ♯ M› ‹x ♯ yvec1› ‹x ≠ y'› ‹x ♯ yvec2› ‹x ♯ N›
by(intro Scope) auto
moreover have "(Ψ, ⦇νx⦈P', ⦇νx⦈P') ∈ Rel" by(rule R1)
ultimately show ?case by blast
qed
next
case(cRes P')
from ‹x ♯ ⦇νy⦈P'› ‹x ≠ y› have "x ♯ P'" by(simp add: abs_fresh)
with ‹Ψ ⊳ ⦇νx⦈P ⟼α ≺ P'› ‹x ♯ Ψ› ‹x ♯ α›
show ?case using ‹bn α ♯* Ψ› ‹bn α ♯* P› ‹bn α ♯* subject α› ‹y ♯ α›
proof(induct rule: resCases'[where C="(x, y)"])
case(cOpen M xvec1 xvec2 x' N P')
from ‹y ♯ M⦇ν*(xvec1@x'#xvec2)⦈⟨N⟩› have "y ≠ x'" and "y ♯ M⦇ν*(xvec1@xvec2)⦈⟨N⟩" by simp+
from ‹Ψ ⊳ ([(x, x')] ∙ P) ⟼M⦇ν*(xvec1@xvec2)⦈⟨N⟩ ≺ P'› ‹y ♯ Ψ› ‹y ♯ M⦇ν*(xvec1@xvec2)⦈⟨N⟩›
have "Ψ ⊳ ⦇νy⦈([(x, x')] ∙ P) ⟼M⦇ν*(xvec1@xvec2)⦈⟨N⟩ ≺ ⦇νy⦈P'"
by(rule Scope)
then have "Ψ ⊳ ⦇νx'⦈(⦇νy⦈([(x, x')] ∙ P)) ⟼M⦇ν*(xvec1@x'#xvec2)⦈⟨N⟩ ≺ ⦇νy⦈P'"
using ‹x' ∈ supp N› ‹x' ♯ Ψ› ‹x' ♯ M› ‹x' ♯ xvec1› ‹x' ♯ xvec2›
by(rule Open)
with ‹y ≠ x'› ‹x ≠ y› ‹x' ♯ P› have "Ψ ⊳ ⦇νx⦈(⦇νy⦈P) ⟼M⦇ν*(xvec1@x'#xvec2)⦈⟨N⟩ ≺ ⦇νy⦈P'"
by(subst alphaRes[where y=x']) (simp add: abs_fresh eqvts calc_atm)+
moreover have "(Ψ, ⦇νy⦈P', ⦇νy⦈P') ∈ Rel" by(rule R1)
ultimately show ?case by blast
next
case(cBrOpen M xvec1 xvec2 x' N P')
from ‹y ♯ ¡M⦇ν*(xvec1@x'#xvec2)⦈⟨N⟩› have "y ≠ x'" and "y ♯ ¡M⦇ν*(xvec1@xvec2)⦈⟨N⟩" by simp+
from ‹Ψ ⊳ ([(x, x')] ∙ P) ⟼¡M⦇ν*(xvec1@xvec2)⦈⟨N⟩ ≺ P'› ‹y ♯ Ψ› ‹y ♯ ¡M⦇ν*(xvec1@xvec2)⦈⟨N⟩›
have "Ψ ⊳ ⦇νy⦈([(x, x')] ∙ P) ⟼¡M⦇ν*(xvec1@xvec2)⦈⟨N⟩ ≺ ⦇νy⦈P'"
by(rule Scope)
then have "Ψ ⊳ ⦇νx'⦈(⦇νy⦈([(x, x')] ∙ P)) ⟼¡M⦇ν*(xvec1@x'#xvec2)⦈⟨N⟩ ≺ ⦇νy⦈P'"
using ‹x' ∈ supp N› ‹x' ♯ Ψ› ‹x' ♯ M› ‹x' ♯ xvec1› ‹x' ♯ xvec2›
by(rule BrOpen)
with ‹y ≠ x'› ‹x ≠ y› ‹x' ♯ P› have "Ψ ⊳ ⦇νx⦈(⦇νy⦈P) ⟼¡M⦇ν*(xvec1@x'#xvec2)⦈⟨N⟩ ≺ ⦇νy⦈P'"
by(subst alphaRes[where y=x']) (simp add: abs_fresh eqvts calc_atm)+
moreover have "(Ψ, ⦇νy⦈P', ⦇νy⦈P') ∈ Rel" by(rule R1)
ultimately show ?case by blast
next
case(cRes P')
from ‹Ψ ⊳ P ⟼α ≺ P'› ‹y ♯ Ψ› ‹y ♯ α›
have "Ψ ⊳ ⦇νy⦈P ⟼α ≺ ⦇νy⦈P'" by(rule Scope)
then have "Ψ ⊳ ⦇νx⦈(⦇νy⦈P) ⟼α ≺ ⦇νx⦈(⦇νy⦈P')" using ‹x ♯ Ψ› ‹x ♯ α›
by(rule Scope)
moreover from ‹x ♯ Ψ› ‹y ♯ Ψ› have "(Ψ, ⦇νx⦈(⦇νy⦈P'), ⦇νy⦈(⦇νx⦈P')) ∈ Rel"
by(rule R2)
ultimately show ?case by blast
next
case(cBrClose M xvec N P')
then show ?case
proof(cases "y ♯ ¡M⦇ν*xvec⦈⟨N⟩")
case True
with ‹Ψ ⊳ P ⟼¡M⦇ν*xvec⦈⟨N⟩ ≺ P'›
have "Ψ ⊳ ⦇νy⦈P ⟼¡M⦇ν*xvec⦈⟨N⟩ ≺ ⦇νy⦈P'" using ‹y ♯ Ψ›
by(intro Scope)
then have "Ψ ⊳ ⦇νx⦈(⦇νy⦈P) ⟼ τ ≺ (⦇νx⦈(⦇ν*xvec⦈⦇νy⦈P'))" using ‹x ∈ supp M› ‹x ♯ Ψ›
by(rule BrClose)
moreover have "(Ψ, (⦇νx⦈(⦇ν*xvec⦈⦇νy⦈P')), ⦇νy⦈(⦇νx⦈(⦇ν*xvec⦈P'))) ∈ Rel"
proof -
have "mset (x#xvec@[y]) = mset (y#x#xvec)"
by simp
moreover have "(x#xvec@[y]) ♯* Ψ" using ‹x ♯ Ψ› ‹y ♯ Ψ› ‹xvec ♯* Ψ›
by simp
ultimately have "(Ψ, ⦇ν*(x#xvec@[y])⦈P', ⦇ν*(y#x#xvec)⦈P') ∈ Rel"
by(metis R3)
then show ?thesis
by(auto simp add: resChainAppend)
qed
ultimately show ?thesis
by blast
next
case False
then have "y ∈ supp(¡M⦇ν*xvec⦈⟨N⟩)" unfolding fresh_def by simp
show ?thesis
proof(cases "y ∈ supp(M)")
case True
with ‹Ψ ⊳ P ⟼¡M⦇ν*xvec⦈⟨N⟩ ≺ P'›
have "Ψ ⊳ ⦇νy⦈P ⟼τ ≺ ⦇νy⦈(⦇ν*xvec⦈P')" using ‹y ♯ Ψ›
by(rule BrClose)
then have "Ψ ⊳ ⦇νx⦈(⦇νy⦈P) ⟼τ ≺ ⦇νx⦈(⦇νy⦈(⦇ν*xvec⦈P'))" using ‹x ♯ Ψ›
by(rule Scope) simp
moreover have "(Ψ, ⦇νx⦈(⦇νy⦈(⦇ν*xvec⦈P')), ⦇νy⦈(⦇νx⦈(⦇ν*xvec⦈P'))) ∈ Rel" using ‹x ♯ Ψ› ‹y ♯ Ψ›
by(metis R2)
ultimately show ?thesis
by blast
next
case False
then have "y ♯ M" by(simp add: fresh_def)
from ‹xvec ♯* (x, y)› have "y ♯ xvec" by simp
with False ‹y ∈ supp(¡M⦇ν*xvec⦈⟨N⟩)›
have "y ∈ supp N"
by(simp add: fresh_def action.supp)
from ‹Ψ ⊳ P ⟼¡M⦇ν*xvec⦈⟨N⟩ ≺ P'› have "Ψ ⊳ P ⟼¡M⦇ν*([]@xvec)⦈⟨N⟩ ≺ P'"
by simp
then have "Ψ ⊳ ⦇νy⦈P ⟼¡M⦇ν*([]@y#xvec)⦈⟨N⟩ ≺ P'" using ‹y ∈ supp N› ‹y ♯ Ψ› ‹y ♯ M› ‹y ♯ xvec›
by(intro BrOpen) (assumption|simp)+
then have "Ψ ⊳ ⦇νy⦈P ⟼¡M⦇ν*(y#xvec)⦈⟨N⟩ ≺ P'"
by simp
then have "Ψ ⊳ ⦇νx⦈(⦇νy⦈P) ⟼τ ≺ ⦇νx⦈(⦇νy⦈(⦇ν*xvec⦈P'))" using ‹x ∈ supp M› ‹x ♯ Ψ›
by(auto dest: BrClose)
moreover have "(Ψ, ⦇νx⦈(⦇νy⦈(⦇ν*xvec⦈P')), ⦇νy⦈(⦇νx⦈(⦇ν*xvec⦈P'))) ∈ Rel" using ‹x ♯ Ψ› ‹y ♯ Ψ›
by(rule R2)
ultimately show ?thesis by blast
qed
qed
qed
next
case(cBrClose M xvec N P')
from ‹xvec ♯* x› have "x ♯ xvec" by simp
have "x ♯ ⦇νx⦈P" by(simp add: fresh_abs_fun_iff[OF pt_name_inst, OF at_name_inst, OF fin_supp])
have "x ♯ P'"
by(rule broutputFreshDerivativeP) fact+
have "x ♯ N"
by(rule broutputFreshDerivativeN) fact+
moreover from ‹Ψ ⊳ ⦇νx⦈P ⟼ ¡M⦇ν*xvec⦈⟨N⟩ ≺ P'› ‹xvec ♯* M› ‹x ♯ ⦇νx⦈P› have "x ♯ M"
by(rule brOutputFreshSubject)
moreover note ‹x ♯ xvec›
ultimately have "x ♯ ¡M⦇ν*xvec⦈⟨N⟩"
by simp
have "bn (¡M⦇ν*xvec⦈⟨N⟩) ♯* Ψ" using ‹xvec ♯* Ψ› by simp
have "bn (¡M⦇ν*xvec⦈⟨N⟩) ♯* P" using ‹xvec ♯* ⦇νx⦈P› ‹x ♯ xvec› by(simp add: fresh_abs_fun_iff[OF pt_name_inst, OF at_name_inst, OF fin_supp])
have "bn (¡M⦇ν*xvec⦈⟨N⟩) ♯* subject (¡M⦇ν*xvec⦈⟨N⟩)" using ‹xvec ♯* M› by simp
have "y ∈ supp(subject (¡M⦇ν*xvec⦈⟨N⟩))" using ‹y ∈ supp M›
by(simp add: supp_some)
obtain M' xvec' N' where "¡M⦇ν*xvec⦈⟨N⟩ = ¡M'⦇ν*xvec'⦈⟨N'⟩"
by auto
from ‹Ψ ⊳ ⦇νx⦈P ⟼ ¡M⦇ν*xvec⦈⟨N⟩ ≺ P'› ‹x ♯ Ψ› ‹x ♯ ¡M⦇ν*xvec⦈⟨N⟩› ‹x ♯ P'› ‹bn (¡M⦇ν*xvec⦈⟨N⟩) ♯* Ψ› ‹bn (¡M⦇ν*xvec⦈⟨N⟩) ♯* P› ‹bn (¡M⦇ν*xvec⦈⟨N⟩) ♯* subject (¡M⦇ν*xvec⦈⟨N⟩)› ‹¡M⦇ν*xvec⦈⟨N⟩ = ¡M'⦇ν*xvec'⦈⟨N'⟩› ‹y ∈ supp(subject (¡M⦇ν*xvec⦈⟨N⟩))›
have "∃Q'. Ψ ⊳ ⦇νx⦈(⦇νy⦈P) ⟼ τ ≺ Q' ∧ (Ψ, Q', ⦇νy⦈(⦇ν*(bn (¡M⦇ν*xvec⦈⟨N⟩))⦈P')) ∈ Rel"
proof(induct rule: resCases'[where C=y])
case cOpen then show ?case by(simp add: residualInject)
next
case(cBrOpen M xvec yvec z N P')
from ‹y ∈ supp (subject (¡M⦇ν*(xvec @ z # yvec)⦈⟨N⟩))› have "y ∈ supp M"
by(simp add: supp_some)
then have "y ≠ z" using ‹z ♯ M› by(auto simp add: fresh_def)
from ‹Ψ ⊳ [(x, z)] ∙ P ⟼ ¡M⦇ν*(xvec @ yvec)⦈⟨N⟩ ≺ P'› ‹y ∈ supp M› ‹y ♯ Ψ›
have "Ψ ⊳ ⦇νy⦈([(x, z)] ∙ P) ⟼ τ ≺ ⦇νy⦈(⦇ν*(xvec@yvec)⦈P')"
by(rule BrClose)
then have "Ψ ⊳ ⦇νz⦈(⦇νy⦈([(x, z)] ∙ P)) ⟼ τ ≺ ⦇νz⦈(⦇νy⦈(⦇ν*(xvec@yvec)⦈P'))" using ‹z ♯ Ψ›
by(rule Scope) simp
then have "Ψ ⊳ ⦇νx⦈(⦇νy⦈P) ⟼ τ ≺ ⦇νz⦈(⦇νy⦈(⦇ν*(xvec@yvec)⦈P'))" using ‹z ♯ P› ‹x ≠ y› ‹y ≠ z›
apply(subst alphaRes[where x=x and y=z])
apply(simp add: fresh_abs_fun_iff[OF pt_name_inst, OF at_name_inst, OF fin_supp])
apply(simp add: eqvts swap_simps)
done
moreover have "(Ψ, ⦇νz⦈(⦇νy⦈(⦇ν*(xvec @ yvec)⦈P')), ⦇νy⦈(⦇ν*bn (¡M⦇ν*(xvec @ z # yvec)⦈⟨N⟩)⦈P')) ∈ Rel"
proof -
have "mset(z#y#xvec@yvec) = mset(y#xvec@z#yvec)"
by simp
moreover have "(z#y#xvec@yvec) ♯* Ψ" using ‹z ♯ Ψ› ‹y ♯ Ψ› ‹xvec ♯* Ψ› ‹yvec ♯* Ψ›
by simp
ultimately have "(Ψ, ⦇ν*(z#y#xvec@yvec)⦈P', ⦇ν*(y#xvec@z#yvec)⦈P') ∈ Rel"
by(metis R3)
then show ?thesis
by(simp add: resChainAppend)
qed
ultimately show ?case
by blast
next
case(cRes P')
from ‹Ψ ⊳ P ⟼ ¡M⦇ν*xvec⦈⟨N⟩ ≺ P'› ‹y ∈ supp M› ‹y ♯ Ψ›
have "Ψ ⊳ ⦇νy⦈P ⟼ τ ≺ ⦇νy⦈(⦇ν*xvec⦈P')"
by(rule BrClose)
then have "Ψ ⊳ ⦇νx⦈(⦇νy⦈P) ⟼ τ ≺ ⦇νx⦈(⦇νy⦈(⦇ν*xvec⦈P'))" using ‹x ♯ Ψ›
by(rule Scope) simp
moreover have "(Ψ, ⦇νx⦈(⦇νy⦈(⦇ν*xvec⦈P')), ⦇νy⦈(⦇ν*bn (¡M⦇ν*xvec⦈⟨N⟩)⦈⦇νx⦈P')) ∈ Rel"
proof -
have "mset(x#y#xvec) = mset(y#xvec@[x])"
by simp
moreover have "(x#y#xvec) ♯* Ψ" using ‹x ♯ Ψ› ‹y ♯ Ψ› ‹xvec ♯* Ψ›
by simp
ultimately have "(Ψ, ⦇ν*(x#y#xvec)⦈P', ⦇ν*(y#xvec@[x])⦈P') ∈ Rel"
by(metis R3)
then show ?thesis by(simp add: resChainAppend)
qed
ultimately show ?case
by blast
next
case cBrClose then show ?case by simp
qed
then show ?case by simp
qed
qed
qed
lemma parAssocLeft:
fixes Ψ :: 'b
and P :: "('a, 'b, 'c) psi"
and Q :: "('a, 'b, 'c) psi"
and R :: "('a, 'b, 'c) psi"
and Rel :: "('b × ('a, 'b, 'c) psi × ('a, 'b, 'c) psi) set"
assumes "eqvt Rel"
and C1: "⋀Ψ' S T U. (Ψ, (S ∥ T) ∥ U, S ∥ (T ∥ U)) ∈ Rel"
and C2: "⋀xvec Ψ' S T U. ⟦xvec ♯* Ψ'; xvec ♯* S⟧ ⟹ (Ψ', ⦇ν*xvec⦈((S ∥ T) ∥ U), S ∥ (⦇ν*xvec⦈(T ∥ U))) ∈ Rel"
and C3: "⋀xvec Ψ' S T U. ⟦xvec ♯* Ψ'; xvec ♯* U⟧ ⟹ (Ψ', (⦇ν*xvec⦈(S ∥ T)) ∥ U, ⦇ν*xvec⦈(S ∥ (T ∥ U))) ∈ Rel"
and C4: "⋀Ψ' S T xvec. ⟦(Ψ', S, T) ∈ Rel; xvec ♯* Ψ'⟧ ⟹ (Ψ', ⦇ν*xvec⦈S, ⦇ν*xvec⦈T) ∈ Rel"
shows "Ψ ⊳ (P ∥ Q) ∥ R ↝[Rel] P ∥ (Q ∥ R)"
using ‹eqvt Rel›
proof(induct rule: simI[of _ _ _ _ "()"])
case(cSim α PQR)
from ‹bn α ♯* (P ∥ Q ∥ R)› have "bn α ♯* P" and "bn α ♯* Q" and "bn α ♯* R" by simp+
then have "bn α ♯* (Q ∥ R)" by simp
with ‹Ψ ⊳ P ∥ (Q ∥ R) ⟼α ≺ PQR› ‹bn α ♯* Ψ› ‹bn α ♯* P›
show ?case using ‹bn α ♯* subject α›
proof(induct rule: parCases[where C = "(Ψ, P, Q, R, α)"])
case(cPar1 P' A⇩Q⇩R Ψ⇩Q⇩R)
from ‹A⇩Q⇩R ♯* (Ψ, P, Q, R, α)› have "A⇩Q⇩R ♯* Q" and "A⇩Q⇩R ♯* R"
by simp+
with ‹extractFrame(Q ∥ R) = ⟨A⇩Q⇩R, Ψ⇩Q⇩R⟩› ‹distinct A⇩Q⇩R›
obtain A⇩Q Ψ⇩Q A⇩R Ψ⇩R where "A⇩Q⇩R = A⇩Q@A⇩R" and "Ψ⇩Q⇩R = Ψ⇩Q ⊗ Ψ⇩R" and FrQ: "extractFrame Q = ⟨A⇩Q, Ψ⇩Q⟩" and FrR: "extractFrame R = ⟨A⇩R, Ψ⇩R⟩"
and "A⇩Q ♯* Ψ⇩R" and "A⇩R ♯* Ψ⇩Q"
by(auto intro: mergeFrameE dest: extractFrameFreshChain)
from ‹A⇩Q⇩R = A⇩Q@A⇩R› ‹A⇩Q⇩R ♯* Ψ› ‹A⇩Q⇩R ♯* P› ‹A⇩Q⇩R ♯* Q› ‹A⇩Q⇩R ♯* α›
have "A⇩Q ♯* Ψ" and "A⇩R ♯* Ψ" and "A⇩Q ♯* P" and "A⇩R ♯* P" and "A⇩Q ♯* Q" and "A⇩R ♯* Q" and "A⇩Q ♯* α" and "A⇩R ♯* α"
by simp+
from ‹Ψ ⊗ Ψ⇩Q⇩R ⊳ P ⟼α ≺ P'› ‹Ψ⇩Q⇩R = Ψ⇩Q ⊗ Ψ⇩R› have "(Ψ ⊗ Ψ⇩R) ⊗ Ψ⇩Q ⊳ P ⟼α ≺ P'"
by(metis statEqTransition Associativity Commutativity Composition)
then have "Ψ ⊗ Ψ⇩R ⊳ P ∥ Q ⟼α ≺ (P' ∥ Q)" using FrQ ‹bn α ♯* Q› ‹A⇩Q ♯* Ψ› ‹A⇩Q ♯* Ψ⇩R› ‹A⇩Q ♯* P› ‹A⇩Q ♯* α›
by(intro Par1) auto
then have "Ψ ⊳ (P ∥ Q) ∥ R ⟼α ≺ ((P' ∥ Q) ∥ R)" using FrR ‹bn α ♯* R› ‹A⇩R ♯* Ψ› ‹A⇩R ♯* P› ‹A⇩R ♯* Q› ‹A⇩R ♯* α›
by(auto intro: Par1)
moreover have "(Ψ, (P' ∥ Q) ∥ R, P' ∥ (Q ∥ R)) ∈ Rel" by(rule C1)
ultimately show ?case by blast
next
case(cPar2 QR A⇩P Ψ⇩P)
from ‹A⇩P ♯* (Ψ, P, Q, R, α)› have "A⇩P ♯* Q" and "A⇩P ♯* R" and "A⇩P ♯* α"
by simp+
have FrP: "extractFrame P = ⟨A⇩P, Ψ⇩P⟩" by fact
with ‹bn α ♯* P› ‹A⇩P ♯* α› have "bn α ♯* Ψ⇩P" by(auto dest: extractFrameFreshChain)
with ‹bn α ♯* Ψ› have "bn α ♯* (Ψ ⊗ Ψ⇩P)" by force
with ‹Ψ ⊗ Ψ⇩P ⊳ Q ∥ R ⟼α ≺ QR›
show ?case using ‹bn α ♯* Q› ‹bn α ♯* R› ‹bn α ♯* subject α› ‹A⇩P ♯* Q› ‹A⇩P ♯* R›
proof(induct rule: parCasesSubject[where C = "(A⇩P, Ψ⇩P, P, Q, R, Ψ)"])
case(cPar1 Q' A⇩R Ψ⇩R)
from ‹A⇩R ♯* (A⇩P, Ψ⇩P, P, Q, R, Ψ)› have "A⇩R ♯* A⇩P" and "A⇩R ♯* P" and "A⇩R ♯* Q" and "A⇩R ♯* Ψ⇩P" and "A⇩R ♯* Ψ"
by simp+
from ‹A⇩P ♯* R› ‹extractFrame R = ⟨A⇩R, Ψ⇩R⟩› ‹A⇩R ♯* A⇩P› have "A⇩P ♯* Ψ⇩R"
by(auto dest: extractFrameFreshChain)
from ‹(Ψ ⊗ Ψ⇩P) ⊗ Ψ⇩R ⊳ Q ⟼α ≺ Q'› have "(Ψ ⊗ Ψ⇩R) ⊗ Ψ⇩P ⊳ Q ⟼α ≺ Q'"
by(metis statEqTransition Associativity Commutativity Composition)
then have "Ψ ⊗ Ψ⇩R ⊳ P ∥ Q ⟼α ≺ (P ∥ Q')"
using FrP ‹bn α ♯* P› ‹A⇩P ♯* Ψ› ‹A⇩P ♯* Ψ⇩R› ‹A⇩P ♯* Q› ‹A⇩P ♯* α›
by(intro Par2) (assumption | force)+
then have "Ψ ⊳ (P ∥ Q) ∥ R ⟼α ≺ ((P ∥ Q') ∥ R)"
using ‹extractFrame R = ⟨A⇩R, Ψ⇩R⟩› ‹bn α ♯* R› ‹A⇩R ♯* Ψ› ‹A⇩R ♯* P› ‹A⇩R ♯* Q› ‹A⇩R ♯* α›
by(intro Par1) (assumption | simp)+
moreover have "(Ψ, (P ∥ Q') ∥ R, P ∥ (Q' ∥ R)) ∈ Rel" by(rule C1)
ultimately show ?case by blast
next
case(cPar2 R' A⇩Q Ψ⇩Q)
from ‹A⇩Q ♯* (A⇩P, Ψ⇩P, P, Q, R, Ψ)› have "A⇩Q ♯* A⇩P" and "A⇩Q ♯* R" and "A⇩Q ♯* Ψ⇩P" and "A⇩Q ♯* Ψ"
by simp+
have FrQ: "extractFrame Q = ⟨A⇩Q, Ψ⇩Q⟩" by fact
from ‹A⇩P ♯* Q› FrQ ‹A⇩Q ♯* A⇩P› have "A⇩P ♯* Ψ⇩Q"
by(auto dest: extractFrameFreshChain)
from ‹(Ψ ⊗ Ψ⇩P) ⊗ Ψ⇩Q ⊳ R ⟼α ≺ R'›
have "Ψ ⊗ (Ψ⇩P ⊗ Ψ⇩Q) ⊳ R ⟼α ≺ R'"
by(blast intro: statEqTransition Associativity)
moreover from FrP FrQ ‹A⇩Q ♯* A⇩P› ‹A⇩P ♯* Ψ⇩Q› ‹A⇩Q ♯* Ψ⇩P›
have "extractFrame(P ∥ Q) = ⟨(A⇩P@A⇩Q), Ψ⇩P ⊗ Ψ⇩Q⟩ " by simp
moreover from ‹bn α ♯* P› ‹bn α ♯* Q› have "bn α ♯* (P ∥ Q)" by simp
moreover from ‹A⇩P ♯* Ψ› ‹A⇩Q ♯* Ψ› have "(A⇩P@A⇩Q) ♯* Ψ" by simp
moreover from ‹A⇩P ♯* R› ‹A⇩Q ♯* R› have "(A⇩P@A⇩Q) ♯* R" by simp
moreover from ‹A⇩P ♯* α› ‹A⇩Q ♯* α› have "(A⇩P@A⇩Q) ♯* α" by simp
ultimately have "Ψ ⊳ (P ∥ Q) ∥ R ⟼α ≺ ((P ∥ Q) ∥ R')"
by(rule Par2)
moreover have "(Ψ, (P ∥ Q) ∥ R', P ∥ (Q ∥ R')) ∈ Rel" by(rule C1)
ultimately show ?case by blast
next
case(cComm1 Ψ⇩R M N Q' A⇩Q Ψ⇩Q K xvec R' A⇩R)
from ‹A⇩Q ♯* (A⇩P, Ψ⇩P, P, Q, R, Ψ)›
have "A⇩Q ♯* P" and "A⇩Q ♯* Q" and "A⇩Q ♯* R" and "A⇩Q ♯* A⇩P" and "A⇩Q ♯* Ψ⇩P" and "A⇩Q ♯* Ψ" by simp+
from ‹A⇩R ♯* (A⇩P, Ψ⇩P, P, Q, R, Ψ)› have "A⇩R ♯* P" and "A⇩R ♯* Q" and "A⇩R ♯* R" and "A⇩R ♯* A⇩P" and "A⇩R ♯* Ψ" by simp+
from ‹xvec ♯* (A⇩P, Ψ⇩P, P, Q, R, Ψ)› have "xvec ♯* A⇩P" and "xvec ♯* P" and "xvec ♯* Q" and "xvec ♯* Ψ" by simp+
have FrQ: "extractFrame Q = ⟨A⇩Q, Ψ⇩Q⟩" by fact
with ‹A⇩P ♯* Q› ‹A⇩Q ♯* A⇩P› have "A⇩P ♯* Ψ⇩Q"
by(auto dest: extractFrameFreshChain)
have FrR: "extractFrame R = ⟨A⇩R, Ψ⇩R⟩" by fact
with ‹A⇩P ♯* R› ‹A⇩R ♯* A⇩P› have "A⇩P ♯* Ψ⇩R"
by(auto dest: extractFrameFreshChain)
from ‹(Ψ ⊗ Ψ⇩P) ⊗ Ψ⇩Q ⊳ R ⟼K⦇ν*xvec⦈⟨N⟩ ≺ R'› ‹A⇩P ♯* R› ‹xvec ♯* A⇩P› ‹xvec ♯* K› ‹distinct xvec› have "A⇩P ♯* N"
by(auto intro: outputFreshChainDerivative)
from ‹(Ψ ⊗ Ψ⇩P) ⊗ Ψ⇩R ⊳ Q ⟼M⦇N⦈ ≺ Q'› have "(Ψ ⊗ Ψ⇩R) ⊗ Ψ⇩P ⊳ Q ⟼M⦇N⦈ ≺ Q'"
by(metis statEqTransition Associativity Commutativity Composition)
then have "Ψ ⊗ Ψ⇩R ⊳ P ∥ Q ⟼M⦇N⦈ ≺ (P ∥ Q')" using FrP ‹A⇩P ♯* Ψ› ‹A⇩P ♯* Ψ⇩R› ‹A⇩P ♯* Q› ‹A⇩P ♯* M› ‹A⇩P ♯* N›
by(intro Par2) auto
moreover from FrP FrQ ‹A⇩P ♯* Ψ⇩Q› ‹A⇩Q ♯* A⇩P› ‹A⇩Q ♯* Ψ⇩P› have "extractFrame(P ∥ Q) = ⟨(A⇩P@A⇩Q), Ψ⇩P ⊗ Ψ⇩Q⟩"
by simp
moreover from ‹(Ψ ⊗ Ψ⇩P) ⊗ Ψ⇩Q ⊳ R ⟼K⦇ν*xvec⦈⟨N⟩ ≺ R'› have "Ψ ⊗ Ψ⇩P ⊗ Ψ⇩Q ⊳ R ⟼K⦇ν*xvec⦈⟨N⟩ ≺ R'"
by(metis statEqTransition Associativity)
moreover note ‹extractFrame R = ⟨A⇩R, Ψ⇩R⟩›
moreover from ‹(Ψ ⊗ Ψ⇩P) ⊗ Ψ⇩Q ⊗ Ψ⇩R ⊢ M ↔ K› have "Ψ ⊗ (Ψ⇩P ⊗ Ψ⇩Q) ⊗ Ψ⇩R ⊢ M ↔ K"
by(metis statEqEnt Associativity Commutativity Composition)
ultimately have "Ψ ⊳ (P ∥ Q) ∥ R ⟼τ ≺ ⦇ν*xvec⦈((P ∥ Q') ∥ R')"
using ‹A⇩P ♯* Ψ› ‹A⇩Q ♯* Ψ› ‹A⇩R ♯* Ψ› ‹A⇩P ♯* P› ‹A⇩Q ♯* P› ‹A⇩R ♯* P› ‹A⇩P ♯* Q› ‹A⇩Q ♯* Q› ‹A⇩R ♯* Q› ‹A⇩P ♯* R› ‹A⇩Q ♯* R› ‹A⇩R ♯* R›
‹A⇩P ♯* M› ‹A⇩Q ♯* M› ‹A⇩R ♯* K› ‹A⇩R ♯* A⇩P› ‹A⇩Q ♯* A⇩R› ‹xvec ♯* P› ‹xvec ♯* Q›
by(intro Comm1) (assumption | simp)+
moreover from ‹xvec ♯* Ψ› ‹xvec ♯* P› have "(Ψ, ⦇ν*xvec⦈((P ∥ Q') ∥ R'), P ∥ (⦇ν*xvec⦈(Q' ∥ R'))) ∈ Rel"
by(rule C2)
ultimately show ?case by blast
next
case(cComm2 Ψ⇩R M xvec N Q' A⇩Q Ψ⇩Q K R' A⇩R)
from ‹A⇩Q ♯* (A⇩P, Ψ⇩P, P, Q, R, Ψ)›
have "A⇩Q ♯* P" and "A⇩Q ♯* Q" and "A⇩Q ♯* R" and "A⇩Q ♯* A⇩P" and "A⇩Q ♯* Ψ" and "A⇩Q ♯* Ψ⇩P" by simp+
from ‹A⇩R ♯* (A⇩P, Ψ⇩P, P, Q, R, Ψ)› have "A⇩R ♯* P" and "A⇩R ♯* Q" and "A⇩R ♯* R" and "A⇩R ♯* A⇩P"and "A⇩R ♯* Ψ" by simp+
from ‹xvec ♯* (A⇩P, Ψ⇩P, P, Q, R, Ψ)› have "xvec ♯* A⇩P" and "xvec ♯* P" and "xvec ♯* Q" and "xvec ♯* Ψ" by simp+
have FrQ: "extractFrame Q = ⟨A⇩Q, Ψ⇩Q⟩" by fact
with ‹A⇩P ♯* Q› ‹A⇩Q ♯* A⇩P› have "A⇩P ♯* Ψ⇩Q"
by(auto dest: extractFrameFreshChain)
have FrR: "extractFrame R = ⟨A⇩R, Ψ⇩R⟩" by fact
with ‹A⇩P ♯* R› ‹A⇩R ♯* A⇩P› have "A⇩P ♯* Ψ⇩R"
by(auto dest: extractFrameFreshChain)
from ‹(Ψ ⊗ Ψ⇩P) ⊗ Ψ⇩R ⊳ Q ⟼M⦇ν*xvec⦈⟨N⟩ ≺ Q'› ‹A⇩P ♯* Q› ‹xvec ♯* A⇩P› ‹xvec ♯* M› ‹distinct xvec› have "A⇩P ♯* N"
by(auto intro: outputFreshChainDerivative)
from ‹(Ψ ⊗ Ψ⇩P) ⊗ Ψ⇩R ⊳ Q ⟼M⦇ν*xvec⦈⟨N⟩ ≺ Q'› have "(Ψ ⊗ Ψ⇩R) ⊗ Ψ⇩P ⊳ Q ⟼M⦇ν*xvec⦈⟨N⟩ ≺ Q'"
by(metis statEqTransition Associativity Commutativity Composition)
then have "Ψ ⊗ Ψ⇩R ⊳ P ∥ Q ⟼M⦇ν*xvec⦈⟨N⟩ ≺ (P ∥ Q')" using FrP ‹A⇩P ♯* Ψ› ‹A⇩P ♯* Ψ⇩R› ‹A⇩P ♯* Q› ‹A⇩P ♯* M› ‹A⇩P ♯* N› ‹xvec ♯* P› ‹xvec ♯* A⇩P›
by(intro Par2) auto
moreover from FrP FrQ ‹A⇩P ♯* Ψ⇩Q› ‹A⇩Q ♯* A⇩P› ‹A⇩Q ♯* Ψ⇩P› have "extractFrame(P ∥ Q) = ⟨(A⇩P@A⇩Q), Ψ⇩P ⊗ Ψ⇩Q⟩"
by simp+
moreover from ‹(Ψ ⊗ Ψ⇩P) ⊗ Ψ⇩Q ⊳ R ⟼K⦇N⦈ ≺ R'› have "Ψ ⊗ Ψ⇩P ⊗ Ψ⇩Q ⊳ R ⟼K⦇N⦈ ≺ R'"
by(metis statEqTransition Associativity)
moreover note ‹extractFrame R = ⟨A⇩R, Ψ⇩R⟩›
moreover from ‹(Ψ ⊗ Ψ⇩P) ⊗ Ψ⇩Q ⊗ Ψ⇩R ⊢ M ↔ K› have "Ψ ⊗ (Ψ⇩P ⊗ Ψ⇩Q) ⊗ Ψ⇩R ⊢ M ↔ K"
by(metis statEqEnt Associativity Commutativity Composition)
ultimately have "Ψ ⊳ (P ∥ Q) ∥ R ⟼τ ≺ ⦇ν*xvec⦈((P ∥ Q') ∥ R')"
using ‹A⇩P ♯* Ψ› ‹A⇩Q ♯* Ψ› ‹A⇩R ♯* Ψ› ‹A⇩P ♯* P› ‹A⇩Q ♯* P› ‹A⇩R ♯* P› ‹A⇩P ♯* Q› ‹A⇩Q ♯* Q› ‹A⇩R ♯* Q› ‹A⇩P ♯* R› ‹A⇩Q ♯* R› ‹A⇩R ♯* R›
‹A⇩P ♯* M› ‹A⇩Q ♯* M› ‹A⇩R ♯* K› ‹A⇩R ♯* A⇩P› ‹A⇩Q ♯* A⇩R› ‹xvec ♯* R›
by(intro Comm2) (assumption | simp)+
moreover from ‹xvec ♯* Ψ› ‹xvec ♯* P› have "(Ψ, ⦇ν*xvec⦈((P ∥ Q') ∥ R'), P ∥ (⦇ν*xvec⦈(Q' ∥ R'))) ∈ Rel"
by(rule C2)
ultimately show ?case by blast
next
case(cBrMerge Ψ⇩R M N Q' A⇩Q Ψ⇩Q R' A⇩R)
from ‹A⇩P ♯* α› ‹α = ¿M⦇N⦈› have "A⇩P ♯* M" and "A⇩P ♯* N" by simp+
from ‹A⇩Q ♯* (A⇩P, Ψ⇩P, P, Q, R, Ψ)›
have "A⇩Q ♯* P" and "A⇩Q ♯* Q" and "A⇩Q ♯* R" and "A⇩Q ♯* A⇩P" and "A⇩Q ♯* Ψ⇩P" and "A⇩Q ♯* Ψ" by simp+
from ‹A⇩R ♯* (A⇩P, Ψ⇩P, P, Q, R, Ψ)› have "A⇩R ♯* P" and "A⇩R ♯* Q" and "A⇩R ♯* R" and "A⇩R ♯* A⇩P" and "A⇩R ♯* Ψ" by simp+
have FrQ: "extractFrame Q = ⟨A⇩Q, Ψ⇩Q⟩" by fact
with ‹A⇩P ♯* Q› ‹A⇩Q ♯* A⇩P› have "A⇩P ♯* Ψ⇩Q"
by(auto dest: extractFrameFreshChain)
have FrR: "extractFrame R = ⟨A⇩R, Ψ⇩R⟩" by fact
with ‹A⇩P ♯* R› ‹A⇩R ♯* A⇩P› have "A⇩P ♯* Ψ⇩R"
by(auto dest: extractFrameFreshChain)
from ‹(Ψ ⊗ Ψ⇩P) ⊗ Ψ⇩R ⊳ Q ⟼¿M⦇N⦈ ≺ Q'› have "(Ψ ⊗ Ψ⇩R) ⊗ Ψ⇩P ⊳ Q ⟼¿M⦇N⦈ ≺ Q'"
by(metis statEqTransition Associativity Commutativity Composition)
then have "Ψ ⊗ Ψ⇩R ⊳ P ∥ Q ⟼¿M⦇N⦈ ≺ (P ∥ Q')" using FrP ‹A⇩P ♯* Ψ› ‹A⇩P ♯* Ψ⇩R› ‹A⇩P ♯* Q› ‹A⇩P ♯* M› ‹A⇩P ♯* N›
by(intro Par2) auto
moreover from FrP FrQ ‹A⇩P ♯* Ψ⇩Q› ‹A⇩Q ♯* A⇩P› ‹A⇩Q ♯* Ψ⇩P› have "extractFrame(P ∥ Q) = ⟨(A⇩P@A⇩Q), Ψ⇩P ⊗ Ψ⇩Q⟩"
by simp
moreover from ‹(Ψ ⊗ Ψ⇩P) ⊗ Ψ⇩Q ⊳ R ⟼¿M⦇N⦈ ≺ R'› have "Ψ ⊗ Ψ⇩P ⊗ Ψ⇩Q ⊳ R ⟼¿M⦇N⦈ ≺ R'"
by(metis statEqTransition Associativity)
moreover note ‹extractFrame R = ⟨A⇩R, Ψ⇩R⟩›
ultimately have "Ψ ⊳ (P ∥ Q) ∥ R ⟼¿M⦇N⦈ ≺ (P ∥ Q') ∥ R'"
using ‹A⇩P ♯* Ψ› ‹A⇩Q ♯* Ψ› ‹A⇩R ♯* Ψ› ‹A⇩P ♯* P› ‹A⇩Q ♯* P› ‹A⇩R ♯* P› ‹A⇩P ♯* Q› ‹A⇩Q ♯* Q› ‹A⇩R ♯* Q› ‹A⇩P ♯* R› ‹A⇩Q ♯* R› ‹A⇩R ♯* R›
‹A⇩P ♯* M› ‹A⇩Q ♯* M› ‹A⇩R ♯* M› ‹A⇩R ♯* A⇩P› ‹A⇩Q ♯* A⇩R›
by(auto intro: BrMerge)
moreover have "(Ψ, (P ∥ Q') ∥ R', P ∥ (Q' ∥ R')) ∈ Rel"
by(rule C1)
ultimately show ?case by blast
next
case(cBrComm1 Ψ⇩R M N Q' A⇩Q Ψ⇩Q xvec R' A⇩R)
from ‹¡M⦇ν*xvec⦈⟨N⟩ = α› have "xvec = bn α"
by(auto simp add: action.inject)
from ‹¡M⦇ν*xvec⦈⟨N⟩ = α› ‹A⇩P ♯* α›
have "A⇩P ♯* xvec" and "A⇩P ♯* M" and "A⇩P ♯* N" by auto
from ‹xvec = bn α› ‹bn α ♯* P› ‹bn α ♯* Ψ› ‹bn α ♯* Ψ⇩P› have "xvec ♯* P" and "xvec ♯* Ψ" and "xvec ♯* Ψ⇩P"
by simp+
from ‹A⇩Q ♯* (A⇩P, Ψ⇩P, P, Q, R, Ψ)›
have "A⇩Q ♯* P" and "A⇩Q ♯* Q" and "A⇩Q ♯* R" and "A⇩Q ♯* A⇩P" and "A⇩Q ♯* Ψ⇩P" and "A⇩Q ♯* Ψ" by simp+
from ‹A⇩R ♯* (A⇩P, Ψ⇩P, P, Q, R, Ψ)› have "A⇩R ♯* P" and "A⇩R ♯* Q" and "A⇩R ♯* R" and "A⇩R ♯* A⇩P" and "A⇩R ♯* Ψ" by simp+
have FrQ: "extractFrame Q = ⟨A⇩Q, Ψ⇩Q⟩" by fact
with ‹A⇩P ♯* Q› ‹A⇩Q ♯* A⇩P› have "A⇩P ♯* Ψ⇩Q"
by(auto dest: extractFrameFreshChain)
have FrR: "extractFrame R = ⟨A⇩R, Ψ⇩R⟩" by fact
with ‹A⇩P ♯* R› ‹A⇩R ♯* A⇩P› have "A⇩P ♯* Ψ⇩R"
by(auto dest: extractFrameFreshChain)
from ‹(Ψ ⊗ Ψ⇩P) ⊗ Ψ⇩R ⊳ Q ⟼¿M⦇N⦈ ≺ Q'› have "(Ψ ⊗ Ψ⇩R) ⊗ Ψ⇩P ⊳ Q ⟼¿M⦇N⦈ ≺ Q'"
by(metis statEqTransition Associativity Commutativity Composition)
then have "Ψ ⊗ Ψ⇩R ⊳ P ∥ Q ⟼¿M⦇N⦈ ≺ (P ∥ Q')" using FrP ‹A⇩P ♯* Ψ› ‹A⇩P ♯* Ψ⇩R› ‹A⇩P ♯* Q› ‹A⇩P ♯* M› ‹A⇩P ♯* N›
by(intro Par2) auto
moreover from FrP FrQ ‹A⇩P ♯* Ψ⇩Q› ‹A⇩Q ♯* A⇩P› ‹A⇩Q ♯* Ψ⇩P› have "extractFrame(P ∥ Q) = ⟨(A⇩P@A⇩Q), Ψ⇩P ⊗ Ψ⇩Q⟩"
by simp
moreover from ‹(Ψ ⊗ Ψ⇩P) ⊗ Ψ⇩Q ⊳ R ⟼¡M⦇ν*xvec⦈⟨N⟩ ≺ R'› have "Ψ ⊗ Ψ⇩P ⊗ Ψ⇩Q ⊳ R ⟼¡M⦇ν*xvec⦈⟨N⟩ ≺ R'"
by(metis statEqTransition Associativity)
moreover note ‹extractFrame R = ⟨A⇩R, Ψ⇩R⟩›
ultimately have "Ψ ⊳ (P ∥ Q) ∥ R ⟼¡M⦇ν*xvec⦈⟨N⟩ ≺ ((P ∥ Q') ∥ R')"
using ‹A⇩P ♯* Ψ› ‹A⇩Q ♯* Ψ› ‹A⇩R ♯* Ψ› ‹A⇩P ♯* P› ‹A⇩Q ♯* P› ‹A⇩R ♯* P› ‹A⇩P ♯* Q› ‹A⇩Q ♯* Q› ‹A⇩R ♯* Q› ‹A⇩P ♯* R› ‹A⇩Q ♯* R› ‹A⇩R ♯* R›
‹A⇩P ♯* M› ‹A⇩Q ♯* M› ‹A⇩R ♯* M› ‹A⇩R ♯* A⇩P› ‹A⇩Q ♯* A⇩R› ‹xvec ♯* P› ‹xvec ♯* Q›
by(intro BrComm1) (assumption | simp)+
moreover have "(Ψ, ((P ∥ Q') ∥ R'), (P ∥ (Q' ∥ R'))) ∈ Rel"
by(rule C1)
ultimately show ?case by blast
next
case(cBrComm2 Ψ⇩R M xvec N Q' A⇩Q Ψ⇩Q R' A⇩R)
from ‹¡M⦇ν*xvec⦈⟨N⟩ = α› have "xvec = bn α"
by(auto simp add: action.inject)
from ‹¡M⦇ν*xvec⦈⟨N⟩ = α› ‹A⇩P ♯* α›
have "A⇩P ♯* xvec" and "A⇩P ♯* M" and "A⇩P ♯* N" by auto
from ‹xvec = bn α› ‹bn α ♯* P› ‹bn α ♯* Ψ› ‹bn α ♯* Ψ⇩P› have "xvec ♯* P" and "xvec ♯* Ψ" and "xvec ♯* Ψ⇩P"
by simp+
from ‹A⇩Q ♯* (A⇩P, Ψ⇩P, P, Q, R, Ψ)›
have "A⇩Q ♯* P" and "A⇩Q ♯* Q" and "A⇩Q ♯* R" and "A⇩Q ♯* A⇩P" and "A⇩Q ♯* Ψ" and "A⇩Q ♯* Ψ⇩P" by simp+
from ‹A⇩R ♯* (A⇩P, Ψ⇩P, P, Q, R, Ψ)› have "A⇩R ♯* P" and "A⇩R ♯* Q" and "A⇩R ♯* R" and "A⇩R ♯* A⇩P"and "A⇩R ♯* Ψ" by simp+
have FrQ: "extractFrame Q = ⟨A⇩Q, Ψ⇩Q⟩" by fact
with ‹A⇩P ♯* Q› ‹A⇩Q ♯* A⇩P› have "A⇩P ♯* Ψ⇩Q"
by(auto dest: extractFrameFreshChain)
have FrR: "extractFrame R = ⟨A⇩R, Ψ⇩R⟩" by fact
with ‹A⇩P ♯* R› ‹A⇩R ♯* A⇩P› have "A⇩P ♯* Ψ⇩R"
by(auto dest: extractFrameFreshChain)
from ‹(Ψ ⊗ Ψ⇩P) ⊗ Ψ⇩R ⊳ Q ⟼¡M⦇ν*xvec⦈⟨N⟩ ≺ Q'› have "(Ψ ⊗ Ψ⇩R) ⊗ Ψ⇩P ⊳ Q ⟼¡M⦇ν*xvec⦈⟨N⟩ ≺ Q'"
by(metis statEqTransition Associativity Commutativity Composition)
then have "Ψ ⊗ Ψ⇩R ⊳ P ∥ Q ⟼¡M⦇ν*xvec⦈⟨N⟩ ≺ (P ∥ Q')" using FrP ‹A⇩P ♯* Ψ› ‹A⇩P ♯* Ψ⇩R› ‹A⇩P ♯* Q› ‹A⇩P ♯* M› ‹A⇩P ♯* N› ‹xvec ♯* P› ‹A⇩P ♯* xvec›
by(intro Par2) auto
moreover from FrP FrQ ‹A⇩P ♯* Ψ⇩Q› ‹A⇩Q ♯* A⇩P› ‹A⇩Q ♯* Ψ⇩P› have "extractFrame(P ∥ Q) = ⟨(A⇩P@A⇩Q), Ψ⇩P ⊗ Ψ⇩Q⟩"
by simp+
moreover from ‹(Ψ ⊗ Ψ⇩P) ⊗ Ψ⇩Q ⊳ R ⟼¿M⦇N⦈ ≺ R'› have "Ψ ⊗ Ψ⇩P ⊗ Ψ⇩Q ⊳ R ⟼¿M⦇N⦈ ≺ R'"
by(metis statEqTransition Associativity)
moreover note ‹extractFrame R = ⟨A⇩R, Ψ⇩R⟩›
ultimately have "Ψ ⊳ (P ∥ Q) ∥ R ⟼¡M⦇ν*xvec⦈⟨N⟩ ≺ ((P ∥ Q') ∥ R')"
using ‹A⇩P ♯* Ψ› ‹A⇩Q ♯* Ψ› ‹A⇩R ♯* Ψ› ‹A⇩P ♯* P› ‹A⇩Q ♯* P› ‹A⇩R ♯* P› ‹A⇩P ♯* Q› ‹A⇩Q ♯* Q› ‹A⇩R ♯* Q› ‹A⇩P ♯* R› ‹A⇩Q ♯* R› ‹A⇩R ♯* R›
‹A⇩P ♯* M› ‹A⇩Q ♯* M› ‹A⇩R ♯* M› ‹A⇩R ♯* A⇩P› ‹A⇩Q ♯* A⇩R› ‹xvec ♯* R›
by(auto intro: BrComm2)
moreover have "(Ψ, ((P ∥ Q') ∥ R'), (P ∥ (Q' ∥ R'))) ∈ Rel"
by(rule C1)
ultimately show ?case by blast
qed
next
case(cComm1 Ψ⇩Q⇩R M N P' A⇩P Ψ⇩P K xvec QR' A⇩Q⇩R)
from ‹xvec ♯* (Ψ, P, Q, R, α)› have "xvec ♯* Q" and "xvec ♯* R" by simp+
from ‹A⇩Q⇩R ♯* (Ψ, P, Q, R, α)› have "A⇩Q⇩R ♯* Q" and "A⇩Q⇩R ♯* R" and "A⇩Q⇩R ♯* Ψ" by simp+
from ‹A⇩P ♯* (Q ∥ R)› have "A⇩P ♯* Q" and "A⇩P ♯* R" by simp+
have PTrans: "Ψ ⊗ Ψ⇩Q⇩R ⊳ P ⟼M⦇N⦈ ≺ P'" and FrP: "extractFrame P = ⟨A⇩P, Ψ⇩P⟩" and MeqK: "Ψ ⊗ Ψ⇩P ⊗ Ψ⇩Q⇩R ⊢ M ↔ K" by fact+
note ‹Ψ ⊗ Ψ⇩P ⊳ Q ∥ R ⟼K⦇ν*xvec⦈⟨N⟩ ≺ QR'›
moreover from ‹xvec ♯* Ψ› ‹xvec ♯* Ψ⇩P› have "xvec ♯* (Ψ ⊗ Ψ⇩P)" by force
moreover note ‹xvec ♯* Q›‹xvec ♯* R› ‹xvec ♯* K›
‹extractFrame(Q ∥ R) = ⟨A⇩Q⇩R, Ψ⇩Q⇩R⟩› ‹distinct A⇩Q⇩R›
moreover from ‹A⇩Q⇩R ♯* Ψ› ‹A⇩Q⇩R ♯* Ψ⇩P› have "A⇩Q⇩R ♯* (Ψ ⊗ Ψ⇩P)" by force
ultimately show ?case using ‹A⇩Q⇩R ♯* Q› ‹A⇩Q⇩R ♯* R› ‹A⇩Q⇩R ♯* K›
proof(induct rule: parCasesOutputFrame)
case(cPar1 Q' A⇩Q Ψ⇩Q A⇩R Ψ⇩R)
have Aeq: "A⇩Q⇩R = A⇩Q@A⇩R" and Ψeq: "Ψ⇩Q⇩R = Ψ⇩Q ⊗ Ψ⇩R" by fact+
from PTrans Ψeq have "(Ψ ⊗ Ψ⇩R) ⊗ Ψ⇩Q ⊳ P ⟼M⦇N⦈ ≺ P'"
by(metis statEqTransition Associativity Commutativity Composition)
moreover note FrP
moreover from ‹(Ψ ⊗ Ψ⇩P) ⊗ Ψ⇩R ⊳ Q ⟼K⦇ν*xvec⦈⟨N⟩ ≺ Q'› have "(Ψ ⊗ Ψ⇩R) ⊗ Ψ⇩P ⊳ Q ⟼K⦇ν*xvec⦈⟨N⟩ ≺ Q'"
by(metis statEqTransition Associativity Commutativity Composition)
moreover note ‹extractFrame Q = ⟨A⇩Q, Ψ⇩Q⟩›
moreover from MeqK Ψeq have "(Ψ ⊗ Ψ⇩R) ⊗ Ψ⇩P ⊗ Ψ⇩Q ⊢ M ↔ K"
by(metis statEqEnt Associativity Commutativity Composition)
moreover from ‹A⇩P ♯* R› ‹A⇩P ♯* A⇩Q⇩R› Aeq ‹extractFrame R = ⟨A⇩R, Ψ⇩R⟩› have "A⇩P ♯* A⇩Q" and "A⇩P ♯* Ψ⇩R"
by(auto dest: extractFrameFreshChain)
moreover from ‹A⇩Q⇩R ♯* P› ‹A⇩Q⇩R ♯* Ψ› Aeq have "A⇩Q ♯* P" and "A⇩Q ♯* Ψ" by simp+
ultimately have "Ψ ⊗ Ψ⇩R ⊳ P ∥ Q ⟼τ ≺ ⦇ν*xvec⦈(P' ∥ Q')"
using ‹A⇩P ♯* Ψ› ‹A⇩P ♯* P› ‹A⇩P ♯* Q› ‹A⇩P ♯* M› ‹A⇩P ♯* A⇩Q› ‹A⇩Q ♯* Ψ› ‹A⇩Q ♯* Ψ⇩R› ‹A⇩Q ♯* P› ‹A⇩Q ♯* Q› ‹A⇩Q ♯* K› ‹xvec ♯* P›
by(intro Comm1) (assumption | force)+
moreover from ‹A⇩Q⇩R ♯* Ψ› Aeq have "A⇩R ♯* Ψ" by simp
moreover from ‹A⇩Q⇩R ♯* P› Aeq have "A⇩R ♯* P" by simp
ultimately have "Ψ ⊳ (P ∥ Q) ∥ R ⟼τ ≺ (⦇ν*xvec⦈(P' ∥ Q')) ∥ R" using ‹extractFrame R = ⟨A⇩R, Ψ⇩R⟩› ‹A⇩R ♯* Q›
by(intro Par1) (assumption | simp)+
moreover from ‹xvec ♯* Ψ› ‹xvec ♯* R› have "(Ψ, (⦇ν*xvec⦈(P' ∥ Q')) ∥ R, ⦇ν*xvec⦈(P' ∥ (Q' ∥ R))) ∈ Rel"
by(rule C3)
ultimately show ?case by blast
next
case(cPar2 R' A⇩Q Ψ⇩Q A⇩R Ψ⇩R)
have Aeq: "A⇩Q⇩R = A⇩Q@A⇩R" and Ψeq: "Ψ⇩Q⇩R = Ψ⇩Q ⊗ Ψ⇩R" by fact+
from ‹A⇩Q⇩R ♯* Ψ› ‹A⇩Q⇩R ♯* P› ‹A⇩Q⇩R ♯* Ψ⇩P› ‹A⇩P ♯* A⇩Q⇩R› Aeq have "A⇩R ♯* Ψ" and "A⇩R ♯* Ψ⇩P" and "A⇩P ♯* A⇩R" and "A⇩P ♯* A⇩Q" and "A⇩R ♯* P" by simp+
from ‹A⇩Q⇩R ♯* Ψ› Aeq have "A⇩Q ♯* Ψ" by simp
from ‹A⇩Q⇩R ♯* P› ‹A⇩P ♯* A⇩Q⇩R› Aeq FrP have "A⇩Q ♯* Ψ⇩P" by(auto dest: extractFrameFreshChain)
from ‹A⇩P ♯* A⇩Q⇩R› ‹extractFrame R = ⟨A⇩R, Ψ⇩R⟩› ‹extractFrame Q = ⟨A⇩Q, Ψ⇩Q⟩› Aeq ‹A⇩P ♯* Q› ‹A⇩P ♯* R› have "A⇩P ♯* Ψ⇩Q" and "A⇩P ♯* Ψ⇩R" by(auto dest: extractFrameFreshChain)
have RTrans: "(Ψ ⊗ Ψ⇩P) ⊗ Ψ⇩Q ⊳ R ⟼K⦇ν*xvec⦈⟨N⟩ ≺ R'" and FrR: "extractFrame R = ⟨A⇩R, Ψ⇩R⟩" by fact+
then have "(Ψ ⊗ Ψ⇩P) ⊗ Ψ⇩Q ⊳ R ⟼ ROut K (⦇ν*xvec⦈N ≺' R')" by(simp add: residualInject)
then obtain K' where KeqK': "((Ψ ⊗ Ψ⇩P) ⊗ Ψ⇩Q) ⊗ Ψ⇩R ⊢ K ↔ K'" and "A⇩P ♯* K'" and "A⇩Q ♯* K'"
using ‹A⇩P ♯* R› ‹A⇩Q ♯* R› ‹A⇩R ♯* Ψ› ‹A⇩R ♯* Ψ⇩P› ‹A⇩R ♯* Ψ⇩Q› ‹A⇩Q ♯* A⇩R› ‹A⇩P ♯* A⇩R› ‹A⇩R ♯* R› ‹A⇩R ♯* R› ‹A⇩R ♯* K› ‹distinct A⇩R› ‹xvec ♯* K› ‹distinct xvec› FrR
using outputObtainPrefix[where B="A⇩P@A⇩Q"]
by (smt (verit, ccfv_threshold) freshChainAppend freshChainSym freshCompChain(1))
from PTrans Ψeq have "(Ψ ⊗ Ψ⇩R) ⊗ Ψ⇩Q ⊳ P ⟼M⦇N⦈ ≺ P'"
by(metis statEqTransition Associativity Commutativity Composition)
moreover from MeqK KeqK' Ψeq have MeqK': "((Ψ ⊗ Ψ⇩R) ⊗ Ψ⇩Q) ⊗ Ψ⇩P ⊢ M ↔ K'"
by(metis statEqEnt Associativity Commutativity Composition chanEqTrans)
ultimately have "(Ψ ⊗ Ψ⇩R) ⊗ Ψ⇩Q ⊳ P ⟼K'⦇N⦈ ≺ P'" using FrP ‹distinct A⇩P› ‹A⇩P ♯* Ψ› ‹A⇩P ♯* P› ‹A⇩P ♯* M› ‹A⇩P ♯* K'› ‹A⇩P ♯* Ψ⇩Q› ‹A⇩P ♯* Ψ⇩R›
by(auto intro: inputRenameSubject)
moreover from ‹A⇩Q⇩R ♯* P› ‹A⇩Q⇩R ♯* N› Aeq have "A⇩Q ♯* P" and "A⇩Q ♯* N" by simp+
ultimately have "Ψ ⊗ Ψ⇩R ⊳ P ∥ Q ⟼K'⦇N⦈ ≺ P' ∥ Q" using ‹extractFrame Q = ⟨A⇩Q, Ψ⇩Q⟩› ‹A⇩Q ♯* Ψ⇩R› ‹A⇩Q ♯* K'› ‹A⇩Q ♯* Ψ›
by(intro Par1) (assumption | force)+
moreover from FrP ‹extractFrame Q = ⟨A⇩Q, Ψ⇩Q⟩› ‹A⇩P ♯* A⇩Q› ‹A⇩P ♯* Ψ⇩Q› ‹A⇩Q ♯* Ψ⇩P›
have "extractFrame(P ∥ Q) = ⟨(A⇩P@A⇩Q), Ψ⇩P ⊗ Ψ⇩Q⟩" by simp+
moreover from RTrans have "Ψ ⊗ (Ψ⇩P ⊗ Ψ⇩Q) ⊳ R ⟼K⦇ν*xvec⦈⟨N⟩ ≺ R'" by(metis Associativity statEqTransition)
moreover note FrR
moreover from MeqK' KeqK' have "Ψ ⊗ (Ψ⇩P ⊗ Ψ⇩Q) ⊗ Ψ⇩R ⊢ K' ↔ K"
by(metis statEqEnt Associativity Commutativity Composition chanEqTrans chanEqSym)
ultimately have "Ψ ⊳ (P ∥ Q) ∥ R ⟼τ ≺ ⦇ν*xvec⦈((P' ∥ Q) ∥ R')"
using ‹A⇩P ♯* Ψ› ‹A⇩Q ♯* Ψ› ‹A⇩P ♯* P› ‹A⇩Q ♯* P› ‹A⇩P ♯* Q› ‹A⇩Q ♯* Q› ‹A⇩P ♯* R› ‹A⇩Q ♯* R› ‹A⇩P ♯* K'› ‹A⇩Q ♯* K'› ‹A⇩P ♯* A⇩R› ‹A⇩Q ♯* A⇩R›
‹A⇩R ♯* Ψ› ‹A⇩R ♯* P› ‹A⇩R ♯* Q› ‹A⇩R ♯* R› ‹A⇩R ♯* K› ‹xvec ♯* P› ‹xvec ♯* Q›
by(intro Comm1) (assumption | simp)+
moreover from ‹xvec ♯* Ψ› have "(Ψ, ⦇ν*xvec⦈((P' ∥ Q) ∥ R'), ⦇ν*xvec⦈(P' ∥ (Q ∥ R'))) ∈ Rel"
by(metis C1 C4)
ultimately show ?case by blast
qed
next
case(cComm2 Ψ⇩Q⇩R M xvec N P' A⇩P Ψ⇩P K QR' A⇩Q⇩R)
from ‹A⇩Q⇩R ♯* (Ψ, P, Q, R, α)› have "A⇩Q⇩R ♯* Q" and "A⇩Q⇩R ♯* R" and "A⇩Q⇩R ♯* Ψ" by simp+
from ‹A⇩P ♯* (Q ∥ R)› ‹xvec ♯* (Q ∥ R)› have "A⇩P ♯* Q" and "A⇩P ♯* R" and "xvec ♯* Q" and "xvec ♯* R" by simp+
have PTrans: "Ψ ⊗ Ψ⇩Q⇩R ⊳ P ⟼M⦇ν*xvec⦈⟨N⟩ ≺ P'" and FrP: "extractFrame P = ⟨A⇩P, Ψ⇩P⟩" and MeqK: "Ψ ⊗ Ψ⇩P ⊗ Ψ⇩Q⇩R ⊢ M ↔ K" by fact+
note ‹Ψ ⊗ Ψ⇩P ⊳ Q ∥ R ⟼K⦇N⦈ ≺ QR'› ‹extractFrame(Q ∥ R) = ⟨A⇩Q⇩R, Ψ⇩Q⇩R⟩› ‹distinct A⇩Q⇩R›
moreover from ‹A⇩Q⇩R ♯* Ψ› ‹A⇩Q⇩R ♯* Ψ⇩P› have "A⇩Q⇩R ♯* (Ψ ⊗ Ψ⇩P)" by force
ultimately show ?case using ‹A⇩Q⇩R ♯* Q› ‹A⇩Q⇩R ♯* R› ‹A⇩Q⇩R ♯* K›
proof(induct rule: parCasesInputFrame)
case(cPar1 Q' A⇩Q Ψ⇩Q A⇩R Ψ⇩R)
have Aeq: "A⇩Q⇩R = A⇩Q@A⇩R" and Ψeq: "Ψ⇩Q⇩R = Ψ⇩Q ⊗ Ψ⇩R" by fact+
from PTrans Ψeq have "(Ψ ⊗ Ψ⇩R) ⊗ Ψ⇩Q ⊳ P ⟼M⦇ν*xvec⦈⟨N⟩ ≺ P'"
by(metis statEqTransition Associativity Commutativity Composition)
moreover note FrP
moreover from ‹(Ψ ⊗ Ψ⇩P) ⊗ Ψ⇩R ⊳ Q ⟼K⦇N⦈ ≺ Q'› have "(Ψ ⊗ Ψ⇩R) ⊗ Ψ⇩P ⊳ Q ⟼K⦇N⦈ ≺ Q'"
by(metis statEqTransition Associativity Commutativity Composition)
moreover note ‹extractFrame Q = ⟨A⇩Q, Ψ⇩Q⟩›
moreover from MeqK Ψeq have "(Ψ ⊗ Ψ⇩R) ⊗ Ψ⇩P ⊗ Ψ⇩Q ⊢ M ↔ K"
by(metis statEqEnt Associativity Commutativity Composition)
moreover from ‹A⇩P ♯* Q› ‹A⇩P ♯* R› ‹A⇩P ♯* A⇩Q⇩R› Aeq ‹extractFrame Q = ⟨A⇩Q, Ψ⇩Q⟩› ‹extractFrame R = ⟨A⇩R, Ψ⇩R⟩›
have "A⇩P ♯* A⇩Q" and "A⇩P ♯* Ψ⇩R" by(auto dest: extractFrameFreshChain)
moreover from ‹A⇩Q⇩R ♯* Ψ› ‹A⇩Q⇩R ♯* P› Aeq have "A⇩Q ♯* Ψ" and "A⇩Q ♯* P" by simp+
ultimately have "Ψ ⊗ Ψ⇩R ⊳ P ∥ Q ⟼τ ≺ ⦇ν*xvec⦈(P' ∥ Q')"
using ‹A⇩P ♯* Ψ› ‹A⇩P ♯* P› ‹A⇩P ♯* Q› ‹A⇩P ♯* M› ‹A⇩P ♯* A⇩Q› ‹A⇩Q ♯* Ψ› ‹A⇩Q ♯* Ψ⇩R› ‹A⇩Q ♯* P› ‹A⇩Q ♯* Q› ‹A⇩Q ♯* K› ‹xvec ♯* Q›
by(intro Comm2) (assumption | force)+
moreover from ‹A⇩Q⇩R ♯* Ψ› ‹A⇩Q⇩R ♯* P› Aeq have "A⇩R ♯* Ψ" and "A⇩R ♯* P" by simp+
ultimately have "Ψ ⊳ (P ∥ Q) ∥ R ⟼τ ≺ (⦇ν*xvec⦈(P' ∥ Q')) ∥ R" using ‹extractFrame R = ⟨A⇩R, Ψ⇩R⟩› ‹A⇩R ♯* Q›
by(intro Par1) (assumption | simp)+
moreover from ‹xvec ♯* Ψ› ‹xvec ♯* R› have "(Ψ, (⦇ν*xvec⦈(P' ∥ Q')) ∥ R, ⦇ν*xvec⦈(P' ∥ (Q' ∥ R))) ∈ Rel"
by(rule C3)
ultimately show ?case by blast
next
case(cPar2 R' A⇩Q Ψ⇩Q A⇩R Ψ⇩R)
have Aeq: "A⇩Q⇩R = A⇩Q@A⇩R" and Ψeq: "Ψ⇩Q⇩R = Ψ⇩Q ⊗ Ψ⇩R" by fact+
from ‹A⇩Q⇩R ♯* Ψ› ‹A⇩Q⇩R ♯* P› ‹A⇩Q⇩R ♯* Ψ⇩P› ‹A⇩P ♯* A⇩Q⇩R› Aeq
have "A⇩R ♯* Ψ" and "A⇩R ♯* Ψ⇩P" and "A⇩P ♯* A⇩R" and "A⇩P ♯* A⇩Q" and "A⇩R ♯* P" and "A⇩Q ♯* Ψ" and "A⇩Q ♯* Ψ⇩P" by simp+
have RTrans: "(Ψ ⊗ Ψ⇩P) ⊗ Ψ⇩Q ⊳ R ⟼K⦇N⦈ ≺ R'" and FrR: "extractFrame R = ⟨A⇩R, Ψ⇩R⟩" by fact+
then obtain K' where KeqK': "((Ψ ⊗ Ψ⇩P) ⊗ Ψ⇩Q) ⊗ Ψ⇩R ⊢ K ↔ K'" and "A⇩P ♯* K'" and "A⇩Q ♯* K'"
using ‹A⇩P ♯* R› ‹A⇩Q ♯* R› ‹A⇩R ♯* Ψ› ‹A⇩R ♯* Ψ⇩P› ‹A⇩R ♯* Ψ⇩Q› ‹A⇩Q ♯* A⇩R› ‹A⇩P ♯* A⇩R› ‹A⇩R ♯* R› ‹A⇩R ♯* R› ‹A⇩R ♯* K› ‹distinct A⇩R›
using inputObtainPrefix[where B="A⇩P@A⇩Q"]
by (smt (verit, ccfv_threshold) freshChainAppend freshChainSym freshCompChain(1))
from PTrans Ψeq have "(Ψ ⊗ Ψ⇩R) ⊗ Ψ⇩Q ⊳ P ⟼M⦇ν*xvec⦈⟨N⟩ ≺ P'"
by(metis statEqTransition Associativity Commutativity Composition)
moreover from MeqK KeqK' Ψeq have MeqK': "((Ψ ⊗ Ψ⇩R) ⊗ Ψ⇩Q) ⊗ Ψ⇩P ⊢ M ↔ K'"
by(metis statEqEnt Associativity Commutativity Composition chanEqTrans)
moreover from ‹A⇩P ♯* R› ‹A⇩P ♯* Q› ‹A⇩P ♯* A⇩Q⇩R› FrR ‹extractFrame Q = ⟨A⇩Q, Ψ⇩Q⟩› Aeq have "A⇩P ♯* Ψ⇩Q" and "A⇩P ♯* Ψ⇩R"
by(auto dest: extractFrameFreshChain)
ultimately have "(Ψ ⊗ Ψ⇩R) ⊗ Ψ⇩Q ⊳ P ⟼K'⦇ν*xvec⦈⟨N⟩ ≺ P'" using FrP ‹distinct A⇩P› ‹A⇩P ♯* Ψ› ‹A⇩P ♯* P› ‹A⇩P ♯* M› ‹A⇩P ♯* K'›
by(auto intro: outputRenameSubject)
moreover from ‹A⇩Q⇩R ♯* P› ‹A⇩Q⇩R ♯* N› ‹A⇩Q⇩R ♯* xvec› Aeq have "A⇩Q ♯* P" and "A⇩Q ♯* N" and "A⇩Q ♯* xvec" by simp+
ultimately have "Ψ ⊗ Ψ⇩R ⊳ P ∥ Q ⟼K'⦇ν*xvec⦈⟨N⟩ ≺ (P' ∥ Q)" using ‹extractFrame Q = ⟨A⇩Q, Ψ⇩Q⟩› ‹A⇩Q ♯* Ψ⇩R› ‹A⇩Q ♯* K'› ‹xvec ♯* Q› ‹A⇩Q ♯* Ψ›
by(intro Par1) (assumption | force)+
moreover from FrP ‹extractFrame Q = ⟨A⇩Q, Ψ⇩Q⟩› ‹A⇩P ♯* A⇩Q› ‹A⇩P ♯* Ψ⇩Q› ‹A⇩Q ♯* Ψ⇩P›
have "extractFrame(P ∥ Q) = ⟨(A⇩P@A⇩Q), Ψ⇩P ⊗ Ψ⇩Q⟩" by simp+
moreover from RTrans have "Ψ ⊗ (Ψ⇩P ⊗ Ψ⇩Q) ⊳ R ⟼K⦇N⦈ ≺ R'" by(metis Associativity statEqTransition)
moreover note FrR
moreover from MeqK' KeqK' have "Ψ ⊗ (Ψ⇩P ⊗ Ψ⇩Q) ⊗ Ψ⇩R ⊢ K' ↔ K"
by(metis statEqEnt Associativity Commutativity Composition chanEqTrans chanEqSym)
ultimately have "Ψ ⊳ (P ∥ Q) ∥ R ⟼τ ≺ ⦇ν*xvec⦈((P' ∥ Q) ∥ R')"
using ‹A⇩P ♯* Ψ› ‹A⇩Q ♯* Ψ› ‹A⇩P ♯* P› ‹A⇩Q ♯* P› ‹A⇩P ♯* Q› ‹A⇩Q ♯* Q› ‹A⇩P ♯* R› ‹A⇩Q ♯* R› ‹A⇩P ♯* K'› ‹A⇩Q ♯* K'› ‹A⇩P ♯* A⇩R› ‹A⇩Q ♯* A⇩R›
‹A⇩R ♯* Ψ› ‹A⇩R ♯* P› ‹A⇩R ♯* Q› ‹A⇩R ♯* R› ‹A⇩R ♯* K› ‹xvec ♯* R›
by(intro Comm2) (assumption | simp)+
moreover from ‹xvec ♯* Ψ› have "(Ψ, ⦇ν*xvec⦈((P' ∥ Q) ∥ R'), ⦇ν*xvec⦈(P' ∥ (Q ∥ R'))) ∈ Rel"
by(metis C1 C4)
ultimately show ?case by blast
qed
next
case(cBrMerge Ψ⇩Q⇩R M N P' A⇩P Ψ⇩P QR' A⇩Q⇩R)
from ‹A⇩Q⇩R ♯* (Ψ, P, Q, R, α)› have "A⇩Q⇩R ♯* Q" and "A⇩Q⇩R ♯* R" and "A⇩Q⇩R ♯* Ψ" by simp+
from ‹A⇩P ♯* (Q ∥ R)› have "A⇩P ♯* Q" and "A⇩P ♯* R" by simp+
have PTrans: "Ψ ⊗ Ψ⇩Q⇩R ⊳ P ⟼¿M⦇N⦈ ≺ P'" and FrP: "extractFrame P = ⟨A⇩P, Ψ⇩P⟩" by fact+
note ‹Ψ ⊗ Ψ⇩P ⊳ Q ∥ R ⟼¿M⦇N⦈ ≺ QR'› ‹extractFrame(Q ∥ R) = ⟨A⇩Q⇩R, Ψ⇩Q⇩R⟩› ‹distinct A⇩Q⇩R›
moreover from ‹A⇩Q⇩R ♯* Ψ› ‹A⇩Q⇩R ♯* Ψ⇩P› have "A⇩Q⇩R ♯* (Ψ ⊗ Ψ⇩P)" by force
ultimately show ?case using ‹A⇩Q⇩R ♯* Q› ‹A⇩Q⇩R ♯* R› ‹A⇩Q⇩R ♯* M› ‹A⇩Q⇩R ♯* N› ‹A⇩Q⇩R ♯* P› ‹A⇩Q⇩R ♯* M› ‹A⇩Q⇩R ♯* Ψ›
proof(induct rule: parCasesBrInputFrame)
case(cPar1 Q' A⇩Q Ψ⇩Q A⇩R Ψ⇩R)
from ‹A⇩Q⇩R ♯* N› ‹A⇩Q⇩R = A⇩Q @ A⇩R› have "A⇩Q ♯* N" and "A⇩R ♯* N"
by simp+
have Aeq: "A⇩Q⇩R = A⇩Q@A⇩R" and Ψeq: "Ψ⇩Q⇩R = Ψ⇩Q ⊗ Ψ⇩R" by fact+
from PTrans Ψeq have "(Ψ ⊗ Ψ⇩R) ⊗ Ψ⇩Q ⊳ P ⟼¿M⦇N⦈ ≺ P'"
by(metis statEqTransition Associativity Commutativity Composition)
moreover note FrP
moreover from ‹(Ψ ⊗ Ψ⇩P) ⊗ Ψ⇩R ⊳ Q ⟼¿M⦇N⦈ ≺ Q'› have "(Ψ ⊗ Ψ⇩R) ⊗ Ψ⇩P ⊳ Q ⟼¿M⦇N⦈ ≺ Q'"
by(metis statEqTransition Associativity Commutativity Composition)
moreover note ‹extractFrame Q = ⟨A⇩Q, Ψ⇩Q⟩›
moreover from ‹A⇩P ♯* Q› ‹A⇩P ♯* R› ‹A⇩P ♯* A⇩Q⇩R› Aeq ‹extractFrame Q = ⟨A⇩Q, Ψ⇩Q⟩› ‹extractFrame R = ⟨A⇩R, Ψ⇩R⟩›
have "A⇩P ♯* A⇩Q" and "A⇩P ♯* Ψ⇩R" by(auto dest: extractFrameFreshChain)
moreover from ‹A⇩Q⇩R ♯* Ψ› ‹A⇩Q⇩R ♯* P› Aeq have "A⇩Q ♯* Ψ" and "A⇩Q ♯* P" by simp+
ultimately have "Ψ ⊗ Ψ⇩R ⊳ P ∥ Q ⟼¿M⦇N⦈ ≺ (P' ∥ Q')"
using ‹A⇩P ♯* Ψ› ‹A⇩P ♯* P› ‹A⇩P ♯* Q› ‹A⇩P ♯* M› ‹A⇩P ♯* A⇩Q› ‹A⇩Q ♯* Ψ› ‹A⇩Q ♯* Ψ⇩R› ‹A⇩Q ♯* P› ‹A⇩Q ♯* Q› ‹A⇩Q ♯* M›
by(intro BrMerge) (assumption | force)+
moreover from ‹A⇩Q⇩R ♯* Ψ› ‹A⇩Q⇩R ♯* P› Aeq have "A⇩R ♯* Ψ" and "A⇩R ♯* P" by simp+
ultimately have "Ψ ⊳ (P ∥ Q) ∥ R ⟼¿M⦇N⦈ ≺ (P' ∥ Q') ∥ R" using ‹extractFrame R = ⟨A⇩R, Ψ⇩R⟩› ‹A⇩R ♯* Q› ‹A⇩R ♯* M› ‹A⇩R ♯* N›
by(intro Par1) (assumption | simp)+
moreover have "(Ψ, (P' ∥ Q') ∥ R, (P' ∥ (Q' ∥ R))) ∈ Rel"
by(rule C1)
ultimately show ?case by blast
next
case(cPar2 R' A⇩Q Ψ⇩Q A⇩R Ψ⇩R)
from ‹A⇩Q⇩R ♯* N› ‹A⇩Q⇩R = A⇩Q @ A⇩R› have "A⇩Q ♯* N" and "A⇩R ♯* N"
by simp+
have Aeq: "A⇩Q⇩R = A⇩Q@A⇩R" and Ψeq: "Ψ⇩Q⇩R = Ψ⇩Q ⊗ Ψ⇩R" by fact+
from ‹A⇩Q⇩R ♯* Ψ› ‹A⇩Q⇩R ♯* P› ‹A⇩Q⇩R ♯* Ψ⇩P› ‹A⇩P ♯* A⇩Q⇩R› Aeq
have "A⇩R ♯* Ψ" and "A⇩R ♯* Ψ⇩P" and "A⇩P ♯* A⇩R" and "A⇩P ♯* A⇩Q" and "A⇩R ♯* P" and "A⇩Q ♯* Ψ" and "A⇩Q ♯* Ψ⇩P" by simp+
have RTrans: "(Ψ ⊗ Ψ⇩P) ⊗ Ψ⇩Q ⊳ R ⟼¿M⦇N⦈ ≺ R'" and FrR: "extractFrame R = ⟨A⇩R, Ψ⇩R⟩" by fact+
from PTrans Ψeq have "(Ψ ⊗ Ψ⇩R) ⊗ Ψ⇩Q ⊳ P ⟼¿M⦇N⦈ ≺ P'"
by(metis statEqTransition Associativity Commutativity Composition)
moreover from ‹A⇩P ♯* R› ‹A⇩P ♯* Q› ‹A⇩P ♯* A⇩Q⇩R› FrR ‹extractFrame Q = ⟨A⇩Q, Ψ⇩Q⟩› Aeq have "A⇩P ♯* Ψ⇩Q" and "A⇩P ♯* Ψ⇩R"
by(auto dest: extractFrameFreshChain)
moreover from ‹A⇩Q⇩R ♯* P› ‹A⇩Q⇩R ♯* N› Aeq have "A⇩Q ♯* P" and "A⇩Q ♯* N" by simp+
ultimately have "Ψ ⊗ Ψ⇩R ⊳ P ∥ Q ⟼¿M⦇N⦈ ≺ (P' ∥ Q)" using ‹extractFrame Q = ⟨A⇩Q, Ψ⇩Q⟩› ‹A⇩Q ♯* Ψ⇩R› ‹A⇩Q ♯* M› ‹A⇩Q ♯* Ψ›
by(intro Par1) (assumption | force)+
moreover from FrP ‹extractFrame Q = ⟨A⇩Q, Ψ⇩Q⟩› ‹A⇩P ♯* A⇩Q› ‹A⇩P ♯* Ψ⇩Q› ‹A⇩Q ♯* Ψ⇩P›
have "extractFrame(P ∥ Q) = ⟨(A⇩P@A⇩Q), Ψ⇩P ⊗ Ψ⇩Q⟩" by simp+
moreover from RTrans have "Ψ ⊗ (Ψ⇩P ⊗ Ψ⇩Q) ⊳ R ⟼¿M⦇N⦈ ≺ R'" by(metis Associativity statEqTransition)
moreover note FrR
ultimately have "Ψ ⊳ (P ∥ Q) ∥ R ⟼¿M⦇N⦈ ≺ ((P' ∥ Q) ∥ R')"
using ‹A⇩P ♯* Ψ› ‹A⇩Q ♯* Ψ› ‹A⇩P ♯* P› ‹A⇩Q ♯* P› ‹A⇩P ♯* Q› ‹A⇩Q ♯* Q› ‹A⇩P ♯* R› ‹A⇩Q ♯* R› ‹A⇩P ♯* M› ‹A⇩Q ♯* M› ‹A⇩P ♯* A⇩R› ‹A⇩Q ♯* A⇩R›
‹A⇩R ♯* Ψ› ‹A⇩R ♯* P› ‹A⇩R ♯* Q› ‹A⇩R ♯* R› ‹A⇩R ♯* M›
by(auto intro: BrMerge)
moreover have "(Ψ, ((P' ∥ Q) ∥ R'), (P' ∥ (Q ∥ R'))) ∈ Rel"
by(rule C1)
ultimately show ?case by blast
next
case(cBrMerge Ψ⇩R Q' A⇩Q Ψ⇩Q R' A⇩R)
have Aeq: "A⇩Q⇩R = A⇩Q@A⇩R" and Ψeq: "Ψ⇩Q⇩R = Ψ⇩Q ⊗ Ψ⇩R" by fact+
from ‹A⇩Q⇩R ♯* N› ‹A⇩Q⇩R ♯* M› ‹A⇩Q⇩R ♯* P› ‹A⇩Q⇩R ♯* Ψ› ‹A⇩Q⇩R ♯* Ψ⇩P› ‹A⇩P ♯* A⇩Q⇩R› Aeq Ψeq
have "A⇩Q ♯* N" and "A⇩R ♯* N" and "A⇩Q ♯* M" and "A⇩R ♯* M"
and "A⇩Q ♯* Ψ" and "A⇩R ♯* Ψ" and "A⇩P ♯* A⇩Q" and "A⇩P ♯* A⇩R"
and "A⇩Q ♯* Ψ⇩P" and "A⇩R ♯* Ψ⇩P"
and "A⇩Q ♯* P" and "A⇩R ♯* P"
by simp+
from ‹extractFrame Q = ⟨A⇩Q, Ψ⇩Q⟩› ‹A⇩P ♯* Q› ‹A⇩P ♯* A⇩Q›
have "A⇩P ♯* Ψ⇩Q"
by (metis extractFrameFreshChain freshFrameDest)
from Aeq Ψeq PTrans have "Ψ ⊗ (Ψ⇩Q ⊗ Ψ⇩R) ⊳ P ⟼ ¿M⦇N⦈ ≺ P'" by simp
then have "Ψ ⊗ (Ψ⇩R ⊗ Ψ⇩Q) ⊳ P ⟼ ¿M⦇N⦈ ≺ P'"
by (metis Commutativity Ψeq compositionSym statEqTransition)
then have "(Ψ ⊗ Ψ⇩R) ⊗ Ψ⇩Q ⊳ P ⟼ ¿M⦇N⦈ ≺ P'"
by (metis AssertionStatEqSym Associativity statEqTransition)
moreover note FrP
moreover from ‹(Ψ ⊗ Ψ⇩P) ⊗ Ψ⇩R ⊳ Q ⟼¿M⦇N⦈ ≺ Q'› have "(Ψ ⊗ Ψ⇩R) ⊗ Ψ⇩P ⊳ Q ⟼¿M⦇N⦈ ≺ Q'"
by (metis associativitySym statEqTransition)
moreover note ‹extractFrame Q = ⟨A⇩Q, Ψ⇩Q⟩›
moreover note ‹A⇩P ♯* Ψ› ‹A⇩Q ♯* Ψ› ‹A⇩Q ♯* Ψ⇩R›
moreover from ‹extractFrame R = ⟨A⇩R, Ψ⇩R⟩› ‹A⇩P ♯* A⇩R› ‹A⇩P ♯* R› have "A⇩P ♯* Ψ⇩R"
by (metis extractFrameFreshChain freshFrameDest)
moreover note ‹A⇩P ♯* P› ‹A⇩P ♯* Q› ‹A⇩P ♯* M› ‹A⇩P ♯* A⇩Q› ‹A⇩Q ♯* P› ‹A⇩Q ♯* Q› ‹A⇩Q ♯* M›
ultimately have "(Ψ ⊗ Ψ⇩R) ⊳ (P ∥ Q) ⟼¿M⦇N⦈ ≺ (P' ∥ Q')"
by(intro BrMerge) (assumption | force)+
from ‹(Ψ ⊗ Ψ⇩P) ⊗ Ψ⇩Q ⊳ R ⟼¿M⦇N⦈ ≺ R'› have "Ψ ⊗ (Ψ⇩P ⊗ Ψ⇩Q) ⊳ R ⟼¿M⦇N⦈ ≺ R'"
by (metis Associativity statEqTransition)
note ‹(Ψ ⊗ Ψ⇩R) ⊳ (P ∥ Q) ⟼¿M⦇N⦈ ≺ (P' ∥ Q')›
moreover from FrP ‹extractFrame Q = ⟨A⇩Q, Ψ⇩Q⟩› ‹A⇩P ♯* A⇩Q› ‹A⇩P ♯* Ψ⇩Q› ‹A⇩Q ♯* Ψ⇩P›
have "extractFrame(P ∥ Q) = ⟨(A⇩P@A⇩Q), (Ψ⇩P ⊗ Ψ⇩Q)⟩" by simp
moreover note ‹Ψ ⊗ (Ψ⇩P ⊗ Ψ⇩Q) ⊳ R ⟼¿M⦇N⦈ ≺ R'› ‹extractFrame R = ⟨A⇩R, Ψ⇩R⟩›
moreover note ‹A⇩P ♯* Ψ› ‹A⇩Q ♯* Ψ› ‹A⇩R ♯* Ψ› ‹A⇩P ♯* P› ‹A⇩Q ♯* P› ‹A⇩P ♯* Q› ‹A⇩Q ♯* Q›
‹A⇩P ♯* R› ‹A⇩Q ♯* R› ‹A⇩P ♯* M› ‹A⇩Q ♯* M› ‹A⇩P ♯* A⇩R› ‹A⇩Q ♯* A⇩R› ‹A⇩R ♯* P› ‹A⇩R ♯* Q› ‹A⇩R ♯* R› ‹A⇩R ♯* M›
ultimately have "Ψ ⊳ (P ∥ Q) ∥ R ⟼¿M⦇N⦈ ≺ (P' ∥ Q') ∥ R'"
by(auto intro: BrMerge)
moreover have "(Ψ, (P' ∥ Q') ∥ R', (P' ∥ (Q' ∥ R'))) ∈ Rel"
by(rule C1)
ultimately show ?case by blast
qed
next
case(cBrComm1 Ψ⇩Q⇩R M N P' A⇩P Ψ⇩P xvec QR' A⇩Q⇩R)
from ‹¡M⦇ν*xvec⦈⟨N⟩ = α› ‹bn α ♯* Q› ‹bn α ♯* R› have "xvec ♯* Q" and "xvec ♯* R" by auto
from ‹A⇩Q⇩R ♯* (Ψ, P, Q, R, α)› have "A⇩Q⇩R ♯* Q" and "A⇩Q⇩R ♯* R" and "A⇩Q⇩R ♯* Ψ" by simp+
from ‹A⇩P ♯* (Q ∥ R)› have "A⇩P ♯* Q" and "A⇩P ♯* R" by simp+
have PTrans: "Ψ ⊗ Ψ⇩Q⇩R ⊳ P ⟼¿M⦇N⦈ ≺ P'" and FrP: "extractFrame P = ⟨A⇩P, Ψ⇩P⟩" by fact+
note ‹Ψ ⊗ Ψ⇩P ⊳ Q ∥ R ⟼¡M⦇ν*xvec⦈⟨N⟩ ≺ QR'›
moreover from ‹xvec ♯* Ψ› ‹xvec ♯* Ψ⇩P› have "xvec ♯* (Ψ ⊗ Ψ⇩P)" by force
moreover note ‹xvec ♯* Q› ‹xvec ♯* R› ‹xvec ♯* M›
‹extractFrame(Q ∥ R) = ⟨A⇩Q⇩R, Ψ⇩Q⇩R⟩› ‹distinct A⇩Q⇩R›
moreover from ‹A⇩Q⇩R ♯* Ψ› ‹A⇩Q⇩R ♯* Ψ⇩P› have "A⇩Q⇩R ♯* (Ψ ⊗ Ψ⇩P)" by force
ultimately show ?case using ‹A⇩Q⇩R ♯* Q› ‹A⇩Q⇩R ♯* R› ‹A⇩Q⇩R ♯* M›
proof(induct rule: parCasesBrOutputFrame)
case(cPar1 Q' A⇩Q Ψ⇩Q A⇩R Ψ⇩R)
from ‹A⇩Q⇩R ♯* N› ‹A⇩Q⇩R = A⇩Q @ A⇩R› have "A⇩Q ♯* N" and "A⇩R ♯* N"
by simp+
from ‹A⇩Q⇩R ♯* xvec› ‹A⇩Q⇩R = (A⇩Q@A⇩R)› have "A⇩Q ♯* xvec" and "A⇩R ♯* xvec"
by simp+
have Aeq: "A⇩Q⇩R = A⇩Q@A⇩R" and Ψeq: "Ψ⇩Q⇩R = Ψ⇩Q ⊗ Ψ⇩R" by fact+
from PTrans Ψeq have "(Ψ ⊗ Ψ⇩R) ⊗ Ψ⇩Q ⊳ P ⟼¿M⦇N⦈ ≺ P'"
by(metis statEqTransition Associativity Commutativity Composition)
moreover note FrP
moreover from ‹(Ψ ⊗ Ψ⇩P) ⊗ Ψ⇩R ⊳ Q ⟼¡M⦇ν*xvec⦈⟨N⟩ ≺ Q'› have "(Ψ ⊗ Ψ⇩R) ⊗ Ψ⇩P ⊳ Q ⟼¡M⦇ν*xvec⦈⟨N⟩ ≺ Q'"
by(metis statEqTransition Associativity Commutativity Composition)
moreover note ‹extractFrame Q = ⟨A⇩Q, Ψ⇩Q⟩›
moreover from ‹A⇩P ♯* Q› ‹A⇩P ♯* R› ‹A⇩P ♯* A⇩Q⇩R› Aeq ‹extractFrame Q = ⟨A⇩Q, Ψ⇩Q⟩› ‹extractFrame R = ⟨A⇩R, Ψ⇩R⟩›
have "A⇩P ♯* A⇩Q" and "A⇩P ♯* Ψ⇩R" by(auto dest: extractFrameFreshChain)
moreover from ‹A⇩Q⇩R ♯* Ψ› ‹A⇩Q⇩R ♯* P› Aeq have "A⇩Q ♯* Ψ" and "A⇩Q ♯* P" by simp+
ultimately have "Ψ ⊗ Ψ⇩R ⊳ P ∥ Q ⟼¡M⦇ν*xvec⦈⟨N⟩ ≺ (P' ∥ Q')"
using ‹A⇩P ♯* Ψ› ‹A⇩P ♯* P› ‹A⇩P ♯* Q› ‹A⇩P ♯* M› ‹A⇩P ♯* A⇩Q› ‹A⇩Q ♯* Ψ› ‹A⇩Q ♯* Ψ⇩R› ‹A⇩Q ♯* P› ‹A⇩Q ♯* Q› ‹A⇩Q ♯* M› ‹xvec ♯* P›
by(intro BrComm1) (assumption | force)+
moreover from ‹A⇩Q⇩R ♯* Ψ› ‹A⇩Q⇩R ♯* P› Aeq have "A⇩R ♯* Ψ" and "A⇩R ♯* P" by simp+
ultimately have "Ψ ⊳ (P ∥ Q) ∥ R ⟼¡M⦇ν*xvec⦈⟨N⟩ ≺ (P' ∥ Q') ∥ R" using ‹extractFrame R = ⟨A⇩R, Ψ⇩R⟩› ‹A⇩R ♯* Q› ‹A⇩R ♯* M› ‹A⇩R ♯* N› ‹xvec ♯* R› ‹A⇩R ♯* xvec›
by(intro Par1) (assumption | simp)+
moreover have "(Ψ, (P' ∥ Q') ∥ R, (P' ∥ (Q' ∥ R))) ∈ Rel"
by(rule C1)
ultimately show ?case by blast
next
case(cPar2 R' A⇩Q Ψ⇩Q A⇩R Ψ⇩R)
from ‹A⇩Q⇩R ♯* N› ‹A⇩Q⇩R = A⇩Q @ A⇩R› have "A⇩Q ♯* N" and "A⇩R ♯* N"
by simp+
have Aeq: "A⇩Q⇩R = A⇩Q@A⇩R" and Ψeq: "Ψ⇩Q⇩R = Ψ⇩Q ⊗ Ψ⇩R" by fact+
from ‹A⇩Q⇩R ♯* Ψ› ‹A⇩Q⇩R ♯* P› ‹A⇩Q⇩R ♯* Ψ⇩P› ‹A⇩P ♯* A⇩Q⇩R› Aeq
have "A⇩R ♯* Ψ" and "A⇩R ♯* Ψ⇩P" and "A⇩P ♯* A⇩R" and "A⇩P ♯* A⇩Q" and "A⇩R ♯* P" and "A⇩Q ♯* Ψ" and "A⇩Q ♯* Ψ⇩P" by simp+
have RTrans: "(Ψ ⊗ Ψ⇩P) ⊗ Ψ⇩Q ⊳ R ⟼¡M⦇ν*xvec⦈⟨N⟩ ≺ R'" and FrR: "extractFrame R = ⟨A⇩R, Ψ⇩R⟩" by fact+
from PTrans Ψeq have "(Ψ ⊗ Ψ⇩R) ⊗ Ψ⇩Q ⊳ P ⟼¿M⦇N⦈ ≺ P'"
by(metis statEqTransition Associativity Commutativity Composition)
moreover from ‹A⇩P ♯* R› ‹A⇩P ♯* Q› ‹A⇩P ♯* A⇩Q⇩R› FrR ‹extractFrame Q = ⟨A⇩Q, Ψ⇩Q⟩› Aeq have "A⇩P ♯* Ψ⇩Q" and "A⇩P ♯* Ψ⇩R"
by(auto dest: extractFrameFreshChain)
moreover from ‹A⇩Q⇩R ♯* P› ‹A⇩Q⇩R ♯* N› Aeq have "A⇩Q ♯* P" and "A⇩Q ♯* N" by simp+
ultimately have "Ψ ⊗ Ψ⇩R ⊳ P ∥ Q ⟼¿M⦇N⦈ ≺ (P' ∥ Q)" using ‹extractFrame Q = ⟨A⇩Q, Ψ⇩Q⟩› ‹A⇩Q ♯* Ψ⇩R› ‹A⇩Q ♯* M› ‹A⇩Q ♯* Ψ›
by(intro Par1) (assumption | force)+
moreover from FrP ‹extractFrame Q = ⟨A⇩Q, Ψ⇩Q⟩› ‹A⇩P ♯* A⇩Q› ‹A⇩P ♯* Ψ⇩Q› ‹A⇩Q ♯* Ψ⇩P›
have "extractFrame(P ∥ Q) = ⟨(A⇩P@A⇩Q), Ψ⇩P ⊗ Ψ⇩Q⟩" by simp+
moreover from RTrans have "Ψ ⊗ (Ψ⇩P ⊗ Ψ⇩Q) ⊳ R ⟼¡M⦇ν*xvec⦈⟨N⟩ ≺ R'" by(metis Associativity statEqTransition)
moreover note FrR
ultimately have "Ψ ⊳ (P ∥ Q) ∥ R ⟼¡M⦇ν*xvec⦈⟨N⟩ ≺ ((P' ∥ Q) ∥ R')"
using ‹A⇩P ♯* Ψ› ‹A⇩Q ♯* Ψ› ‹A⇩P ♯* P› ‹A⇩Q ♯* P› ‹A⇩P ♯* Q› ‹A⇩Q ♯* Q› ‹A⇩P ♯* R› ‹A⇩Q ♯* R› ‹A⇩P ♯* M› ‹A⇩Q ♯* M› ‹A⇩P ♯* A⇩R› ‹A⇩Q ♯* A⇩R›
‹A⇩R ♯* Ψ› ‹A⇩R ♯* P› ‹A⇩R ♯* Q› ‹A⇩R ♯* R› ‹A⇩R ♯* M› ‹xvec ♯* P› ‹xvec ♯* Q›
by(intro BrComm1) (assumption | simp)+
moreover have "(Ψ, ((P' ∥ Q) ∥ R'), (P' ∥ (Q ∥ R'))) ∈ Rel"
by(rule C1)
ultimately show ?case by blast
next
case(cBrComm1 Ψ⇩R Q' A⇩Q Ψ⇩Q R' A⇩R)
from ‹A⇩Q⇩R ♯* N› ‹A⇩Q⇩R = A⇩Q @ A⇩R› have "A⇩Q ♯* N" and "A⇩R ♯* N"
by simp+
from ‹A⇩Q⇩R ♯* xvec› ‹A⇩Q⇩R = (A⇩Q@A⇩R)› have "A⇩Q ♯* xvec" and "A⇩R ♯* xvec"
by simp+
have Aeq: "A⇩Q⇩R = A⇩Q@A⇩R" and Ψeq: "Ψ⇩Q⇩R = Ψ⇩Q ⊗ Ψ⇩R" by fact+
from Aeq ‹A⇩Q⇩R ♯* P› have "A⇩Q ♯* P" and "A⇩R ♯* P" by simp+
from Aeq ‹A⇩Q⇩R ♯* Ψ› have "A⇩Q ♯* Ψ" and "A⇩R ♯* Ψ" by simp+
from Aeq ‹A⇩P ♯* A⇩Q⇩R› have "A⇩P ♯* A⇩Q" and "A⇩P ♯* A⇩R" by simp+
with ‹A⇩P ♯* Q› ‹extractFrame Q = ⟨A⇩Q, Ψ⇩Q⟩› have "A⇩P ♯* Ψ⇩Q"
by (metis extractFrameFreshChain freshChainSym freshFrameDest)
from PTrans Ψeq have "(Ψ ⊗ Ψ⇩R) ⊗ Ψ⇩Q ⊳ P ⟼¿M⦇N⦈ ≺ P'"
by(metis statEqTransition Associativity Commutativity Composition)
moreover note FrP
moreover from ‹(Ψ ⊗ Ψ⇩P) ⊗ Ψ⇩R ⊳ Q ⟼¿M⦇N⦈ ≺ Q'› have "(Ψ ⊗ Ψ⇩R) ⊗ Ψ⇩P ⊳ Q ⟼¿M⦇N⦈ ≺ Q'"
by(metis statEqTransition Associativity Commutativity Composition)
moreover note ‹extractFrame Q = ⟨A⇩Q, Ψ⇩Q⟩›
moreover from ‹A⇩P ♯* Q› ‹A⇩P ♯* R› ‹A⇩P ♯* A⇩Q⇩R› Aeq ‹extractFrame Q = ⟨A⇩Q, Ψ⇩Q⟩› ‹extractFrame R = ⟨A⇩R, Ψ⇩R⟩›
have "A⇩P ♯* A⇩Q" and "A⇩P ♯* Ψ⇩R" by(auto dest: extractFrameFreshChain)
moreover from ‹A⇩Q⇩R ♯* Ψ› ‹A⇩Q⇩R ♯* P› Aeq have "A⇩Q ♯* Ψ" and "A⇩Q ♯* P" by simp+
ultimately have PQTrans: "Ψ ⊗ Ψ⇩R ⊳ P ∥ Q ⟼¿M⦇N⦈ ≺ (P' ∥ Q')"
using ‹A⇩P ♯* Ψ› ‹A⇩P ♯* P› ‹A⇩P ♯* Q› ‹A⇩P ♯* M› ‹A⇩P ♯* A⇩Q› ‹A⇩Q ♯* Ψ› ‹A⇩Q ♯* Ψ⇩R› ‹A⇩Q ♯* P› ‹A⇩Q ♯* Q› ‹A⇩Q ♯* M›
by(intro BrMerge) (assumption | force)+
have RTrans: "(Ψ ⊗ Ψ⇩P) ⊗ Ψ⇩Q ⊳ R ⟼¡M⦇ν*xvec⦈⟨N⟩ ≺ R'" and FrR: "extractFrame R = ⟨A⇩R, Ψ⇩R⟩" by fact+
note PQTrans
moreover from FrP ‹extractFrame Q = ⟨A⇩Q, Ψ⇩Q⟩› ‹A⇩P ♯* A⇩Q› ‹A⇩P ♯* Ψ⇩Q› ‹A⇩Q⇩R ♯* Ψ⇩P› Aeq
have "extractFrame(P ∥ Q) = ⟨(A⇩P@A⇩Q), Ψ⇩P ⊗ Ψ⇩Q⟩" by simp+
moreover from RTrans have "Ψ ⊗ (Ψ⇩P ⊗ Ψ⇩Q) ⊳ R ⟼¡M⦇ν*xvec⦈⟨N⟩ ≺ R'"
by (metis Associativity statEqTransition)
moreover note FrR
moreover note ‹A⇩P ♯* Ψ› ‹A⇩Q ♯* Ψ› ‹A⇩P ♯* P› ‹A⇩P ♯* Q› ‹A⇩Q ♯* P› ‹A⇩Q ♯* Q› ‹A⇩P ♯* R› ‹A⇩Q ♯* R›
‹A⇩P ♯* M› ‹A⇩Q ♯* M› ‹A⇩P ♯* A⇩R› ‹A⇩Q ♯* A⇩R› ‹A⇩R ♯* Ψ› ‹A⇩R ♯* P› ‹A⇩R ♯* Q› ‹A⇩R ♯* R› ‹A⇩R ♯* M› ‹xvec ♯* P› ‹xvec ♯* Q›
ultimately have "Ψ ⊳ (P ∥ Q) ∥ R ⟼¡M⦇ν*xvec⦈⟨N⟩ ≺ (P' ∥ Q') ∥ R'"
by(intro BrComm1) (assumption | force)+
moreover have "(Ψ, (P' ∥ Q') ∥ R', (P' ∥ (Q' ∥ R'))) ∈ Rel"
by(rule C1)
ultimately show ?case by blast
next
case(cBrComm2 Ψ⇩R Q' A⇩Q Ψ⇩Q R' A⇩R)
from ‹A⇩Q⇩R ♯* N› ‹A⇩Q⇩R = A⇩Q @ A⇩R› have "A⇩Q ♯* N" and "A⇩R ♯* N"
by simp+
from ‹A⇩Q⇩R ♯* xvec› ‹A⇩Q⇩R = (A⇩Q@A⇩R)› have "A⇩Q ♯* xvec" and "A⇩R ♯* xvec"
by simp+
have Aeq: "A⇩Q⇩R = A⇩Q@A⇩R" and Ψeq: "Ψ⇩Q⇩R = Ψ⇩Q ⊗ Ψ⇩R" by fact+
from Aeq ‹A⇩Q⇩R ♯* P› have "A⇩Q ♯* P" and "A⇩R ♯* P" by simp+
from Aeq ‹A⇩Q⇩R ♯* Ψ› have "A⇩Q ♯* Ψ" and "A⇩R ♯* Ψ" by simp+
from Aeq ‹A⇩P ♯* A⇩Q⇩R› have "A⇩P ♯* A⇩Q" and "A⇩P ♯* A⇩R" by simp+
with ‹A⇩P ♯* Q› ‹extractFrame Q = ⟨A⇩Q, Ψ⇩Q⟩› have "A⇩P ♯* Ψ⇩Q"
by (metis extractFrameFreshChain freshChainSym freshFrameDest)
from PTrans Ψeq have "(Ψ ⊗ Ψ⇩R) ⊗ Ψ⇩Q ⊳ P ⟼¿M⦇N⦈ ≺ P'"
by(metis statEqTransition Associativity Commutativity Composition)
moreover note FrP
moreover from ‹(Ψ ⊗ Ψ⇩P) ⊗ Ψ⇩R ⊳ Q ⟼¡M⦇ν*xvec⦈⟨N⟩ ≺ Q'› have "(Ψ ⊗ Ψ⇩R) ⊗ Ψ⇩P ⊳ Q ⟼¡M⦇ν*xvec⦈⟨N⟩ ≺ Q'"
by(metis statEqTransition Associativity Commutativity Composition)
moreover note ‹extractFrame Q = ⟨A⇩Q, Ψ⇩Q⟩›
moreover from ‹A⇩P ♯* Q› ‹A⇩P ♯* R› ‹A⇩P ♯* A⇩Q⇩R› Aeq ‹extractFrame Q = ⟨A⇩Q, Ψ⇩Q⟩› ‹extractFrame R = ⟨A⇩R, Ψ⇩R⟩›
have "A⇩P ♯* A⇩Q" and "A⇩P ♯* Ψ⇩R" by(auto dest: extractFrameFreshChain)
moreover from ‹A⇩Q⇩R ♯* Ψ› ‹A⇩Q⇩R ♯* P› Aeq have "A⇩Q ♯* Ψ" and "A⇩Q ♯* P" by simp+
ultimately have PQTrans: "Ψ ⊗ Ψ⇩R ⊳ P ∥ Q ⟼¡M⦇ν*xvec⦈⟨N⟩ ≺ (P' ∥ Q')"
using ‹A⇩P ♯* Ψ› ‹A⇩P ♯* P› ‹A⇩P ♯* Q› ‹A⇩P ♯* M› ‹A⇩P ♯* A⇩Q› ‹A⇩Q ♯* Ψ› ‹A⇩Q ♯* Ψ⇩R› ‹A⇩Q ♯* P› ‹A⇩Q ♯* Q› ‹A⇩Q ♯* M› ‹xvec ♯* P›
by(intro BrComm1) (assumption | force)+
have RTrans: "(Ψ ⊗ Ψ⇩P) ⊗ Ψ⇩Q ⊳ R ⟼¿M⦇N⦈ ≺ R'" and FrR: "extractFrame R = ⟨A⇩R, Ψ⇩R⟩" by fact+
note PQTrans
moreover from FrP ‹extractFrame Q = ⟨A⇩Q, Ψ⇩Q⟩› ‹A⇩P ♯* A⇩Q› ‹A⇩P ♯* Ψ⇩Q› ‹A⇩Q⇩R ♯* Ψ⇩P› Aeq
have "extractFrame(P ∥ Q) = ⟨(A⇩P@A⇩Q), Ψ⇩P ⊗ Ψ⇩Q⟩" by simp+
moreover from RTrans have "Ψ ⊗ (Ψ⇩P ⊗ Ψ⇩Q) ⊳ R ⟼¿M⦇N⦈ ≺ R'"
by (metis Associativity statEqTransition)
moreover note FrR
moreover note ‹A⇩P ♯* Ψ› ‹A⇩Q ♯* Ψ› ‹A⇩P ♯* P› ‹A⇩P ♯* Q› ‹A⇩Q ♯* P› ‹A⇩Q ♯* Q› ‹A⇩P ♯* R› ‹A⇩Q ♯* R›
‹A⇩P ♯* M› ‹A⇩Q ♯* M› ‹A⇩P ♯* A⇩R› ‹A⇩Q ♯* A⇩R› ‹A⇩R ♯* Ψ› ‹A⇩R ♯* P› ‹A⇩R ♯* Q› ‹A⇩R ♯* R› ‹A⇩R ♯* M›
‹xvec ♯* P› ‹xvec ♯* Q› ‹xvec ♯* R›
ultimately have "Ψ ⊳ (P ∥ Q) ∥ R ⟼¡M⦇ν*xvec⦈⟨N⟩ ≺ (P' ∥ Q') ∥ R'"
by(auto intro: BrComm2)
moreover have "(Ψ, (P' ∥ Q') ∥ R', (P' ∥ (Q' ∥ R'))) ∈ Rel"
by(rule C1)
ultimately show ?case by blast
qed
next
case(cBrComm2 Ψ⇩Q⇩R M xvec N P' A⇩P Ψ⇩P QR' A⇩Q⇩R)
from ‹¡M⦇ν*xvec⦈⟨N⟩ = α› ‹bn α ♯* Q› ‹bn α ♯* R› have "xvec ♯* Q" and "xvec ♯* R" by auto
from ‹A⇩Q⇩R ♯* (Ψ, P, Q, R, α)› have "A⇩Q⇩R ♯* Q" and "A⇩Q⇩R ♯* R" and "A⇩Q⇩R ♯* Ψ" by simp+
from ‹A⇩P ♯* (Q ∥ R)› have "A⇩P ♯* Q" and "A⇩P ♯* R" by simp+
have PTrans: "Ψ ⊗ Ψ⇩Q⇩R ⊳ P ⟼¡M⦇ν*xvec⦈⟨N⟩ ≺ P'" and FrP: "extractFrame P = ⟨A⇩P, Ψ⇩P⟩" by fact+
note ‹Ψ ⊗ Ψ⇩P ⊳ Q ∥ R ⟼¿M⦇N⦈ ≺ QR'› ‹extractFrame(Q ∥ R) = ⟨A⇩Q⇩R, Ψ⇩Q⇩R⟩› ‹distinct A⇩Q⇩R›
moreover from ‹A⇩Q⇩R ♯* Ψ› ‹A⇩Q⇩R ♯* Ψ⇩P› have "A⇩Q⇩R ♯* (Ψ ⊗ Ψ⇩P)" by force
ultimately show ?case using ‹A⇩Q⇩R ♯* Q› ‹A⇩Q⇩R ♯* R› ‹A⇩Q⇩R ♯* M›
proof(induct rule: parCasesBrInputFrame)
case(cPar1 Q' A⇩Q Ψ⇩Q A⇩R Ψ⇩R)
from ‹A⇩Q⇩R ♯* N› ‹A⇩Q⇩R = A⇩Q @ A⇩R› have "A⇩Q ♯* N" and "A⇩R ♯* N"
by simp+
from ‹A⇩Q⇩R ♯* xvec› ‹A⇩Q⇩R = (A⇩Q@A⇩R)› have "A⇩Q ♯* xvec" and "A⇩R ♯* xvec"
by simp+
have Aeq: "A⇩Q⇩R = A⇩Q@A⇩R" and Ψeq: "Ψ⇩Q⇩R = Ψ⇩Q ⊗ Ψ⇩R" by fact+
from PTrans Ψeq have "(Ψ ⊗ Ψ⇩R) ⊗ Ψ⇩Q ⊳ P ⟼¡M⦇ν*xvec⦈⟨N⟩ ≺ P'"
by(metis statEqTransition Associativity Commutativity Composition)
moreover note FrP
moreover from ‹(Ψ ⊗ Ψ⇩P) ⊗ Ψ⇩R ⊳ Q ⟼¿M⦇N⦈ ≺ Q'› have "(Ψ ⊗ Ψ⇩R) ⊗ Ψ⇩P ⊳ Q ⟼¿M⦇N⦈ ≺ Q'"
by(metis statEqTransition Associativity Commutativity Composition)
moreover note ‹extractFrame Q = ⟨A⇩Q, Ψ⇩Q⟩›
moreover from ‹A⇩P ♯* Q› ‹A⇩P ♯* R› ‹A⇩P ♯* A⇩Q⇩R› Aeq ‹extractFrame Q = ⟨A⇩Q, Ψ⇩Q⟩› ‹extractFrame R = ⟨A⇩R, Ψ⇩R⟩›
have "A⇩P ♯* A⇩Q" and "A⇩P ♯* Ψ⇩R" by(auto dest: extractFrameFreshChain)
moreover from ‹A⇩Q⇩R ♯* Ψ› ‹A⇩Q⇩R ♯* P› Aeq have "A⇩Q ♯* Ψ" and "A⇩Q ♯* P" by simp+
ultimately have "Ψ ⊗ Ψ⇩R ⊳ P ∥ Q ⟼¡M⦇ν*xvec⦈⟨N⟩ ≺ (P' ∥ Q')"
using ‹A⇩P ♯* Ψ› ‹A⇩P ♯* P› ‹A⇩P ♯* Q› ‹A⇩P ♯* M› ‹A⇩P ♯* A⇩Q› ‹A⇩Q ♯* Ψ› ‹A⇩Q ♯* Ψ⇩R› ‹A⇩Q ♯* P› ‹A⇩Q ♯* Q› ‹A⇩Q ♯* M› ‹xvec ♯* Q›
by(intro BrComm2) (assumption | force)+
moreover from ‹A⇩Q⇩R ♯* Ψ› ‹A⇩Q⇩R ♯* P› Aeq have "A⇩R ♯* Ψ" and "A⇩R ♯* P" by simp+
ultimately have "Ψ ⊳ (P ∥ Q) ∥ R ⟼¡M⦇ν*xvec⦈⟨N⟩ ≺ (P' ∥ Q') ∥ R" using ‹extractFrame R = ⟨A⇩R, Ψ⇩R⟩› ‹A⇩R ♯* Q› ‹A⇩R ♯* M› ‹A⇩R ♯* N› ‹xvec ♯* R› ‹A⇩R ♯* xvec›
by(intro Par1) (assumption | simp)+
moreover have "(Ψ, (P' ∥ Q') ∥ R, (P' ∥ (Q' ∥ R))) ∈ Rel"
by(rule C1)
ultimately show ?case by blast
next
case(cPar2 R' A⇩Q Ψ⇩Q A⇩R Ψ⇩R)
from ‹A⇩Q⇩R ♯* N› ‹A⇩Q⇩R ♯* xvec› ‹A⇩Q⇩R = A⇩Q @ A⇩R› have "A⇩Q ♯* N" and "A⇩R ♯* N" and "A⇩Q ♯* xvec" and "A⇩R ♯* xvec"
by simp+
have Aeq: "A⇩Q⇩R = A⇩Q@A⇩R" and Ψeq: "Ψ⇩Q⇩R = Ψ⇩Q ⊗ Ψ⇩R" by fact+
from ‹A⇩Q⇩R ♯* Ψ› ‹A⇩Q⇩R ♯* P› ‹A⇩Q⇩R ♯* Ψ⇩P› ‹A⇩P ♯* A⇩Q⇩R› Aeq
have "A⇩R ♯* Ψ" and "A⇩R ♯* Ψ⇩P" and "A⇩P ♯* A⇩R" and "A⇩P ♯* A⇩Q" and "A⇩R ♯* P" and "A⇩Q ♯* Ψ" and "A⇩Q ♯* Ψ⇩P" by simp+
have RTrans: "(Ψ ⊗ Ψ⇩P) ⊗ Ψ⇩Q ⊳ R ⟼¿M⦇N⦈ ≺ R'" and FrR: "extractFrame R = ⟨A⇩R, Ψ⇩R⟩" by fact+
from PTrans Ψeq have "(Ψ ⊗ Ψ⇩R) ⊗ Ψ⇩Q ⊳ P ⟼¡M⦇ν*xvec⦈⟨N⟩ ≺ P'"
by(metis statEqTransition Associativity Commutativity Composition)
moreover from ‹A⇩P ♯* R› ‹A⇩P ♯* Q› ‹A⇩P ♯* A⇩Q⇩R› FrR ‹extractFrame Q = ⟨A⇩Q, Ψ⇩Q⟩› Aeq have "A⇩P ♯* Ψ⇩Q" and "A⇩P ♯* Ψ⇩R"
by(auto dest: extractFrameFreshChain)
moreover from ‹A⇩Q⇩R ♯* P› ‹A⇩Q⇩R ♯* N› Aeq have "A⇩Q ♯* P" and "A⇩Q ♯* N" by simp+
ultimately have "Ψ ⊗ Ψ⇩R ⊳ P ∥ Q ⟼¡M⦇ν*xvec⦈⟨N⟩ ≺ (P' ∥ Q)" using ‹extractFrame Q = ⟨A⇩Q, Ψ⇩Q⟩› ‹A⇩Q ♯* Ψ⇩R›
‹A⇩Q ♯* M› ‹A⇩Q ♯* Ψ› ‹xvec ♯* Q› ‹A⇩Q ♯* xvec›
by(intro Par1) (assumption | force)+
moreover from FrP ‹extractFrame Q = ⟨A⇩Q, Ψ⇩Q⟩› ‹A⇩P ♯* A⇩Q› ‹A⇩P ♯* Ψ⇩Q› ‹A⇩Q ♯* Ψ⇩P›
have "extractFrame(P ∥ Q) = ⟨(A⇩P@A⇩Q), Ψ⇩P ⊗ Ψ⇩Q⟩" by simp+
moreover from RTrans have "Ψ ⊗ (Ψ⇩P ⊗ Ψ⇩Q) ⊳ R ⟼¿M⦇N⦈ ≺ R'" by(metis Associativity statEqTransition)
moreover note FrR
ultimately have "Ψ ⊳ (P ∥ Q) ∥ R ⟼¡M⦇ν*xvec⦈⟨N⟩ ≺ ((P' ∥ Q) ∥ R')"
using ‹A⇩P ♯* Ψ› ‹A⇩Q ♯* Ψ› ‹A⇩P ♯* P› ‹A⇩Q ♯* P› ‹A⇩P ♯* Q› ‹A⇩Q ♯* Q› ‹A⇩P ♯* R› ‹A⇩Q ♯* R› ‹A⇩P ♯* M› ‹A⇩Q ♯* M› ‹A⇩P ♯* A⇩R› ‹A⇩Q ♯* A⇩R›
‹A⇩R ♯* Ψ› ‹A⇩R ♯* P› ‹A⇩R ♯* Q› ‹A⇩R ♯* R› ‹A⇩R ♯* M› ‹xvec ♯* P› ‹xvec ♯* R›
by(auto intro: BrComm2)
moreover have "(Ψ, ((P' ∥ Q) ∥ R'), (P' ∥ (Q ∥ R'))) ∈ Rel"
by(rule C1)
ultimately show ?case by blast
next
case(cBrMerge Ψ⇩R Q' A⇩Q Ψ⇩Q R' A⇩R)
from ‹A⇩Q⇩R ♯* N› ‹A⇩Q⇩R = A⇩Q @ A⇩R› have "A⇩Q ♯* N" and "A⇩R ♯* N"
by simp+
from ‹A⇩Q⇩R ♯* xvec› ‹A⇩Q⇩R = (A⇩Q@A⇩R)› have "A⇩Q ♯* xvec" and "A⇩R ♯* xvec"
by simp+
have Aeq: "A⇩Q⇩R = A⇩Q@A⇩R" and Ψeq: "Ψ⇩Q⇩R = Ψ⇩Q ⊗ Ψ⇩R" by fact+
from Aeq ‹A⇩Q⇩R ♯* P› have "A⇩Q ♯* P" and "A⇩R ♯* P" by simp+
from Aeq ‹A⇩Q⇩R ♯* M› have "A⇩Q ♯* M" and "A⇩R ♯* M" by simp+
from Aeq ‹A⇩Q⇩R ♯* Ψ› have "A⇩Q ♯* Ψ" and "A⇩R ♯* Ψ" by simp+
from Aeq ‹A⇩P ♯* A⇩Q⇩R› have "A⇩P ♯* A⇩Q" and "A⇩P ♯* A⇩R" by simp+
with ‹A⇩P ♯* Q› ‹extractFrame Q = ⟨A⇩Q, Ψ⇩Q⟩› have "A⇩P ♯* Ψ⇩Q"
by (metis extractFrameFreshChain freshChainSym freshFrameDest)
from PTrans Ψeq have "(Ψ ⊗ Ψ⇩R) ⊗ Ψ⇩Q ⊳ P ⟼¡M⦇ν*xvec⦈⟨N⟩ ≺ P'"
by(metis statEqTransition Associativity Commutativity Composition)
moreover note FrP
moreover from ‹(Ψ ⊗ Ψ⇩P) ⊗ Ψ⇩R ⊳ Q ⟼¿M⦇N⦈ ≺ Q'› have "(Ψ ⊗ Ψ⇩R) ⊗ Ψ⇩P ⊳ Q ⟼¿M⦇N⦈ ≺ Q'"
by(metis statEqTransition Associativity Commutativity Composition)
moreover note ‹extractFrame Q = ⟨A⇩Q, Ψ⇩Q⟩›
moreover from ‹A⇩P ♯* Q› ‹A⇩P ♯* R› ‹A⇩P ♯* A⇩Q⇩R› Aeq ‹extractFrame Q = ⟨A⇩Q, Ψ⇩Q⟩› ‹extractFrame R = ⟨A⇩R, Ψ⇩R⟩›
have "A⇩P ♯* A⇩Q" and "A⇩P ♯* Ψ⇩R" by(auto dest: extractFrameFreshChain)
moreover from ‹A⇩Q⇩R ♯* Ψ› ‹A⇩Q⇩R ♯* P› Aeq have "A⇩Q ♯* Ψ" and "A⇩Q ♯* P" by simp+
ultimately have PQTrans: "Ψ ⊗ Ψ⇩R ⊳ P ∥ Q ⟼¡M⦇ν*xvec⦈⟨N⟩ ≺ (P' ∥ Q')"
using ‹A⇩P ♯* Ψ› ‹A⇩P ♯* P› ‹A⇩P ♯* Q› ‹A⇩P ♯* M› ‹A⇩P ♯* A⇩Q› ‹A⇩Q ♯* Ψ› ‹A⇩Q ♯* Ψ⇩R› ‹A⇩Q ♯* P› ‹A⇩Q ♯* Q› ‹xvec ♯* Q› ‹A⇩Q ♯* M›
by(intro BrComm2) (assumption | force)+
have RTrans: "(Ψ ⊗ Ψ⇩P) ⊗ Ψ⇩Q ⊳ R ⟼¿M⦇N⦈ ≺ R'" and FrR: "extractFrame R = ⟨A⇩R, Ψ⇩R⟩" by fact+
note PQTrans
moreover from FrP ‹extractFrame Q = ⟨A⇩Q, Ψ⇩Q⟩› ‹A⇩P ♯* A⇩Q› ‹A⇩P ♯* Ψ⇩Q› ‹A⇩Q⇩R ♯* Ψ⇩P› Aeq
have "extractFrame(P ∥ Q) = ⟨(A⇩P@A⇩Q), Ψ⇩P ⊗ Ψ⇩Q⟩" by simp+
moreover from RTrans have "Ψ ⊗ (Ψ⇩P ⊗ Ψ⇩Q) ⊳ R ⟼¿M⦇N⦈ ≺ R'"
by (metis Associativity statEqTransition)
moreover note FrR
moreover note ‹A⇩P ♯* Ψ› ‹A⇩Q ♯* Ψ› ‹A⇩P ♯* P› ‹A⇩P ♯* Q› ‹A⇩Q ♯* P› ‹A⇩Q ♯* Q› ‹A⇩P ♯* R› ‹A⇩Q ♯* R›
‹A⇩P ♯* M› ‹A⇩Q ♯* M› ‹A⇩P ♯* A⇩R› ‹A⇩Q ♯* A⇩R› ‹A⇩R ♯* Ψ› ‹A⇩R ♯* P› ‹A⇩R ♯* Q› ‹A⇩R ♯* R› ‹A⇩R ♯* M› ‹xvec ♯* P› ‹xvec ♯* R›
ultimately have "Ψ ⊳ (P ∥ Q) ∥ R ⟼¡M⦇ν*xvec⦈⟨N⟩ ≺ (P' ∥ Q') ∥ R'"
by(auto intro: BrComm2)
moreover have "(Ψ, (P' ∥ Q') ∥ R', (P' ∥ (Q' ∥ R'))) ∈ Rel"
by(rule C1)
ultimately show ?case by blast
qed
qed
qed
lemma parNilLeft:
fixes Ψ :: 'b
and P :: "('a, 'b, 'c) psi"
and Rel :: "('b × ('a, 'b, 'c) psi × ('a, 'b, 'c) psi) set"
assumes "eqvt Rel"
and C1: "⋀Q. (Ψ, Q ∥ 𝟬, Q) ∈ Rel"
shows "Ψ ⊳ (P ∥ 𝟬) ↝[Rel] P"
using ‹eqvt Rel›
proof(induct rule: simI[of _ _ _ _ "()"])
case(cSim α P')
from ‹Ψ ⊳ P ⟼α ≺ P'› have "Ψ ⊗ SBottom' ⊳ P ⟼α ≺ P'"
by(metis statEqTransition Identity AssertionStatEqSym)
then have "Ψ ⊳ (P ∥ 𝟬) ⟼α ≺ (P' ∥ 𝟬)"
by(rule Par1) auto
moreover have "(Ψ, P' ∥ 𝟬, P') ∈ Rel" by(rule C1)
ultimately show ?case by blast
qed
lemma parNilRight:
fixes Ψ :: 'b
and P :: "('a, 'b, 'c) psi"
and Rel :: "('b × ('a, 'b, 'c) psi × ('a, 'b, 'c) psi) set"
assumes "eqvt Rel"
and C1: "⋀Q. (Ψ, Q, (Q ∥ 𝟬)) ∈ Rel"
shows "Ψ ⊳ P ↝[Rel] (P ∥ 𝟬)"
using ‹eqvt Rel›
proof(induct rule: simI[of _ _ _ _ "()"])
case(cSim α P')
note ‹Ψ ⊳ P ∥ 𝟬 ⟼α ≺ P'› ‹bn α ♯* Ψ› ‹bn α ♯* P›
moreover have "bn α ♯* 𝟬" by simp
ultimately show ?case using ‹bn α ♯* subject α›
proof(induct rule: parCases[where C="()"])
case(cPar1 P' A⇩Q Ψ⇩Q)
from ‹extractFrame(𝟬) = ⟨A⇩Q, Ψ⇩Q⟩› have "Ψ⇩Q = SBottom'" by auto
with ‹Ψ ⊗ Ψ⇩Q ⊳ P ⟼α ≺ P'› have "Ψ ⊳ P ⟼α ≺ P'"
by(metis statEqTransition Identity)
moreover have "(Ψ, P', P' ∥ 𝟬) ∈ Rel" by(rule C1)
ultimately show ?case by blast
next
case(cPar2 Q' A⇩P Ψ⇩P)
from ‹Ψ ⊗ Ψ⇩P ⊳ 𝟬 ⟼α ≺ Q'› have False
by auto
then show ?case by simp
next
case(cComm1 Ψ⇩Q M N P' A⇩P Ψ⇩P K xvec Q' A⇩Q)
from ‹Ψ ⊗ Ψ⇩P ⊳ 𝟬 ⟼K⦇ν*xvec⦈⟨N⟩ ≺ Q'› have False by auto
then show ?case by simp
next
case(cComm2 Ψ⇩Q M xvec N P' A⇩P Ψ⇩P K Q' A⇩Q)
from ‹Ψ ⊗ Ψ⇩P ⊳ 𝟬 ⟼K⦇N⦈ ≺ Q'› have False
by auto
then show ?case by simp
next
case(cBrMerge Ψ⇩Q M N P' A⇩P Ψ⇩P Q' A⇩Q)
from ‹Ψ ⊗ Ψ⇩P ⊳ 𝟬 ⟼ ¿M⦇N⦈ ≺ Q'› have False
by auto
then show ?case by simp
next
case(cBrComm1 Ψ⇩Q M N P' A⇩P Ψ⇩P xvec Q' A⇩Q)
from ‹Ψ ⊗ Ψ⇩P ⊳ 𝟬 ⟼ ¡M⦇ν*xvec⦈⟨N⟩ ≺ Q'› have False
by auto
then show ?case by simp
next
case(cBrComm2 Ψ⇩Q M xvec N P' A⇩P Ψ⇩P Q' A⇩Q)
from ‹Ψ ⊗ Ψ⇩P ⊳ 𝟬 ⟼ ¿M⦇N⦈ ≺ Q'› have False
by auto
then show ?case by simp
qed
qed
lemma resNilLeft:
fixes x :: name
and Ψ :: 'b
and Rel :: "('b × ('a, 'b, 'c) psi × ('a, 'b, 'c) psi) set"
shows "Ψ ⊳ ⦇νx⦈𝟬 ↝[Rel] 𝟬"
by(auto simp add: simulation_def)
lemma resNilRight:
fixes x :: name
and Ψ :: 'b
and Rel :: "('b × ('a, 'b, 'c) psi × ('a, 'b, 'c) psi) set"
shows "Ψ ⊳ 𝟬 ↝[Rel] ⦇νx⦈𝟬"
apply(clarsimp simp add: simulation_def)
by(cases rule: semantics.cases) (auto simp add: psi.inject alpha')
lemma inputPushResLeft:
fixes Ψ :: 'b
and x :: name
and M :: 'a
and xvec :: "name list"
and N :: 'a
and P :: "('a, 'b, 'c) psi"
assumes "eqvt Rel"
and "x ♯ Ψ"
and "x ♯ M"
and "x ♯ xvec"
and "x ♯ N"
and C1: "⋀Q. (Ψ, Q, Q) ∈ Rel"
shows "Ψ ⊳ ⦇νx⦈(M⦇λ*xvec N⦈.P) ↝[Rel] M⦇λ*xvec N⦈.⦇νx⦈P"
proof -
note ‹eqvt Rel› ‹x ♯ Ψ›
moreover have "x ♯ ⦇νx⦈(M⦇λ*xvec N⦈.P)" by(simp add: abs_fresh)
moreover from ‹x ♯ M› ‹x ♯ N› have "x ♯ M⦇λ*xvec N⦈.⦇νx⦈P"
by(auto simp add: inputChainFresh abs_fresh)
ultimately show ?thesis
proof(induct rule: simIFresh[of _ _ _ _ _ "()"])
case(cSim α P')
from ‹Ψ ⊳ M⦇λ*xvec N⦈.⦇νx⦈P ⟼α ≺ P'› ‹x ♯ α› show ?case
proof(induct rule: inputCases)
case(cInput K Tvec)
from ‹x ♯ K⦇N[xvec::=Tvec]⦈› have "x ♯ K" and "x ♯ N[xvec::=Tvec]" by simp+
from ‹Ψ ⊢ M ↔ K› ‹distinct xvec› ‹set xvec ⊆ supp N› ‹length xvec = length Tvec›
have "Ψ ⊳ M⦇λ*xvec N⦈.P ⟼K⦇(N[xvec::=Tvec])⦈ ≺ P[xvec::=Tvec]"
by(rule Input)
then have "Ψ ⊳ ⦇νx⦈(M⦇λ*xvec N⦈.P) ⟼K⦇(N[xvec::=Tvec])⦈ ≺ ⦇νx⦈(P[xvec::=Tvec])" using ‹x ♯ Ψ› ‹x ♯ K› ‹x ♯ N[xvec::=Tvec]›
by(intro Scope) auto
moreover from ‹length xvec = length Tvec› ‹distinct xvec› ‹set xvec ⊆ supp N› ‹x ♯ N[xvec::=Tvec]› have "x ♯ Tvec"
by(rule substTerm.subst3)
with ‹x ♯ xvec› have "(Ψ, ⦇νx⦈(P[xvec::=Tvec]), (⦇νx⦈P)[xvec::=Tvec]) ∈ Rel"
by(force intro: C1)
ultimately show ?case by blast
next
case(cBrInput K Tvec)
from ‹x ♯ ¿K⦇N[xvec::=Tvec]⦈› have "x ♯ K" and "x ♯ N[xvec::=Tvec]" by simp+
from ‹Ψ ⊢ K ≽ M› ‹distinct xvec› ‹set xvec ⊆ supp N› ‹length xvec = length Tvec›
have "Ψ ⊳ M⦇λ*xvec N⦈.P ⟼¿K⦇(N[xvec::=Tvec])⦈ ≺ P[xvec::=Tvec]"
by(rule BrInput)
then have "Ψ ⊳ ⦇νx⦈(M⦇λ*xvec N⦈.P) ⟼¿K⦇(N[xvec::=Tvec])⦈ ≺ ⦇νx⦈(P[xvec::=Tvec])" using ‹x ♯ Ψ› ‹x ♯ K› ‹x ♯ N[xvec::=Tvec]›
by(intro Scope) auto
moreover from ‹length xvec = length Tvec› ‹distinct xvec› ‹set xvec ⊆ supp N› ‹x ♯ N[xvec::=Tvec]› have "x ♯ Tvec"
by(rule substTerm.subst3)
with ‹x ♯ xvec› have "(Ψ, ⦇νx⦈(P[xvec::=Tvec]), (⦇νx⦈P)[xvec::=Tvec]) ∈ Rel"
by(force intro: C1)
ultimately show ?case by blast
qed
qed
qed
lemma inputPushResRight:
fixes Ψ :: 'b
and x :: name
and M :: 'a
and xvec :: "name list"
and N :: 'a
and P :: "('a, 'b, 'c) psi"
assumes "eqvt Rel"
and "x ♯ Ψ"
and "x ♯ M"
and "x ♯ xvec"
and "x ♯ N"
and C1: "⋀Q. (Ψ, Q, Q) ∈ Rel"
shows "Ψ ⊳ M⦇λ*xvec N⦈.⦇νx⦈P ↝[Rel] ⦇νx⦈(M⦇λ*xvec N⦈.P)"
proof -
note ‹eqvt Rel› ‹x ♯ Ψ›
moreover from ‹x ♯ M› ‹x ♯ N› have "x ♯ M⦇λ*xvec N⦈.⦇νx⦈P"
by(auto simp add: inputChainFresh abs_fresh)
moreover have "x ♯ ⦇νx⦈(M⦇λ*xvec N⦈.P)" by(simp add: abs_fresh)
ultimately show ?thesis
proof(induct rule: simIFresh[of _ _ _ _ _ "()"])
case(cSim α P')
note ‹Ψ ⊳ ⦇νx⦈(M⦇λ*xvec N⦈.P) ⟼α ≺ P'› ‹x ♯ Ψ› ‹x ♯ α› ‹x ♯ P'› ‹bn α ♯* Ψ›
moreover from ‹bn α ♯* (⦇νx⦈(M⦇λ*xvec N⦈.P))› ‹x ♯ α› have "bn α ♯* (M⦇λ*xvec N⦈.P)"
by simp
ultimately show ?case using ‹bn α ♯* subject α›
proof(induct rule: resCases[where C="()"])
case(cRes P')
from ‹Ψ ⊳ M⦇λ*xvec N⦈.P ⟼α ≺ P'› ‹x ♯ α› show ?case
proof(induct rule: inputCases)
case(cInput K Tvec)
from ‹x ♯ K⦇N[xvec::=Tvec]⦈› have "x ♯ K" and "x ♯ N[xvec::=Tvec]" by simp+
from ‹Ψ ⊢ M ↔ K› ‹distinct xvec› ‹set xvec ⊆ supp N› ‹length xvec = length Tvec›
have "Ψ ⊳ M⦇λ*xvec N⦈.(⦇νx⦈P) ⟼K⦇(N[xvec::=Tvec])⦈ ≺ (⦇νx⦈P)[xvec::=Tvec]"
by(rule Input)
moreover from ‹length xvec = length Tvec› ‹distinct xvec› ‹set xvec ⊆ supp N› ‹x ♯ N[xvec::=Tvec]› have "x ♯ Tvec"
by(rule substTerm.subst3)
with ‹x ♯ xvec› have "(Ψ, (⦇νx⦈P)[xvec::=Tvec], ⦇νx⦈(P[xvec::=Tvec])) ∈ Rel"
by(force intro: C1)
ultimately show ?case by blast
next
case(cBrInput K Tvec)
from ‹x ♯ ¿K⦇N[xvec::=Tvec]⦈› have "x ♯ K" and "x ♯ N[xvec::=Tvec]" by simp+
from ‹Ψ ⊢ K ≽ M› ‹distinct xvec› ‹set xvec ⊆ supp N› ‹length xvec = length Tvec›
have "Ψ ⊳ M⦇λ*xvec N⦈.(⦇νx⦈P) ⟼¿K⦇(N[xvec::=Tvec])⦈ ≺ (⦇νx⦈P)[xvec::=Tvec]"
by(rule BrInput)
moreover from ‹length xvec = length Tvec› ‹distinct xvec› ‹set xvec ⊆ supp N› ‹x ♯ N[xvec::=Tvec]› have "x ♯ Tvec"
by(rule substTerm.subst3)
with ‹x ♯ xvec› have "(Ψ, (⦇νx⦈P)[xvec::=Tvec], ⦇νx⦈(P[xvec::=Tvec])) ∈ Rel"
by(force intro: C1)
ultimately show ?case by blast
qed
next
case cOpen
then have False by auto
then show ?case by simp
next
case cBrOpen
then have False by auto
then show ?case by simp
next
case (cBrClose M xvec N P')
then have False by auto
then show ?case by simp
qed
qed
qed
lemma outputPushResLeft:
fixes Ψ :: 'b
and x :: name
and M :: 'a
and N :: 'a
and P :: "('a, 'b, 'c) psi"
assumes "eqvt Rel"
and "x ♯ Ψ"
and "x ♯ M"
and "x ♯ N"
and C1: "⋀Q. (Ψ, Q, Q) ∈ Rel"
shows "Ψ ⊳ ⦇νx⦈(M⟨N⟩.P) ↝[Rel] M⟨N⟩.⦇νx⦈P"
proof -
note ‹eqvt Rel› ‹x ♯ Ψ›
moreover have "x ♯ ⦇νx⦈(M⟨N⟩.P)" by(simp add: abs_fresh)
moreover from ‹x ♯ M› ‹x ♯ N› have "x ♯ M⟨N⟩.⦇νx⦈P"
by(auto simp add: abs_fresh)
ultimately show ?thesis
proof(induct rule: simIFresh[of _ _ _ _ _ "()"])
case(cSim α P')
from ‹Ψ ⊳ M⟨N⟩.⦇νx⦈P ⟼α ≺ P'› ‹x ♯ α›
show ?case
proof(induct rule: outputCases)
case(cOutput K)
from ‹Ψ ⊢ M ↔ K› have "Ψ ⊳ M⟨N⟩.P ⟼K⟨N⟩ ≺ P"
by(rule Output)
then have "Ψ ⊳ ⦇νx⦈(M⟨N⟩.P) ⟼K⟨N⟩ ≺ ⦇νx⦈P" using ‹x ♯ Ψ› ‹x ♯ K⟨N⟩›
by(rule Scope)
moreover have "(Ψ, ⦇νx⦈P, ⦇νx⦈P) ∈ Rel" by(rule C1)
ultimately show ?case by blast
next
case(cBrOutput K)
from ‹Ψ ⊢ M ≼ K› have "Ψ ⊳ M⟨N⟩.P ⟼¡K⟨N⟩ ≺ P"
by(rule BrOutput)
then have "Ψ ⊳ ⦇νx⦈(M⟨N⟩.P) ⟼¡K⟨N⟩ ≺ ⦇νx⦈P" using ‹x ♯ Ψ› ‹x ♯ ¡K⟨N⟩›
by(rule Scope)
moreover have "(Ψ, ⦇νx⦈P, ⦇νx⦈P) ∈ Rel" by(rule C1)
ultimately show ?case by blast
qed
qed
qed
lemma broutputNoBind:
fixes Ψ :: 'b
and M :: 'a
and N :: 'a
and P :: "('a, 'b, 'c) psi"
and α :: "'a action"
and P' :: "('a, 'b, 'c) psi"
assumes "Ψ ⊳ M⟨N⟩.P ⟼(¡K⦇ν*xvec⦈⟨N'⟩) ≺ P'"
shows "xvec = []"
proof -
from assms have "bn(¡K⦇ν*xvec⦈⟨N'⟩) = []"
by(induct rule: outputCases) auto
then show ?thesis by simp
qed
lemma broutputObjectEq:
fixes Ψ :: 'b
and M :: 'a
and N :: 'a
and P :: "('a, 'b, 'c) psi"
and α :: "'a action"
and P' :: "('a, 'b, 'c) psi"
assumes "Ψ ⊳ M⟨N⟩.P ⟼(¡K⦇ν*xvec⦈⟨N'⟩) ≺ P'"
shows "N = N'"
proof -
from assms have "object(¡K⦇ν*xvec⦈⟨N'⟩) = Some N"
by(induct rule: outputCases) auto
then show ?thesis
by simp
qed
lemma brOutputOutputCases[consumes 1, case_names cBrOutput]:
fixes Ψ :: 'b
and M :: 'a
and N :: 'a
and P :: "('a, 'b, 'c) psi"
and α :: "'a action"
and P' :: "('a, 'b, 'c) psi"
assumes Trans: "Ψ ⊳ M⟨N⟩.P ⟼(¡K⦇ν*xvec⦈⟨N'⟩) ≺ P'"
and rBrOutput: "⋀K. Ψ ⊢ M ≼ K ⟹ Prop (¡K⟨N⟩) P"
shows "Prop (¡K⦇ν*xvec⦈⟨N'⟩) P'"
proof -
have "xvec = []" using Trans by(rule broutputNoBind)
then obtain K' N'' where eq: "(¡K⦇ν*xvec⦈⟨N'⟩)= ¡K'⟨N''⟩"
by blast
have "N = N'" using Trans by(rule broutputObjectEq)
from Trans ‹(¡K⦇ν*xvec⦈⟨N'⟩)= ¡K'⟨N''⟩›
show ?thesis unfolding ‹N = N'›[symmetric]
proof(induct rule: outputCases)
case (cOutput K) then show ?case by simp
next
case (cBrOutput K) then show ?case
by(intro rBrOutput)
qed
qed
lemma outputPushResRight:
fixes Ψ :: 'b
and x :: name
and M :: 'a
and N :: 'a
and P :: "('a, 'b, 'c) psi"
assumes "eqvt Rel"
and "x ♯ Ψ"
and "x ♯ M"
and "x ♯ N"
and C1: "⋀Q. (Ψ, Q, Q) ∈ Rel"
shows "Ψ ⊳ M⟨N⟩.⦇νx⦈P ↝[Rel] ⦇νx⦈(M⟨N⟩.P)"
proof -
note ‹eqvt Rel› ‹x ♯ Ψ›
moreover from ‹x ♯ M› ‹x ♯ N› have "x ♯ M⟨N⟩.⦇νx⦈P"
by(auto simp add: abs_fresh)
moreover have "x ♯ ⦇νx⦈(M⟨N⟩.P)" by(simp add: abs_fresh)
ultimately show ?thesis
proof(induct rule: simIFresh[of _ _ _ _ _ "(M, N)"])
case(cSim α P')
note ‹Ψ ⊳ ⦇νx⦈(M⟨N⟩.P) ⟼α ≺ P'› ‹x ♯ Ψ› ‹x ♯ α› ‹x ♯ P'› ‹bn α ♯* Ψ›
moreover from ‹bn α ♯* (⦇νx⦈(M⟨N⟩.P))› ‹x ♯ α› have "bn α ♯* (M⟨N⟩.P)" by simp
ultimately show ?case using ‹bn α ♯* subject α› ‹bn α ♯* (M, N)› ‹x ♯ α› ‹x ♯ M› ‹x ♯ N›
proof(induct rule: resCases[where C="()"])
case(cOpen K xvec1 xvec2 y N' P')
from ‹bn(K⦇ν*(xvec1@y#xvec2)⦈⟨N'⟩) ♯* (M, N)› have "y ♯ N" by simp+
from ‹Ψ ⊳ M⟨N⟩.P ⟼K⦇ν*(xvec1@xvec2)⦈⟨([(x, y)] ∙ N')⟩ ≺ ([(x, y)] ∙ P')›
have "N = ([(x, y)] ∙ N')"
apply -
by(ind_cases "Ψ ⊳ M⟨N⟩.P ⟼K⦇ν*(xvec1@xvec2)⦈⟨([(x, y)] ∙ N')⟩ ≺ ([(x, y)] ∙ P')")
(auto simp add: residualInject psi.inject)
with ‹x ♯ N› ‹y ♯ N› ‹x ≠ y› have "N = N'"
by(subst pt_bij[OF pt_name_inst, OF at_name_inst, symmetric, where pi="[(x, y)]"])
(simp add: fresh_left calc_atm)
with ‹y ∈ supp N'› ‹y ♯ N› have False by(simp add: fresh_def)
then show ?case by simp
next
case(cBrOpen K xvec1 xvec2 y N' P')
from ‹bn(¡K⦇ν*(xvec1@y#xvec2)⦈⟨N'⟩) ♯* (M, N)› have "y ♯ N" by simp+
from ‹Ψ ⊳ M⟨N⟩.P ⟼¡K⦇ν*(xvec1@xvec2)⦈⟨([(x, y)] ∙ N')⟩ ≺ ([(x, y)] ∙ P')›
have "N = ([(x, y)] ∙ N')"
apply -
by(ind_cases "Ψ ⊳ M⟨N⟩.P ⟼¡K⦇ν*(xvec1@xvec2)⦈⟨([(x, y)] ∙ N')⟩ ≺ ([(x, y)] ∙ P')")
(auto simp add: residualInject psi.inject)
with ‹x ♯ N› ‹y ♯ N› ‹x ≠ y› have "N = N'"
by(subst pt_bij[OF pt_name_inst, OF at_name_inst, symmetric, where pi="[(x, y)]"])
(simp add: fresh_left calc_atm)
with ‹y ∈ supp N'› ‹y ♯ N› have False by(simp add: fresh_def)
then show ?case by simp
next
case(cRes P')
from ‹Ψ ⊳ M⟨N⟩.P ⟼α ≺ P'› show ?case
proof(induct rule: outputCases)
case(cOutput K)
from ‹Ψ ⊢ M ↔ K› have "Ψ ⊳ M⟨N⟩.(⦇νx⦈P) ⟼K⟨N⟩ ≺ ⦇νx⦈P"
by(rule Output)
moreover have "(Ψ, ⦇νx⦈P, ⦇νx⦈P) ∈ Rel" by(rule C1)
ultimately show ?case by force
next
case(cBrOutput K)
from ‹Ψ ⊢ M ≼ K› have "Ψ ⊳ M⟨N⟩.(⦇νx⦈P) ⟼¡K⟨N⟩ ≺ ⦇νx⦈P"
by(rule BrOutput)
moreover have "(Ψ, ⦇νx⦈P, ⦇νx⦈P) ∈ Rel" by(rule C1)
ultimately show ?case by force
qed
next
case(cBrClose K xvec N' P')
from ‹Ψ ⊳ M⟨N⟩.P ⟼ ¡K⦇ν*xvec⦈⟨N'⟩ ≺ P'›
have "Ψ ⊢ M ≼ the(subject(¡K⦇ν*xvec⦈⟨N'⟩))"
by(rule brOutputOutputCases) simp
then have "Ψ ⊢ M ≼ K" by simp
then have "supp K ⊆ (supp M:: name set)" by(rule chanOutConSupp)
then have False using ‹x ∈ supp K› ‹x ♯ M› unfolding fresh_def
by blast
then show ?case by(rule FalseE)
qed
qed
qed
lemma casePushResLeft:
fixes Ψ :: 'b
and x :: name
and Cs :: "('c × ('a, 'b, 'c) psi) list"
assumes "eqvt Rel"
and "x ♯ Ψ"
and "x ♯ map fst Cs"
and C1: "⋀Q. (Ψ, Q, Q) ∈ Rel"
shows "Ψ ⊳ ⦇νx⦈(Cases Cs) ↝[Rel] Cases (map (λ(φ, P). (φ, ⦇νx⦈P)) Cs)"
proof -
note ‹eqvt Rel› ‹x ♯ Ψ›
moreover have "x ♯ ⦇νx⦈(Cases Cs)" by(simp add: abs_fresh)
moreover from ‹x ♯ map fst Cs› have "x ♯ Cases (map (λ(φ, P). (φ, ⦇νx⦈P)) Cs)"
by(induct Cs) (auto simp add: abs_fresh)
ultimately show ?thesis
proof(induct rule: simIFresh[of _ _ _ _ _ Cs])
case(cSim α P'')
from ‹Ψ ⊳ Cases (map (λ(φ, P). (φ, ⦇νx⦈P)) Cs) ⟼α ≺ P''›
show ?case
proof(induct rule: caseCases)
case(cCase φ P')
from ‹(φ, P') ∈ set (map (λ(φ, P). (φ, ⦇νx⦈P)) Cs)›
obtain P where "(φ, P) ∈ set Cs" and "P' = ⦇νx⦈P"
by(induct Cs) auto
from ‹guarded P'› ‹P' = ⦇νx⦈P› have "guarded P" by simp
from ‹Ψ ⊳ P' ⟼α ≺ P''› ‹P' = ⦇νx⦈P› have "Ψ ⊳ ⦇νx⦈P ⟼α ≺ P''"
by simp
moreover note ‹x ♯ Ψ› ‹x ♯ α› ‹x ♯ P''› ‹bn α ♯* Ψ›
moreover from ‹bn α ♯* Cs› ‹(φ, P) ∈ set Cs›
have "bn α ♯* P" by(auto dest: memFreshChain)
ultimately show ?case using ‹bn α ♯* subject α› ‹x ♯ α› ‹bn α ♯* Cs›
proof(induct rule: resCases[where C="()"])
case(cOpen M xvec1 xvec2 y N P')
from ‹x ♯ M⦇ν*(xvec1@y#xvec2)⦈⟨N⟩› have "x ♯ xvec1" and "x ♯ xvec2" and "x ♯ M" by simp+
from ‹bn(M⦇ν*(xvec1@y#xvec2)⦈⟨N⟩) ♯* Cs› have "y ♯ Cs" by simp
from ‹Ψ ⊳ P ⟼M⦇ν*(xvec1@xvec2)⦈⟨([(x, y)] ∙ N)⟩ ≺ ([(x, y)] ∙ P')› ‹(φ, P) ∈ set Cs› ‹Ψ ⊢ φ› ‹guarded P›
have "Ψ ⊳ Cases Cs ⟼M⦇ν*(xvec1@xvec2)⦈⟨([(x, y)] ∙ N)⟩ ≺ ([(x, y)] ∙ P')" by(rule Case)
then have "([(x, y)] ∙ Ψ) ⊳ ([(x, y)] ∙ (Cases Cs)) ⟼([(x, y)] ∙ (M⦇ν*(xvec1@xvec2)⦈⟨([(x, y)] ∙ N)⟩ ≺ ([(x, y)] ∙ P')))"
by(rule semantics.eqvt)
with ‹x ♯ Ψ› ‹x ♯ M› ‹y ♯ xvec1› ‹y ♯ xvec2› ‹y ♯ Ψ› ‹y ♯ M› ‹x ♯ xvec1› ‹x ♯ xvec2›
have "Ψ ⊳ ([(x, y)] ∙ (Cases Cs)) ⟼M⦇ν*(xvec1@xvec2)⦈⟨N⟩ ≺ P'" by(simp add: eqvts)
then have "Ψ ⊳ ⦇νy⦈([(x, y)] ∙ (Cases Cs)) ⟼M⦇ν*(xvec1@y#xvec2)⦈⟨N⟩ ≺ P'" using ‹y ∈ supp N› ‹y ♯ Ψ› ‹y ♯ M› ‹y ♯ xvec1› ‹y ♯ xvec2›
by(rule Open)
then have "Ψ ⊳ ⦇νx⦈(Cases Cs) ⟼M⦇ν*(xvec1@y#xvec2)⦈⟨N⟩ ≺ P'" using ‹y ♯ Cs›
by(simp add: alphaRes)
moreover have "(Ψ, P', P') ∈ Rel" by(rule C1)
ultimately show ?case by blast
next
case(cBrOpen M xvec1 xvec2 y N P')
from ‹x ♯ ¡M⦇ν*(xvec1@y#xvec2)⦈⟨N⟩› have "x ♯ xvec1" and "x ♯ xvec2" and "x ♯ M" by simp+
from ‹bn(¡M⦇ν*(xvec1@y#xvec2)⦈⟨N⟩) ♯* Cs› have "y ♯ Cs" by simp
from ‹Ψ ⊳ P ⟼¡M⦇ν*(xvec1@xvec2)⦈⟨([(x, y)] ∙ N)⟩ ≺ ([(x, y)] ∙ P')› ‹(φ, P) ∈ set Cs› ‹Ψ ⊢ φ› ‹guarded P›
have "Ψ ⊳ Cases Cs ⟼¡M⦇ν*(xvec1@xvec2)⦈⟨([(x, y)] ∙ N)⟩ ≺ ([(x, y)] ∙ P')" by(rule Case)
then have "([(x, y)] ∙ Ψ) ⊳ ([(x, y)] ∙ (Cases Cs)) ⟼([(x, y)] ∙ (¡M⦇ν*(xvec1@xvec2)⦈⟨([(x, y)] ∙ N)⟩ ≺ ([(x, y)] ∙ P')))"
by(rule semantics.eqvt)
with ‹x ♯ Ψ› ‹x ♯ M› ‹y ♯ xvec1› ‹y ♯ xvec2› ‹y ♯ Ψ› ‹y ♯ M› ‹x ♯ xvec1› ‹x ♯ xvec2›
have "Ψ ⊳ ([(x, y)] ∙ (Cases Cs)) ⟼¡M⦇ν*(xvec1@xvec2)⦈⟨N⟩ ≺ P'" by(simp add: eqvts)
then have "Ψ ⊳ ⦇νy⦈([(x, y)] ∙ (Cases Cs)) ⟼¡M⦇ν*(xvec1@y#xvec2)⦈⟨N⟩ ≺ P'" using ‹y ∈ supp N› ‹y ♯ Ψ› ‹y ♯ M› ‹y ♯ xvec1› ‹y ♯ xvec2›
by(rule BrOpen)
then have "Ψ ⊳ ⦇νx⦈(Cases Cs) ⟼¡M⦇ν*(xvec1@y#xvec2)⦈⟨N⟩ ≺ P'" using ‹y ♯ Cs›
by(simp add: alphaRes)
moreover have "(Ψ, P', P') ∈ Rel" by(rule C1)
ultimately show ?case by blast
next
case(cRes P')
from ‹Ψ ⊳ P ⟼α ≺ P'› ‹(φ, P) ∈ set Cs› ‹Ψ ⊢ φ› ‹guarded P›
have "Ψ ⊳ Cases Cs ⟼α ≺ P'" by(rule Case)
then have "Ψ ⊳ ⦇νx⦈(Cases Cs) ⟼α ≺ ⦇νx⦈P'" using ‹x ♯ Ψ› ‹x ♯ α›
by(rule Scope)
moreover have "(Ψ, ⦇νx⦈P', ⦇νx⦈P') ∈ Rel" by(rule C1)
ultimately show ?case by blast
next
case(cBrClose M xvec N P')
from ‹Ψ ⊳ P ⟼ ¡M⦇ν*xvec⦈⟨N⟩ ≺ P'› ‹x ∈ supp M› ‹x ♯ Ψ›
have "Ψ ⊳ ⦇νx⦈P ⟼ τ ≺ ⦇νx⦈(⦇ν*xvec⦈P')"
by(rule BrClose)
from ‹Ψ ⊳ P ⟼ ¡M⦇ν*xvec⦈⟨N⟩ ≺ P'› ‹(φ, P) ∈ set Cs› ‹Ψ ⊢ φ› ‹guarded P›
have "Ψ ⊳ Cases Cs ⟼ ¡M⦇ν*xvec⦈⟨N⟩ ≺ P'"
by(rule Case)
then have "Ψ ⊳ ⦇νx⦈(Cases Cs) ⟼ τ ≺ ⦇νx⦈(⦇ν*xvec⦈P')" using ‹x ∈ supp M› ‹x ♯ Ψ›
by(rule BrClose)
moreover have "(Ψ,⦇νx⦈(⦇ν*xvec⦈P'),⦇νx⦈(⦇ν*xvec⦈P')) ∈ Rel" by fact
ultimately show ?case by blast
qed
qed
qed
qed
lemma casePushResRight:
fixes Ψ :: 'b
and x :: name
and Cs :: "('c × ('a, 'b, 'c) psi) list"
assumes "eqvt Rel"
and "x ♯ Ψ"
and "x ♯ map fst Cs"
and C1: "⋀Q. (Ψ, Q, Q) ∈ Rel"
shows "Ψ ⊳ Cases (map (λ(φ, P). (φ, ⦇νx⦈P)) Cs) ↝[Rel] ⦇νx⦈(Cases Cs)"
proof -
note ‹eqvt Rel› ‹x ♯ Ψ›
moreover from ‹x ♯ map fst Cs› have "x ♯ Cases (map (λ(φ, P). (φ, ⦇νx⦈P)) Cs)"
by(induct Cs) (auto simp add: abs_fresh)
moreover have "x ♯ ⦇νx⦈(Cases Cs)" by(simp add: abs_fresh)
ultimately show ?thesis
proof(induct rule: simIFresh[of _ _ _ _ _ Cs])
case(cSim α P'')
note ‹Ψ ⊳ ⦇νx⦈(Cases Cs) ⟼α ≺ P''› ‹x ♯ Ψ› ‹x ♯ α› ‹x ♯ P''› ‹bn α ♯* Ψ›
moreover from ‹bn α ♯* ⦇νx⦈(Cases Cs)› ‹x ♯ α› have "bn α ♯* (Cases Cs)" by simp
ultimately show ?case using ‹bn α ♯* subject α› ‹x ♯ α› ‹bn α ♯* Cs›
proof(induct rule: resCases[where C="()"])
case(cOpen M xvec1 xvec2 y N P')
from ‹x ♯ M⦇ν*(xvec1@y#xvec2)⦈⟨N⟩› have "x ♯ xvec1" and "x ♯ xvec2" and "x ♯ M" by simp+
from ‹bn(M⦇ν*(xvec1@y#xvec2)⦈⟨N⟩) ♯* Cs› have "y ♯ Cs" by simp
from ‹Ψ ⊳ Cases Cs ⟼M⦇ν*(xvec1@xvec2)⦈⟨([(x, y)] ∙ N)⟩ ≺ ([(x, y)] ∙ P')›
show ?case
proof(induct rule: caseCases)
case(cCase φ P)
from ‹Ψ ⊳ P ⟼M⦇ν*(xvec1@xvec2)⦈⟨([(x, y)] ∙ N)⟩ ≺ ([(x, y)] ∙ P')›
have "([(x, y)] ∙ Ψ) ⊳ ([(x, y)] ∙ P) ⟼([(x, y)] ∙ (M⦇ν*(xvec1@xvec2)⦈⟨([(x, y)] ∙ N)⟩ ≺ ([(x, y)] ∙ P')))"
by(rule semantics.eqvt)
with ‹x ♯ Ψ› ‹x ♯ M› ‹y ♯ xvec1› ‹y ♯ xvec2› ‹y ♯ Ψ› ‹y ♯ M› ‹x ♯ xvec1› ‹x ♯ xvec2›
have "Ψ ⊳ ([(x, y)] ∙ P) ⟼M⦇ν*(xvec1@xvec2)⦈⟨N⟩ ≺ P'" by(simp add: eqvts)
then have "Ψ ⊳ ⦇νy⦈([(x, y)] ∙ P) ⟼M⦇ν*(xvec1@y#xvec2)⦈⟨N⟩ ≺ P'" using ‹y ∈ supp N› ‹y ♯ Ψ› ‹y ♯ M› ‹y ♯ xvec1› ‹y ♯ xvec2›
by(rule Open)
then have "Ψ ⊳ ⦇νx⦈P ⟼M⦇ν*(xvec1@y#xvec2)⦈⟨N⟩ ≺ P'" using ‹y ♯ Cs› ‹(φ, P) ∈ set Cs›
by(subst alphaRes, auto dest: memFresh)
moreover from ‹(φ, P) ∈ set Cs› have "(φ, ⦇νx⦈P) ∈ set (map (λ(φ, P). (φ, ⦇νx⦈P)) Cs)"
by(induct Cs) auto
moreover note ‹Ψ ⊢ φ›
moreover from ‹guarded P› have "guarded(⦇νx⦈P)" by simp
ultimately have "Ψ ⊳ (Cases (map (λ(φ, P). (φ, ⦇νx⦈P)) Cs)) ⟼M⦇ν*(xvec1@y#xvec2)⦈⟨N⟩ ≺ P'"
by(rule Case)
moreover have "(Ψ, P', P') ∈ Rel" by(rule C1)
ultimately show ?case by blast
qed
next
case(cBrOpen M xvec1 xvec2 y N P')
from ‹x ♯ ¡M⦇ν*(xvec1@y#xvec2)⦈⟨N⟩› have "x ♯ xvec1" and "x ♯ xvec2" and "x ♯ M" by simp+
from ‹bn(¡M⦇ν*(xvec1@y#xvec2)⦈⟨N⟩) ♯* Cs› have "y ♯ Cs" by simp
from ‹Ψ ⊳ Cases Cs ⟼¡M⦇ν*(xvec1@xvec2)⦈⟨([(x, y)] ∙ N)⟩ ≺ ([(x, y)] ∙ P')›
show ?case
proof(induct rule: caseCases)
case(cCase φ P)
from ‹Ψ ⊳ P ⟼¡M⦇ν*(xvec1@xvec2)⦈⟨([(x, y)] ∙ N)⟩ ≺ ([(x, y)] ∙ P')›
have "([(x, y)] ∙ Ψ) ⊳ ([(x, y)] ∙ P) ⟼([(x, y)] ∙ (¡M⦇ν*(xvec1@xvec2)⦈⟨([(x, y)] ∙ N)⟩ ≺ ([(x, y)] ∙ P')))"
by(rule semantics.eqvt)
with ‹x ♯ Ψ› ‹x ♯ M› ‹y ♯ xvec1› ‹y ♯ xvec2› ‹y ♯ Ψ› ‹y ♯ M› ‹x ♯ xvec1› ‹x ♯ xvec2›
have "Ψ ⊳ ([(x, y)] ∙ P) ⟼¡M⦇ν*(xvec1@xvec2)⦈⟨N⟩ ≺ P'" by(simp add: eqvts)
then have "Ψ ⊳ ⦇νy⦈([(x, y)] ∙ P) ⟼¡M⦇ν*(xvec1@y#xvec2)⦈⟨N⟩ ≺ P'" using ‹y ∈ supp N› ‹y ♯ Ψ› ‹y ♯ M› ‹y ♯ xvec1› ‹y ♯ xvec2›
by(rule BrOpen)
then have "Ψ ⊳ ⦇νx⦈P ⟼¡M⦇ν*(xvec1@y#xvec2)⦈⟨N⟩ ≺ P'" using ‹y ♯ Cs› ‹(φ, P) ∈ set Cs›
by(subst alphaRes, auto dest: memFresh)
moreover from ‹(φ, P) ∈ set Cs› have "(φ, ⦇νx⦈P) ∈ set (map (λ(φ, P). (φ, ⦇νx⦈P)) Cs)"
by(induct Cs) auto
moreover note ‹Ψ ⊢ φ›
moreover from ‹guarded P› have "guarded(⦇νx⦈P)" by simp
ultimately have "Ψ ⊳ (Cases (map (λ(φ, P). (φ, ⦇νx⦈P)) Cs)) ⟼¡M⦇ν*(xvec1@y#xvec2)⦈⟨N⟩ ≺ P'"
by(rule Case)
moreover have "(Ψ, P', P') ∈ Rel" by(rule C1)
ultimately show ?case by blast
qed
next
case(cRes P')
from ‹Ψ ⊳ Cases Cs ⟼α ≺ P'›
show ?case
proof(induct rule: caseCases)
case(cCase φ P)
from ‹Ψ ⊳ P ⟼α ≺ P'› ‹x ♯ Ψ› ‹x ♯ α›
have "Ψ ⊳ ⦇νx⦈P ⟼α ≺ ⦇νx⦈P'" by(rule Scope)
moreover from ‹(φ, P) ∈ set Cs› have "(φ, ⦇νx⦈P) ∈ set (map (λ(φ, P). (φ, ⦇νx⦈P)) Cs)"
by(induct Cs) auto
moreover note ‹Ψ ⊢ φ›
moreover from ‹guarded P› have "guarded(⦇νx⦈P)" by simp
ultimately have "Ψ ⊳ (Cases (map (λ(φ, P). (φ, ⦇νx⦈P)) Cs)) ⟼α ≺ ⦇νx⦈P'"
by(rule Case)
moreover have "(Ψ, ⦇νx⦈P', ⦇νx⦈P') ∈ Rel" by(rule C1)
ultimately show ?case by blast
qed
next
case(cBrClose M xvec N P')
then show ?case
proof(induct rule: caseCases)
case(cCase φ P)
from ‹Ψ ⊳ P ⟼¡M⦇ν*xvec⦈⟨N⟩ ≺ P'› ‹x ♯ Ψ› ‹x ∈ supp M›
have "Ψ ⊳ ⦇νx⦈P ⟼ τ ≺ ⦇νx⦈(⦇ν*xvec⦈P')"
by(intro BrClose)
moreover from ‹(φ, P) ∈ set Cs› have "(φ, ⦇νx⦈P) ∈ set (map (λ(φ, P). (φ, ⦇νx⦈P)) Cs)"
by(induct Cs) auto
moreover from ‹guarded P› have "guarded(⦇νx⦈P)" by simp
moreover note ‹Ψ ⊢ φ›
ultimately have "Ψ ⊳ Cases map (λ(φ, P). (φ, ⦇νx⦈P)) Cs ⟼ τ ≺ ⦇νx⦈(⦇ν*xvec⦈P')"
by(intro Case)
moreover have "(Ψ,⦇νx⦈(⦇ν*xvec⦈P'),⦇νx⦈(⦇ν*xvec⦈P')) ∈ Rel"
by fact
ultimately show ?case
by blast
qed
qed
qed
qed
lemma resInputCases[consumes 5, case_names cRes]:
fixes Ψ :: 'b
and x :: name
and P :: "('a, 'b, 'c) psi"
and M :: 'a
and N :: 'a
and P' :: "('a, 'b, 'c) psi"
and C :: "'d::fs_name"
assumes Trans: "Ψ ⊳ ⦇νx⦈P ⟼M⦇N⦈ ≺ P'"
and "x ♯ Ψ"
and "x ♯ M"
and "x ♯ N"
and "x ♯ P'"
and rScope: "⋀P'. ⟦Ψ ⊳ P ⟼M⦇N⦈ ≺ P'⟧ ⟹ Prop (⦇νx⦈P')"
shows "Prop P'"
proof -
note Trans ‹x ♯ Ψ›
moreover from ‹x ♯ M› ‹x ♯ N› have "x ♯ M⦇N⦈" by simp
moreover note ‹x ♯ P'›
ultimately show ?thesis using assms
by(induct rule: resInputCases') simp
qed
lemma resBrInputCases[consumes 5, case_names cRes]:
fixes Ψ :: 'b
and x :: name
and P :: "('a, 'b, 'c) psi"
and M :: 'a
and N :: 'a
and P' :: "('a, 'b, 'c) psi"
and C :: "'d::fs_name"
assumes Trans: "Ψ ⊳ ⦇νx⦈P ⟼¿M⦇N⦈ ≺ P'"
and "x ♯ Ψ"
and "x ♯ M"
and "x ♯ N"
and "x ♯ P'"
and rScope: "⋀P'. ⟦Ψ ⊳ P ⟼¿M⦇N⦈ ≺ P'⟧ ⟹ Prop (⦇νx⦈P')"
shows "Prop P'"
proof -
note Trans ‹x ♯ Ψ›
moreover from ‹x ♯ M› ‹x ♯ N› have "x ♯ ¿M⦇N⦈" by simp
moreover note ‹x ♯ P'›
ultimately show ?thesis using assms
by(induct rule: resBrInputCases') simp
qed
lemma swap_supp:
fixes x :: name
and y :: name
and N :: 'a
assumes "y ∈ supp N"
shows "x ∈ supp ([(x, y)] ∙ N)"
using assms
by (metis fresh_bij fresh_def swap_simps(2))
lemma swap_supp':
fixes x :: name
and y :: name
and N :: 'a
assumes "x ∈ supp N"
shows "y ∈ supp ([(x, y)] ∙ N)"
using assms
by (metis fresh_bij fresh_def swap_simps(1))
lemma brOutputFreshSubjectChain:
fixes Ψ :: 'b
and Q :: "('a, 'b, 'c) psi"
and M :: 'a
and xvec :: "name list"
and N :: 'a
and Q' :: "('a, 'b, 'c) psi"
and yvec :: "name list"
assumes "Ψ ⊳ Q ⟼ ¡M⦇ν*xvec⦈⟨N⟩ ≺ Q'"
and "xvec ♯* M"
and "yvec ♯* Q"
shows "yvec ♯* M"
using assms
proof(induct yvec)
case Nil
then show ?case by simp
next
case(Cons a yvec)
then have "yvec ♯* M" by simp
from ‹(a # yvec) ♯* Q› have "a ♯ Q" by simp
from ‹xvec ♯* M› ‹a ♯ Q› ‹Ψ ⊳ Q ⟼ ¡M⦇ν*xvec⦈⟨N⟩ ≺ Q'›
have "a ♯ M"
by(simp add: brOutputFreshSubject)
with ‹yvec ♯* M› show ?case
by simp
qed
lemma scopeExtLeft:
fixes x :: name
and P :: "('a, 'b, 'c) psi"
and Ψ :: 'b
and Q :: "('a, 'b, 'c) psi"
and Rel :: "('b × ('a, 'b, 'c) psi × ('a, 'b, 'c) psi) set"
assumes "x ♯ P"
and "x ♯ Ψ"
and "eqvt Rel"
and C1: "⋀Ψ' R. (Ψ', R, R) ∈ Rel"
and C2: "⋀y Ψ' R S zvec. ⟦y ♯ Ψ'; y ♯ R; zvec ♯* Ψ'⟧ ⟹ (Ψ', ⦇νy⦈(⦇ν*zvec⦈(R ∥ S)), ⦇ν*zvec⦈(R ∥ ⦇νy⦈S)) ∈ Rel"
and C3: "⋀Ψ' zvec R y. ⟦y ♯ Ψ'; zvec ♯* Ψ'⟧ ⟹ (Ψ', ⦇νy⦈(⦇ν*zvec⦈R), ⦇ν*zvec⦈(⦇νy⦈R)) ∈ Rel"
and C4: "⋀Ψ' R S zvec. ⟦zvec ♯* R; zvec ♯* Ψ'⟧ ⟹ (Ψ', (⦇ν*zvec⦈(R ∥ S)), (R ∥ (⦇ν*zvec⦈S))) ∈ Rel"
shows "Ψ ⊳ ⦇νx⦈(P ∥ Q) ↝[Rel] P ∥ ⦇νx⦈Q"
proof -
note ‹eqvt Rel› ‹x ♯ Ψ›
moreover have "x ♯ ⦇νx⦈(P ∥ Q)" by(simp add: abs_fresh)
moreover from ‹x ♯ P› have "x ♯ P ∥ ⦇νx⦈Q" by(simp add: abs_fresh)
ultimately show ?thesis
proof(induct rule: simIFresh[of _ _ _ _ _ x])
case(cSim α PQ)
from ‹x ♯ α› ‹bn α ♯* (P ∥ ⦇νx⦈Q)› have "bn α ♯* Q" by simp
note ‹Ψ ⊳ P ∥ ⦇νx⦈Q ⟼α ≺ PQ› ‹bn α ♯* Ψ›
moreover from ‹bn α ♯* (P ∥ ⦇νx⦈Q)› have "bn α ♯* P" and "bn α ♯* ⦇νx⦈Q" by simp+
ultimately show ?case using ‹bn α ♯* subject α› ‹x ♯ PQ›
proof(induct rule: parCases[where C=x])
case(cPar1 P' A⇩Q Ψ⇩Q)
from ‹x ♯ P' ∥ ⦇νx⦈Q› have "x ♯ P'" by simp
have PTrans: "Ψ ⊗ Ψ⇩Q ⊳ P ⟼α ≺ P'" by fact
from ‹extractFrame(⦇νx⦈Q) = ⟨A⇩Q, Ψ⇩Q⟩› have FrxQ: "⦇νx⦈(extractFrame Q) = ⟨A⇩Q, Ψ⇩Q⟩" by simp
then obtain y A⇩Q' where A: "A⇩Q = y#A⇩Q'" by(cases A⇩Q) auto
with ‹A⇩Q ♯* Ψ› ‹A⇩Q ♯* P› ‹A⇩Q ♯* α›
have "A⇩Q' ♯* Ψ" and "A⇩Q' ♯* P" and "A⇩Q' ♯* α"
and "y ♯ Ψ" and "y ♯ P" and "y ♯ α"
by simp+
from PTrans ‹y ♯ P› ‹y ♯ α› ‹bn α ♯* subject α› ‹distinct(bn α)› have "y ♯ P'"
by(auto intro: freeFreshDerivative)
note PTrans
moreover from A ‹A⇩Q ♯* x› FrxQ have "extractFrame([(y, x)] ∙ Q) = ⟨A⇩Q', Ψ⇩Q⟩"
by(simp add: frame.inject alpha' fresh_list_cons eqvts)
moreover from ‹bn α ♯* Q› have "([(y, x)] ∙ (bn α)) ♯* ([(y, x)] ∙ Q)"
by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
with ‹x ♯ α› ‹A⇩Q ♯* α› A have "bn α ♯* ([(y, x)] ∙ Q)" by simp
ultimately have "Ψ ⊳ P ∥ ([(y, x)] ∙ Q) ⟼α ≺ (P' ∥ ([(y, x)] ∙ Q))"
using ‹A⇩Q' ♯* Ψ› ‹A⇩Q' ♯* P› ‹A⇩Q' ♯* α›
by(rule Par1)
then have "Ψ ⊳ ⦇νy⦈(P ∥ ([(y, x)] ∙ Q)) ⟼α ≺ ⦇νy⦈(P' ∥ ([(y, x)] ∙ Q))"
using ‹y ♯ Ψ› ‹y ♯ α›
by(rule Scope)
then have "([(y, x)] ∙ Ψ) ⊳ ([(y, x)] ∙ (⦇νy⦈(P ∥ ([(y, x)] ∙ Q)))) ⟼([(y, x)] ∙ (α ≺ ⦇νy⦈(P' ∥ ([(y, x)] ∙ Q))))"
by(rule semantics.eqvt)
with ‹x ♯ Ψ› ‹y ♯ Ψ› ‹x ♯ P› ‹y ♯ P› ‹x ♯ α› ‹y ♯ α› ‹x ♯ P'› ‹y ♯ P'›
have "Ψ ⊳ ⦇νx⦈(P ∥ Q) ⟼α ≺ ⦇νx⦈(P' ∥ Q)"
by(simp add: eqvts calc_atm)
moreover from ‹x ♯ Ψ› ‹x ♯ P'› have "(Ψ, ⦇νx⦈(⦇ν*[]⦈(P' ∥ Q)), ⦇ν*[]⦈(P' ∥ ⦇νx⦈Q)) ∈ Rel"
by(rule C2) auto
ultimately show ?case
by force
next
case(cPar2 xQ' A⇩P Ψ⇩P)
from ‹A⇩P ♯* ⦇νx⦈Q› ‹A⇩P ♯* x› have "A⇩P ♯* Q" by simp
note ‹Ψ ⊗ Ψ⇩P ⊳ ⦇νx⦈Q ⟼α ≺ xQ'›
moreover have FrP: "extractFrame P = ⟨A⇩P, Ψ⇩P⟩" by fact
with ‹x ♯ P› ‹A⇩P ♯* x› have "x ♯ Ψ⇩P" and "x ♯ A⇩P"
by(force dest: extractFrameFresh)+
with ‹x ♯ Ψ› have "x ♯ Ψ ⊗ Ψ⇩P" by force
moreover note ‹x ♯ α›
moreover from ‹x ♯ P ∥ xQ'› have "x ♯ xQ'" by simp
moreover from FrP ‹bn α ♯* P› ‹A⇩P ♯* α› have "bn α ♯* Ψ⇩P"
by(auto dest: extractFrameFreshChain)
with ‹bn α ♯* Ψ› have "bn α ♯* (Ψ ⊗ Ψ⇩P)" by force
ultimately show ?case using ‹bn α ♯* Q› ‹bn α ♯* subject α› ‹x ♯ α› ‹bn α ♯* P› ‹A⇩P ♯* α› ‹bn α ♯* Ψ› ‹x ♯ P›
proof(induct rule: resCases'[where C="(P, A⇩P, Ψ)"])
case(cOpen M xvec1 xvec2 y N Q')
from ‹x ♯ M⦇ν*(xvec1@y#xvec2)⦈⟨N⟩› have "x ♯ xvec1" and "x ≠ y" and "x ♯ xvec2" by simp+
from ‹bn(M⦇ν*(xvec1@y#xvec2)⦈⟨N⟩) ♯* Ψ› have "y ♯ Ψ" by simp
note ‹Ψ ⊗ Ψ⇩P ⊳ ([(x, y)] ∙ Q) ⟼M⦇ν*(xvec1@xvec2)⦈⟨N⟩ ≺ Q'› FrP
moreover from ‹bn(M⦇ν*(xvec1@y#xvec2)⦈⟨N⟩) ♯* P› have "(xvec1@xvec2) ♯* P" and "y ♯ P" by simp+
moreover from ‹A⇩P ♯* (M⦇ν*(xvec1@y#xvec2)⦈⟨N⟩)› have "A⇩P ♯* (M⦇ν*(xvec1@xvec2)⦈⟨N⟩)" and "y ♯ A⇩P" by simp+
moreover from ‹A⇩P ♯* Q› ‹x ♯ A⇩P› ‹y ♯ A⇩P› have "A⇩P ♯* ([(x, y)] ∙ Q)" by simp
ultimately have "Ψ ⊳ P ∥ ([(x, y)] ∙ Q) ⟼M⦇ν*(xvec1@xvec2)⦈⟨N⟩ ≺ (P ∥ Q')"
using ‹A⇩P ♯* Ψ›
by(intro Par2) (assumption | simp)+
then have "Ψ ⊳ ⦇νy⦈(P ∥ ([(x, y)] ∙ Q)) ⟼M⦇ν*(xvec1@y#xvec2)⦈⟨N⟩ ≺ (P ∥ Q')"
using ‹y ∈ supp N› ‹y ♯ Ψ› ‹y ♯ M› ‹y ♯ xvec1› ‹y ♯ xvec2›
by(rule Open)
with ‹x ♯ P› ‹y ♯ P› ‹y ♯ Q› have "Ψ ⊳ ⦇νx⦈(P ∥ Q) ⟼M⦇ν*(xvec1@y#xvec2)⦈⟨N⟩ ≺ (P ∥ Q')"
by(subst alphaRes[where y=y]) (simp add: fresh_left calc_atm eqvts)+
moreover have "(Ψ, P ∥ Q', P ∥ Q') ∈ Rel" by(rule C1)
ultimately show ?case by blast
next
case(cBrOpen M xvec1 xvec2 y N Q')
from ‹x ♯ ¡M⦇ν*(xvec1@y#xvec2)⦈⟨N⟩› have "x ♯ xvec1" and "x ≠ y" and "x ♯ xvec2" by simp+
from ‹bn(¡M⦇ν*(xvec1@y#xvec2)⦈⟨N⟩) ♯* Ψ› have "y ♯ Ψ" by simp
note ‹Ψ ⊗ Ψ⇩P ⊳ ([(x, y)] ∙ Q) ⟼¡M⦇ν*(xvec1@xvec2)⦈⟨N⟩ ≺ Q'› FrP
moreover from ‹bn(¡M⦇ν*(xvec1@y#xvec2)⦈⟨N⟩) ♯* P› have "(xvec1@xvec2) ♯* P" and "y ♯ P" by simp+
moreover from ‹A⇩P ♯* (¡M⦇ν*(xvec1@y#xvec2)⦈⟨N⟩)› have "A⇩P ♯* (¡M⦇ν*(xvec1@xvec2)⦈⟨N⟩)" and "y ♯ A⇩P" by simp+
moreover from ‹A⇩P ♯* Q› ‹x ♯ A⇩P› ‹y ♯ A⇩P› have "A⇩P ♯* ([(x, y)] ∙ Q)" by simp
ultimately have "Ψ ⊳ P ∥ ([(x, y)] ∙ Q) ⟼¡M⦇ν*(xvec1@xvec2)⦈⟨N⟩ ≺ (P ∥ Q')"
using ‹A⇩P ♯* Ψ›
by(intro Par2) (assumption | simp)+
then have "Ψ ⊳ ⦇νy⦈(P ∥ ([(x, y)] ∙ Q)) ⟼¡M⦇ν*(xvec1@y#xvec2)⦈⟨N⟩ ≺ (P ∥ Q')"
using ‹y ∈ supp N› ‹y ♯ Ψ› ‹y ♯ M› ‹y ♯ xvec1› ‹y ♯ xvec2›
by(rule BrOpen)
with ‹x ♯ P› ‹y ♯ P› ‹y ♯ Q› have "Ψ ⊳ ⦇νx⦈(P ∥ Q) ⟼¡M⦇ν*(xvec1@y#xvec2)⦈⟨N⟩ ≺ (P ∥ Q')"
by(subst alphaRes[where y=y]) (simp add: fresh_left calc_atm eqvts)+
moreover have "(Ψ, P ∥ Q', P ∥ Q') ∈ Rel" by(rule C1)
ultimately show ?case by blast
next
case(cRes Q')
from ‹Ψ ⊗ Ψ⇩P ⊳ Q ⟼α ≺ Q'› FrP ‹bn α ♯* P›
have "Ψ ⊳ P ∥ Q ⟼α ≺ (P ∥ Q')" using ‹A⇩P ♯* Ψ› ‹A⇩P ♯* Q› ‹A⇩P ♯* α›
by(rule Par2)
then have "Ψ ⊳ ⦇νx⦈(P ∥ Q) ⟼α ≺ ⦇νx⦈(P ∥ Q')" using ‹x ♯ Ψ› ‹x ♯ α›
by(rule Scope)
moreover from ‹x ♯ Ψ› ‹x ♯ P› have "(Ψ, ⦇νx⦈(⦇ν*[]⦈(P ∥ Q')), ⦇ν*[]⦈(P ∥ ⦇νx⦈Q')) ∈ Rel"
by(rule C2) auto
ultimately show ?case
by force
next
case(cBrClose M xvec N Q')
from ‹xvec ♯* (P, A⇩P, Ψ)›
have "xvec ♯* P" and "A⇩P ♯* xvec" and "xvec ♯* Ψ"
by simp+
from ‹Ψ ⊗ Ψ⇩P ⊳ Q ⟼ ¡M⦇ν*xvec⦈⟨N⟩ ≺ Q'› ‹A⇩P ♯* Q›
‹xvec ♯* M›
have "A⇩P ♯* M"
by(simp add: brOutputFreshSubjectChain)
from ‹Ψ ⊗ Ψ⇩P ⊳ Q ⟼ ¡M⦇ν*xvec⦈⟨N⟩ ≺ Q'› ‹xvec ♯* M› ‹distinct xvec› ‹A⇩P ♯* Q› ‹A⇩P ♯* xvec›
have "A⇩P ♯* N" and "A⇩P ♯* Q'" by(simp add: broutputFreshChainDerivative)+
from ‹A⇩P ♯* M› ‹A⇩P ♯* xvec› ‹A⇩P ♯* N› have "A⇩P ♯* (¡M⦇ν*xvec⦈⟨N⟩)"
by simp
from ‹Ψ ⊗ Ψ⇩P ⊳ Q ⟼ ¡M⦇ν*xvec⦈⟨N⟩ ≺ Q'› ‹extractFrame P = ⟨A⇩P, Ψ⇩P⟩› ‹xvec ♯* P› ‹A⇩P ♯* Ψ› ‹A⇩P ♯* Q› ‹A⇩P ♯* (¡M⦇ν*xvec⦈⟨N⟩)›
have "Ψ ⊳ P ∥ Q ⟼ ¡M⦇ν*xvec⦈⟨N⟩ ≺ P ∥ Q'"
by(simp add: Par2)
from ‹x ♯ Ψ› ‹x ♯ P› ‹xvec ♯* P› ‹xvec ♯* Ψ›
have "(x#xvec) ♯* P" and "(x#xvec) ♯* Ψ" by simp+
then have "(Ψ, (⦇ν*(x#xvec)⦈(P ∥ Q')), P ∥ (⦇ν*(x#xvec)⦈Q')) ∈ Rel"
by(rule C4)
then have "(Ψ, ⦇νx⦈(⦇ν*xvec⦈(P ∥ Q')), P ∥ ⦇νx⦈(⦇ν*xvec⦈Q')) ∈ Rel"
by simp
moreover from ‹Ψ ⊳ P ∥ Q ⟼ ¡M⦇ν*xvec⦈⟨N⟩ ≺ P ∥ Q'› ‹x ∈ supp M› ‹x ♯ Ψ›
have "Ψ ⊳ ⦇νx⦈(P ∥ Q) ⟼ τ ≺ ⦇νx⦈(⦇ν*xvec⦈(P ∥ Q'))"
by(rule BrClose)
ultimately show ?case
by force
qed
next
case(cComm1 Ψ⇩Q M N P' A⇩P Ψ⇩P K xvec xQ' A⇩Q)
have QTrans: "Ψ ⊗ Ψ⇩P ⊳ ⦇νx⦈Q ⟼K⦇ν*xvec⦈⟨N⟩ ≺ xQ'" and FrQ: "extractFrame(⦇νx⦈Q) = ⟨A⇩Q, Ψ⇩Q⟩" by fact+
have PTrans: "Ψ ⊗ Ψ⇩Q ⊳ P ⟼M⦇N⦈ ≺ P'" and FrP: "extractFrame P = ⟨A⇩P, Ψ⇩P⟩" by fact+
have "x ♯ ⦇νx⦈Q" by(simp add: abs_fresh)
with QTrans have "x ♯ N" and "x ♯ xQ'" using ‹xvec ♯* x› ‹xvec ♯* K› ‹distinct xvec›
by(force intro: outputFreshDerivative)+
from PTrans ‹x ♯ P› ‹x ♯ N› have "x ♯ P'" by(rule inputFreshDerivative)
from ‹x ♯ ⦇νx⦈Q› FrQ ‹A⇩Q ♯* x› have "x ♯ Ψ⇩Q"
by(drule_tac extractFrameFresh) auto
from ‹x ♯ P› FrP ‹A⇩P ♯* x› have "x ♯ Ψ⇩P"
by(drule_tac extractFrameFresh) auto
from ‹A⇩P ♯* ⦇νx⦈Q› ‹A⇩P ♯* x› have "A⇩P ♯* Q" by simp
from ‹A⇩Q ♯* ⦇νx⦈Q› ‹A⇩Q ♯* x› have "A⇩Q ♯* Q" by simp
from PTrans FrP ‹distinct A⇩P› ‹x ♯ P› ‹A⇩Q ♯* P› ‹xvec ♯* P› ‹A⇩P ♯* Ψ› ‹A⇩P ♯* Ψ⇩Q› ‹A⇩P ♯* x› ‹A⇩P ♯* A⇩Q› ‹A⇩P ♯* P› ‹A⇩P ♯* M› ‹A⇩P ♯* xvec›
obtain M' where "(Ψ ⊗ Ψ⇩Q) ⊗ Ψ⇩P ⊢ M ↔ M'" and "x ♯ M'" and "A⇩Q ♯* M'" and "xvec ♯* M'"
by(elim inputObtainPrefix[where B="x#xvec@A⇩Q"]) (assumption | force simp add: fresh_star_list_cons)+
then have MeqM': "Ψ ⊗ Ψ⇩P ⊗ Ψ⇩Q ⊢ M ↔ M'" by(metis statEqEnt Associativity Commutativity Composition)
with ‹Ψ ⊗ Ψ⇩P ⊗ Ψ⇩Q ⊢ M ↔ K› have "Ψ ⊗ Ψ⇩P ⊗ Ψ⇩Q ⊢ K ↔ M'"
by(blast intro: chanEqTrans chanEqSym)
then have "(Ψ ⊗ Ψ⇩P) ⊗ Ψ⇩Q ⊢ K ↔ M'" by(metis statEqEnt Associativity Commutativity Composition)
with QTrans FrQ ‹distinct A⇩Q› ‹A⇩Q ♯* Ψ› ‹A⇩Q ♯* Ψ⇩P› ‹A⇩Q ♯* (⦇νx⦈Q)› ‹A⇩Q ♯* K› ‹A⇩Q ♯* M'›
have "Ψ ⊗ Ψ⇩P ⊳ ⦇νx⦈Q ⟼M'⦇ν*xvec⦈⟨N⟩ ≺ xQ'"
by(force intro: outputRenameSubject)
moreover from ‹x ♯ Ψ› ‹x ♯ Ψ⇩P› have "x ♯ Ψ ⊗ Ψ⇩P" by force
moreover from ‹xvec ♯* x› have "x ♯ xvec" by simp
with ‹x ♯ M'› ‹x ♯ N› have "x ♯ M'⦇ν*xvec⦈⟨N⟩" by simp
moreover note ‹x ♯ xQ'›
moreover from ‹xvec ♯* Ψ› ‹xvec ♯* Ψ⇩P› have "xvec ♯* (Ψ ⊗ Ψ⇩P)" by force
moreover from ‹xvec ♯* ⦇νx⦈Q› ‹x ♯ xvec› have "xvec ♯* Q" by simp
moreover note ‹xvec ♯* M'›
moreover from ‹xvec ♯* Ψ› ‹xvec ♯* Ψ⇩P› have "bn(M'⦇ν*xvec⦈⟨N⟩) ♯* (Ψ ⊗ Ψ⇩P)" by force
moreover from ‹xvec ♯* ⦇νx⦈Q› ‹x ♯ xvec› have "bn(M'⦇ν*xvec⦈⟨N⟩) ♯* Q" by simp
moreover from ‹xvec ♯* P› have "bn(M'⦇ν*xvec⦈⟨N⟩) ♯* P" by simp
from ‹xvec ♯* Ψ› have "bn(M'⦇ν*xvec⦈⟨N⟩) ♯* Ψ" by simp
from ‹A⇩Q ♯* xvec› ‹A⇩Q ♯* M'› ‹A⇩Q ♯* N› have "A⇩Q ♯* (M'⦇ν*xvec⦈⟨N⟩)" by simp
have "object(M'⦇ν*xvec⦈⟨N⟩) = Some N" by simp
have "bn(M'⦇ν*xvec⦈⟨N⟩) = xvec" by simp
have "subject(M'⦇ν*xvec⦈⟨N⟩) = Some M'" by simp
from ‹xvec ♯* M'› have "bn(M'⦇ν*xvec⦈⟨N⟩) ♯* subject(M'⦇ν*xvec⦈⟨N⟩)" by simp
ultimately show ?case
using ‹x ♯ M'⦇ν*xvec⦈⟨N⟩› ‹bn(M'⦇ν*xvec⦈⟨N⟩) ♯* P› ‹bn(M'⦇ν*xvec⦈⟨N⟩) ♯* Ψ› ‹object(M'⦇ν*xvec⦈⟨N⟩) = Some N›
‹bn(M'⦇ν*xvec⦈⟨N⟩) = xvec› ‹subject(M'⦇ν*xvec⦈⟨N⟩) = Some M'› ‹A⇩Q ♯* (M'⦇ν*xvec⦈⟨N⟩)›
proof(induct rule: resOutputCases''''')
case(cOpen M'' xvec1 xvec2 y N' Q')
then have Eq: "M'⦇ν*xvec⦈⟨N⟩ = M''⦇ν*(xvec1 @ y # xvec2)⦈⟨N'⟩" by simp
from ‹x ♯ M'⦇ν*xvec⦈⟨N⟩› Eq have "x ♯ xvec1" and "x ≠ y" and "x ♯ xvec2" and "x ♯ M''"
by simp+
from ‹bn(M'⦇ν*xvec⦈⟨N⟩) ♯* P› Eq have "(xvec1@xvec2) ♯* P" and "y ♯ P" by simp+
from ‹A⇩Q ♯* (M'⦇ν*xvec⦈⟨N⟩)› Eq have "(xvec1@xvec2) ♯* A⇩Q" and "y ♯ A⇩Q" and "A⇩Q ♯* M''" by simp+
from ‹bn(M'⦇ν*xvec⦈⟨N⟩) ♯* Ψ› Eq have "(xvec1@xvec2) ♯* Ψ" and "y ♯ Ψ" by simp+
from Eq have "N = N'" and "xvec = xvec1@y#xvec2" and "M' = M''" by(simp add: action.inject)+
from ‹x ♯ P› ‹y ♯ P› ‹x ≠ y› ‹Ψ ⊗ Ψ⇩Q ⊳ P ⟼M⦇N⦈ ≺ P'›
have "Ψ ⊗ Ψ⇩Q ⊳ P ⟼M⦇([(y, x)] ∙ N)⦈ ≺ ([(y, x)] ∙ P')"
by(elim inputAlpha[where xvec="[y]"]) (auto simp add: calc_atm)
then have PTrans: "Ψ ⊗ Ψ⇩Q ⊳ P ⟼M⦇([(x, y)] ∙ N)⦈ ≺ ([(x, y)] ∙ P')"
by(simp add: name_swap)
from ‹Ψ ⊗ Ψ⇩P ⊳ [(x, y)] ∙ Q ⟼ M''⦇ν*(xvec1 @ xvec2)⦈⟨N'⟩ ≺ Q'›
have "[(x, y)] ∙ (Ψ ⊗ Ψ⇩P ⊳ [(x, y)] ∙ Q ⟼ M''⦇ν*(xvec1 @ xvec2)⦈⟨N'⟩ ≺ Q')"
by simp
then have "[(x, y)] ∙ (Ψ ⊗ Ψ⇩P) ⊳ ([(x, y)] ∙ ([(x, y)] ∙ Q)) ⟼ ([(x, y)] ∙ M'')⦇ν*([(x, y)] ∙ (xvec1 @ xvec2))⦈⟨([(x, y)] ∙ N')⟩ ≺ [(x, y)] ∙ Q'"
by(simp add: eqvts)
with ‹x ♯ (Ψ ⊗ Ψ⇩P)› ‹y ♯ (Ψ ⊗ Ψ⇩P)› ‹x ♯ xvec1› ‹y ♯ xvec1› ‹x ♯ xvec2› ‹y ♯ xvec2› ‹x ♯ M''› ‹y ♯ M''›
have QTrans: "Ψ ⊗ Ψ⇩P ⊳ Q ⟼ M''⦇ν*(xvec1 @ xvec2)⦈⟨([(x, y)] ∙ N')⟩ ≺ [(x, y)] ∙ Q'"
by simp
with ‹A⇩Q ♯* x› ‹y ♯ A⇩Q› ‹distinct xvec1› ‹distinct xvec2› ‹xvec1 ♯* xvec2› ‹xvec1 ♯* M''› ‹xvec2 ♯* M''›
‹(xvec1@xvec2) ♯* A⇩Q›
have "A⇩Q ♯* ([(x, y)] ∙ Q')" using ‹A⇩Q ♯* Q›
by(elim outputFreshChainDerivative(2)) (assumption | simp)+
from ‹extractFrame(⦇νx⦈Q) = ⟨A⇩Q, Ψ⇩Q⟩› have FrxQ: "⦇νx⦈(extractFrame Q) = ⟨A⇩Q, Ψ⇩Q⟩" by simp
then obtain z A⇩Q' where A: "A⇩Q = z#A⇩Q'" by(cases A⇩Q) auto
with ‹A⇩Q ♯* Ψ› ‹A⇩Q ♯* P› ‹A⇩Q ♯* P'› ‹A⇩Q ♯* Ψ⇩P› ‹A⇩Q ♯* Q› ‹(xvec1@xvec2) ♯* A⇩Q› ‹A⇩Q ♯* M''› ‹A⇩Q ♯* ([(x, y)] ∙ Q')› ‹y ♯ A⇩Q› ‹A⇩Q ♯* N›
have "A⇩Q' ♯* Ψ" and "A⇩Q' ♯* P" and "A⇩Q' ♯* Ψ⇩P" and "A⇩Q' ♯* Q"
and "z ♯ Ψ" and "z ♯ P" and "z ♯ P'" and "z ♯ Ψ⇩P" and "z ♯ Q" and "z ♯ xvec1" and "z ♯ xvec2"
and "z ♯ M''" and "z ♯ ([(x, y)] ∙ Q')" and "A⇩Q' ♯* M''" and "z ≠ y" and "z ♯ (xvec1@xvec2)"
by auto
from A ‹A⇩P ♯* A⇩Q› have "A⇩P ♯* A⇩Q'" and "z ♯ A⇩P" by simp+
from A ‹A⇩Q ♯* x› have "x ≠ z" and "x ♯ A⇩Q'" by simp+
from ‹distinct A⇩Q› A have "z ♯ A⇩Q'"
by(induct A⇩Q') (auto simp add: fresh_list_nil fresh_list_cons)
from PTrans ‹x ♯ P› ‹z ♯ P› ‹x ≠ z› have "Ψ ⊗ Ψ⇩Q ⊳ P ⟼M⦇([(x, z)] ∙ [(x, y)] ∙ N)⦈ ≺ ([(x, z)] ∙ [(x, y)] ∙ P')"
by(elim inputAlpha[where xvec="[x]"]) (auto simp add: calc_atm)
moreover note FrP
moreover from QTrans have "([(x, z)] ∙ (Ψ ⊗ Ψ⇩P)) ⊳ ([(x, z)] ∙ Q) ⟼([(x, z)] ∙ (M''⦇ν*(xvec1@xvec2)⦈⟨([(x, y)] ∙ N')⟩ ≺ ([(x, y)] ∙ Q')))"
by(rule semantics.eqvt)
with ‹x ♯ Ψ› ‹z ♯ Ψ› ‹x ♯ Ψ⇩P› ‹z ♯ Ψ⇩P› ‹x ♯ M''› ‹z ♯ M''› ‹x ♯ xvec1› ‹x ♯ xvec2› ‹z ♯ xvec1› ‹z ♯ xvec2›
have "Ψ ⊗ Ψ⇩P ⊳ ([(x, z)] ∙ Q) ⟼M''⦇ν*(xvec1@xvec2)⦈⟨([(x, z)] ∙ [(x, y)] ∙ N')⟩ ≺ ([(x, z)] ∙ [(x, y)] ∙ Q')"
by(simp add: eqvts)
moreover from A ‹A⇩Q ♯* x› FrxQ have "extractFrame([(x, z)] ∙ Q) = ⟨A⇩Q', Ψ⇩Q⟩"
by(clarsimp simp add: alpha' eqvts frame.inject fresh_list_cons name_swap)
moreover from ‹A⇩P ♯* Q› have "([(x, z)] ∙ A⇩P) ♯* ([(x, z)] ∙ Q)"
by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
with ‹A⇩P ♯* x› ‹z ♯ A⇩P› have "A⇩P ♯* ([(x, z)] ∙ Q)" by simp
moreover from ‹A⇩Q' ♯* Q› have "([(x, z)] ∙ A⇩Q') ♯* ([(x, z)] ∙ Q)"
by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
with ‹x ♯ A⇩Q'› ‹z ♯ A⇩Q'› have "A⇩Q' ♯* ([(x, z)] ∙ Q)" by simp
ultimately have "Ψ ⊳ (P ∥ ([(x, z)] ∙ Q)) ⟼τ ≺ ⦇ν*(xvec1@xvec2)⦈(([(x, z)] ∙ [(x, y)] ∙ P') ∥ ([(x, z)] ∙ [(x, y)] ∙ Q'))"
using MeqM' ‹M'=M''› ‹N=N'› ‹A⇩P ♯* Ψ› ‹A⇩P ♯* P› ‹A⇩P ♯* A⇩Q'› ‹A⇩Q' ♯* Ψ› ‹A⇩Q' ♯* P› ‹(xvec1@xvec2) ♯* P› ‹A⇩P ♯* M› ‹A⇩Q' ♯* M''›
by(intro Comm1) (assumption | simp)+
with‹z ♯ Ψ› have "Ψ ⊳ ⦇νz⦈(P ∥ ([(x, z)] ∙ Q)) ⟼τ ≺ ⦇νz⦈(⦇ν*(xvec1@xvec2)⦈(([(x, z)] ∙ [(x, y)] ∙ P') ∥ ([(x, z)] ∙ [(x, y)] ∙ Q')))"
by(intro Scope) auto
moreover from ‹x ♯ P› ‹z ♯ P› ‹z ♯ Q› have "⦇νz⦈(P ∥ ([(x, z)] ∙ Q)) = ⦇νx⦈([(x, z)] ∙ (P ∥ ([(x, z)] ∙ Q)))"
by(subst alphaRes[of x]) (auto simp add: calc_atm fresh_left name_swap)
with ‹x ♯ P› ‹z ♯ P› have "⦇νz⦈(P ∥ ([(x, z)] ∙ Q)) = ⦇νx⦈(P ∥ Q)"
by(simp add: eqvts)
moreover from ‹z ≠ y› ‹x ≠ z› ‹z ♯ P'› ‹z ♯ [(x, y)] ∙ Q'› have "⦇νz⦈(⦇ν*(xvec1@xvec2)⦈(([(x, z)] ∙ [(x, y)] ∙ P') ∥ ([(x, z)] ∙ [(x, y)] ∙ Q'))) = ⦇νx⦈([(x, z)] ∙ (⦇ν*(xvec1@xvec2)⦈(([(x, z)] ∙ [(x, y)] ∙ P') ∥ ([(x, z)] ∙ [(x, y)] ∙ Q'))))"
by(subst alphaRes[of x]) (auto simp add: resChainFresh fresh_left calc_atm name_swap)
with ‹x ♯ xvec1› ‹x ♯ xvec2› ‹z ♯ xvec1› ‹z ♯ xvec2› have "⦇νz⦈(⦇ν*(xvec1@xvec2)⦈(([(x, z)] ∙ [(x, y)] ∙ P') ∥ ([(x, z)] ∙ [(x, y)] ∙ Q'))) = ⦇νx⦈(⦇ν*(xvec1@xvec2)⦈(([(x, y)] ∙ P') ∥ ([(x, y)] ∙ Q')))"
by(simp add: eqvts)
moreover from ‹x ♯ P'› ‹x ♯ Q'› ‹x ♯ xvec1› ‹x ♯ xvec2› ‹y ♯ xvec1› ‹y ♯ xvec2›
have "⦇νx⦈(⦇ν*(xvec1@xvec2)⦈(([(x, y)] ∙ P') ∥ ([(x, y)] ∙ Q'))) = ⦇νy⦈(⦇ν*(xvec1@xvec2)⦈(P' ∥ Q'))"
by(subst alphaRes[of y]) (auto simp add: resChainFresh calc_atm eqvts fresh_left name_swap)
ultimately have "Ψ ⊳ ⦇νx⦈(P ∥ Q) ⟼τ ≺ ⦇νy⦈(⦇ν*(xvec1@xvec2)⦈(P' ∥ Q'))"
by simp
moreover from ‹y ♯ Ψ› ‹(xvec1@xvec2) ♯* Ψ› ‹xvec=xvec1@y#xvec2›
have "(Ψ, ⦇νy⦈(⦇ν*(xvec1@xvec2)⦈(P' ∥ Q')), ⦇ν*xvec⦈(P' ∥ Q')) ∈ Rel"
by(force intro: C3 simp add: resChainAppend)
ultimately show ?case by blast
next
case(cRes Q')
have QTrans: "Ψ ⊗ Ψ⇩P ⊳ Q ⟼M'⦇ν*xvec⦈⟨N⟩ ≺ Q'" by fact
with ‹A⇩Q ♯* Q› ‹A⇩Q ♯* xvec› ‹xvec ♯* M'› ‹distinct xvec› have "A⇩Q ♯* Q'"
by(force dest: outputFreshChainDerivative)
with ‹extractFrame(⦇νx⦈Q) = ⟨A⇩Q, Ψ⇩Q⟩› have FrxQ: "⦇νx⦈(extractFrame Q) = ⟨A⇩Q, Ψ⇩Q⟩" by simp
then obtain y A⇩Q' where A: "A⇩Q = y#A⇩Q'" by(cases A⇩Q) auto
with ‹A⇩Q ♯* Ψ› ‹A⇩Q ♯* P› ‹A⇩Q ♯* P'› ‹A⇩Q ♯* Ψ⇩P› ‹A⇩Q ♯* Q› ‹A⇩Q ♯* xvec› ‹A⇩Q ♯* M'› ‹A⇩Q ♯* Q'›
have "A⇩Q' ♯* Ψ" and "A⇩Q' ♯* P" and "A⇩Q' ♯* Ψ⇩P" and "A⇩Q' ♯* Q" and "A⇩Q ♯* xvec" and "A⇩Q ♯* Q'"
and "y ♯ Ψ" and "y ♯ P" and "y ♯ P'" and "y ♯ Ψ⇩P" and "y ♯ Q" and "y ♯ xvec" and "y ♯ M'" and "y ♯ Q'"
and "A⇩Q' ♯* M'"
by(simp)+
from A ‹A⇩P ♯* A⇩Q› have "A⇩P ♯* A⇩Q'" and "y ♯ A⇩P" by(simp add: fresh_star_list_cons)+
from A ‹A⇩Q ♯* x› have "x ≠ y" and "x ♯ A⇩Q'" by(simp add: fresh_list_cons)+
with A ‹distinct A⇩Q› have "y ♯ A⇩Q'"
by(induct A⇩Q') (auto simp add: fresh_list_nil fresh_list_cons)
from ‹x ♯ P› ‹y ♯ P› ‹x ≠ y› ‹Ψ ⊗ Ψ⇩Q ⊳ P ⟼M⦇N⦈ ≺ P'›
have "Ψ ⊗ Ψ⇩Q ⊳ P ⟼M⦇([(y, x)] ∙ N)⦈ ≺ [(y, x)] ∙ P'"
by(intro inputAlpha[where xvec="[y]"]) (auto simp add: calc_atm)
then have "Ψ ⊗ Ψ⇩Q ⊳ P ⟼M⦇([(x, y)] ∙ N)⦈ ≺ [(x, y)] ∙ P'"
by(simp add: name_swap)
moreover note FrP
moreover from ‹Ψ ⊗ Ψ⇩P ⊳ Q ⟼M'⦇ν*xvec⦈⟨N⟩ ≺ Q'› have "([(x, y)] ∙ (Ψ ⊗ Ψ⇩P)) ⊳ ([(x, y)] ∙ Q) ⟼([(x, y)] ∙ (M'⦇ν*xvec⦈⟨N⟩ ≺ Q'))"
by(rule semantics.eqvt)
with ‹x ♯ Ψ› ‹y ♯ Ψ› ‹x ♯ Ψ⇩P› ‹y ♯ Ψ⇩P› ‹x ♯ M'› ‹y ♯ M'› ‹x ♯ xvec› ‹y ♯ xvec›
have "Ψ ⊗ Ψ⇩P ⊳ ([(x, y)] ∙ Q) ⟼M'⦇ν*xvec⦈⟨([(x, y)] ∙ N)⟩ ≺ ([(x, y)] ∙ Q')"
by(simp add: eqvts)
moreover from A ‹A⇩Q ♯* x› FrxQ have FrQ: "extractFrame([(x, y)] ∙ Q) = ⟨A⇩Q', Ψ⇩Q⟩"
by(clarsimp simp add: alpha' eqvts frame.inject fresh_list_cons name_swap)
moreover from ‹A⇩P ♯* Q› have "([(x, y)] ∙ A⇩P) ♯* ([(x, y)] ∙ Q)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
with ‹A⇩P ♯* x› ‹y ♯ A⇩P› have "A⇩P ♯* ([(x, y)] ∙ Q)" by simp
moreover from ‹A⇩Q' ♯* Q› have "([(x, y)] ∙ A⇩Q') ♯* ([(x, y)] ∙ Q)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
with ‹x ♯ A⇩Q'› ‹y ♯ A⇩Q'› have "A⇩Q' ♯* ([(x, y)] ∙ Q)" by simp
ultimately have "Ψ ⊳ (P ∥ ([(x, y)] ∙ Q)) ⟼τ ≺ ⦇ν*xvec⦈([(x, y)] ∙ P') ∥ ([(x, y)] ∙ Q')"
using MeqM' ‹A⇩P ♯* Ψ› ‹A⇩P ♯* P› ‹A⇩P ♯* A⇩Q'› ‹A⇩Q' ♯* Ψ› ‹A⇩Q' ♯* P› ‹xvec ♯* P› ‹A⇩P ♯* M› ‹A⇩Q' ♯* M'›
by(intro Comm1) (assumption | simp)+
with‹y ♯ Ψ› have "Ψ ⊳ ⦇νy⦈(P ∥ ([(x, y)] ∙ Q)) ⟼τ ≺ ⦇νy⦈(⦇ν*xvec⦈(([(x, y)] ∙ P') ∥ ([(x, y)] ∙ Q')))"
by(intro Scope) auto
moreover from ‹x ♯ P› ‹y ♯ P› ‹y ♯ Q› have "⦇νy⦈(P ∥ ([(x, y)] ∙ Q)) = ⦇νx⦈([(x, y)] ∙ (P ∥ ([(x, y)] ∙ Q)))"
by(subst alphaRes[of x]) (auto simp add: calc_atm fresh_left name_swap)
with ‹x ♯ P› ‹y ♯ P› have "⦇νy⦈(P ∥ ([(x, y)] ∙ Q)) = ⦇νx⦈(P ∥ Q)"
by(simp add: eqvts)
moreover from ‹y ♯ P'› ‹y ♯ Q'› ‹x ♯ xvec› ‹y ♯ xvec› have "⦇νy⦈(⦇ν*xvec⦈(([(x, y)] ∙ P') ∥ ([(x, y)] ∙ Q'))) = ⦇νx⦈(⦇ν*xvec⦈(P' ∥ Q'))"
by(subst alphaRes[of y]) (auto simp add: resChainFresh calc_atm eqvts fresh_left name_swap)
ultimately have "Ψ ⊳ ⦇νx⦈(P ∥ Q) ⟼τ ≺ ⦇νx⦈(⦇ν*xvec⦈(P' ∥ Q'))"
by simp
moreover from ‹x ♯ Ψ› ‹x ♯ P'› ‹xvec ♯* Ψ› have "(Ψ, ⦇νx⦈(⦇ν*xvec⦈(P' ∥ Q')), ⦇ν*xvec⦈(P' ∥ ⦇νx⦈Q')) ∈ Rel"
by(rule C2)
ultimately show ?case by blast
qed
next
case(cComm2 Ψ⇩Q M xvec N P' A⇩P Ψ⇩P K xQ' A⇩Q)
have QTrans: "Ψ ⊗ Ψ⇩P ⊳ ⦇νx⦈Q ⟼K⦇N⦈ ≺ xQ'" and FrQ: "extractFrame(⦇νx⦈Q) = ⟨A⇩Q, Ψ⇩Q⟩" by fact+
have PTrans: "Ψ ⊗ Ψ⇩Q ⊳ P ⟼M⦇ν*xvec⦈⟨N⟩ ≺ P'" and FrP: "extractFrame P = ⟨A⇩P, Ψ⇩P⟩" by fact+
from PTrans ‹x ♯ P› have "x ♯ N" and "x ♯ P'" using ‹xvec ♯* x› ‹xvec ♯* M› ‹distinct xvec›
by(force intro: outputFreshDerivative)+
have "x ♯ ⦇νx⦈Q" by(simp add: abs_fresh)
with FrQ ‹A⇩Q ♯* x› have "x ♯ Ψ⇩Q"
by(drule_tac extractFrameFresh) auto
from ‹x ♯ P› FrP ‹A⇩P ♯* x› have "x ♯ Ψ⇩P"
by(drule_tac extractFrameFresh) auto
from ‹A⇩P ♯* ⦇νx⦈Q› ‹A⇩P ♯* x› have "A⇩P ♯* Q" by simp
from ‹A⇩Q ♯* ⦇νx⦈Q› ‹A⇩Q ♯* x› have "A⇩Q ♯* Q" by simp
from ‹xvec ♯* x› ‹xvec ♯* ⦇νx⦈Q› have "xvec ♯* Q" by simp
from ‹Ψ ⊗ Ψ⇩Q ⊳ P ⟼ M⦇ν*xvec⦈⟨N⟩ ≺ P'› have PResTrans: "Ψ ⊗ Ψ⇩Q ⊳ P ⟼ ROut M (⦇ν*xvec⦈N ≺' P')"
by(simp add: residualInject)
from PResTrans FrP ‹distinct A⇩P› ‹x ♯ P› ‹A⇩Q ♯* P› ‹A⇩P ♯* Ψ› ‹A⇩P ♯* Ψ⇩Q› ‹A⇩P ♯* x› ‹A⇩P ♯* A⇩Q› ‹A⇩P ♯* P› ‹A⇩P ♯* M› ‹xvec ♯* M› ‹distinct xvec›
obtain M' where "(Ψ ⊗ Ψ⇩Q) ⊗ Ψ⇩P ⊢ M ↔ M'" and "x ♯ M'" and "A⇩Q ♯* M'"
by(elim outputObtainPrefix[where B="x#A⇩Q"]) (assumption | force simp add: fresh_star_list_cons)+
then have MeqM': "Ψ ⊗ Ψ⇩P ⊗ Ψ⇩Q ⊢ M ↔ M'"
by(metis statEqEnt Associativity Commutativity Composition)
with ‹Ψ ⊗ Ψ⇩P ⊗ Ψ⇩Q ⊢ M ↔ K› have "Ψ ⊗ Ψ⇩P ⊗ Ψ⇩Q ⊢ K ↔ M'"
by(blast intro: chanEqTrans chanEqSym)
then have "(Ψ ⊗ Ψ⇩P) ⊗ Ψ⇩Q ⊢ K ↔ M'"
by(metis statEqEnt Associativity Commutativity Composition)
with QTrans FrQ ‹distinct A⇩Q› ‹A⇩Q ♯* Ψ› ‹A⇩Q ♯* Ψ⇩P› ‹A⇩Q ♯* (⦇νx⦈Q)› ‹A⇩Q ♯* K› ‹A⇩Q ♯* M'›
have "Ψ ⊗ Ψ⇩P ⊳ ⦇νx⦈Q ⟼M'⦇N⦈ ≺ xQ'" by(force intro: inputRenameSubject)
moreover from ‹x ♯ Ψ› ‹x ♯ Ψ⇩P› have "x ♯ Ψ ⊗ Ψ⇩P" by force
moreover note ‹x ♯ M'› ‹x ♯ N›
moreover from QTrans ‹x ♯ N› have "x ♯ xQ'" by(force dest: inputFreshDerivative simp add: abs_fresh)
ultimately show ?case
proof(induct rule: resInputCases)
case(cRes Q')
have QTrans: "Ψ ⊗ Ψ⇩P ⊳ Q ⟼M'⦇N⦈ ≺ Q'" by fact
with ‹A⇩Q ♯* Q› ‹A⇩Q ♯* N› have "A⇩Q ♯* Q'"
by(elim inputFreshChainDerivative)
with ‹extractFrame(⦇νx⦈Q) = ⟨A⇩Q, Ψ⇩Q⟩› have FrxQ: "⦇νx⦈(extractFrame Q) = ⟨A⇩Q, Ψ⇩Q⟩" by simp
then obtain y A⇩Q' where A: "A⇩Q = y#A⇩Q'" by(cases A⇩Q) auto
with ‹A⇩Q ♯* Ψ› ‹A⇩Q ♯* P› ‹A⇩Q ♯* P'› ‹A⇩Q ♯* Ψ⇩P› ‹A⇩Q ♯* Q› ‹A⇩Q ♯* xvec› ‹A⇩Q ♯* M'› ‹A⇩Q ♯* Q'› ‹A⇩Q ♯* N›
have "A⇩Q' ♯* Ψ" and "A⇩Q' ♯* P" and "A⇩Q' ♯* Ψ⇩P" and "A⇩Q' ♯* Q" and "A⇩Q ♯* xvec" and "A⇩Q ♯* Q'"
and "y ♯ Ψ" and "y ♯ P" and "y ♯ P'" and "y ♯ Ψ⇩P" and "y ♯ Q" and "y ♯ xvec" and "y ♯ M'" and "y ♯ Q'" and "y ♯ N"
and "A⇩Q' ♯* M'"
by(simp)+
from A ‹A⇩P ♯* A⇩Q› have "A⇩P ♯* A⇩Q'" and "y ♯ A⇩P" by(simp add: fresh_star_list_cons)+
from A ‹A⇩Q ♯* x› have "x ≠ y" and "x ♯ A⇩Q'" by(simp add: fresh_list_cons)+
with A ‹distinct A⇩Q› have "y ♯ A⇩Q'"
by(induct A⇩Q') (auto simp add: fresh_list_nil fresh_list_cons)
note PTrans FrP
moreover from ‹Ψ ⊗ Ψ⇩P ⊳ Q ⟼M'⦇N⦈ ≺ Q'› have "([(x, y)] ∙ (Ψ ⊗ Ψ⇩P)) ⊳ ([(x, y)] ∙ Q) ⟼([(x, y)] ∙ (M'⦇N⦈ ≺ Q'))"
by(rule semantics.eqvt)
with ‹x ♯ Ψ› ‹y ♯ Ψ› ‹x ♯ Ψ⇩P› ‹y ♯ Ψ⇩P› ‹x ♯ M'› ‹y ♯ M'› ‹x ♯ N› ‹y ♯ N›
have "Ψ ⊗ Ψ⇩P ⊳ ([(x, y)] ∙ Q) ⟼M'⦇N⦈ ≺ ([(x, y)] ∙ Q')"
by(simp add: eqvts)
moreover from A ‹A⇩Q ♯* x› FrxQ have FrQ: "extractFrame([(x, y)] ∙ Q) = ⟨A⇩Q', Ψ⇩Q⟩" and "y ♯ extractFrame Q"
by(clarsimp simp add: alpha' eqvts frame.inject fresh_list_cons name_swap)+
moreover from ‹A⇩P ♯* Q› have "([(x, y)] ∙ A⇩P) ♯* ([(x, y)] ∙ Q)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
with ‹A⇩P ♯* x› ‹y ♯ A⇩P› have "A⇩P ♯* ([(x, y)] ∙ Q)" by simp
moreover from ‹A⇩Q' ♯* Q› have "([(x, y)] ∙ A⇩Q') ♯* ([(x, y)] ∙ Q)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
with ‹x ♯ A⇩Q'› ‹y ♯ A⇩Q'› have "A⇩Q' ♯* ([(x, y)] ∙ Q)" by simp
moreover from ‹xvec ♯* Q› have "([(x, y)] ∙ xvec) ♯* ([(x, y)] ∙ Q)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
with ‹xvec ♯* x› ‹y ♯ xvec› have "xvec ♯* ([(x, y)] ∙ Q)" by simp
ultimately have "Ψ ⊳ (P ∥ ([(x, y)] ∙ Q)) ⟼τ ≺ ⦇ν*xvec⦈(P' ∥ ([(x, y)] ∙ Q'))"
using MeqM' ‹A⇩P ♯* Ψ› ‹A⇩P ♯* P› ‹A⇩P ♯* A⇩Q'› ‹A⇩Q' ♯* Ψ› ‹A⇩Q' ♯* P› ‹A⇩P ♯* M› ‹A⇩Q' ♯* M'›
by(intro Comm2) (assumption | simp)+
with‹y ♯ Ψ› have "Ψ ⊳ ⦇νy⦈(P ∥ ([(x, y)] ∙ Q)) ⟼τ ≺ ⦇νy⦈(⦇ν*xvec⦈(P' ∥ ([(x, y)] ∙ Q')))"
by(intro Scope) auto
moreover from ‹x ♯ P› ‹y ♯ P› ‹y ♯ Q› have "⦇νy⦈(P ∥ ([(x, y)] ∙ Q)) = ⦇νx⦈([(x, y)] ∙ (P ∥ ([(x, y)] ∙ Q)))"
by(subst alphaRes[of x]) (auto simp add: calc_atm fresh_left name_swap)
with ‹x ♯ P› ‹y ♯ P› have "⦇νy⦈(P ∥ ([(x, y)] ∙ Q)) = ⦇νx⦈(P ∥ Q)"
by(simp add: eqvts)
moreover from ‹x ♯ P'› ‹y ♯ P'› ‹y ♯ Q'› ‹xvec ♯* x› ‹y ♯ xvec› have "⦇νy⦈(⦇ν*xvec⦈(P' ∥ ([(x, y)] ∙ Q'))) = ⦇νx⦈(⦇ν*xvec⦈(P' ∥ Q'))"
by(subst alphaRes[of y]) (auto simp add: resChainFresh calc_atm eqvts fresh_left name_swap)
ultimately have "Ψ ⊳ ⦇νx⦈(P ∥ Q) ⟼τ ≺ ⦇νx⦈(⦇ν*xvec⦈(P' ∥ Q'))"
by simp
moreover from ‹x ♯ Ψ› ‹x ♯ P'› ‹xvec ♯* Ψ› have "(Ψ, ⦇νx⦈(⦇ν*xvec⦈(P' ∥ Q')), ⦇ν*xvec⦈(P' ∥ ⦇νx⦈Q')) ∈ Rel"
by(rule C2)
ultimately show ?case by blast
qed
next
case(cBrMerge Ψ⇩Q M N P' A⇩P Ψ⇩P xQ' A⇩Q)
have QTrans: "Ψ ⊗ Ψ⇩P ⊳ ⦇νx⦈Q ⟼¿M⦇N⦈ ≺ xQ'" and FrQ: "extractFrame(⦇νx⦈Q) = ⟨A⇩Q, Ψ⇩Q⟩" by fact+
have PTrans: "Ψ ⊗ Ψ⇩Q ⊳ P ⟼¿M⦇N⦈ ≺ P'" and FrP: "extractFrame P = ⟨A⇩P, Ψ⇩P⟩" by fact+
from ‹x ♯ α› ‹α = ¿M⦇N⦈› have "x ♯ M" and "x ♯ N" by simp+
from ‹x ♯ P' ∥ xQ'› have "x ♯ xQ'" by simp
from FrP ‹A⇩P ♯* x› ‹x ♯ P› have "x ♯ Ψ⇩P"
by(drule_tac extractFrameFresh) auto
from PTrans ‹x ♯ P› ‹x ♯ N› have "x ♯ P'"
by(rule brinputFreshDerivative)
have "x ♯ ⦇νx⦈Q" by(simp add: abs_fresh)
with FrQ ‹A⇩Q ♯* x› have "x ♯ Ψ⇩Q"
by(drule_tac extractFrameFresh) auto
from ‹A⇩P ♯* ⦇νx⦈Q› ‹A⇩P ♯* x› have "A⇩P ♯* Q" by simp
from ‹A⇩Q ♯* ⦇νx⦈Q› ‹A⇩Q ♯* x› have "A⇩Q ♯* Q" by simp
from ‹x ♯ Ψ› ‹x ♯ Ψ⇩P›
have "x ♯ (Ψ ⊗ Ψ⇩P)" by force
from ‹Ψ ⊗ Ψ⇩P ⊳ ⦇νx⦈Q ⟼ ¿M⦇N⦈ ≺ xQ'› ‹x ♯ (Ψ ⊗ Ψ⇩P)› ‹x ♯ M› ‹x ♯ N› ‹x ♯ xQ'›
show ?case
proof(induct rule: resBrInputCases)
case(cRes Q')
have QTrans: "Ψ ⊗ Ψ⇩P ⊳ Q ⟼¿M⦇N⦈ ≺ Q'" by fact
with ‹A⇩Q ♯* Q› ‹A⇩Q ♯* N› have "A⇩Q ♯* Q'"
by(elim brinputFreshChainDerivative)
with ‹extractFrame(⦇νx⦈Q) = ⟨A⇩Q, Ψ⇩Q⟩› have FrxQ: "⦇νx⦈(extractFrame Q) = ⟨A⇩Q, Ψ⇩Q⟩" by simp
then obtain y A⇩Q' where A: "A⇩Q = y#A⇩Q'" by(cases A⇩Q) auto
with ‹A⇩Q ♯* Ψ› ‹A⇩Q ♯* P› ‹A⇩Q ♯* P'› ‹A⇩Q ♯* Ψ⇩P› ‹A⇩Q ♯* Q› ‹A⇩Q ♯* M› ‹A⇩Q ♯* Q'› ‹A⇩Q ♯* N›
have "A⇩Q' ♯* Ψ" and "A⇩Q' ♯* P" and "A⇩Q' ♯* Ψ⇩P" and "A⇩Q' ♯* Q" and "A⇩Q ♯* Q'"
and "y ♯ Ψ" and "y ♯ P" and "y ♯ P'" and "y ♯ Ψ⇩P" and "y ♯ Q" and "y ♯ M" and "y ♯ Q'" and "y ♯ N"
and "A⇩Q' ♯* M"
by(simp)+
from A ‹A⇩P ♯* A⇩Q› have "A⇩P ♯* A⇩Q'" and "y ♯ A⇩P" by(simp add: fresh_star_list_cons)+
from A ‹A⇩Q ♯* x› have "x ≠ y" and "x ♯ A⇩Q'" by(simp add: fresh_list_cons)+
with A ‹distinct A⇩Q› have "y ♯ A⇩Q'"
by(induct A⇩Q') (auto simp add: fresh_list_nil fresh_list_cons)
note PTrans FrP
moreover from ‹Ψ ⊗ Ψ⇩P ⊳ Q ⟼¿M⦇N⦈ ≺ Q'› have "([(x, y)] ∙ (Ψ ⊗ Ψ⇩P)) ⊳ ([(x, y)] ∙ Q) ⟼([(x, y)] ∙ (¿M⦇N⦈ ≺ Q'))"
by(rule semantics.eqvt)
with ‹x ♯ Ψ› ‹y ♯ Ψ› ‹x ♯ Ψ⇩P› ‹y ♯ Ψ⇩P› ‹x ♯ M› ‹y ♯ M› ‹x ♯ N› ‹y ♯ N›
have "Ψ ⊗ Ψ⇩P ⊳ ([(x, y)] ∙ Q) ⟼¿M⦇N⦈ ≺ ([(x, y)] ∙ Q')"
by(simp add: eqvts)
moreover from A ‹A⇩Q ♯* x› FrxQ have FrQ: "extractFrame([(x, y)] ∙ Q) = ⟨A⇩Q', Ψ⇩Q⟩" and "y ♯ extractFrame Q"
by(clarsimp simp add: alpha' eqvts frame.inject fresh_list_cons name_swap)+
moreover from ‹A⇩P ♯* Q› have "([(x, y)] ∙ A⇩P) ♯* ([(x, y)] ∙ Q)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
with ‹A⇩P ♯* x› ‹y ♯ A⇩P› have "A⇩P ♯* ([(x, y)] ∙ Q)" by simp
moreover from ‹A⇩Q' ♯* Q› have "([(x, y)] ∙ A⇩Q') ♯* ([(x, y)] ∙ Q)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
with ‹x ♯ A⇩Q'› ‹y ♯ A⇩Q'› have "A⇩Q' ♯* ([(x, y)] ∙ Q)" by simp
ultimately have "Ψ ⊳ (P ∥ ([(x, y)] ∙ Q)) ⟼¿M⦇N⦈ ≺ (P' ∥ ([(x, y)] ∙ Q'))"
using ‹A⇩P ♯* Ψ› ‹A⇩P ♯* P› ‹A⇩P ♯* A⇩Q'› ‹A⇩Q' ♯* Ψ› ‹A⇩Q' ♯* P› ‹A⇩P ♯* M› ‹A⇩Q' ♯* M›
by(intro BrMerge)
with‹y ♯ Ψ› ‹y ♯ M› ‹y ♯ N› have "Ψ ⊳ ⦇νy⦈(P ∥ ([(x, y)] ∙ Q)) ⟼¿M⦇N⦈ ≺ ⦇νy⦈(P' ∥ ([(x, y)] ∙ Q'))"
by(intro Scope) auto
moreover from ‹x ♯ P› ‹y ♯ P› ‹y ♯ Q› have "⦇νy⦈(P ∥ ([(x, y)] ∙ Q)) = ⦇νx⦈([(x, y)] ∙ (P ∥ ([(x, y)] ∙ Q)))"
by(subst alphaRes[of x]) (auto simp add: calc_atm fresh_left name_swap)
moreover with ‹x ♯ P› ‹y ♯ P› have "⦇νy⦈(P ∥ ([(x, y)] ∙ Q)) = ⦇νx⦈(P ∥ Q)"
by(simp add: eqvts)
moreover from ‹x ♯ P'› ‹y ♯ P'› ‹y ♯ Q'› have "⦇νy⦈(P' ∥ ([(x, y)] ∙ Q')) = ⦇νx⦈(P' ∥ Q')"
by(subst alphaRes[of y]) (auto simp add: resChainFresh calc_atm eqvts fresh_left name_swap)
ultimately have finTrans: "Ψ ⊳ ⦇νx⦈(P ∥ Q) ⟼¿M⦇N⦈ ≺ ⦇νx⦈(P' ∥ Q')"
by simp
from ‹x ♯ Ψ› ‹x ♯ P'› have "[x] ♯* P'" and "[x] ♯* Ψ" by simp+
then have "(Ψ, ⦇ν*[x]⦈(P' ∥ Q'), (P' ∥ (⦇ν*[x]⦈Q'))) ∈ Rel"
by(rule C4)
then have "(Ψ, ⦇νx⦈(P' ∥ Q'), (P' ∥ ⦇νx⦈Q')) ∈ Rel"
by simp
with finTrans show ?case by blast
qed
next
case(cBrComm1 Ψ⇩Q M N P' A⇩P Ψ⇩P xvec xQ' A⇩Q)
have QTrans: "Ψ ⊗ Ψ⇩P ⊳ ⦇νx⦈Q ⟼¡M⦇ν*xvec⦈⟨N⟩ ≺ xQ'" and FrQ: "extractFrame(⦇νx⦈Q) = ⟨A⇩Q, Ψ⇩Q⟩" by fact+
have PTrans: "Ψ ⊗ Ψ⇩Q ⊳ P ⟼¿M⦇N⦈ ≺ P'" and FrP: "extractFrame P = ⟨A⇩P, Ψ⇩P⟩" by fact+
from ‹bn α ♯* x› ‹¡M⦇ν*xvec⦈⟨N⟩ = α› have "x ♯ xvec" by force
have "x ♯ ⦇νx⦈Q" by(simp add: abs_fresh)
with QTrans have "x ♯ N" and "x ♯ xQ'" using ‹x ♯ xvec› ‹xvec ♯* M› ‹distinct xvec›
by(force intro: broutputFreshDerivative)+
from PTrans ‹x ♯ P› ‹x ♯ N› have "x ♯ P'" by(rule brinputFreshDerivative)
from ‹x ♯ ⦇νx⦈Q› FrQ ‹A⇩Q ♯* x› have "x ♯ Ψ⇩Q"
by(drule_tac extractFrameFresh) auto
from ‹x ♯ P› FrP ‹A⇩P ♯* x› have "x ♯ Ψ⇩P"
by(drule_tac extractFrameFresh) auto
from ‹A⇩P ♯* ⦇νx⦈Q› ‹A⇩P ♯* x› have "A⇩P ♯* Q" by simp
from ‹A⇩Q ♯* ⦇νx⦈Q› ‹A⇩Q ♯* x› have "A⇩Q ♯* Q" by simp
note QTrans
moreover from ‹x ♯ Ψ› ‹x ♯ Ψ⇩P› have "x ♯ Ψ ⊗ Ψ⇩P" by force
moreover from ‹x ♯ xvec› have "xvec ♯* x" by simp
from ‹x ♯ α› ‹¡M⦇ν*xvec⦈⟨N⟩ = α› have "x ♯ ¡M⦇ν*xvec⦈⟨N⟩" by simp
moreover note ‹x ♯ xQ'›
moreover from ‹xvec ♯* Ψ› ‹xvec ♯* Ψ⇩P› have "xvec ♯* (Ψ ⊗ Ψ⇩P)" by force
moreover from ‹xvec ♯* ⦇νx⦈Q› ‹x ♯ xvec› have "xvec ♯* Q" by simp
moreover note ‹xvec ♯* M›
moreover from ‹xvec ♯* Ψ› ‹xvec ♯* Ψ⇩P› have "bn(¡M⦇ν*xvec⦈⟨N⟩) ♯* (Ψ ⊗ Ψ⇩P)" by force
moreover from ‹xvec ♯* ⦇νx⦈Q› ‹x ♯ xvec› have "bn(¡M⦇ν*xvec⦈⟨N⟩) ♯* Q" by simp
moreover from ‹xvec ♯* P› have "bn(¡M⦇ν*xvec⦈⟨N⟩) ♯* P" by simp
from ‹xvec ♯* Ψ› have "bn(¡M⦇ν*xvec⦈⟨N⟩) ♯* Ψ" by simp
from ‹A⇩Q ♯* xvec› ‹A⇩Q ♯* M› ‹A⇩Q ♯* N› have "A⇩Q ♯* (¡M⦇ν*xvec⦈⟨N⟩)" by simp
have "object(¡M⦇ν*xvec⦈⟨N⟩) = Some N" by simp
have "bn(¡M⦇ν*xvec⦈⟨N⟩) = xvec" by simp
have "subject(¡M⦇ν*xvec⦈⟨N⟩) = Some M" by simp
from ‹xvec ♯* M› have "bn(¡M⦇ν*xvec⦈⟨N⟩) ♯* subject(¡M⦇ν*xvec⦈⟨N⟩)" by simp
ultimately show ?case
using ‹x ♯ ¡M⦇ν*xvec⦈⟨N⟩› ‹bn(¡M⦇ν*xvec⦈⟨N⟩) ♯* P› ‹bn(¡M⦇ν*xvec⦈⟨N⟩) ♯* Ψ› ‹object(¡M⦇ν*xvec⦈⟨N⟩) = Some N›
‹bn(¡M⦇ν*xvec⦈⟨N⟩) = xvec› ‹subject(¡M⦇ν*xvec⦈⟨N⟩) = Some M› ‹A⇩Q ♯* (¡M⦇ν*xvec⦈⟨N⟩)›
proof(induct rule: resBrOutputCases')
case(cBrOpen M'' xvec1 xvec2 y N' Q')
then have Eq: "¡M⦇ν*xvec⦈⟨N⟩ = ¡M''⦇ν*(xvec1 @ y # xvec2)⦈⟨N'⟩" by simp
from ‹x ♯ ¡M⦇ν*xvec⦈⟨N⟩› Eq have "x ♯ xvec1" and "x ≠ y" and "x ♯ xvec2" and "x ♯ M''"
by simp+
from ‹bn(¡M⦇ν*xvec⦈⟨N⟩) ♯* P› Eq have "(xvec1@xvec2) ♯* P" and "y ♯ P" by simp+
from ‹A⇩Q ♯* (¡M⦇ν*xvec⦈⟨N⟩)› Eq have "(xvec1@xvec2) ♯* A⇩Q" and "y ♯ A⇩Q" and "A⇩Q ♯* M''" by simp+
from ‹bn(¡M⦇ν*xvec⦈⟨N⟩) ♯* Ψ› Eq have "(xvec1@xvec2) ♯* Ψ" and "y ♯ Ψ" by simp+
from Eq have "N = N'" and "xvec = xvec1@y#xvec2" and "M = M''" by(simp add: action.inject)+
from ‹x ♯ P› ‹y ♯ P› ‹x ≠ y› ‹Ψ ⊗ Ψ⇩Q ⊳ P ⟼¿M⦇N⦈ ≺ P'›
have "Ψ ⊗ Ψ⇩Q ⊳ P ⟼¿M⦇([(y, x)] ∙ N)⦈ ≺ ([(y, x)] ∙ P')"
by(intro brinputAlpha[where xvec="[y]"]) (auto simp add: calc_atm)
then have PTrans: "Ψ ⊗ Ψ⇩Q ⊳ P ⟼¿M⦇([(x, y)] ∙ N)⦈ ≺ ([(x, y)] ∙ P')"
by(simp add: name_swap)
from ‹Ψ ⊗ Ψ⇩P ⊳ [(x, y)] ∙ Q ⟼ ¡M''⦇ν*(xvec1 @ xvec2)⦈⟨N'⟩ ≺ Q'›
have "[(x, y)] ∙ (Ψ ⊗ Ψ⇩P ⊳ [(x, y)] ∙ Q ⟼ ¡M''⦇ν*(xvec1 @ xvec2)⦈⟨N'⟩ ≺ Q')"
by simp
then have "[(x, y)] ∙ (Ψ ⊗ Ψ⇩P) ⊳ ([(x, y)] ∙ ([(x, y)] ∙ Q)) ⟼ ¡([(x, y)] ∙ M'')⦇ν*([(x, y)] ∙ (xvec1 @ xvec2))⦈⟨([(x, y)] ∙ N')⟩ ≺ [(x, y)] ∙ Q'"
by(simp add: eqvts)
with ‹x ♯ (Ψ ⊗ Ψ⇩P)› ‹y ♯ (Ψ ⊗ Ψ⇩P)› ‹x ♯ xvec1› ‹y ♯ xvec1› ‹x ♯ xvec2› ‹y ♯ xvec2› ‹x ♯ M''› ‹y ♯ M''›
have QTrans: "Ψ ⊗ Ψ⇩P ⊳ Q ⟼ ¡M''⦇ν*(xvec1 @ xvec2)⦈⟨([(x, y)] ∙ N')⟩ ≺ [(x, y)] ∙ Q'"
by simp
with ‹A⇩Q ♯* x› ‹y ♯ A⇩Q› ‹distinct xvec1› ‹distinct xvec2› ‹xvec1 ♯* xvec2› ‹xvec1 ♯* M''› ‹xvec2 ♯* M''›
‹(xvec1@xvec2) ♯* A⇩Q›
have "A⇩Q ♯* ([(x, y)] ∙ Q')" using ‹A⇩Q ♯* Q›
by(elim broutputFreshChainDerivative(2)) (assumption | simp)+
from ‹extractFrame(⦇νx⦈Q) = ⟨A⇩Q, Ψ⇩Q⟩› have FrxQ: "⦇νx⦈(extractFrame Q) = ⟨A⇩Q, Ψ⇩Q⟩" by simp
then obtain z A⇩Q' where A: "A⇩Q = z#A⇩Q'" by(cases A⇩Q) auto
with ‹A⇩Q ♯* Ψ› ‹A⇩Q ♯* P› ‹A⇩Q ♯* P'› ‹A⇩Q ♯* Ψ⇩P› ‹A⇩Q ♯* Q› ‹(xvec1@xvec2) ♯* A⇩Q› ‹A⇩Q ♯* M''› ‹A⇩Q ♯* ([(x, y)] ∙ Q')› ‹y ♯ A⇩Q› ‹A⇩Q ♯* N›
have "A⇩Q' ♯* Ψ" and "A⇩Q' ♯* P" and "A⇩Q' ♯* Ψ⇩P" and "A⇩Q' ♯* Q"
and "z ♯ Ψ" and "z ♯ P" and "z ♯ P'" and "z ♯ Ψ⇩P" and "z ♯ Q" and "z ♯ xvec1" and "z ♯ xvec2"
and "z ♯ M''" and "z ♯ ([(x, y)] ∙ Q')" and "A⇩Q' ♯* M''" and "z ≠ y" and "z ♯ (xvec1@xvec2)"
by auto
from A ‹A⇩P ♯* A⇩Q› have "A⇩P ♯* A⇩Q'" and "z ♯ A⇩P" by simp+
from A ‹A⇩Q ♯* x› have "x ≠ z" and "x ♯ A⇩Q'" by simp+
from ‹A⇩Q ♯* (¡M⦇ν*xvec⦈⟨N⟩)› A have "z ♯ M" and "z ♯ xvec" and "z ♯ N" by simp+
from ‹y ∈ supp N'› ‹N = N'› have "y ∈ supp N" by simp
then have "x ∈ supp ([(x, y)] ∙ N)"
by (rule swap_supp)
then have zsupp: "z ∈ supp ([(x, z)] ∙ [(x, y)] ∙ N)"
by (rule swap_supp')
from ‹distinct A⇩Q› A have "z ♯ A⇩Q'"
by(induct A⇩Q') (auto simp add: fresh_list_nil fresh_list_cons)
from PTrans ‹x ♯ P› ‹z ♯ P› ‹x ≠ z› have "Ψ ⊗ Ψ⇩Q ⊳ P ⟼¿M⦇([(x, z)] ∙ [(x, y)] ∙ N)⦈ ≺ ([(x, z)] ∙ [(x, y)] ∙ P')"
by(elim brinputAlpha[where xvec="[x]"]) (auto simp add: calc_atm)
moreover note FrP
moreover from QTrans have "([(x, z)] ∙ (Ψ ⊗ Ψ⇩P)) ⊳ ([(x, z)] ∙ Q) ⟼([(x, z)] ∙ (¡M''⦇ν*(xvec1@xvec2)⦈⟨([(x, y)] ∙ N')⟩ ≺ ([(x, y)] ∙ Q')))"
by(rule semantics.eqvt)
with ‹x ♯ Ψ› ‹z ♯ Ψ› ‹x ♯ Ψ⇩P› ‹z ♯ Ψ⇩P› ‹x ♯ M''› ‹z ♯ M''› ‹x ♯ xvec1› ‹x ♯ xvec2› ‹z ♯ xvec1› ‹z ♯ xvec2›
have "Ψ ⊗ Ψ⇩P ⊳ ([(x, z)] ∙ Q) ⟼¡M''⦇ν*(xvec1@xvec2)⦈⟨([(x, z)] ∙ [(x, y)] ∙ N')⟩ ≺ ([(x, z)] ∙ [(x, y)] ∙ Q')"
by(simp add: eqvts)
moreover from A ‹A⇩Q ♯* x› FrxQ have "extractFrame([(x, z)] ∙ Q) = ⟨A⇩Q', Ψ⇩Q⟩"
by(clarsimp simp add: alpha' eqvts frame.inject fresh_list_cons name_swap)
moreover from ‹A⇩P ♯* Q› have "([(x, z)] ∙ A⇩P) ♯* ([(x, z)] ∙ Q)"
by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
with ‹A⇩P ♯* x› ‹z ♯ A⇩P› have "A⇩P ♯* ([(x, z)] ∙ Q)" by simp
moreover from ‹A⇩Q' ♯* Q› have "([(x, z)] ∙ A⇩Q') ♯* ([(x, z)] ∙ Q)"
by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
with ‹x ♯ A⇩Q'› ‹z ♯ A⇩Q'› have "A⇩Q' ♯* ([(x, z)] ∙ Q)" by simp
ultimately have "Ψ ⊳ (P ∥ ([(x, z)] ∙ Q)) ⟼¡M⦇ν*(xvec1@xvec2)⦈⟨([(x, z)] ∙ [(x, y)] ∙ N)⟩ ≺ (([(x, z)] ∙ [(x, y)] ∙ P') ∥ ([(x, z)] ∙ [(x, y)] ∙ Q'))"
using ‹M=M''› ‹N=N'› ‹A⇩P ♯* Ψ› ‹A⇩P ♯* P› ‹A⇩P ♯* A⇩Q'› ‹A⇩Q' ♯* Ψ› ‹A⇩Q' ♯* P› ‹(xvec1@xvec2) ♯* P› ‹A⇩P ♯* M› ‹A⇩Q' ♯* M''›
by(intro BrComm1) (assumption | simp)+
then have permTrans: "Ψ ⊳ ⦇νz⦈(P ∥ ([(x, z)] ∙ Q)) ⟼¡M⦇ν*(xvec1@z#xvec2)⦈⟨([(x, z)] ∙ [(x, y)] ∙ N)⟩ ≺ (([(x, z)] ∙ [(x, y)] ∙ P') ∥ ([(x, z)] ∙ [(x, y)] ∙ Q'))"
using zsupp ‹z ♯ Ψ› ‹z ♯ M› ‹z ♯ xvec1› ‹z ♯ xvec2›
by(rule BrOpen)
from ‹x ♯ P› ‹z ♯ P› ‹z ♯ Q› have "⦇νz⦈(P ∥ ([(x, z)] ∙ Q)) = ⦇νx⦈([(x, z)] ∙ (P ∥ ([(x, z)] ∙ Q)))"
by(subst alphaRes[of x]) (auto simp add: calc_atm fresh_left name_swap)
with ‹x ♯ P› ‹z ♯ P› have "⦇νz⦈(P ∥ ([(x, z)] ∙ Q)) = ⦇νx⦈(P ∥ Q)"
by(simp add: eqvts)
with permTrans have permTrans2: "Ψ ⊳ ⦇νx⦈(P ∥ Q) ⟼¡M⦇ν*(xvec1@z#xvec2)⦈⟨([(x, z)] ∙ [(x, y)] ∙ N)⟩ ≺ (([(x, z)] ∙ [(x, y)] ∙ P') ∥ ([(x, z)] ∙ [(x, y)] ∙ Q'))"
by simp
then have "Ψ ⊳ ⦇νx⦈(P ∥ Q) ⟼ RBrOut M (⦇ν*(xvec1@z#xvec2)⦈([(x, z)] ∙ [(x, y)] ∙ N) ≺' (([(x, z)] ∙ [(x, y)] ∙ P') ∥ ([(x, z)] ∙ [(x, y)] ∙ Q')))"
by(simp add: residualInject)
then have permTrans3: "Ψ ⊳ ⦇νx⦈(P ∥ Q) ⟼ RBrOut M (⦇ν*(xvec1@z#xvec2)⦈ (([(x, z)] ∙ [(x, y)] ∙ N) ≺' ([(x, z)] ∙ [(x, y)] ∙ (P' ∥ Q'))))"
by simp
from ‹z ♯ N› ‹x ≠ z› ‹z ≠ y›
have "z ♯ ([(x, y)] ∙ N)"
by (metis calc_atm(2) calc_atm(3) fresh_right name_prm_name.simps(2) name_prm_name_def singleton_rev_conv swap_name_def)
from ‹z ♯ P'› ‹x ≠ z› ‹z ≠ y›
have "z ♯ ([(x, y)] ∙ P')"
by (metis calc_atm(2) calc_atm(3) fresh_right name_prm_name.simps(2) name_prm_name_def singleton_rev_conv swap_name_def)
from ‹z ♯ ([(x, y)] ∙ N)› have "x ♯ ([(x, z)] ∙ [(x, y)] ∙ N)"
by (metis calc_atm(2) calc_atm(3) fresh_right name_prm_name.simps(2) name_prm_name_def singleton_rev_conv swap_name_def)
from ‹z ♯ ([(x, y)] ∙ P')› have "x ♯ ([(x, z)] ∙ [(x, y)] ∙ P')"
by (metis calc_atm(2) calc_atm(3) fresh_right name_prm_name.simps(2) name_prm_name_def singleton_rev_conv swap_name_def)
moreover from ‹z ♯ [(x, y)] ∙ Q'›
have "x ♯ ([(x, z)] ∙ [(x, y)] ∙ Q')"
by (metis calc_atm(2) calc_atm(3) fresh_right name_prm_name.simps(2) name_prm_name_def singleton_rev_conv swap_name_def)
ultimately have "x ♯ ([(x, z)] ∙ [(x, y)] ∙ (P' ∥ Q'))"
by simp
have "z ∈ set ((xvec1@z#xvec2))"
by simp
from ‹x ♯ ([(x, z)] ∙ [(x, y)] ∙ N)› ‹x ♯ ([(x, z)] ∙ [(x, y)] ∙ (P' ∥ Q'))› ‹z ∈ set ((xvec1@z#xvec2))›
have eq1: "(⦇ν*(xvec1@z#xvec2)⦈ (([(x, z)] ∙ [(x, y)] ∙ N) ≺' ([(x, z)] ∙ [(x, y)] ∙ (P' ∥ Q')))) = (⦇ν*([(z, x)] ∙ (xvec1@z#xvec2))⦈ (([(z, x)] ∙ [(x, z)] ∙ [(x, y)] ∙ N) ≺' ([(z, x)] ∙ [(x, z)] ∙ [(x, y)] ∙ (P' ∥ Q'))))"
by(rule boundOutputChainSwap)
have eq2: "(⦇ν*([(z, x)] ∙ (xvec1@z#xvec2))⦈ (([(z, x)] ∙ [(x, z)] ∙ [(x, y)] ∙ N) ≺' ([(z, x)] ∙ [(x, z)] ∙ [(x, y)] ∙ (P' ∥ Q')))) = (⦇ν*([(z, x)] ∙ (xvec1@z#xvec2))⦈ (([(x, y)] ∙ N) ≺' ([(x, y)] ∙ (P' ∥ Q'))))"
by auto (metis perm_swap(2))+
from ‹x ♯ xvec1› ‹x ♯ xvec2› ‹z ♯ xvec1› ‹z ♯ xvec2›
have "([(z, x)] ∙ (xvec1@z#xvec2)) = (xvec1@x#xvec2)"
by(simp add: eqvts swap_simps)
then have eq3: "(⦇ν*([(z, x)] ∙ (xvec1@z#xvec2))⦈ (([(x, y)] ∙ N) ≺' ([(x, y)] ∙ (P' ∥ Q')))) = (⦇ν*(xvec1@x#xvec2)⦈ (([(x, y)] ∙ N) ≺' ([(x, y)] ∙ (P' ∥ Q'))))"
by simp
from permTrans3 eq1 eq2 eq3 have noXZTrans: "Ψ ⊳ ⦇νx⦈(P ∥ Q) ⟼ RBrOut M (⦇ν*(xvec1@x#xvec2)⦈ (([(x, y)] ∙ N) ≺' ([(x, y)] ∙ (P' ∥ Q'))))"
by simp
from ‹x ♯ N›
have "y ♯ ([(x, y)] ∙ N)"
by (metis calc_atm(2) fresh_right name_prm_name.simps(2) name_prm_name_def singleton_rev_conv swap_name_def swap_simps(2))
moreover from ‹x ♯ P'› ‹x ♯ Q'›
have "y ♯ ([(x, y)] ∙ (P' ∥ Q'))"
by (metis fresh_left psi.fresh(5) singleton_rev_conv swap_simps(2))
moreover have "x ∈ set (xvec1@x#xvec2)"
by simp
ultimately have eq1: "⦇ν*(xvec1@x#xvec2)⦈([(x, y)] ∙ N) ≺' ([(x, y)] ∙ (P' ∥ Q')) = ⦇ν*([(x, y)] ∙ (xvec1@x#xvec2))⦈([(x, y)] ∙ [(x, y)] ∙ N) ≺' ([(x, y)] ∙ [(x, y)] ∙ (P' ∥ Q'))"
by(rule boundOutputChainSwap)
have eq2: "⦇ν*([(x, y)] ∙ (xvec1@x#xvec2))⦈([(x, y)] ∙ [(x, y)] ∙ N) ≺' ([(x, y)] ∙ [(x, y)] ∙ (P' ∥ Q')) = ⦇ν*([(x, y)] ∙ (xvec1@x#xvec2))⦈N ≺' (P' ∥ Q')"
by simp
from ‹x ♯ xvec1› ‹x ♯ xvec2› ‹y ♯ xvec1› ‹y ♯ xvec2›
have eq3: "([(x, y)] ∙ (xvec1@x#xvec2)) = (xvec1@y#xvec2)"
by(simp add: eqvts swap_simps)
from eq1 eq2 eq3 noXZTrans have "Ψ ⊳ ⦇νx⦈(P ∥ Q) ⟼ RBrOut M (⦇ν*(xvec1@y#xvec2)⦈N ≺' (P' ∥ Q'))"
by simp
then have "Ψ ⊳ ⦇νx⦈(P ∥ Q) ⟼¡M⦇ν*(xvec1@y#xvec2)⦈⟨N⟩ ≺ (P' ∥ Q')"
by(simp add: residualInject)
with ‹xvec = xvec1@y#xvec2›
have "Ψ ⊳ ⦇νx⦈(P ∥ Q) ⟼¡M⦇ν*xvec⦈⟨N⟩ ≺ (P' ∥ Q')"
by simp
moreover have "(Ψ, P' ∥ Q', P' ∥ Q') ∈ Rel"
by(rule C1)
ultimately show ?case
by force
next
case(cRes Q')
from ‹x ♯ ¡M⦇ν*xvec⦈⟨N⟩›
have "x ♯ M" and "x ♯ xvec" and "x ♯ N"
by simp+
have QTrans: "Ψ ⊗ Ψ⇩P ⊳ Q ⟼¡M⦇ν*xvec⦈⟨N⟩ ≺ Q'" by fact
with ‹A⇩Q ♯* Q› ‹A⇩Q ♯* xvec› ‹xvec ♯* M› ‹distinct xvec› have "A⇩Q ♯* Q'"
by(force dest: broutputFreshChainDerivative)
with ‹extractFrame(⦇νx⦈Q) = ⟨A⇩Q, Ψ⇩Q⟩› have FrxQ: "⦇νx⦈(extractFrame Q) = ⟨A⇩Q, Ψ⇩Q⟩" by simp
then obtain y A⇩Q' where A: "A⇩Q = y#A⇩Q'" by(cases A⇩Q) auto
with ‹A⇩Q ♯* Ψ› ‹A⇩Q ♯* P› ‹A⇩Q ♯* P'› ‹A⇩Q ♯* Ψ⇩P› ‹A⇩Q ♯* Q› ‹A⇩Q ♯* xvec› ‹A⇩Q ♯* M› ‹A⇩Q ♯* Q'› ‹A⇩Q ♯* N›
have "A⇩Q' ♯* Ψ" and "A⇩Q' ♯* P" and "A⇩Q' ♯* Ψ⇩P" and "A⇩Q' ♯* Q" and "A⇩Q ♯* xvec" and "A⇩Q ♯* Q'"
and "y ♯ Ψ" and "y ♯ P" and "y ♯ P'" and "y ♯ Ψ⇩P" and "y ♯ Q" and "y ♯ xvec" and "y ♯ M" and "y ♯ Q'"
and "A⇩Q' ♯* M" and "y ♯ N"
by(simp)+
from A ‹A⇩P ♯* A⇩Q› have "A⇩P ♯* A⇩Q'" and "y ♯ A⇩P" by(simp add: fresh_star_list_cons)+
from A ‹A⇩Q ♯* x› have "x ≠ y" and "x ♯ A⇩Q'" by(simp add: fresh_list_cons)+
with A ‹distinct A⇩Q› have "y ♯ A⇩Q'"
by(induct A⇩Q') (auto simp add: fresh_list_nil fresh_list_cons)
from ‹x ♯ N› have "y ♯ [(x, y)] ∙ N"
by (metis calc_atm(2) fresh_right name_prm_name.simps(2) name_prm_name_def singleton_rev_conv swap_name_def swap_simps(2))
from ‹x ♯ P› ‹y ♯ P› ‹x ≠ y› ‹Ψ ⊗ Ψ⇩Q ⊳ P ⟼¿M⦇N⦈ ≺ P'›
have "Ψ ⊗ Ψ⇩Q ⊳ P ⟼¿M⦇([(y, x)] ∙ N)⦈ ≺ [(y, x)] ∙ P'"
by(elim brinputAlpha[where xvec="[y]"]) (auto simp add: calc_atm)
then have "Ψ ⊗ Ψ⇩Q ⊳ P ⟼¿M⦇([(x, y)] ∙ N)⦈ ≺ [(x, y)] ∙ P'"
by(simp add: name_swap)
moreover note FrP
moreover from ‹Ψ ⊗ Ψ⇩P ⊳ Q ⟼¡M⦇ν*xvec⦈⟨N⟩ ≺ Q'› have "([(x, y)] ∙ (Ψ ⊗ Ψ⇩P)) ⊳ ([(x, y)] ∙ Q) ⟼([(x, y)] ∙ (¡M⦇ν*xvec⦈⟨N⟩ ≺ Q'))"
by(rule semantics.eqvt)
with ‹x ♯ Ψ› ‹y ♯ Ψ› ‹x ♯ Ψ⇩P› ‹y ♯ Ψ⇩P› ‹x ♯ M› ‹y ♯ M› ‹x ♯ xvec› ‹y ♯ xvec›
have "Ψ ⊗ Ψ⇩P ⊳ ([(x, y)] ∙ Q) ⟼¡M⦇ν*xvec⦈⟨([(x, y)] ∙ N)⟩ ≺ ([(x, y)] ∙ Q')"
by(simp add: eqvts)
moreover from A ‹A⇩Q ♯* x› FrxQ have FrQ: "extractFrame([(x, y)] ∙ Q) = ⟨A⇩Q', Ψ⇩Q⟩"
by(clarsimp simp add: alpha' eqvts frame.inject fresh_list_cons name_swap)
moreover from ‹A⇩P ♯* Q› have "([(x, y)] ∙ A⇩P) ♯* ([(x, y)] ∙ Q)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
with ‹A⇩P ♯* x› ‹y ♯ A⇩P› have "A⇩P ♯* ([(x, y)] ∙ Q)" by simp
moreover from ‹A⇩Q' ♯* Q› have "([(x, y)] ∙ A⇩Q') ♯* ([(x, y)] ∙ Q)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
with ‹x ♯ A⇩Q'› ‹y ♯ A⇩Q'› have "A⇩Q' ♯* ([(x, y)] ∙ Q)" by simp
ultimately have "Ψ ⊳ (P ∥ ([(x, y)] ∙ Q)) ⟼¡M⦇ν*xvec⦈⟨([(x, y)] ∙ N)⟩ ≺ ([(x, y)] ∙ P') ∥ ([(x, y)] ∙ Q')"
using ‹A⇩P ♯* Ψ› ‹A⇩P ♯* P› ‹A⇩P ♯* A⇩Q'› ‹A⇩Q' ♯* Ψ› ‹A⇩Q' ♯* P› ‹xvec ♯* P› ‹A⇩P ♯* M› ‹A⇩Q' ♯* M›
by(intro BrComm1) (assumption | simp)+
with‹y ♯ Ψ› ‹y ♯ M› ‹y ♯ xvec› ‹y ♯ [(x, y)] ∙ N› have "Ψ ⊳ ⦇νy⦈(P ∥ ([(x, y)] ∙ Q)) ⟼¡M⦇ν*xvec⦈⟨([(x, y)] ∙ N)⟩ ≺ ⦇νy⦈(([(x, y)] ∙ P') ∥ ([(x, y)] ∙ Q'))"
by(intro Scope) auto
moreover from ‹x ♯ P› ‹y ♯ P› ‹y ♯ Q› have "⦇νy⦈(P ∥ ([(x, y)] ∙ Q)) = ⦇νx⦈([(x, y)] ∙ (P ∥ ([(x, y)] ∙ Q)))"
by(subst alphaRes[of x]) (auto simp add: calc_atm fresh_left name_swap)
with ‹x ♯ P› ‹y ♯ P› have "⦇νy⦈(P ∥ ([(x, y)] ∙ Q)) = ⦇νx⦈(P ∥ Q)"
by(simp add: eqvts)
moreover from ‹y ♯ P'› ‹y ♯ Q'› ‹x ♯ xvec› ‹y ♯ xvec› have "⦇νy⦈(([(x, y)] ∙ P') ∥ ([(x, y)] ∙ Q')) = ⦇νx⦈(P' ∥ Q')"
by(subst alphaRes[of y]) (auto simp add: resChainFresh calc_atm eqvts fresh_left name_swap)
ultimately have "Ψ ⊳ ⦇νx⦈(P ∥ Q) ⟼¡M⦇ν*xvec⦈⟨([(x, y)] ∙ N)⟩ ≺ ⦇νx⦈(P' ∥ Q')"
by simp
with ‹x ♯ N› ‹y ♯ N›
have Trans: "Ψ ⊳ ⦇νx⦈(P ∥ Q) ⟼¡M⦇ν*xvec⦈⟨N⟩ ≺ ⦇νx⦈(P' ∥ Q')"
by simp
from ‹x ♯ P'› ‹x ♯ Ψ›
have "(Ψ, ⦇ν*[x]⦈(P' ∥ Q'), (P' ∥ (⦇ν*[x]⦈Q'))) ∈ Rel"
by(intro C4) simp+
then have Relation: "(Ψ, ⦇νx⦈(P' ∥ Q'), (P' ∥ (⦇νx⦈Q'))) ∈ Rel"
by simp
from Trans Relation show ?case
by blast
qed
next
case(cBrComm2 Ψ⇩Q M xvec N P' A⇩P Ψ⇩P xQ' A⇩Q)
from ‹x ♯ α› ‹¡M⦇ν*xvec⦈⟨N⟩ = α›
have "x ♯ M" "x ♯ xvec" "x ♯ N"
by force+
have QTrans: "Ψ ⊗ Ψ⇩P ⊳ ⦇νx⦈Q ⟼¿M⦇N⦈ ≺ xQ'" and FrQ: "extractFrame(⦇νx⦈Q) = ⟨A⇩Q, Ψ⇩Q⟩" by fact+
have PTrans: "Ψ ⊗ Ψ⇩Q ⊳ P ⟼¡M⦇ν*xvec⦈⟨N⟩ ≺ P'" and FrP: "extractFrame P = ⟨A⇩P, Ψ⇩P⟩" by fact+
from PTrans ‹x ♯ P› have "x ♯ N" and "x ♯ P'" using ‹x ♯ xvec› ‹xvec ♯* M› ‹distinct xvec›
by(force intro: broutputFreshDerivative)+
have "x ♯ ⦇νx⦈Q" by(simp add: abs_fresh)
with FrQ ‹A⇩Q ♯* x› have "x ♯ Ψ⇩Q"
by(drule_tac extractFrameFresh) auto
from ‹x ♯ P› FrP ‹A⇩P ♯* x› have "x ♯ Ψ⇩P"
by(drule_tac extractFrameFresh) auto
from ‹A⇩P ♯* ⦇νx⦈Q› ‹A⇩P ♯* x› have "A⇩P ♯* Q" by simp
from ‹A⇩Q ♯* ⦇νx⦈Q› ‹A⇩Q ♯* x› have "A⇩Q ♯* Q" by simp
from ‹x ♯ xvec› ‹xvec ♯* ⦇νx⦈Q› have "xvec ♯* Q" by simp
from ‹Ψ ⊗ Ψ⇩Q ⊳ P ⟼ ¡M⦇ν*xvec⦈⟨N⟩ ≺ P'› have PResTrans: "Ψ ⊗ Ψ⇩Q ⊳ P ⟼ RBrOut M (⦇ν*xvec⦈N ≺' P')"
by(simp add: residualInject)
note QTrans
moreover from ‹x ♯ Ψ› ‹x ♯ Ψ⇩P› have "x ♯ Ψ ⊗ Ψ⇩P" by force
moreover note ‹x ♯ M› ‹x ♯ N›
moreover from QTrans ‹x ♯ N› have "x ♯ xQ'" by(force dest: brinputFreshDerivative simp add: abs_fresh)
ultimately show ?case
proof(induct rule: resBrInputCases)
case(cRes Q')
have QTrans: "Ψ ⊗ Ψ⇩P ⊳ Q ⟼¿M⦇N⦈ ≺ Q'" by fact
with ‹A⇩Q ♯* Q› ‹A⇩Q ♯* N› have "A⇩Q ♯* Q'"
by(elim brinputFreshChainDerivative)
with ‹extractFrame(⦇νx⦈Q) = ⟨A⇩Q, Ψ⇩Q⟩› have FrxQ: "⦇νx⦈(extractFrame Q) = ⟨A⇩Q, Ψ⇩Q⟩" by simp
then obtain y A⇩Q' where A: "A⇩Q = y#A⇩Q'" by(cases A⇩Q) auto
with ‹A⇩Q ♯* Ψ› ‹A⇩Q ♯* P› ‹A⇩Q ♯* P'› ‹A⇩Q ♯* Ψ⇩P› ‹A⇩Q ♯* Q› ‹A⇩Q ♯* xvec› ‹A⇩Q ♯* M› ‹A⇩Q ♯* Q'› ‹A⇩Q ♯* N›
have "A⇩Q' ♯* Ψ" and "A⇩Q' ♯* P" and "A⇩Q' ♯* Ψ⇩P" and "A⇩Q' ♯* Q" and "A⇩Q ♯* xvec" and "A⇩Q ♯* Q'"
and "y ♯ Ψ" and "y ♯ P" and "y ♯ P'" and "y ♯ Ψ⇩P" and "y ♯ Q" and "y ♯ xvec" and "y ♯ M" and "y ♯ Q'" and "y ♯ N"
and "A⇩Q' ♯* M"
by(simp)+
from A ‹A⇩P ♯* A⇩Q› have "A⇩P ♯* A⇩Q'" and "y ♯ A⇩P" by(simp add: fresh_star_list_cons)+
from A ‹A⇩Q ♯* x› have "x ≠ y" and "x ♯ A⇩Q'" by(simp add: fresh_list_cons)+
with A ‹distinct A⇩Q› have "y ♯ A⇩Q'"
by(induct A⇩Q') (auto simp add: fresh_list_nil fresh_list_cons)
note PTrans FrP
moreover from ‹Ψ ⊗ Ψ⇩P ⊳ Q ⟼¿M⦇N⦈ ≺ Q'› have "([(x, y)] ∙ (Ψ ⊗ Ψ⇩P)) ⊳ ([(x, y)] ∙ Q) ⟼([(x, y)] ∙ (¿M⦇N⦈ ≺ Q'))"
by(rule semantics.eqvt)
with ‹x ♯ Ψ› ‹y ♯ Ψ› ‹x ♯ Ψ⇩P› ‹y ♯ Ψ⇩P› ‹x ♯ M› ‹y ♯ M› ‹x ♯ N› ‹y ♯ N›
have "Ψ ⊗ Ψ⇩P ⊳ ([(x, y)] ∙ Q) ⟼¿M⦇N⦈ ≺ ([(x, y)] ∙ Q')"
by(simp add: eqvts)
moreover from A ‹A⇩Q ♯* x› FrxQ have FrQ: "extractFrame([(x, y)] ∙ Q) = ⟨A⇩Q', Ψ⇩Q⟩" and "y ♯ extractFrame Q"
by(clarsimp simp add: alpha' eqvts frame.inject fresh_list_cons name_swap)+
moreover from ‹A⇩P ♯* Q› have "([(x, y)] ∙ A⇩P) ♯* ([(x, y)] ∙ Q)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
with ‹A⇩P ♯* x› ‹y ♯ A⇩P› have "A⇩P ♯* ([(x, y)] ∙ Q)" by simp
moreover from ‹A⇩Q' ♯* Q› have "([(x, y)] ∙ A⇩Q') ♯* ([(x, y)] ∙ Q)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
with ‹x ♯ A⇩Q'› ‹y ♯ A⇩Q'› have "A⇩Q' ♯* ([(x, y)] ∙ Q)" by simp
moreover from ‹xvec ♯* Q› have "([(x, y)] ∙ xvec) ♯* ([(x, y)] ∙ Q)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
with ‹x ♯ xvec› ‹y ♯ xvec› have "xvec ♯* ([(x, y)] ∙ Q)" by simp
ultimately have "Ψ ⊳ (P ∥ ([(x, y)] ∙ Q)) ⟼¡M⦇ν*xvec⦈⟨N⟩ ≺ (P' ∥ ([(x, y)] ∙ Q'))"
using ‹A⇩P ♯* Ψ› ‹A⇩P ♯* P› ‹A⇩P ♯* A⇩Q'› ‹A⇩Q' ♯* Ψ› ‹A⇩Q' ♯* P› ‹A⇩P ♯* M› ‹A⇩Q' ♯* M›
by(intro BrComm2) (assumption | simp)+
with ‹y ♯ Ψ› ‹y ♯ M› ‹y ♯ xvec› ‹y ♯ N› have "Ψ ⊳ ⦇νy⦈(P ∥ ([(x, y)] ∙ Q)) ⟼¡M⦇ν*xvec⦈⟨N⟩ ≺ ⦇νy⦈(P' ∥ ([(x, y)] ∙ Q'))"
by(intro Scope) auto
moreover from ‹x ♯ P› ‹y ♯ P› ‹y ♯ Q› have "⦇νy⦈(P ∥ ([(x, y)] ∙ Q)) = ⦇νx⦈([(x, y)] ∙ (P ∥ ([(x, y)] ∙ Q)))"
by(subst alphaRes[of x]) (auto simp add: calc_atm fresh_left name_swap)
with ‹x ♯ P› ‹y ♯ P› have "⦇νy⦈(P ∥ ([(x, y)] ∙ Q)) = ⦇νx⦈(P ∥ Q)"
by(simp add: eqvts)
moreover from ‹x ♯ P'› ‹y ♯ P'› ‹y ♯ Q'› ‹x ♯ xvec› ‹y ♯ xvec› have "⦇νy⦈(P' ∥ ([(x, y)] ∙ Q')) = ⦇νx⦈(P' ∥ Q')"
by(subst alphaRes[of y]) (auto simp add: resChainFresh calc_atm eqvts fresh_left name_swap)
ultimately have "Ψ ⊳ ⦇νx⦈(P ∥ Q) ⟼¡M⦇ν*xvec⦈⟨N⟩ ≺ ⦇νx⦈(P' ∥ Q')"
by simp
moreover from ‹x ♯ Ψ› ‹x ♯ P'› have "(Ψ, ⦇ν*[x]⦈(P' ∥ Q'), (P' ∥ (⦇ν*[x]⦈Q'))) ∈ Rel"
by(intro C4) simp+
moreover then have "(Ψ, ⦇νx⦈(P' ∥ Q'), (P' ∥ (⦇νx⦈Q'))) ∈ Rel"
by simp
ultimately show ?case by blast
qed
qed
qed
qed
lemma scopeExtRight:
fixes x :: name
and P :: "('a, 'b, 'c) psi"
and Ψ :: 'b
and Q :: "('a, 'b, 'c) psi"
and Rel :: "('b × ('a, 'b, 'c) psi × ('a, 'b, 'c) psi) set"
assumes "x ♯ P"
and "x ♯ Ψ"
and "eqvt Rel"
and C1: "⋀Ψ' R. (Ψ, R, R) ∈ Rel"
and C2: "⋀y Ψ' R S zvec. ⟦y ♯ Ψ'; y ♯ R; zvec ♯* Ψ'⟧ ⟹ (Ψ', ⦇ν*zvec⦈(R ∥ ⦇νy⦈S), ⦇νy⦈(⦇ν*zvec⦈(R ∥ S))) ∈ Rel"
and C3: "⋀Ψ' R S zvec. ⟦zvec ♯* R; zvec ♯* Ψ'⟧ ⟹ (Ψ', (R ∥ (⦇ν*zvec⦈S)), (⦇ν*zvec⦈(R ∥ S))) ∈ Rel"
shows "Ψ ⊳ P ∥ ⦇νx⦈Q ↝[Rel] ⦇νx⦈(P ∥ Q)"
proof -
note ‹eqvt Rel› ‹x ♯ Ψ›
moreover from ‹x ♯ P› have "x ♯ P ∥ ⦇νx⦈Q" by(simp add: abs_fresh)
moreover from ‹x ♯ P› have "x ♯ ⦇νx⦈(P ∥ Q)" by(simp add: abs_fresh)
ultimately show ?thesis
proof(induct rule: simIFresh[of _ _ _ _ _ "()"])
case(cSim α xPQ)
from ‹bn α ♯* (P ∥ ⦇νx⦈Q)› ‹x ♯ α› have "bn α ♯* P" and "bn α ♯* Q" by simp+
note ‹Ψ ⊳ ⦇νx⦈(P ∥ Q) ⟼α ≺ xPQ› ‹x ♯ Ψ› ‹x ♯ α› ‹x ♯ xPQ› ‹bn α ♯* Ψ›
moreover from ‹bn α ♯* P› ‹bn α ♯* Q› have "bn α ♯* (P ∥ Q)" by simp
ultimately show ?case using ‹bn α ♯* subject α› ‹x ♯ α›
proof(induct rule: resCases[where C="()"])
case(cOpen M xvec1 xvec2 y N PQ)
from ‹x ♯ M⦇ν*(xvec1@y#xvec2)⦈⟨N⟩› have "x ♯ xvec1" and "x ≠ y" and "x ♯ xvec2" and "x ♯ M" by simp+
from ‹xvec1 ♯* (P ∥ Q)› ‹xvec2 ♯* (P ∥ Q)› ‹y ♯ (P ∥ Q)›
have "(xvec1@xvec2) ♯* P" and "(xvec1@xvec2) ♯* Q" and "y ♯ P" and "y ♯ Q"
by simp+
from ‹Ψ ⊳ P ∥ Q ⟼ M⦇ν*(xvec1@xvec2)⦈⟨([(x, y)] ∙ N)⟩ ≺ ([(x, y)] ∙ PQ)›
have "([(x, y)] ∙ Ψ) ⊳ ([(x, y)] ∙ (P ∥ Q)) ⟼ ([(x, y)] ∙ (M⦇ν*(xvec1@xvec2)⦈⟨([(x, y)] ∙ N)⟩ ≺ ([(x, y)] ∙ PQ)))"
by(rule semantics.eqvt)
with ‹x ♯ Ψ› ‹y ♯ Ψ› ‹x ♯ P› ‹y ♯ P› ‹x ♯ M› ‹y ♯ M› ‹x ♯ xvec1› ‹x ♯ xvec2› ‹y ♯ xvec1› ‹y ♯ xvec2›
have "Ψ ⊳ P ∥ ([(x, y)] ∙ Q) ⟼ M⦇ν*(xvec1@xvec2)⦈⟨N⟩ ≺ PQ"
by(simp add: eqvts)
moreover from ‹xvec1 ♯* Ψ› ‹xvec2 ♯* Ψ› have "(xvec1@xvec2) ♯* Ψ" by simp
moreover note ‹(xvec1@xvec2) ♯* P›
moreover from ‹(xvec1@xvec2) ♯* Q› have "([(x, y)] ∙ (xvec1@xvec2)) ♯* ([(x, y)] ∙ Q)"
by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
with ‹x ♯ xvec1› ‹x ♯ xvec2› ‹y ♯ xvec1› ‹y ♯ xvec2› have "(xvec1@xvec2) ♯* ([(x, y)] ∙ Q)"
by(auto simp add: eqvts)
moreover from ‹xvec1 ♯* M› ‹xvec2 ♯* M› have "(xvec1@xvec2) ♯* M" by simp
ultimately show ?case
proof(induct rule: parOutputCases[where C=y])
case(cPar1 P' A⇩Q Ψ⇩Q)
from ‹y ♯ xvec1› ‹y ♯ xvec2› have "y ♯ (xvec1@xvec2)" by(auto simp add: fresh_list_append)
with ‹Ψ ⊗ Ψ⇩Q ⊳ P ⟼M⦇ν*(xvec1@xvec2)⦈⟨N⟩ ≺ P'› ‹(xvec1@xvec2) ♯* M› ‹y ♯ P›
‹distinct xvec1› ‹distinct xvec2› ‹xvec1 ♯* xvec2›
have "y ♯ N" by(force intro: outputFreshDerivative)
with ‹y ∈ supp N› have False by(simp add: fresh_def)
then show ?case by simp
next
case(cPar2 Q' A⇩P Ψ⇩P)
have FrP: "extractFrame P = ⟨A⇩P, Ψ⇩P⟩" by fact
with ‹y ♯ P› ‹A⇩P ♯* y› have "y ♯ Ψ⇩P"
apply(drule_tac extractFrameFresh)
by(simp add: frameResChainFresh) (simp add: fresh_def name_list_supp)
from ‹Ψ ⊗ Ψ⇩P ⊳ ([(x, y)] ∙ Q) ⟼M⦇ν*(xvec1@xvec2)⦈⟨N⟩ ≺ Q'› ‹y ∈ supp N› ‹y ♯ Ψ› ‹y ♯ Ψ⇩P› ‹y ♯ M› ‹y ♯ xvec1› ‹y ♯ xvec2›
have "Ψ ⊗ Ψ⇩P ⊳ ⦇νy⦈(([(x, y)] ∙ Q)) ⟼M⦇ν*(xvec1@y#xvec2)⦈⟨N⟩ ≺ Q'" by(force intro: Open)
with ‹y ♯ Q› have "Ψ ⊗ Ψ⇩P ⊳ ⦇νx⦈Q ⟼M⦇ν*(xvec1@y#xvec2)⦈⟨N⟩ ≺ Q'"
by(simp add: alphaRes)
moreover from ‹A⇩P ♯* ([(x, y)] ∙ Q)› have "A⇩P ♯* ⦇νy⦈([(x, y)] ∙ Q)" by(auto simp add: fresh_star_def abs_fresh)
with ‹y ♯ Q› have "A⇩P ♯* ⦇νx⦈Q" by(simp add: alphaRes)
ultimately have "Ψ ⊳ P ∥ (⦇νx⦈Q) ⟼M⦇ν*(xvec1@y#xvec2)⦈⟨N⟩ ≺ (P ∥ Q')"
using FrP ‹(xvec1@xvec2) ♯* P› ‹A⇩P ♯* Ψ› ‹A⇩P ♯* M› ‹y ♯ P› ‹A⇩P ♯* (xvec1@xvec2)› ‹A⇩P ♯* y› ‹A⇩P ♯* N›
by(intro Par2) auto
moreover have "(Ψ, P ∥ Q', P ∥ Q') ∈ Rel" by(rule C1)
ultimately show ?case by blast
qed
next
case(cBrOpen M xvec1 xvec2 y N PQ)
from ‹x ♯ ¡M⦇ν*(xvec1@y#xvec2)⦈⟨N⟩› have "x ♯ xvec1" and "x ≠ y" and "x ♯ xvec2" and "x ♯ M" by simp+
from ‹xvec1 ♯* (P ∥ Q)› ‹xvec2 ♯* (P ∥ Q)› ‹y ♯ (P ∥ Q)›
have "(xvec1@xvec2) ♯* P" and "(xvec1@xvec2) ♯* Q" and "y ♯ P" and "y ♯ Q"
by simp+
from ‹Ψ ⊳ P ∥ Q ⟼ ¡M⦇ν*(xvec1@xvec2)⦈⟨([(x, y)] ∙ N)⟩ ≺ ([(x, y)] ∙ PQ)›
have "([(x, y)] ∙ Ψ) ⊳ ([(x, y)] ∙ (P ∥ Q)) ⟼ ([(x, y)] ∙ (¡M⦇ν*(xvec1@xvec2)⦈⟨([(x, y)] ∙ N)⟩ ≺ ([(x, y)] ∙ PQ)))"
by(rule semantics.eqvt)
with ‹x ♯ Ψ› ‹y ♯ Ψ› ‹x ♯ P› ‹y ♯ P› ‹x ♯ M› ‹y ♯ M› ‹x ♯ xvec1› ‹x ♯ xvec2› ‹y ♯ xvec1› ‹y ♯ xvec2›
have "Ψ ⊳ P ∥ ([(x, y)] ∙ Q) ⟼ ¡M⦇ν*(xvec1@xvec2)⦈⟨N⟩ ≺ PQ"
by(simp add: eqvts)
moreover from ‹xvec1 ♯* Ψ› ‹xvec2 ♯* Ψ› have "(xvec1@xvec2) ♯* Ψ" by simp
moreover note ‹(xvec1@xvec2) ♯* P›
moreover from ‹(xvec1@xvec2) ♯* Q› have "([(x, y)] ∙ (xvec1@xvec2)) ♯* ([(x, y)] ∙ Q)"
by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
with ‹x ♯ xvec1› ‹x ♯ xvec2› ‹y ♯ xvec1› ‹y ♯ xvec2› have "(xvec1@xvec2) ♯* ([(x, y)] ∙ Q)"
by(auto simp add: eqvts)
moreover from ‹xvec1 ♯* M› ‹xvec2 ♯* M› have "(xvec1@xvec2) ♯* M" by simp
ultimately show ?case
proof(induct rule: parBrOutputCases[where C=y])
case(cPar1 P' A⇩Q Ψ⇩Q)
from ‹y ♯ xvec1› ‹y ♯ xvec2› have "y ♯ (xvec1@xvec2)" by(auto simp add: fresh_list_append)
with ‹Ψ ⊗ Ψ⇩Q ⊳ P ⟼¡M⦇ν*(xvec1@xvec2)⦈⟨N⟩ ≺ P'› ‹(xvec1@xvec2) ♯* M› ‹y ♯ P›
‹distinct xvec1› ‹distinct xvec2› ‹xvec1 ♯* xvec2›
have "y ♯ N" by(force intro: broutputFreshDerivative)
with ‹y ∈ supp N› have False by(simp add: fresh_def)
then show ?case by simp
next
case(cPar2 Q' A⇩P Ψ⇩P)
have FrP: "extractFrame P = ⟨A⇩P, Ψ⇩P⟩" by fact
with ‹y ♯ P› ‹A⇩P ♯* y› have "y ♯ Ψ⇩P"
apply(drule_tac extractFrameFresh)
by(simp add: frameResChainFresh) (simp add: fresh_def name_list_supp)
from ‹Ψ ⊗ Ψ⇩P ⊳ ([(x, y)] ∙ Q) ⟼¡M⦇ν*(xvec1@xvec2)⦈⟨N⟩ ≺ Q'› ‹y ∈ supp N› ‹y ♯ Ψ› ‹y ♯ Ψ⇩P› ‹y ♯ M› ‹y ♯ xvec1› ‹y ♯ xvec2›
have "Ψ ⊗ Ψ⇩P ⊳ ⦇νy⦈(([(x, y)] ∙ Q)) ⟼¡M⦇ν*(xvec1@y#xvec2)⦈⟨N⟩ ≺ Q'" by(force intro: BrOpen)
with ‹y ♯ Q› have "Ψ ⊗ Ψ⇩P ⊳ ⦇νx⦈Q ⟼¡M⦇ν*(xvec1@y#xvec2)⦈⟨N⟩ ≺ Q'"
by(simp add: alphaRes)
moreover from ‹A⇩P ♯* ([(x, y)] ∙ Q)› have "A⇩P ♯* ⦇νy⦈([(x, y)] ∙ Q)" by(auto simp add: fresh_star_def abs_fresh)
with ‹y ♯ Q› have "A⇩P ♯* ⦇νx⦈Q" by(simp add: alphaRes)
ultimately have "Ψ ⊳ P ∥ (⦇νx⦈Q) ⟼¡M⦇ν*(xvec1@y#xvec2)⦈⟨N⟩ ≺ (P ∥ Q')"
using FrP ‹(xvec1@xvec2) ♯* P› ‹A⇩P ♯* Ψ› ‹A⇩P ♯* M› ‹y ♯ P› ‹A⇩P ♯* (xvec1@xvec2)› ‹A⇩P ♯* y› ‹A⇩P ♯* N›
by(intro Par2) auto
moreover have "(Ψ, P ∥ Q', P ∥ Q') ∈ Rel" by(rule C1)
ultimately show ?case by blast
next
case(cBrComm1 Ψ⇩Q P' A⇩P Ψ⇩P Q' A⇩Q)
have FrP: "extractFrame P = ⟨A⇩P, Ψ⇩P⟩" by fact
with ‹y ♯ P› ‹A⇩P ♯* y› have "y ♯ Ψ⇩P"
apply(drule_tac extractFrameFresh)
by(simp add: frameResChainFresh) (simp add: fresh_def name_list_supp)
from ‹extractFrame ([(x, y)] ∙ Q) = ⟨A⇩Q, Ψ⇩Q⟩›
have "extractFrame (⦇νy⦈([(x, y)] ∙ Q)) = ⟨(y#A⇩Q), Ψ⇩Q⟩"
by simp
with ‹y ♯ Q›
have FrQ: "extractFrame (⦇νx⦈Q) = ⟨(y#A⇩Q), Ψ⇩Q⟩"
by(simp add: alphaRes)
from ‹Ψ ⊗ Ψ⇩P ⊳ [(x, y)] ∙ Q ⟼ ¡M⦇ν*(xvec1 @ xvec2)⦈⟨N⟩ ≺ Q'› ‹y ∈ supp N› ‹y ♯ Ψ› ‹y ♯ Ψ⇩P› ‹y ♯ M› ‹y ♯ xvec1› ‹y ♯ xvec2›
have "Ψ ⊗ Ψ⇩P ⊳ ⦇νy⦈([(x, y)] ∙ Q) ⟼ ¡M⦇ν*(xvec1@y#xvec2)⦈⟨N⟩ ≺ Q'"
by(force intro: BrOpen)
with ‹y ♯ Q› have QTrans: "Ψ ⊗ Ψ⇩P ⊳ ⦇νx⦈Q ⟼¡M⦇ν*(xvec1@y#xvec2)⦈⟨N⟩ ≺ Q'"
by(simp add: alphaRes)
from ‹A⇩P ♯* ([(x, y)] ∙ Q)› have "A⇩P ♯* ⦇νy⦈([(x, y)] ∙ Q)" by(auto simp add: fresh_star_def abs_fresh)
with ‹y ♯ Q› have "A⇩P ♯* ⦇νx⦈Q" by(simp add: alphaRes)
from ‹A⇩Q ♯* ([(x, y)] ∙ Q)› have "A⇩Q ♯* ⦇νy⦈([(x, y)] ∙ Q)" by(auto simp add: fresh_star_def abs_fresh)
with ‹y ♯ Q› have "A⇩Q ♯* ⦇νx⦈Q" by(simp add: alphaRes)
moreover from ‹y ♯ Q› have "y ♯ ⦇νx⦈Q"
by (metis resChain.base resChain.step resChainFresh)
ultimately have "(y # A⇩Q) ♯* (⦇νx⦈Q)" by simp
from ‹Ψ ⊗ Ψ⇩Q ⊳ P ⟼ ¿M⦇N⦈ ≺ P'› FrP QTrans FrQ
‹A⇩P ♯* Ψ› ‹A⇩P ♯* P› ‹A⇩P ♯* ⦇νx⦈Q› ‹A⇩P ♯* M› ‹A⇩P ♯* y› ‹A⇩P ♯* A⇩Q›
‹y ♯ Ψ› ‹A⇩Q ♯* Ψ› ‹y ♯ P› ‹A⇩Q ♯* P› ‹(y#A⇩Q) ♯* ⦇νx⦈Q›
‹y ♯ M› ‹A⇩Q ♯* M› ‹y ♯ P› ‹(xvec1@xvec2) ♯* P›
have "Ψ ⊳ P ∥ (⦇νx⦈Q) ⟼¡M⦇ν*(xvec1@y#xvec2)⦈⟨N⟩ ≺ (P' ∥ Q')"
by(intro BrComm1) (assumption | simp)+
moreover have "(Ψ, P' ∥ Q', P' ∥ Q') ∈ Rel" by(rule C1)
ultimately show ?case by blast
next
case(cBrComm2 Ψ⇩Q P' A⇩P Ψ⇩P Q' A⇩Q)
from ‹y ♯ xvec1› ‹y ♯ xvec2› have "y ♯ (xvec1@xvec2)" by(auto simp add: fresh_list_append)
with ‹Ψ ⊗ Ψ⇩Q ⊳ P ⟼¡M⦇ν*(xvec1@xvec2)⦈⟨N⟩ ≺ P'› ‹(xvec1@xvec2) ♯* M› ‹y ♯ P›
‹distinct xvec1› ‹distinct xvec2› ‹xvec1 ♯* xvec2›
have "y ♯ N" by(force intro: broutputFreshDerivative)
with ‹y ∈ supp N› have False by(simp add: fresh_def)
then show ?case by simp
qed
next
case(cRes PQ)
from ‹Ψ ⊳ P ∥ Q ⟼α ≺ PQ› ‹bn α ♯* Ψ› ‹bn α ♯* P› ‹bn α ♯* Q› ‹bn α ♯* subject α›
show ?case
proof(induct rule: parCases[where C=x])
case(cPar1 P' A⇩Q Ψ⇩Q)
note ‹Ψ ⊗ Ψ⇩Q ⊳ P ⟼α ≺ P'›
moreover with ‹x ♯ P› ‹x ♯ α› ‹bn α ♯* subject α› ‹distinct(bn α)›
have "x ♯ P'" by(force dest: freeFreshDerivative)
with ‹extractFrame Q = ⟨A⇩Q, Ψ⇩Q⟩› have "extractFrame(⦇νx⦈Q) = ⟨(x#A⇩Q), Ψ⇩Q⟩"
by simp
moreover from ‹bn α ♯* Q› have "bn α ♯* (⦇νx⦈Q)" by(simp add: fresh_star_def abs_fresh)
moreover from ‹x ♯ Ψ› ‹A⇩Q ♯* Ψ› have "(x#A⇩Q) ♯* Ψ" by simp
moreover from ‹x ♯ P› ‹A⇩Q ♯* P› have "(x#A⇩Q) ♯* P" by simp
moreover from ‹x ♯ α› ‹A⇩Q ♯* α› have "(x#A⇩Q) ♯* α" by simp
ultimately have "Ψ ⊳ P ∥ ⦇νx⦈Q ⟼α ≺ (P' ∥ ⦇νx⦈Q)"
by(rule Par1)
moreover from ‹x ♯ P'› ‹x ♯ Ψ› have "(Ψ, ⦇ν*[]⦈(P' ∥ (⦇νx⦈Q)), ⦇νx⦈(⦇ν*[]⦈(P' ∥ Q))) ∈ Rel"
by(intro C2) auto
ultimately show ?case
by force
next
case(cPar2 Q' A⇩P Ψ⇩P)
have FrP: "extractFrame P = ⟨A⇩P, Ψ⇩P⟩" by fact
with ‹x ♯ P› ‹A⇩P ♯* x› have "x ♯ Ψ⇩P"
apply(drule_tac extractFrameFresh)
by(simp add: frameResChainFresh) (simp add: fresh_def name_list_supp)
from ‹Ψ ⊗ Ψ⇩P ⊳ Q ⟼α ≺ Q'› ‹x ♯ Ψ› ‹x ♯ Ψ⇩P› ‹x ♯ α›
have "Ψ ⊗ Ψ⇩P ⊳ ⦇νx⦈Q ⟼α ≺ ⦇νx⦈Q'"
by(intro Scope) auto
moreover note FrP ‹bn α ♯* P› ‹A⇩P ♯* Ψ›
moreover from ‹A⇩P ♯* Q› have "A⇩P ♯* ⦇νx⦈Q" by(simp add: fresh_star_def abs_fresh)
ultimately have "Ψ ⊳ P ∥ ⦇νx⦈Q ⟼α ≺ (P ∥ ⦇νx⦈Q')" using ‹A⇩P ♯* α›
by(rule Par2)
moreover from ‹x ♯ P› ‹x ♯ Ψ› have "(Ψ, ⦇ν*[]⦈(P ∥ (⦇νx⦈Q')), ⦇νx⦈(⦇ν*[]⦈(P ∥ Q'))) ∈ Rel"
by(intro C2) auto
ultimately show ?case
by force
next
case(cComm1 Ψ⇩Q M N P' A⇩P Ψ⇩P K xvec Q' A⇩Q)
have PTrans: "Ψ ⊗ Ψ⇩Q ⊳ P ⟼M⦇N⦈ ≺ P'" and FrP: "extractFrame P = ⟨A⇩P, Ψ⇩P⟩" by fact+
have QTrans: "Ψ ⊗ Ψ⇩P ⊳ Q ⟼K⦇ν*xvec⦈⟨N⟩ ≺ Q'" and FrQ: "extractFrame Q = ⟨A⇩Q, Ψ⇩Q⟩" by fact+
from FrP ‹x ♯ P› have "x ♯ ⟨A⇩P, Ψ⇩P⟩"
by(auto dest: extractFrameFresh)
with ‹A⇩P ♯* x› have "x ♯ Ψ⇩P" by(simp add: frameResChainFresh) (simp add: fresh_def name_list_supp)
from PTrans FrP ‹distinct A⇩P› ‹x ♯ P› ‹A⇩P ♯* Ψ› ‹A⇩P ♯* Ψ⇩Q› ‹A⇩P ♯* x› ‹A⇩P ♯* P› ‹A⇩P ♯* M› ‹A⇩Q ♯* P› ‹A⇩P ♯* A⇩Q›
obtain M' where MeqM': "(Ψ ⊗ Ψ⇩Q) ⊗ Ψ⇩P ⊢ M ↔ M'" and "x ♯ M'" and "A⇩Q ♯* M'"
by(elim inputObtainPrefix[where B="x#A⇩Q"]) (assumption | force simp add: fresh_star_list_cons)+
from MeqM' have MeqM': "Ψ ⊗ Ψ⇩P ⊗ Ψ⇩Q ⊢ M ↔ M'"
by(metis statEqEnt Associativity Composition Commutativity)
with ‹Ψ ⊗ Ψ⇩P ⊗ Ψ⇩Q ⊢ M ↔ K› have "Ψ ⊗ Ψ⇩P ⊗ Ψ⇩Q ⊢ K ↔ M'"
by(blast intro: chanEqTrans chanEqSym)
then have "(Ψ ⊗ Ψ⇩P) ⊗ Ψ⇩Q ⊢ K ↔ M'"
by(metis statEqEnt Associativity Composition Commutativity)
with QTrans FrQ ‹distinct A⇩Q› have QTrans: "Ψ ⊗ Ψ⇩P ⊳ Q ⟼M'⦇ν*xvec⦈⟨N⟩ ≺ Q'" using ‹A⇩Q ♯* Ψ› ‹A⇩Q ♯* Ψ⇩P› ‹A⇩Q ♯* Q› ‹A⇩Q ♯* K› ‹A⇩Q ♯* M'›
by(elim outputRenameSubject) (assumption | force)+
show ?case
proof(cases "x ∈ supp N")
note PTrans FrP
moreover assume "x ∈ supp N"
then have "Ψ ⊗ Ψ⇩P ⊳ ⦇νx⦈Q ⟼M'⦇ν*([]@(x#xvec))⦈⟨N⟩ ≺ Q'" using QTrans ‹x ♯ Ψ› ‹x ♯ Ψ⇩P› ‹x ♯ M'› ‹xvec ♯* x›
by(intro Open) (assumption | force simp add: fresh_list_nil)+
then have QTrans: "Ψ ⊗ Ψ⇩P ⊳ ⦇νx⦈Q ⟼M'⦇ν*(x#xvec)⦈⟨N⟩ ≺ Q'" by simp
moreover from ‹extractFrame Q = ⟨A⇩Q, Ψ⇩Q⟩› have "extractFrame(⦇νx⦈Q) = ⟨(x#A⇩Q), Ψ⇩Q⟩"
by simp
moreover note MeqM' ‹A⇩P ♯* Ψ› ‹A⇩P ♯* P›
moreover from ‹A⇩P ♯* Q› have "A⇩P ♯* (⦇νx⦈Q)" by(simp add: fresh_star_def abs_fresh)
moreover from ‹A⇩P ♯* A⇩Q› ‹A⇩P ♯* x› have "A⇩P ♯* (x#A⇩Q)"
by(simp add: fresh_star_def fresh_list_cons) (auto simp add: fresh_def name_list_supp)
moreover from ‹x ♯ Ψ› ‹A⇩Q ♯* Ψ› have "(x#A⇩Q) ♯* Ψ" by simp
moreover from ‹x ♯ P› ‹A⇩Q ♯* P› have "(x#A⇩Q) ♯* P" by simp
moreover from ‹x ♯ M'› ‹A⇩Q ♯* M'› have "(x#A⇩Q) ♯* M'" by simp
moreover from ‹A⇩Q ♯* Q› have "(x#A⇩Q) ♯* (⦇νx⦈Q)" by(simp add: fresh_star_def abs_fresh)
moreover from ‹x ♯ P› ‹xvec ♯* P› have "(x#xvec) ♯* P" by(simp)
ultimately have "Ψ ⊳ P ∥ ⦇νx⦈Q ⟼τ ≺ ⦇ν*(x#xvec)⦈(P' ∥ Q')" using ‹A⇩P ♯* M›
by(intro Comm1) (assumption | simp)+
moreover have "(Ψ, ⦇νx⦈(⦇ν*xvec⦈(P' ∥ Q')), ⦇νx⦈(⦇ν*xvec⦈(P' ∥ Q'))) ∈ Rel" by(rule C1)
ultimately show ?case by force
next
note PTrans FrP
moreover assume "x ∉ supp N"
then have "x ♯ N" by(simp add: fresh_def)
with QTrans ‹x ♯ Ψ› ‹x ♯ Ψ⇩P› ‹x ♯ M'› ‹xvec ♯* x›
have QTrans: "Ψ ⊗ Ψ⇩P ⊳ ⦇νx⦈Q ⟼M'⦇ν*xvec⦈⟨N⟩ ≺ ⦇νx⦈Q'"
by(intro Scope) (assumption | force)+
moreover from PTrans ‹x ♯ P› ‹x ♯ N› have "x ♯ P'" by(rule inputFreshDerivative)
with ‹extractFrame Q = ⟨A⇩Q, Ψ⇩Q⟩› have "extractFrame(⦇νx⦈Q) = ⟨(x#A⇩Q), Ψ⇩Q⟩"
by simp
moreover note MeqM' ‹A⇩P ♯* Ψ› ‹A⇩P ♯* P›
moreover from ‹A⇩P ♯* Q› have "A⇩P ♯* (⦇νx⦈Q)" by(simp add: fresh_star_def abs_fresh)
moreover from ‹A⇩P ♯* A⇩Q› ‹A⇩P ♯* x› have "A⇩P ♯* (x#A⇩Q)"
by(simp add: fresh_star_def fresh_list_cons) (auto simp add: fresh_def name_list_supp)
moreover from ‹x ♯ Ψ› ‹A⇩Q ♯* Ψ› have "(x#A⇩Q) ♯* Ψ" by simp
moreover from ‹x ♯ P› ‹A⇩Q ♯* P› have "(x#A⇩Q) ♯* P" by simp
moreover from ‹x ♯ M'› ‹A⇩Q ♯* M'› have "(x#A⇩Q) ♯* M'" by simp
moreover from ‹A⇩Q ♯* Q› have "(x#A⇩Q) ♯* (⦇νx⦈Q)" by(simp add: fresh_star_def abs_fresh)
moreover from ‹x ♯ P› ‹xvec ♯* P› have "(x#xvec) ♯* P" by(simp)
ultimately have "Ψ ⊳ P ∥ ⦇νx⦈Q ⟼τ ≺ ⦇ν*xvec⦈(P' ∥ ⦇νx⦈Q')" using ‹A⇩P ♯* M›
by(intro Comm1) (assumption | simp)+
moreover from ‹x ♯ Ψ› ‹x ♯ P'› ‹xvec ♯* Ψ› have "(Ψ, ⦇ν*xvec⦈(P' ∥ ⦇νx⦈Q'), ⦇νx⦈(⦇ν*xvec⦈(P' ∥ Q'))) ∈ Rel" by(rule C2)
ultimately show ?case by blast
qed
next
case(cComm2 Ψ⇩Q M xvec N P' A⇩P Ψ⇩P K Q' A⇩Q)
have PTrans: "Ψ ⊗ Ψ⇩Q ⊳ P ⟼M⦇ν*xvec⦈⟨N⟩ ≺ P'" and FrP: "extractFrame P = ⟨A⇩P, Ψ⇩P⟩" by fact+
have QTrans: "Ψ ⊗ Ψ⇩P ⊳ Q ⟼K⦇N⦈ ≺ Q'" and FrQ: "extractFrame Q = ⟨A⇩Q, Ψ⇩Q⟩" by fact+
from FrP ‹x ♯ P› have "x ♯ ⟨A⇩P, Ψ⇩P⟩"
by(auto dest: extractFrameFresh)
with ‹A⇩P ♯* x› have "x ♯ Ψ⇩P" by(simp add: frameResChainFresh) (simp add: fresh_def name_list_supp)
from PTrans
have "Ψ ⊗ Ψ⇩Q ⊳ P ⟼ ROut M (⦇ν*xvec⦈N ≺' P')"
by(simp add: residualInject)
with FrP ‹distinct A⇩P› ‹x ♯ P› ‹A⇩P ♯* Ψ› ‹A⇩P ♯* Ψ⇩Q› ‹A⇩P ♯* x› ‹A⇩P ♯* P› ‹A⇩P ♯* M› ‹A⇩Q ♯* P› ‹A⇩P ♯* A⇩Q› ‹xvec ♯* M› ‹distinct xvec›
obtain M' where MeqM': "(Ψ ⊗ Ψ⇩Q) ⊗ Ψ⇩P ⊢ M ↔ M'" and "x ♯ M'" and "A⇩Q ♯* M'"
by(elim outputObtainPrefix[where B="x#A⇩Q"]) (assumption | force simp add: fresh_star_list_cons)+
from MeqM' have MeqM': "Ψ ⊗ Ψ⇩P ⊗ Ψ⇩Q ⊢ M ↔ M'"
by(metis statEqEnt Associativity Commutativity Composition)
with ‹Ψ ⊗ Ψ⇩P ⊗ Ψ⇩Q ⊢ M ↔ K› have "Ψ ⊗ Ψ⇩P ⊗ Ψ⇩Q ⊢ K ↔ M'"
by(blast intro: chanEqTrans chanEqSym)
then have "(Ψ ⊗ Ψ⇩P) ⊗ Ψ⇩Q ⊢ K ↔ M'"
by(metis statEqEnt Associativity Commutativity Composition)
with QTrans FrQ ‹distinct A⇩Q› have QTrans: "Ψ ⊗ Ψ⇩P ⊳ Q ⟼M'⦇N⦈ ≺ Q'" using ‹A⇩Q ♯* Ψ› ‹A⇩Q ♯* Ψ⇩P› ‹A⇩Q ♯* Q› ‹A⇩Q ♯* K› ‹A⇩Q ♯* M'›
by(auto intro: inputRenameSubject)
from PTrans ‹x ♯ P› ‹xvec ♯* x› ‹distinct xvec› ‹xvec ♯* M›
have "x ♯ N" and "x ♯ P'" by(force intro: outputFreshDerivative)+
from QTrans ‹x ♯ Ψ› ‹x ♯ Ψ⇩P› ‹x ♯ M'› ‹x ♯ N› have QTrans: "Ψ ⊗ Ψ⇩P ⊳ ⦇νx⦈Q ⟼M'⦇N⦈ ≺ ⦇νx⦈Q'"
by(intro Scope) (assumption | force)+
note PTrans FrP QTrans
moreover with ‹extractFrame Q = ⟨A⇩Q, Ψ⇩Q⟩› have "extractFrame(⦇νx⦈Q) = ⟨(x#A⇩Q), Ψ⇩Q⟩"
by simp
moreover note MeqM' ‹A⇩P ♯* Ψ› ‹A⇩P ♯* P›
moreover from ‹A⇩P ♯* Q› have "A⇩P ♯* (⦇νx⦈Q)" by(simp add: fresh_star_def abs_fresh)
moreover from ‹A⇩P ♯* A⇩Q› ‹A⇩P ♯* x› have "A⇩P ♯* (x#A⇩Q)"
by(simp add: fresh_star_def fresh_list_cons) (auto simp add: fresh_def name_list_supp)
moreover from ‹x ♯ Ψ› ‹A⇩Q ♯* Ψ› have "(x#A⇩Q) ♯* Ψ" by simp
moreover from ‹x ♯ P› ‹A⇩Q ♯* P› have "(x#A⇩Q) ♯* P" by simp
moreover from ‹x ♯ M'› ‹A⇩Q ♯* M'› have "(x#A⇩Q) ♯* M'" by simp
moreover from ‹A⇩Q ♯* Q› have "(x#A⇩Q) ♯* (⦇νx⦈Q)" by(simp add: fresh_star_def abs_fresh)
moreover from ‹xvec ♯* Q› have "(x#xvec) ♯* (⦇νx⦈Q)" by(simp add: abs_fresh fresh_star_def)
ultimately have "Ψ ⊳ P ∥ ⦇νx⦈Q ⟼τ ≺ ⦇ν*xvec⦈(P' ∥ ⦇νx⦈Q')" using ‹A⇩P ♯* M›
by(intro Comm2) (assumption | simp)+
moreover from ‹x ♯ Ψ› ‹x ♯ P'› ‹xvec ♯* Ψ› have "(Ψ, ⦇ν*xvec⦈(P' ∥ ⦇νx⦈Q'), ⦇νx⦈(⦇ν*xvec⦈(P' ∥ Q'))) ∈ Rel" by(rule C2)
ultimately show ?case by blast
next
case(cBrMerge Ψ⇩Q M N P' A⇩P Ψ⇩P Q' A⇩Q)
have PTrans: "Ψ ⊗ Ψ⇩Q ⊳ P ⟼¿M⦇N⦈ ≺ P'" and FrP: "extractFrame P = ⟨A⇩P, Ψ⇩P⟩" by fact+
have QTrans: "Ψ ⊗ Ψ⇩P ⊳ Q ⟼¿M⦇N⦈ ≺ Q'" and FrQ: "extractFrame Q = ⟨A⇩Q, Ψ⇩Q⟩" by fact+
from FrP ‹x ♯ P› have "x ♯ ⟨A⇩P, Ψ⇩P⟩"
by(auto dest: extractFrameFresh)
with ‹A⇩P ♯* x› have "x ♯ Ψ⇩P"
by(simp add: frameResChainFresh) (simp add: fresh_def name_list_supp)
from ‹x ♯ α› ‹α = ¿M⦇N⦈› have "x ♯ M" and "x ♯ N" by simp+
from PTrans ‹x ♯ P› ‹x ♯ N›
have "x ♯ P'"
by(rule brinputFreshDerivative)
from QTrans ‹x ♯ Ψ› ‹x ♯ Ψ⇩P› ‹x ♯ M› ‹x ♯ N› have QTrans: "Ψ ⊗ Ψ⇩P ⊳ ⦇νx⦈Q ⟼¿M⦇N⦈ ≺ ⦇νx⦈Q'"
by(intro Scope) (assumption | force)+
note PTrans FrP QTrans
moreover with ‹extractFrame Q = ⟨A⇩Q, Ψ⇩Q⟩› have "extractFrame(⦇νx⦈Q) = ⟨(x#A⇩Q), Ψ⇩Q⟩"
by simp
moreover note ‹A⇩P ♯* Ψ› ‹A⇩P ♯* P›
moreover from ‹A⇩P ♯* Q› have "A⇩P ♯* (⦇νx⦈Q)" by(simp add: fresh_star_def abs_fresh)
moreover from ‹A⇩P ♯* A⇩Q› ‹A⇩P ♯* x› have "A⇩P ♯* (x#A⇩Q)"
by(simp add: fresh_star_def fresh_list_cons) (auto simp add: fresh_def name_list_supp)
moreover from ‹x ♯ Ψ› ‹A⇩Q ♯* Ψ› have "(x#A⇩Q) ♯* Ψ" by simp
moreover from ‹x ♯ P› ‹A⇩Q ♯* P› have "(x#A⇩Q) ♯* P" by simp
moreover from ‹x ♯ M› ‹A⇩Q ♯* M› have "(x#A⇩Q) ♯* M" by simp
moreover from ‹A⇩Q ♯* Q› have "(x#A⇩Q) ♯* (⦇νx⦈Q)" by(simp add: fresh_star_def abs_fresh)
ultimately have Trans: "Ψ ⊳ P ∥ ⦇νx⦈Q ⟼¿M⦇N⦈ ≺ (P' ∥ ⦇νx⦈Q')" using ‹A⇩P ♯* M›
by(intro BrMerge) (assumption | simp)+
from ‹x ♯ Ψ› ‹x ♯ P'› have "(Ψ, (P' ∥ (⦇ν*[x]⦈Q')), ⦇ν*[x]⦈(P' ∥ Q')) ∈ Rel"
by(intro C3) simp+
then have Relation: "(Ψ, (P' ∥ (⦇νx⦈Q')), ⦇νx⦈(P' ∥ Q')) ∈ Rel" by simp
with Trans Relation show ?case by blast
next
case(cBrComm1 Ψ⇩Q M N P' A⇩P Ψ⇩P xvec Q' A⇩Q)
from ‹x ♯ α› ‹¡M⦇ν*xvec⦈⟨N⟩ = α› have "x ♯ M" and "x ♯ xvec" and "x ♯ N" by force+
have PTrans: "Ψ ⊗ Ψ⇩Q ⊳ P ⟼¿M⦇N⦈ ≺ P'" and FrP: "extractFrame P = ⟨A⇩P, Ψ⇩P⟩" by fact+
have QTrans: "Ψ ⊗ Ψ⇩P ⊳ Q ⟼¡M⦇ν*xvec⦈⟨N⟩ ≺ Q'" and FrQ: "extractFrame Q = ⟨A⇩Q, Ψ⇩Q⟩" by fact+
from FrP ‹x ♯ P› have "x ♯ ⟨A⇩P, Ψ⇩P⟩"
by(auto dest: extractFrameFresh)
with ‹A⇩P ♯* x› have "x ♯ Ψ⇩P"
by(simp add: frameResChainFresh) (simp add: fresh_def name_list_supp)
show ?case
proof(cases "x ∈ supp N")
note PTrans FrP
moreover assume "x ∈ supp N"
with ‹x ♯ N› have False
by(simp add: fresh_def)
then show ?case
by simp
next
note PTrans FrP
moreover assume "x ∉ supp N"
from QTrans ‹x ♯ Ψ› ‹x ♯ Ψ⇩P› ‹x ♯ M› ‹x ♯ N› ‹x ♯ xvec›
have QTrans: "Ψ ⊗ Ψ⇩P ⊳ ⦇νx⦈Q ⟼¡M⦇ν*xvec⦈⟨N⟩ ≺ ⦇νx⦈Q'"
by(intro Scope) (assumption | force)+
moreover from PTrans ‹x ♯ P› ‹x ♯ N› have "x ♯ P'" by(rule brinputFreshDerivative)
with ‹extractFrame Q = ⟨A⇩Q, Ψ⇩Q⟩› have "extractFrame(⦇νx⦈Q) = ⟨(x#A⇩Q), Ψ⇩Q⟩"
by simp
moreover note ‹A⇩P ♯* Ψ› ‹A⇩P ♯* P›
moreover from ‹A⇩P ♯* Q› have "A⇩P ♯* (⦇νx⦈Q)" by(simp add: fresh_star_def abs_fresh)
moreover from ‹A⇩P ♯* A⇩Q› ‹A⇩P ♯* x› have "A⇩P ♯* (x#A⇩Q)"
by(simp add: fresh_star_def fresh_list_cons) (auto simp add: fresh_def name_list_supp)
moreover from ‹x ♯ Ψ› ‹A⇩Q ♯* Ψ› have "(x#A⇩Q) ♯* Ψ" by simp
moreover from ‹x ♯ P› ‹A⇩Q ♯* P› have "(x#A⇩Q) ♯* P" by simp
moreover from ‹x ♯ M› ‹A⇩Q ♯* M› have "(x#A⇩Q) ♯* M" by simp
moreover from ‹A⇩Q ♯* Q› have "(x#A⇩Q) ♯* (⦇νx⦈Q)" by(simp add: fresh_star_def abs_fresh)
moreover from ‹x ♯ P› ‹xvec ♯* P› have "(x#xvec) ♯* P" by(simp)
ultimately have Trans: "Ψ ⊳ P ∥ ⦇νx⦈Q ⟼¡M⦇ν*xvec⦈⟨N⟩ ≺ (P' ∥ ⦇νx⦈Q')" using ‹A⇩P ♯* M›
by(intro BrComm1) (assumption | simp)+
from ‹x ♯ Ψ› ‹x ♯ P'› have "(Ψ, (P' ∥ (⦇ν*[x]⦈Q')), ⦇ν*[x]⦈(P' ∥ Q')) ∈ Rel"
by(intro C3) simp+
then have Relation: "(Ψ, (P' ∥ (⦇νx⦈Q')), ⦇νx⦈(P' ∥ Q')) ∈ Rel" by simp
from Trans Relation show ?case by blast
qed
next
case(cBrComm2 Ψ⇩Q M xvec N P' A⇩P Ψ⇩P Q' A⇩Q)
from ‹x ♯ α› ‹¡M⦇ν*xvec⦈⟨N⟩ = α› have "x ♯ M" and "x ♯ xvec" and "x ♯ N" by force+
have PTrans: "Ψ ⊗ Ψ⇩Q ⊳ P ⟼¡M⦇ν*xvec⦈⟨N⟩ ≺ P'" and FrP: "extractFrame P = ⟨A⇩P, Ψ⇩P⟩" by fact+
have QTrans: "Ψ ⊗ Ψ⇩P ⊳ Q ⟼¿M⦇N⦈ ≺ Q'" and FrQ: "extractFrame Q = ⟨A⇩Q, Ψ⇩Q⟩" by fact+
from FrP ‹x ♯ P› have "x ♯ ⟨A⇩P, Ψ⇩P⟩"
by(auto dest: extractFrameFresh)
with ‹A⇩P ♯* x› have "x ♯ Ψ⇩P"
by(simp add: frameResChainFresh) (simp add: fresh_def name_list_supp)
from PTrans ‹x ♯ P› ‹x ♯ xvec› ‹distinct xvec› ‹xvec ♯* M›
have "x ♯ N" and "x ♯ P'" by(force intro: broutputFreshDerivative)+
from QTrans ‹x ♯ Ψ› ‹x ♯ Ψ⇩P› ‹x ♯ M› ‹x ♯ N› have QTrans: "Ψ ⊗ Ψ⇩P ⊳ ⦇νx⦈Q ⟼¿M⦇N⦈ ≺ ⦇νx⦈Q'"
by(intro Scope) (assumption | force)+
note PTrans FrP QTrans
moreover with ‹extractFrame Q = ⟨A⇩Q, Ψ⇩Q⟩› have "extractFrame(⦇νx⦈Q) = ⟨(x#A⇩Q), Ψ⇩Q⟩"
by simp
moreover note ‹A⇩P ♯* Ψ› ‹A⇩P ♯* P›
moreover from ‹A⇩P ♯* Q› have "A⇩P ♯* (⦇νx⦈Q)" by(simp add: fresh_star_def abs_fresh)
moreover from ‹A⇩P ♯* A⇩Q› ‹A⇩P ♯* x› have "A⇩P ♯* (x#A⇩Q)"
by(simp add: fresh_star_def fresh_list_cons) (auto simp add: fresh_def name_list_supp)
moreover from ‹x ♯ Ψ› ‹A⇩Q ♯* Ψ› have "(x#A⇩Q) ♯* Ψ" by simp
moreover from ‹x ♯ P› ‹A⇩Q ♯* P› have "(x#A⇩Q) ♯* P" by simp
moreover from ‹x ♯ M› ‹A⇩Q ♯* M› have "(x#A⇩Q) ♯* M" by simp
moreover from ‹A⇩Q ♯* Q› have "(x#A⇩Q) ♯* (⦇νx⦈Q)" by(simp add: fresh_star_def abs_fresh)
moreover from ‹xvec ♯* Q› have "(x#xvec) ♯* (⦇νx⦈Q)" by(simp add: abs_fresh fresh_star_def)
ultimately have Trans: "Ψ ⊳ P ∥ ⦇νx⦈Q ⟼¡M⦇ν*xvec⦈⟨N⟩ ≺ (P' ∥ ⦇νx⦈Q')" using ‹A⇩P ♯* M›
by(intro BrComm2) (assumption | simp)+
from ‹x ♯ Ψ› ‹x ♯ P'› have "(Ψ, (P' ∥ (⦇ν*[x]⦈Q')), ⦇ν*[x]⦈(P' ∥ Q')) ∈ Rel"
by(intro C3) simp+
then have Relation: "(Ψ, (P' ∥ ⦇νx⦈Q'), ⦇νx⦈(P' ∥ Q')) ∈ Rel" by simp
from Trans Relation show ?case by blast
qed
next
case(cBrClose M xvec N PQ)
from ‹xvec ♯* (P ∥ Q)› have "xvec ♯* P" and "xvec ♯* Q" by simp+
note ‹Ψ ⊳ P ∥ Q ⟼ ¡M⦇ν*xvec⦈⟨N⟩ ≺ PQ› ‹xvec ♯* Ψ› ‹xvec ♯* P› ‹xvec ♯* Q› ‹xvec ♯* M›
then show ?case
proof(induct rule: parBrOutputCases[where C=x])
case(cPar1 P' A⇩Q Ψ⇩Q)
from ‹Ψ ⊗ Ψ⇩Q ⊳ P ⟼ ¡M⦇ν*xvec⦈⟨N⟩ ≺ P'› ‹xvec ♯* M› ‹x ♯ P›
have "x ♯ M"
by(rule brOutputFreshSubject)
with ‹x ∈ supp M› have False
by(simp add: fresh_def)
then show ?case
by simp
next
case(cPar2 Q' A⇩P Ψ⇩P)
from ‹A⇩P ♯* x› have "x ♯ A⇩P" by simp
with ‹x ♯ P› ‹extractFrame P = ⟨A⇩P, Ψ⇩P⟩›
have "x ♯ Ψ⇩P"
apply(drule_tac extractFrameFresh)
by(simp add: frameResChainFresh) (simp add: fresh_def name_list_supp)
from ‹x ♯ Ψ› ‹x ♯ Ψ⇩P› have "x ♯ (Ψ ⊗ Ψ⇩P)" by force
from ‹Ψ ⊗ Ψ⇩P ⊳ Q ⟼ ¡M⦇ν*xvec⦈⟨N⟩ ≺ Q'› ‹x ∈ supp M› ‹x ♯ (Ψ ⊗ Ψ⇩P)›
have QTrans: "Ψ ⊗ Ψ⇩P ⊳ ⦇νx⦈Q ⟼ τ ≺ ⦇νx⦈(⦇ν*xvec⦈Q')"
by(rule BrClose)
with ‹extractFrame P = ⟨A⇩P, Ψ⇩P⟩› ‹A⇩P ♯* Ψ› ‹x ♯ A⇩P› ‹A⇩P ♯* Q›
have Trans: "Ψ ⊳ (P ∥ ⦇νx⦈Q) ⟼ τ ≺ (P ∥ ⦇νx⦈(⦇ν*xvec⦈Q'))"
by(intro Par2) (assumption | simp)+
from ‹x ♯ P› ‹xvec ♯* P› have "(x#xvec) ♯* P" by simp
moreover from ‹x ♯ Ψ› ‹xvec ♯* Ψ› have "(x#xvec) ♯* Ψ" by simp
ultimately have "(Ψ, (P ∥ (⦇ν*(x#xvec)⦈Q')), (⦇ν*(x#xvec)⦈P ∥ Q')) ∈ Rel"
by(rule C3)
then have Relation: "(Ψ, (P ∥ ⦇νx⦈(⦇ν*xvec⦈Q')), ⦇νx⦈(⦇ν*xvec⦈P ∥ Q')) ∈ Rel"
by simp
from Trans Relation
show ?case
by blast
next
case(cBrComm1 Ψ⇩Q P' A⇩P Ψ⇩P Q' A⇩Q)
from ‹Ψ ⊗ Ψ⇩Q ⊳ P ⟼ ¿M⦇N⦈ ≺ P'› ‹x ♯ P›
have "x ♯ M"
by(rule brInputFreshSubject)
with ‹x ∈ supp M› have False
by(simp add: fresh_def)
then show ?case
by simp
next
case(cBrComm2 Ψ⇩Q P' A⇩P Ψ⇩P Q' A⇩Q)
from ‹Ψ ⊗ Ψ⇩Q ⊳ P ⟼ ¡M⦇ν*xvec⦈⟨N⟩ ≺ P'› ‹xvec ♯* M› ‹x ♯ P›
have "x ♯ M"
by(rule brOutputFreshSubject)
with ‹x ∈ supp M› have False
by(simp add: fresh_def)
then show ?case
by simp
qed
qed
qed
qed
lemma simParComm:
fixes Ψ :: 'b
and P :: "('a, 'b, 'c) psi"
and Q :: "('a, 'b, 'c) psi"
and Rel :: "('b × ('a, 'b, 'c) psi × ('a, 'b, 'c) psi) set"
assumes "eqvt Rel"
and C1: "⋀Ψ' R S. (Ψ', R ∥ S, S ∥ R) ∈ Rel"
and C2: "⋀Ψ' R S xvec. ⟦(Ψ', R, S) ∈ Rel; xvec ♯* Ψ'⟧ ⟹ (Ψ', ⦇ν*xvec⦈R, ⦇ν*xvec⦈S) ∈ Rel"
shows "Ψ ⊳ P ∥ Q ↝[Rel] Q ∥ P"
using ‹eqvt Rel›
proof(induct rule: simI[of _ _ _ _ "()"])
case(cSim α PQ)
from ‹bn α ♯* (P ∥ Q)› have "bn α ♯* Q" and "bn α ♯* P" by simp+
with ‹Ψ ⊳ Q ∥ P ⟼α ≺ PQ› ‹bn α ♯* Ψ› show ?case using ‹bn α ♯* subject α›
proof(induct rule: parCases[where C="()"])
case(cPar1 Q' A⇩P Ψ⇩P)
from ‹Ψ ⊗ Ψ⇩P ⊳ Q ⟼α≺ Q'› ‹extractFrame P = ⟨A⇩P, Ψ⇩P⟩› ‹bn α ♯* P› ‹A⇩P ♯* Ψ› ‹A⇩P ♯* Q› ‹A⇩P ♯* α›
have "Ψ ⊳ P ∥ Q ⟼α ≺ (P ∥ Q')" by(rule Par2)
moreover have "(Ψ, P ∥ Q', Q' ∥ P) ∈ Rel" by(rule C1)
ultimately show ?case by blast
next
case(cPar2 P' A⇩Q Ψ⇩Q)
from ‹Ψ ⊗ Ψ⇩Q ⊳ P ⟼α ≺ P'› ‹extractFrame Q = ⟨A⇩Q, Ψ⇩Q⟩› ‹bn α ♯* Q› ‹A⇩Q ♯* Ψ› ‹A⇩Q ♯* P› ‹A⇩Q ♯* α›
have "Ψ ⊳ P ∥ Q ⟼α ≺ (P' ∥ Q)" by(rule Par1)
moreover have "(Ψ, P' ∥ Q, Q ∥ P') ∈ Rel" by(rule C1)
ultimately show ?case by blast
next
case(cComm1 Ψ⇩P M N Q' A⇩Q Ψ⇩Q K xvec P' A⇩P)
note ‹Ψ ⊗ Ψ⇩Q ⊳ P ⟼K⦇ν*xvec⦈⟨N⟩ ≺ P'› ‹extractFrame P = ⟨A⇩P, Ψ⇩P⟩›
‹Ψ ⊗ Ψ⇩P ⊳ Q ⟼M⦇N⦈ ≺ Q'› ‹extractFrame Q = ⟨A⇩Q, Ψ⇩Q⟩›
moreover from ‹Ψ ⊗ Ψ⇩Q ⊗ Ψ⇩P ⊢ M ↔ K›
have "Ψ ⊗ Ψ⇩Q ⊗ Ψ⇩P ⊢ K ↔ M"
by(rule chanEqSym)
then have "Ψ ⊗ Ψ⇩P ⊗ Ψ⇩Q ⊢ K ↔ M"
by(blast intro: statEqEnt compositionSym Commutativity)
ultimately have "Ψ ⊳ P ∥ Q ⟼τ ≺ ⦇ν*xvec⦈(P' ∥ Q')"
using ‹A⇩P ♯* Ψ› ‹A⇩P ♯* P› ‹A⇩P ♯* Q› ‹A⇩Q ♯* A⇩P› ‹A⇩Q ♯* Ψ› ‹A⇩Q ♯* P› ‹A⇩Q ♯* Q› ‹xvec ♯* Q› ‹A⇩P ♯* K› ‹A⇩Q ♯* M›
by(intro Comm2) (assumption | simp)+
moreover have "(Ψ, P' ∥ Q', Q' ∥ P') ∈ Rel" by(rule C1)
then have "(Ψ, ⦇ν*xvec⦈(P' ∥ Q'), ⦇ν*xvec⦈(Q' ∥ P')) ∈ Rel" using ‹xvec ♯* Ψ› by(rule C2)
ultimately show ?case by blast
next
case(cComm2 Ψ⇩P M xvec N Q' A⇩Q Ψ⇩Q K P' A⇩P)
note ‹Ψ ⊗ Ψ⇩Q ⊳ P ⟼K⦇N⦈ ≺ P'› ‹extractFrame P = ⟨A⇩P, Ψ⇩P⟩›
‹Ψ ⊗ Ψ⇩P ⊳ Q ⟼M⦇ν*xvec⦈⟨N⟩ ≺ Q'› ‹extractFrame Q = ⟨A⇩Q, Ψ⇩Q⟩›
moreover from ‹Ψ ⊗ Ψ⇩Q ⊗ Ψ⇩P ⊢ M ↔ K›
have "Ψ ⊗ Ψ⇩Q ⊗ Ψ⇩P ⊢ K ↔ M"
by(rule chanEqSym)
then have "Ψ ⊗ Ψ⇩P ⊗ Ψ⇩Q ⊢ K ↔ M"
by(blast intro: statEqEnt compositionSym Commutativity)
ultimately have "Ψ ⊳ P ∥ Q ⟼τ ≺ ⦇ν*xvec⦈(P' ∥ Q')"
using ‹A⇩P ♯* Ψ› ‹A⇩P ♯* P› ‹A⇩P ♯* Q› ‹A⇩Q ♯* A⇩P› ‹A⇩Q ♯* Ψ› ‹A⇩Q ♯* P› ‹A⇩Q ♯* Q› ‹xvec ♯* P› ‹A⇩P ♯* K› ‹A⇩Q ♯* M›
by(intro Comm1) (assumption | simp add: freshChainSym)+
moreover have "(Ψ, P' ∥ Q', Q' ∥ P') ∈ Rel" by(rule C1)
then have "(Ψ, ⦇ν*xvec⦈(P' ∥ Q'), ⦇ν*xvec⦈(Q' ∥ P')) ∈ Rel" using ‹xvec ♯* Ψ› by(rule C2)
ultimately show ?case by blast
next
case(cBrMerge Ψ⇩P M N Q' A⇩Q Ψ⇩Q P' A⇩P)
then have "Ψ ⊳ P ∥ Q ⟼ ¿M⦇N⦈ ≺ P' ∥ Q'"
by(intro BrMerge) (assumption|simp)+
then show ?case by(blast intro: C1)
next
case(cBrComm1 Ψ⇩P M N Q' A⇩Q Ψ⇩Q xvec P' A⇩P)
then have "Ψ ⊳ P ∥ Q ⟼ ¡M⦇ν*xvec⦈⟨N⟩ ≺ P' ∥ Q'"
by(intro BrComm2) (assumption|simp)+
then show ?case by(blast intro: C1)
next
case(cBrComm2 Ψ⇩P M xvec N Q' A⇩Q Ψ⇩Q P' A⇩P)
then have "Ψ ⊳ P ∥ Q ⟼ ¡M⦇ν*xvec⦈⟨N⟩ ≺ P' ∥ Q'"
by(intro BrComm1) (assumption|simp)+
then show ?case by(blast intro: C1)
qed
qed
lemma bangExtLeft:
fixes Ψ :: 'b
and P :: "('a, 'b, 'c) psi"
assumes "guarded P"
and "⋀Ψ' Q. (Ψ', Q, Q) ∈ Rel"
shows "Ψ ⊳ !P ↝[Rel] P ∥ !P"
using assms
by(auto simp add: simulation_def dest: Bang)
lemma bangExtRight:
fixes Ψ :: 'b
and P :: "('a, 'b, 'c) psi"
assumes C1: "⋀Ψ' Q. (Ψ', Q, Q) ∈ Rel"
shows "Ψ ⊳ P ∥ !P ↝[Rel] !P"
proof -
{
fix α P'
assume "Ψ ⊳ !P ⟼α ≺ P'"
then have "Ψ ⊳ P ∥ !P ⟼α ≺ P'"
apply -
by(ind_cases "Ψ ⊳ !P ⟼α ≺ P'") (auto simp add: psi.inject)
moreover have "(Ψ, P', P') ∈ Rel" by(rule C1)
ultimately have "∃P''. Ψ ⊳ P ∥ !P ⟼α ≺ P'' ∧ (Ψ, P'', P') ∈ Rel"
by blast
}
then show ?thesis
by(auto simp add: simulation_def)
qed
end
end