diff --git a/thys/Broadcast_Psi/Bisim_Pres.thy b/thys/Broadcast_Psi/Bisim_Pres.thy new file mode 100644 --- /dev/null +++ b/thys/Broadcast_Psi/Bisim_Pres.thy @@ -0,0 +1,613 @@ +theory Bisim_Pres + imports Bisimulation Sim_Pres +begin + +text \This file is a (heavily modified) variant of the theory {\it Psi\_Calculi.Bisim\_Pres} +from~\cite{DBLP:journals/afp/Bengtson12}.\ + +context env begin + +lemma bisimInputPres: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and Q :: "('a, 'b, 'c) psi" + and M :: 'a + and xvec :: "name list" + and N :: 'a + +assumes "\Tvec. length xvec = length Tvec \ \ \ P[xvec::=Tvec] \ Q[xvec::=Tvec]" + +shows "\ \ M\\*xvec N\.P \ M\\*xvec N\.Q" +proof - + let ?X = "{(\, M\\*xvec N\.P, M\\*xvec N\.Q) | \ M xvec N P Q. \Tvec. length xvec = length Tvec \ \ \ P[xvec::=Tvec] \ Q[xvec::=Tvec]}" + + from assms have "(\, M\\*xvec N\.P, M\\*xvec N\.Q) \ ?X" by blast + then show ?thesis + proof(coinduct rule: bisimCoinduct) + case(cStatEq \ P Q) + then show ?case by auto + next + case(cSim \ P Q) + then show ?case by(blast intro: inputPres) + next + case(cExt \ P Q \') + then show ?case by(blast dest: bisimE) + next + case(cSym \ P Q) + then show ?case by(blast dest: bisimE) + qed +qed + +lemma bisimOutputPres: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and Q :: "('a, 'b, 'c) psi" + and M :: 'a + and N :: 'a + +assumes "\ \ P \ Q" + +shows "\ \ M\N\.P \ M\N\.Q" +proof - + let ?X = "{(\, M\N\.P, M\N\.Q) | \ M N P Q. \ \ P \ Q}" + from \\ \ P \ Q\ have "(\, M\N\.P, M\N\.Q) \ ?X" by auto + then show ?thesis + by(coinduct rule: bisimCoinduct, auto) (blast intro: outputPres dest: bisimE)+ +qed + +lemma bisimCasePres: + fixes \ :: 'b + and CsP :: "('c \ ('a, 'b, 'c) psi) list" + and CsQ :: "('c \ ('a, 'b, 'c) psi) list" + +assumes "\\ P. (\, P) \ set CsP \ \Q. (\, Q) \ set CsQ \ guarded Q \ \ \ P \ Q" + and "\\ Q. (\, Q) \ set CsQ \ \P. (\, P) \ set CsP \ guarded P \ \ \ P \ Q" + +shows "\ \ Cases CsP \ Cases CsQ" +proof - + let ?X = "{(\, Cases CsP, Cases CsQ) | \ CsP CsQ. (\\ P. (\, P) \ set CsP \ (\Q. (\, Q) \ set CsQ \ guarded Q \ \ \ P \ Q)) \ + (\\ Q. (\, Q) \ set CsQ \ (\P. (\, P) \ set CsP \ guarded P \ \ \ P \ Q))}" + from assms have "(\, Cases CsP, Cases CsQ) \ ?X" by auto + then show ?thesis + proof(coinduct rule: bisimCoinduct) + case(cStatEq \ P Q) + then show ?case by auto + next + case(cSim \ CasesP CasesQ) + then obtain CsP CsQ where C1: "\\ Q. (\, Q) \ set CsQ \ \P. (\, P) \ set CsP \ guarded P \ \ \ P \ Q" + and A: "CasesP = Cases CsP" and B: "CasesQ = Cases CsQ" + by auto + note C1 + moreover have "\\ P Q. \ \ P \ Q \ \ \ P \[bisim] Q" by(rule bisimE) + moreover have "bisim \ ?X \ bisim" by blast + ultimately have "\ \ Cases CsP \[(?X \ bisim)] Cases CsQ" + by(rule casePres) + then show ?case using A B by blast + next + case(cExt \ P Q) + then show ?case by(blast dest: bisimE) + next + case(cSym \ P Q) + then show ?case by(blast dest: bisimE) + qed +qed + +lemma bisimResPres: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and Q :: "('a, 'b, 'c) psi" + and x :: name + +assumes "\ \ P \ Q" + and "x \ \" + +shows "\ \ \\x\P \ \\x\Q" +proof - + let ?X = "{(\, \\*xvec\P, \\*xvec\Q) | \ xvec P Q. \ \ P \ Q \ xvec \* \}" + have "eqvt ?X" using bisimClosed pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst] + by(auto simp add: eqvts eqvt_def) blast + have "(\, \\x\P, \\x\Q) \ ?X" + proof - + from \x \ \\ have "[x] \* \" + by auto + with \\ \ P \ Q\ show ?thesis + by (smt mem_Collect_eq resChain.base resChain.step) + qed + then show ?thesis + proof(coinduct rule: bisimCoinduct) + case(cStatEq \ xP xQ) + from \(\, xP, xQ) \ ?X\ obtain xvec P Q where "\ \ P \ Q" and "xvec \* \" and "xP = \\*xvec\P" and "xQ = \\*xvec\Q" + by auto + moreover from \\ \ P \ Q\ have PeqQ: "insertAssertion(extractFrame P) \ \\<^sub>F insertAssertion(extractFrame Q) \" + by(rule bisimE) + ultimately show ?case by(auto intro: frameResChainPres) + next + case(cSim \ xP xQ) + from \(\, xP, xQ) \ ?X\ obtain xvec P Q where "\ \ P \ Q" and "xvec \* \" and "xP = \\*xvec\P" and "xQ = \\*xvec\Q" + by auto + from \\ \ P \ Q\ have "\ \ P \[bisim] Q" by(rule bisimE) + note \eqvt ?X\ + then have "eqvt(?X \ bisim)" by auto + from \xvec \* \\ + have "\ \ \\*xvec\P \[(?X \ bisim)] \\*xvec\Q" + proof(induct xvec) + case Nil + have "\ \ \\*[]\P \[bisim] \\*[]\Q" using \\ \ P \ Q\ + unfolding resChain.simps + by(rule bisimE) + then show ?case by(rule monotonic) auto + next + case(Cons x xvec) + then have "x \ \" and "xvec \* \" + by auto + from \xvec \* \\ have "\ \ \\*xvec\P \[(?X \ bisim)] \\*xvec\Q" + by(rule Cons) + moreover note \eqvt(?X \ bisim)\ + moreover note \x \ \\ + moreover have "?X \ bisim \ ?X \ bisim" by auto + moreover have "\\ P Q xvec. \(\, P, Q) \ ?X \ bisim; xvec \* \\ \ (\, \\*xvec\P, \\*xvec\Q) \ ?X \ bisim" + by (smt (verit, ccfv_threshold) Un_iff freshChainAppend mem_Collect_eq prod.inject resChainAppend) + ultimately have "\ \ \\x\(\\*xvec\P) \[(?X \ bisim)] \\x\(\\*xvec\Q)" + by(rule resPres) + then show ?case unfolding resChain.simps by - + next + qed + with \xP = \\*xvec\P\ \xQ = \\*xvec\Q\ show ?case + by simp + next + case(cExt \ xP xQ \') + from \(\, xP, xQ) \ ?X\ obtain xvec P Q where "\ \ P \ Q" and "xvec \* \" and xpe: "xP = \\*xvec\P" and xqe: "xQ = \\*xvec\Q" + by auto + obtain p::"name prm" where "(p\xvec) \* \" and "(p\xvec) \* \'" and "(p\xvec) \* P" and "(p\xvec) \* Q" and "(p\xvec) \* xvec" and "distinctPerm p" and "set p \ (set xvec) \ (set (p\xvec))" + by(rule name_list_avoiding[where c="(\,\',P,Q,xvec)"]) auto + from \\ \ P \ Q\ have "\ \ (p\\') \ P \ Q" + by(rule bisimE) + moreover have "xvec \* (\ \ (p\\'))" using \xvec \* \\ \(p\xvec) \* \'\ \distinctPerm p\ + apply(intro freshCompChain) + apply assumption + apply(subst perm_bool[where pi=p,symmetric]) + by(simp add: eqvts) + ultimately have "(\ \ (p\\'),\\*xvec\P,\\*xvec\Q) \ ?X" + by auto + then have "(p\(\ \ (p\\'),\\*xvec\P,\\*xvec\Q)) \ ?X" using \eqvt ?X\ + by(intro eqvtI) + then have "(\ \ \',\\*(p\xvec)\(p\P),\\*(p\xvec)\(p\Q)) \ ?X" using \distinctPerm p\ \set p \ (set xvec) \ (set (p\xvec))\ \xvec \* \\ \(p\xvec) \* \\ + by(simp add: eqvts) + then have "(\ \ \',\\*xvec\P,\\*xvec\Q) \ ?X" using \(p \ xvec) \* Q\ \(p \ xvec) \* P\ \set p \ set xvec \ set (p \ xvec)\ \distinctPerm p\ + by(subst (1 2) resChainAlpha[where p=p]) auto + then show ?case unfolding xpe xqe + by blast + next + case(cSym \ P Q) + then show ?case + by(blast dest: bisimE) + qed +qed + +lemma bisimResChainPres: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and Q :: "('a, 'b, 'c) psi" + and xvec :: "name list" + +assumes "\ \ P \ Q" + and "xvec \* \" + +shows "\ \ \\*xvec\P \ \\*xvec\Q" + using assms + by(induct xvec) (auto intro: bisimResPres) + +lemma bisimParPresAux: + fixes \ :: 'b + and \\<^sub>R :: 'b + and P :: "('a, 'b, 'c) psi" + and Q :: "('a, 'b, 'c) psi" + and R :: "('a, 'b, 'c) psi" + and A\<^sub>R :: "name list" + +assumes "\ \ \\<^sub>R \ P \ Q" + and FrR: "extractFrame R = \A\<^sub>R, \\<^sub>R\" + and "A\<^sub>R \* \" + and "A\<^sub>R \* P" + and "A\<^sub>R \* Q" + +shows "\ \ P \ R \ Q \ R" +proof - + let ?X = "{(\, \\*xvec\(P \ R), \\*xvec\(Q \ R)) | xvec \ P Q R. xvec \* \ \ (\A\<^sub>R \\<^sub>R. (extractFrame R = \A\<^sub>R, \\<^sub>R\ \ A\<^sub>R \* \ \ A\<^sub>R \* P \ A\<^sub>R \* Q) \ + \ \ \\<^sub>R \ P \ Q)}" + { + fix xvec :: "name list" + and \ :: 'b + and P :: "('a, 'b, 'c) psi" + and Q :: "('a, 'b, 'c) psi" + and R :: "('a, 'b, 'c) psi" + + assume "xvec \* \" + and "\A\<^sub>R \\<^sub>R. \extractFrame R = \A\<^sub>R, \\<^sub>R\; A\<^sub>R \* \; A\<^sub>R \* P; A\<^sub>R \* Q\ \ \ \ \\<^sub>R \ P \ Q" + + then have "(\, \\*xvec\(P \ R), \\*xvec\(Q \ R)) \ ?X" + by auto blast + } + + note XI = this + { + fix xvec :: "name list" + and \ :: 'b + and P :: "('a, 'b, 'c) psi" + and Q :: "('a, 'b, 'c) psi" + and R :: "('a, 'b, 'c) psi" + and C :: "'d::fs_name" + + assume "xvec \* \" + and A: "\A\<^sub>R \\<^sub>R. \extractFrame R = \A\<^sub>R, \\<^sub>R\; A\<^sub>R \* \; A\<^sub>R \* P; A\<^sub>R \* Q; A\<^sub>R \* C\ \ \ \ \\<^sub>R \ P \ Q" + + from \xvec \* \\ have "(\, \\*xvec\(P \ R), \\*xvec\(Q \ R)) \ ?X" + proof(rule XI) + fix A\<^sub>R \\<^sub>R + assume FrR: "extractFrame R = \A\<^sub>R, \\<^sub>R\" + obtain p::"name prm" where "(p \ A\<^sub>R) \* \" and "(p \ A\<^sub>R) \* P" and "(p \ A\<^sub>R) \* Q" and "(p \ A\<^sub>R) \* R" and "(p \ A\<^sub>R) \* C" + and "(p \ A\<^sub>R) \* \\<^sub>R" and S: "(set p) \ (set A\<^sub>R) \ (set(p \ A\<^sub>R))" and "distinctPerm p" + by(rule name_list_avoiding[where c="(\, P, Q, R, \\<^sub>R, C)"]) auto + from FrR \(p \ A\<^sub>R) \* \\<^sub>R\ S have "extractFrame R = \(p \ A\<^sub>R), p \ \\<^sub>R\" by(simp add: frameChainAlpha') + + moreover assume "A\<^sub>R \* \" + then have "(p \ A\<^sub>R) \* (p \ \)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>R \* \\ \(p \ A\<^sub>R) \* \\ S have "(p \ A\<^sub>R) \* \" by simp + moreover assume "A\<^sub>R \* P" + then have "(p \ A\<^sub>R) \* (p \ P)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>R \* P\ \(p \ A\<^sub>R) \* P\ S have "(p \ A\<^sub>R) \* P" by simp + moreover assume "A\<^sub>R \* Q" + then have "(p \ A\<^sub>R) \* (p \ Q)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>R \* Q\ \(p \ A\<^sub>R) \* Q\ S have "(p \ A\<^sub>R) \* Q" by simp + ultimately have "\ \ (p \ \\<^sub>R) \ P \ Q" using \(p \ A\<^sub>R) \* C\ A by blast + then have "(p \ (\ \ (p \ \\<^sub>R))) \ (p \ P) \ (p \ Q)" by(rule bisimClosed) + with \A\<^sub>R \* \\ \(p \ A\<^sub>R) \* \\ \A\<^sub>R \* P\ \(p \ A\<^sub>R) \* P\ \A\<^sub>R \* Q\ \(p \ A\<^sub>R) \* Q\ S \distinctPerm p\ + show "\ \ \\<^sub>R \ P \ Q" by(simp add: eqvts) + qed + } + note XI' = this + + have "eqvt ?X" + unfolding eqvt_def + proof + fix x + assume "x \ ?X" + then obtain xvec \ P Q R where 1: "x = (\, \\*xvec\P \ R, \\*xvec\Q \ R)" and 2: "xvec \* \" and + 3: "\A\<^sub>R \\<^sub>R. extractFrame R = \A\<^sub>R, \\<^sub>R\ \ A\<^sub>R \* \ \ A\<^sub>R \* P \ A\<^sub>R \* Q \ \ \ \\<^sub>R \ P \ Q" + by blast + show "\p::(name \ name) list. p \ x \ ?X" + proof + fix p :: "(name \ name) list" + from 1 have "p \ x = (p \ \, \\*p \ xvec\(p \ P) \ (p \ R), \\*p \ xvec\(p \ Q) \ (p \ R))" + by (simp add: eqvts) + moreover from 2 have "(p \ xvec) \* (p \ \)" + by (simp add: fresh_star_bij(2)) + moreover have "\A\<^sub>R \\<^sub>R. extractFrame (p \ R) = \A\<^sub>R, \\<^sub>R\ \ A\<^sub>R \* (p \ \) \ A\<^sub>R \* (p \ P) \ A\<^sub>R \* (p \ Q) \ (p \ \) \ \\<^sub>R \ (p \ P) \ (p \ Q)" + proof (rule allI, rule allI, rule impI, (erule conjE)+) + fix A\<^sub>R \\<^sub>R + assume exF: "extractFrame (p \ R) = \A\<^sub>R, \\<^sub>R\" and freshPsi: "A\<^sub>R \* (p \ \)" and freshP: "A\<^sub>R \* (p \ P)" and freshQ: "A\<^sub>R \* (p \ Q)" + from exF have "extractFrame R = \rev p \ A\<^sub>R, rev p \ \\<^sub>R\" + by (metis Chain.simps(5) extractFrameEqvt(1) frame.perm(1) frameResChainEqvt) + moreover from freshPsi have "(rev p \ A\<^sub>R) \* \" + by (metis fresh_star_bij(2) perm_pi_simp(1)) + moreover from freshP have "(rev p \ A\<^sub>R) \* P" + by (metis fresh_star_bij(2) perm_pi_simp(1)) + moreover from freshQ have "(rev p \ A\<^sub>R) \* Q" + by (metis fresh_star_bij(2) perm_pi_simp(1)) + ultimately show "(p \ \) \ \\<^sub>R \ p \ P \ p \ Q" + using 3 by (metis (no_types, opaque_lifting) bisimClosed perm_pi_simp(2) statEqvt') + qed + ultimately show "p \ x \ ?X" + by blast + qed + qed + + moreover have Res: "\\ P Q x. \(\, P, Q) \ ?X \ bisim; x \ \\ \ (\, \\x\P, \\x\Q) \ ?X \ bisim" + proof - + fix \ P Q x + assume "(\, P, Q) \ ?X \ bisim" and "(x::name) \ \" + show "(\, \\x\P, \\x\Q) \ ?X \ bisim" + proof(cases "(\, P, Q) \ ?X") + assume "(\, P, Q) \ ?X" + then obtain xvec P' Q' R where "P = \\*xvec\P' \ R" and "Q = \\*xvec\Q' \ R" and "xvec \* \" + and "\A\<^sub>R \\<^sub>R. extractFrame R = \A\<^sub>R, \\<^sub>R\ \ A\<^sub>R \* \ \ A\<^sub>R \* P' \ A\<^sub>R \* Q' \ \ \ \\<^sub>R \ P' \ Q'" + by blast + moreover have "\\x\(\\*xvec\P' \ R) = \\*(x # xvec)\P' \ R" + by simp + moreover have "\\x\(\\*xvec\Q' \ R) = \\*(x # xvec)\Q' \ R" + by simp + moreover from \x \ \\ \xvec \* \\ have "(x # xvec) \* \" + by simp + ultimately have "(\, \\x\P, \\x\Q) \ ?X" + by blast + then show ?thesis by simp + next + assume "\(\, P, Q) \ ?X" + with \(\, P, Q) \ ?X \ bisim\ have "\ \ P \ Q" + by blast + then have "\ \ \\x\P \ \\x\Q" using \x \ \\ + by(rule bisimResPres) + then show ?thesis + by simp + qed + qed + have ResChain: "\\ P Q xvec. \(\, P, Q) \ ?X \ bisim; xvec \* \\ \ (\, \\*xvec\P, \\*xvec\Q) \ ?X \ bisim" + proof - + fix \ P Q + and xvec::"name list" + assume "(\, P, Q) \ ?X \ bisim" + and "xvec \* \" + then show "(\, \\*xvec\P, \\*xvec\Q) \ ?X \ bisim" + proof(induct xvec) + case Nil then show ?case by simp + next + case(Cons x xvec) + then have "(\, \\*xvec\P, \\*xvec\Q) \ ?X \ bisim" + by simp + moreover have "x \ \" using Cons by simp + ultimately show ?case unfolding resChain.simps + by(rule Res) + qed + qed + have "(\, P \ R, Q \ R) \ ?X" + proof - + { + fix A\<^sub>R' :: "name list" + and \\<^sub>R' :: 'b + + assume FrR': "extractFrame R = \A\<^sub>R', \\<^sub>R'\" + and "A\<^sub>R' \* \" + and "A\<^sub>R' \* P" + and "A\<^sub>R' \* Q" + + obtain p where "(p \ A\<^sub>R') \* A\<^sub>R" and "(p \ A\<^sub>R') \* \\<^sub>R'" and "(p \ A\<^sub>R') \* \" and "(p \ A\<^sub>R') \* P" and "(p \ A\<^sub>R') \* Q" + and Sp: "(set p) \ (set A\<^sub>R') \ (set(p \ A\<^sub>R'))" and "distinctPerm p" + by(rule name_list_avoiding[where c="(A\<^sub>R, \, \\<^sub>R', P, Q)"]) auto + + from FrR' \(p \ A\<^sub>R') \* \\<^sub>R'\ Sp have "extractFrame R = \(p \ A\<^sub>R'), p \ \\<^sub>R'\" + by(simp add: frameChainAlpha eqvts) + with FrR \(p \ A\<^sub>R') \* A\<^sub>R\ obtain q::"name prm" + where Sq: "set q \ set(p \ A\<^sub>R') \ set A\<^sub>R" and "distinctPerm q" and "\\<^sub>R = q \ p \ \\<^sub>R'" + by(force elim: frameChainEq) + + from \\ \ \\<^sub>R \ P \ Q\ \\\<^sub>R = q \ p \ \\<^sub>R'\ have "\ \ (q \ p \ \\<^sub>R') \ P \ Q" by simp + then have "(q \ (\ \ (q \ p \ \\<^sub>R'))) \ (q \ P) \ (q \ Q)" by(rule bisimClosed) + with Sq \A\<^sub>R \* \\ \(p \ A\<^sub>R') \* \\ \A\<^sub>R \* P\ \(p \ A\<^sub>R') \* P\ \A\<^sub>R \* Q\ \(p \ A\<^sub>R') \* Q\ \distinctPerm q\ + have "\ \ (p \ \\<^sub>R') \ P \ Q" by(simp add: eqvts) + then have "(p \ (\ \ (p \ \\<^sub>R'))) \ (p \ P) \ (p \ Q)" by(rule bisimClosed) + with Sp \A\<^sub>R' \* \\ \(p \ A\<^sub>R') \* \\ \A\<^sub>R' \* P\ \(p \ A\<^sub>R') \* P\ \A\<^sub>R' \* Q\ \(p \ A\<^sub>R') \* Q\ \distinctPerm p\ + have "\ \ \\<^sub>R' \ P \ Q" by(simp add: eqvts) + } + then show ?thesis + apply clarsimp + apply(rule exI[where x="[]"]) + by auto blast + qed + then show ?thesis + proof(coinduct rule: bisimCoinduct) + case(cStatEq \ PR QR) + from \(\, PR, QR) \ ?X\ + obtain xvec P Q R where PFrR: "PR = \\*xvec\(P \ R)" and QFrR: "QR = \\*xvec\(Q \ R)" + and "xvec \* \" and *: "\A\<^sub>R \\<^sub>R. extractFrame R = \A\<^sub>R, \\<^sub>R\ \ A\<^sub>R \* \ \ A\<^sub>R \* P \ A\<^sub>R \* Q \ \ \ \\<^sub>R \ P \ Q" + by blast + obtain A\<^sub>R \\<^sub>R where FrR: "extractFrame R = \A\<^sub>R, \\<^sub>R\" and fresh: "A\<^sub>R \* (\, P, Q, R)" + using freshFrame by metis + from fresh have "A\<^sub>R \* \" and "A\<^sub>R \* P" and "A\<^sub>R \* Q" and "A\<^sub>R \* R" + by auto + with FrR have PSimQ: "\ \ \\<^sub>R \ P \ Q" + using * by blast + + obtain A\<^sub>P \\<^sub>P where FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" and "A\<^sub>P \* \" and "A\<^sub>P \* A\<^sub>R" and "A\<^sub>P \* \\<^sub>R" + by(rule freshFrame[where C="(\, A\<^sub>R, \\<^sub>R)"]) auto + obtain A\<^sub>Q \\<^sub>Q where FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" and "A\<^sub>Q \* \" and "A\<^sub>Q \* A\<^sub>R" and "A\<^sub>Q \* \\<^sub>R" + by(rule freshFrame[where C="(\, A\<^sub>R, \\<^sub>R)"]) auto + from \A\<^sub>R \* P\ \A\<^sub>R \* Q\ \A\<^sub>P \* A\<^sub>R\ \A\<^sub>Q \* A\<^sub>R\ FrP FrQ have "A\<^sub>R \* \\<^sub>P" and "A\<^sub>R \* \\<^sub>Q" + by(force dest: extractFrameFreshChain)+ + + have "\(A\<^sub>P@A\<^sub>R), \ \ \\<^sub>P \ \\<^sub>R\ \\<^sub>F \(A\<^sub>Q@A\<^sub>R), \ \ \\<^sub>Q \ \\<^sub>R\" + proof - + have "\A\<^sub>P, \ \ \\<^sub>P \ \\<^sub>R\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>R) \ \\<^sub>P\" + by(metis frameResChainPres frameNilStatEq Associativity Commutativity AssertionStatEqTrans Composition) + moreover from PSimQ have "insertAssertion(extractFrame P) (\ \ \\<^sub>R) \\<^sub>F insertAssertion(extractFrame Q) (\ \ \\<^sub>R)" + by(rule bisimE) + with FrP FrQ freshCompChain \A\<^sub>P \* \\ \A\<^sub>P \* \\<^sub>R\ \A\<^sub>Q \* \\ \A\<^sub>Q \* \\<^sub>R\ have "\A\<^sub>P, (\ \ \\<^sub>R) \ \\<^sub>P\ \\<^sub>F \A\<^sub>Q, (\ \ \\<^sub>R) \ \\<^sub>Q\" + by auto + moreover have "\A\<^sub>Q, (\ \ \\<^sub>R) \ \\<^sub>Q\ \\<^sub>F \A\<^sub>Q, \ \ \\<^sub>Q \ \\<^sub>R\" + by(metis frameResChainPres frameNilStatEq Associativity Commutativity AssertionStatEqTrans Composition) + ultimately have "\A\<^sub>P, \ \ \\<^sub>P \ \\<^sub>R\ \\<^sub>F \A\<^sub>Q, \ \ \\<^sub>Q \ \\<^sub>R\" + by(blast intro: FrameStatEqTrans) + then have "\(A\<^sub>R@A\<^sub>P), \ \ \\<^sub>P \ \\<^sub>R\ \\<^sub>F \(A\<^sub>R@A\<^sub>Q), \ \ \\<^sub>Q \ \\<^sub>R\" + by(drule_tac frameResChainPres) (simp add: frameChainAppend) + then show ?thesis + apply(simp add: frameChainAppend) + by(metis frameResChainComm FrameStatEqTrans) + qed + moreover from \A\<^sub>P \* \\ \A\<^sub>R \* \\ have "(A\<^sub>P@A\<^sub>R) \* \" by simp + moreover from \A\<^sub>Q \* \\ \A\<^sub>R \* \\ have "(A\<^sub>Q@A\<^sub>R) \* \" by simp + ultimately have PFrRQR: "insertAssertion(extractFrame(P \ R)) \ \\<^sub>F insertAssertion(extractFrame(Q \ R)) \" + using FrP FrQ FrR \A\<^sub>P \* A\<^sub>R\ \A\<^sub>P \* \\<^sub>R\ \A\<^sub>Q \* A\<^sub>R\ \A\<^sub>Q \* \\<^sub>R\ \A\<^sub>R \* \\<^sub>P\ \A\<^sub>R \* \\<^sub>Q\ + by simp + + from \xvec \* \\ have "insertAssertion (extractFrame(\\*xvec\P \ R)) \ \\<^sub>F \\*xvec\(insertAssertion (extractFrame(P \ R)) \)" + by(rule insertAssertionExtractFrameFresh) + moreover from PFrRQR have "\\*xvec\(insertAssertion (extractFrame(P \ R)) \) \\<^sub>F \\*xvec\(insertAssertion (extractFrame(Q \ R)) \)" + by(induct xvec) (auto intro: frameResPres) + moreover from \xvec \* \\ have "\\*xvec\(insertAssertion (extractFrame(Q \ R)) \) \\<^sub>F insertAssertion (extractFrame(\\*xvec\Q \ R)) \" + by(rule FrameStatEqSym[OF insertAssertionExtractFrameFresh]) + ultimately show ?case using PFrR QFrR + by(blast intro: FrameStatEqTrans) + next + case(cSim \ PR QR) + { + fix \ P Q R xvec + assume "\A\<^sub>R \\<^sub>R. \extractFrame R = \A\<^sub>R, \\<^sub>R\; A\<^sub>R \* \; A\<^sub>R \* P; A\<^sub>R \* Q\ \ \ \ \\<^sub>R \ P \ Q" + moreover have "eqvt bisim" by simp + moreover from \eqvt ?X\ have "eqvt(?X \ bisim)" by auto + moreover from bisimE(1) have "\\ P Q. \ \ P \ Q \ insertAssertion (extractFrame Q) \ \\<^sub>F insertAssertion (extractFrame P) \" by(simp add: FrameStatEq_def) + moreover note bisimE(2) bisimE(3) + moreover + { + fix \ P Q A\<^sub>R \\<^sub>R R + assume PSimQ: "\ \ \\<^sub>R \ P \ Q" + and FrR: "extractFrame R = \A\<^sub>R, \\<^sub>R\" + and "A\<^sub>R \* \" + and "A\<^sub>R \* P" + and "A\<^sub>R \* Q" + then have "(\, P \ R, Q \ R) \ ?X" + proof - + have "P \ R = \\*[]\(P \ R)" by simp + moreover have "Q \ R = \\*[]\(Q \ R)" by simp + moreover have "([]::name list) \* \" by simp + moreover + { + fix A\<^sub>R' \\<^sub>R' + + assume FrR': "extractFrame R = \A\<^sub>R', \\<^sub>R'\" + and "A\<^sub>R' \* \" + and "A\<^sub>R' \* P" + and "A\<^sub>R' \* Q" + obtain p where "(p \ A\<^sub>R') \* A\<^sub>R" + and "(p \ A\<^sub>R') \* \\<^sub>R'" + and "(p \ A\<^sub>R') \* \" + and "(p \ A\<^sub>R') \* P" + and "(p \ A\<^sub>R') \* Q" + and S: "(set p) \ (set A\<^sub>R') \ (set(p \ A\<^sub>R'))" and "distinctPerm p" + by(rule name_list_avoiding[where c="(A\<^sub>R, \, \\<^sub>R', P, Q)"]) auto + + from \(p \ A\<^sub>R') \* \\<^sub>R'\ S have "\A\<^sub>R', \\<^sub>R'\ = \p \ A\<^sub>R', p \ \\<^sub>R'\" + by(simp add: frameChainAlpha) + + with FrR' have FrR'': "extractFrame R = \p \ A\<^sub>R', p \ \\<^sub>R'\" by simp + with FrR \(p \ A\<^sub>R') \* A\<^sub>R\ + obtain q where "p \ \\<^sub>R' = (q::name prm) \ \\<^sub>R" and S': "set q \ (set A\<^sub>R) \ set(p \ A\<^sub>R')" and "distinctPerm q" + apply clarsimp + apply(drule sym) + apply simp + by(drule frameChainEq) auto + from PSimQ have "(q \ (\ \ \\<^sub>R)) \ (q \ P) \ (q \ Q)" + by(rule bisimClosed) + with \A\<^sub>R \* \\ \A\<^sub>R \* P\ \A\<^sub>R \* Q\ \(p \ A\<^sub>R') \* \\ \(p \ A\<^sub>R') \* P\ \(p \ A\<^sub>R') \* Q\ S' + have "\ \ (q \ \\<^sub>R) \ P \ Q" by(simp add: eqvts) + then have "(p \ (\ \ (q \ \\<^sub>R))) \ (p \ P) \ (p \ Q)" by(rule bisimClosed) + with \A\<^sub>R' \* \\ \A\<^sub>R' \* P\ \A\<^sub>R' \* Q\ \(p \ A\<^sub>R') \* \\ \(p \ A\<^sub>R') \* P\ \(p \ A\<^sub>R') \* Q\ S \distinctPerm p\ \(p \ \\<^sub>R') = q \ \\<^sub>R\ + have "\ \ \\<^sub>R' \ P \ Q" + by(drule_tac sym) (simp add: eqvts) + } + ultimately show ?thesis + by blast + qed + then have "(\, P \ R, Q \ R) \ ?X \ bisim" + by simp + } + moreover have "\\ P Q xvec. \(\, P, Q) \ ?X \ bisim; (xvec::name list) \* \\ \ (\, \\*xvec\P, \\*xvec\Q) \ ?X \ bisim" + proof - + fix \ P Q xvec + assume "(\, P, Q) \ ?X \ bisim" + assume "(xvec::name list) \* \" + then show "(\, \\*xvec\P, \\*xvec\Q) \ ?X \ bisim" + proof(induct xvec) + case Nil + then show ?case using \(\, P, Q) \ ?X \ bisim\ by simp + next + case(Cons x xvec) + then show ?case by(simp only: resChain.simps) (rule Res, auto) + qed + qed + ultimately have "\ \ P \ R \[(?X \ bisim)] Q \ R" using statEqBisim + by(rule parPres) + moreover assume "(xvec::name list) \* \" + moreover from \eqvt ?X\ have "eqvt(?X \ bisim)" by auto + ultimately have "\ \ \\*xvec\(P \ R) \[(?X \ bisim)] \\*xvec\(Q \ R)" using ResChain + by(intro resChainPres) + } + with \(\, PR, QR) \ ?X\ show ?case by blast + next + case(cExt \ PR QR \') + + from \(\, PR, QR) \ ?X\ + obtain xvec P Q R A\<^sub>R \\<^sub>R where PFrR: "PR = \\*xvec\(P \ R)" and QFrR: "QR = \\*xvec\(Q \ R)" + and "xvec \* \" and A: "\A\<^sub>R \\<^sub>R. (extractFrame R = \A\<^sub>R, \\<^sub>R\ \ A\<^sub>R \* \ \ A\<^sub>R \* P \ A\<^sub>R \* Q) \ \ \ \\<^sub>R \ P \ Q" + by auto + + obtain p where "(p \ xvec) \* \" + and "(p \ xvec) \* P" + and "(p \ xvec) \* Q" + and "(p \ xvec) \* R" + and "(p \ xvec) \* \'" + and S: "(set p) \ (set xvec) \ (set(p \ xvec))" and "distinctPerm p" + by(rule name_list_avoiding[where c="(\, P, Q, R, \')"]) auto + + from \(p \ xvec) \* P\ \(p \ xvec) \* R\ S have "\\*xvec\(P \ R) = \\*(p \ xvec)\(p \ (P \ R))" + by(subst resChainAlpha) auto + then have PRAlpha: "\\*xvec\(P \ R) = \\*(p \ xvec)\((p \ P) \ (p \ R))" + by(simp add: eqvts) + + from \(p \ xvec) \* Q\ \(p \ xvec) \* R\ S have "\\*xvec\(Q \ R) = \\*(p \ xvec)\(p \ (Q \ R))" + by(subst resChainAlpha) auto + then have QRAlpha: "\\*xvec\(Q \ R) = \\*(p \ xvec)\((p \ Q) \ (p \ R))" + by(simp add: eqvts) + + have "(\ \ \', \\*(p \ xvec)\((p \ P) \ (p \ R)), \\*(p \ xvec)\((p \ Q) \ (p \ R))) \ ?X" + proof(rule XI'[where C2="(\, (p \ P), (p \ Q), R, \', xvec, p \ xvec)"]) + show "(p \ xvec) \* (\ \ \')" + using \(p \ xvec) \* \\ \(p \ xvec) \* \'\ by auto + next + fix A\<^sub>R \\<^sub>R + assume FrR: "extractFrame (p \ R) = \A\<^sub>R, \\<^sub>R\" and "A\<^sub>R \* (\ \ \')" and "A\<^sub>R \* (p \ P)" and "A\<^sub>R \* (\, p \ P, p \ Q, R, \', xvec, p \ xvec)" + then have "A\<^sub>R \* \" and "A\<^sub>R \* (p \ Q)" + by simp_all + from FrR have "(p \ (extractFrame (p \ R))) = (p \ \A\<^sub>R, \\<^sub>R\)" + by simp + with \distinctPerm p\ have "extractFrame R = \p \ A\<^sub>R, p \ \\<^sub>R\" + by(simp add: eqvts) + moreover from \A\<^sub>R \* \\ have "(p \ A\<^sub>R) \* (p \ \)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \xvec \* \\ \(p \ xvec) \* \\ S have "(p \ A\<^sub>R) \* \" + by simp + moreover from \A\<^sub>R \* (p \ P)\ have "(p \ A\<^sub>R) \* (p \ p \ P)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \distinctPerm p\ have "(p \ A\<^sub>R) \* P" + by simp + moreover from \A\<^sub>R \* (p \ Q)\ have "(p \ A\<^sub>R) \* (p \ p \ Q)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \distinctPerm p\ have "(p \ A\<^sub>R) \* Q" + by simp + ultimately have "\ \ (p \ \\<^sub>R) \ P \ Q" + using A by blast + then have "(\ \ (p \ \\<^sub>R)) \ (p \ \') \ P \ Q" + by(rule bisimE) + moreover have "(\ \ (p \ \\<^sub>R)) \ (p \ \') \ (\ \ (p \ \')) \ (p \ \\<^sub>R)" + by(metis Associativity Commutativity Composition AssertionStatEqTrans AssertionStatEqSym) + ultimately have "(\ \ (p \ \')) \ (p \ \\<^sub>R) \ P \ Q" + by(rule statEqBisim) + then have "(p \ ((\ \ (p \ \')) \ (p \ \\<^sub>R))) \ (p \ P) \ (p \ Q)" + by(rule bisimClosed) + with \distinctPerm p\ \xvec \* \\ \(p \ xvec) \* \\ S show "(\ \ \') \ \\<^sub>R \ (p \ P) \ (p \ Q)" + by(simp add: eqvts) + qed + with PFrR QFrR PRAlpha QRAlpha show ?case by simp + next + case(cSym \ PR QR) + then show ?case by(blast dest: bisimE) + qed +qed + +lemma bisimParPres: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and Q :: "('a, 'b, 'c) psi" + and R :: "('a, 'b, 'c) psi" + +assumes "\ \ P \ Q" + +shows "\ \ P \ R \ Q \ R" +proof - + obtain A\<^sub>R \\<^sub>R where "extractFrame R = \A\<^sub>R, \\<^sub>R\" and "A\<^sub>R \* \" and "A\<^sub>R \* P" and "A\<^sub>R \* Q" + by(rule freshFrame[where C="(\, P, Q)"]) auto + moreover from \\ \ P \ Q\ have "\ \ \\<^sub>R \ P \ Q" by(rule bisimE) + ultimately show ?thesis by(intro bisimParPresAux) +qed + +end + +end diff --git a/thys/Broadcast_Psi/Bisim_Struct_Cong.thy b/thys/Broadcast_Psi/Bisim_Struct_Cong.thy new file mode 100644 --- /dev/null +++ b/thys/Broadcast_Psi/Bisim_Struct_Cong.thy @@ -0,0 +1,2024 @@ +theory Bisim_Struct_Cong + imports Bisim_Pres Sim_Struct_Cong "Psi_Calculi.Structural_Congruence" +begin + +text \This file is a (heavily modified) variant of the theory {\it Psi\_Calculi.Bisim\_Struct\_Cong} +from~\cite{DBLP:journals/afp/Bengtson12}.\ + +context env begin + +lemma bisimParComm: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and Q :: "('a, 'b, 'c) psi" + +shows "\ \ P \ Q \ Q \ P" +proof - + let ?X = "{((\::'b), \\*xvec\((P::('a, 'b, 'c) psi) \ Q), \\*xvec\(Q \ P)) | xvec \ P Q. xvec \* \}" + + have "eqvt ?X" + by(force simp add: eqvt_def pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst] eqvts) + + have "(\, P \ Q, Q \ P) \ ?X" + using resChain.base by(smt (verit) freshSets(1) mem_Collect_eq) + then show ?thesis + proof(coinduct rule: bisimWeakCoinduct) + case(cStatEq \ PQ QP) + from \(\, PQ, QP) \ ?X\ + obtain xvec P Q where PFrQ: "PQ = \\*xvec\(P \ Q)" and QFrP: "QP = \\*xvec\(Q \ P)" and "xvec \* \" + by auto + + obtain A\<^sub>P \\<^sub>P where FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" and "A\<^sub>P \* \" and "A\<^sub>P \* Q" + by(rule freshFrame[where C="(\, Q)"]) auto + obtain A\<^sub>Q \\<^sub>Q where FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" and "A\<^sub>Q \* \" and "A\<^sub>Q \* A\<^sub>P" and "A\<^sub>Q \* \\<^sub>P" + by(rule freshFrame[where C="(\, A\<^sub>P, \\<^sub>P)"]) auto + from FrQ \A\<^sub>Q \* A\<^sub>P\ \A\<^sub>P \* Q\ have "A\<^sub>P \* \\<^sub>Q" by(force dest: extractFrameFreshChain) + have "\(xvec@A\<^sub>P@A\<^sub>Q), \ \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \(xvec@A\<^sub>Q@A\<^sub>P), \ \ \\<^sub>Q \ \\<^sub>P\" + by(simp add: frameChainAppend) + (metis frameResChainPres frameResChainComm frameNilStatEq compositionSym Associativity Commutativity FrameStatEqTrans) + with FrP FrQ PFrQ QFrP \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>Q \* \\<^sub>P\ \A\<^sub>Q \* A\<^sub>P\ \xvec \* \\ \A\<^sub>P \* \\ \A\<^sub>Q \* \\ + show ?case by(auto simp add: frameChainAppend) + next + case(cSim \ PQ QP) + from \(\, PQ, QP) \ ?X\ + obtain xvec P Q where PFrQ: "PQ = \\*xvec\(P \ Q)" and QFrP: "QP = \\*xvec\(Q \ P)" + and "xvec \* \" + by auto + moreover have "\ \ \\*xvec\(P \ Q) \[?X] \\*xvec\(Q \ P)" + proof - + have "\ \ P \ Q \[?X] Q \ P" + proof - + note \eqvt ?X\ + moreover have "\\ P Q. (\, P \ Q, Q \ P) \ ?X" + using resChain.base by(smt (verit) freshSets(1) mem_Collect_eq) + moreover have "\\ P Q xvec. \(\, P, Q) \ ?X; xvec \* \\ \ (\, \\*xvec\P, \\*xvec\Q) \ ?X" + proof (induct xvec) + case Nil + then show ?case + by (smt (verit, ccfv_threshold) Pair_inject freshChainAppend mem_Collect_eq resChainAppend) + next + case (Cons a xvec) + then show ?case + by blast + qed + ultimately show ?thesis by(rule simParComm) + qed + moreover note \eqvt ?X\ \xvec \* \\ + moreover have "\\ P Q xvec. \(\, P, Q) \ ?X; xvec \* \\ \ (\, \\*xvec\P, \\*xvec\Q) \ ?X" + using resChainAppend[symmetric] freshChainAppend by blast + ultimately show ?thesis + by(rule resChainPres) + qed + ultimately show ?case by simp + next + case(cExt \ PQ QP \') + from \(\, PQ, QP) \ ?X\ + obtain xvec P Q where PFrQ: "PQ = \\*xvec\(P \ Q)" and QFrP: "QP = \\*xvec\(Q \ P)" + and "xvec \* \" + by auto + + obtain p where "(p \ xvec) \* \" + and "(p \ xvec) \* P" + and "(p \ xvec) \* Q" + and "(p \ xvec) \* \'" + and S: "(set p) \ (set xvec) \ (set(p \ xvec))" and "distinctPerm p" + by(rule name_list_avoiding[where c="(\, P, Q, \')"]) auto + + from \(p \ xvec) \* P\ \(p \ xvec) \* Q\ S have "\\*xvec\(P \ Q) = \\*(p \ xvec)\(p \ (P \ Q))" + by(subst resChainAlpha) auto + then have PQAlpha: "\\*xvec\(P \ Q) = \\*(p \ xvec)\((p \ P) \ (p \ Q))" + by(simp add: eqvts) + + from \(p \ xvec) \* P\ \(p \ xvec) \* Q\ S have "\\*xvec\(Q \ P) = \\*(p \ xvec)\(p \ (Q \ P))" + by(subst resChainAlpha) auto + then have QPAlpha: "\\*xvec\(Q \ P) = \\*(p \ xvec)\((p \ Q) \ (p \ P))" + by(simp add: eqvts) + + from \(p \ xvec) \* \\ \(p \ xvec) \* \'\ have "(\ \ \', \\*(p \ xvec)\((p \ P) \ (p \ Q)), \\*(p \ xvec)\((p \ Q) \ (p \ P))) \ ?X" + by auto + with PFrQ QFrP PQAlpha QPAlpha show ?case by simp + next + case(cSym \ PR QR) + then show ?case by blast + qed +qed + +inductive_set resCommRel :: "('b \ ('a,'b,'c) psi \ ('a,'b,'c) psi) set" + where + resCommRel_refl : "(\,P,P) \ resCommRel" + | resCommRel_swap : "\a \ \; b \ \\ \ (\,\\a\(\\b\P),\\b\(\\a\P)) \ resCommRel" + | resCommRel_res : "\(\,P,Q) \ resCommRel; a \ \\ \ (\,\\a\P,\\a\Q) \ resCommRel" + +lemma eqvtResCommRel: "eqvt resCommRel" +proof - + { + fix \ P Q + and p::"name prm" + assume "(\,P,Q) \ resCommRel" + then have "(p\\,p\P,p\Q) \ resCommRel" + proof(induct rule: resCommRel.inducts) + case resCommRel_refl then show ?case by(rule resCommRel.intros) + next + case(resCommRel_swap a \ b P) + then have "(p\a) \ p\\" and "(p\b) \ p\\" + apply - + by(subst (asm) (1 2) perm_bool[where pi=p,symmetric], simp add: eqvts)+ + then show ?case unfolding eqvts + by(rule resCommRel.intros) + next + case(resCommRel_res \ P Q a) + from \a \ \\ have "(p\a) \ p\\" + apply - + by(subst (asm) perm_bool[where pi=p,symmetric], simp add: eqvts) + then show ?case using \(p \ \, p \ P, p \ Q) \ resCommRel\ + unfolding eqvts + by(intro resCommRel.intros) + qed + } + then show ?thesis unfolding eqvt_def + by auto +qed + +lemma resCommRelStarRes: + assumes "(\,P,Q) \ resCommRel\<^sup>\" + and "a \ \" + shows "(\,\\a\P,\\a\Q) \ resCommRel\<^sup>\" + using assms +proof(induct rule: rel_trancl.induct) + case r_into_rel_trancl then show ?case by(auto intro: resCommRel_res) +next + case(rel_trancl_into_rel_trancl \ P Q R) + then show ?case + by(auto dest: resCommRel_res intro: rel_trancl.intros) +qed + +lemma resCommRelStarResChain: + assumes "(\,P,Q) \ resCommRel\<^sup>\" + and "xvec \* \" + shows "(\,\\*xvec\P,\\*xvec\Q) \ resCommRel\<^sup>\" + using assms + by(induct xvec) (auto simp add: resCommRelStarRes) + +lemma length_induct1[consumes 0, case_names Nil Cons]: + assumes b: "P []" + and s: "\ x xvec. \\ yvec. length xvec=length yvec\P yvec\ \ P(x#xvec)" + shows "P xvec" +proof(induct xvec rule: length_induct) + case(1 xvec) + then show ?case + proof(cases xvec) + case Nil then show ?thesis by(simp add: b) + next + case(Cons y yvec) + with \\ys. length ys < length xvec \ P ys\ have "\ys. length ys = length yvec \ P ys" by simp + then show ?thesis unfolding Cons + using s by auto + qed +qed + +lemma oneStepPerm_in_rel: + assumes "xvec \* \" + shows "(\,\\*xvec\P,\\*(rotate1 xvec)\P) \ resCommRel\<^sup>\" + using assms +proof(induct xvec rule: length_induct1) + case Nil then show ?case by(auto intro: resCommRel_refl) +next + case(Cons x xvec) + note Cons1 = this + show ?case + proof(cases xvec) + case Nil then show ?thesis + by(auto intro: resCommRel_refl) + next + case(Cons y yvec) + then have "x \ \" and "y \ \" and "xvec \* \" using \(x # xvec) \* \\ + by simp+ + have "(\, \\*(x#y#yvec)\P, \\*(y#x#yvec)\P) \ resCommRel\<^sup>\" using \x \ \\ \y \ \\ + by(auto intro: resCommRel_swap) + moreover have "(\, \\*(y#x#yvec)\P, \\*(y#rotate1(x#yvec))\P) \ resCommRel\<^sup>\" + proof - + have "(\, \\*(x#yvec)\P, \\*rotate1(x#yvec)\P) \ resCommRel\<^sup>\" using \xvec \* \\ Cons \x \ \\ + by(intro Cons1) auto + then show ?thesis using \y \ \\ + unfolding resChain.simps + by(rule resCommRelStarRes) + qed + ultimately show ?thesis unfolding Cons + by(auto intro: rel_trancl_transitive) + qed +qed + +lemma fresh_same_multiset: + fixes xvec::"name list" + and yvec::"name list" + assumes "mset xvec = mset yvec" + and "xvec \* X" + shows "yvec \* X" +proof - + from \mset xvec = mset yvec\ have "set xvec = set yvec" + using mset_eq_setD by blast + then show ?thesis using \xvec \* X\ + unfolding fresh_def fresh_star_def name_list_supp + by auto +qed + +lemma nStepPerm_in_rel: + assumes "xvec \* \" + shows "(\,\\*xvec\P,\\*(rotate n xvec)\P) \ resCommRel\<^sup>\" + using assms +proof(induct n) + case 0 then show ?case by(auto intro: resCommRel_refl) +next + case(Suc n) + then have "(\, \\*xvec\P, \\*rotate n xvec\P) \ resCommRel\<^sup>\" + by simp + moreover have "(\, \\*rotate n xvec\P,\\*rotate (Suc n) xvec\P) \ resCommRel\<^sup>\" + proof - + { + fix xvec::"name list" + assume "xvec \* \" + have "rotate1 xvec \* \" using \xvec \* \\ + proof(induct xvec) + case Nil then show ?case by simp + next + case Cons then show ?case by simp + qed + } + note f1 = this + have "rotate n xvec \* \" using \xvec \* \\ + by(induct n) (auto simp add: f1) + then show ?thesis + by(auto simp add: oneStepPerm_in_rel) + qed + ultimately show ?case + by(rule rel_trancl_transitive) +qed + +lemma any_perm_in_rel: + assumes "xvec \* \" + and "mset xvec = mset yvec" + shows "(\,\\*xvec\P,\\*yvec\P) \ resCommRel\<^sup>\" + using assms +proof(induct xvec arbitrary: yvec rule: length_induct1) + case Nil then show ?case by(auto intro: resCommRel.intros) +next + case(Cons x xvec) + obtain yvec1 yvec2 where yveceq: "yvec=yvec1@x#yvec2" + proof - + assume 1: "\yvec1 yvec2. yvec = yvec1 @ x # yvec2 \ thesis" + { + note \mset (x # xvec) = mset yvec\ + then have "\ yvec1 yvec2. yvec=yvec1@x#yvec2" + proof(induct xvec arbitrary: yvec rule: length_induct1) + case Nil + from \mset [x] = mset yvec\ have "yvec = [x]" + apply(induct yvec) + apply simp + apply(simp add: single_is_union) + done + then show ?case by blast + next + case(Cons x' xvec) + have less: "{#x'#} \# mset yvec" unfolding \mset (x # x' # xvec) = mset yvec\[symmetric] + by simp + from \mset (x # x' # xvec) = mset yvec\ + have "mset (x # xvec) = mset (remove1 x' yvec)" + by (metis mset_remove1 remove1.simps(2)) + then have "\yvec1 yvec2. (remove1 x' yvec) = yvec1 @ x # yvec2" + by(intro Cons) simp+ + then obtain yvec1 yvec2 where "(remove1 x' yvec) = yvec1 @ x # yvec2" + by blast + then show ?case + proof(induct yvec arbitrary: yvec1 yvec2) + case Nil then show ?case by simp + next + case(Cons y yvec) then show ?case + proof(cases "x'=y") + case True + with \remove1 x' (y # yvec) = yvec1 @ x # yvec2\ + have "yvec = yvec1 @ x # yvec2" + by simp + then show ?thesis + by (metis append_Cons) + next + case False + then have "remove1 x' (y#yvec) = y # remove1 x' (yvec)" + by simp + note Cons2=Cons + show ?thesis + proof(cases yvec1) + case Nil + then show ?thesis using Cons2 False + by auto + next + case(Cons y1 yvec1a) + from \remove1 x' (y # yvec) = yvec1 @ x # yvec2\ \yvec1 = y1 # yvec1a\ False have "y1=y" + by simp + from \remove1 x' (y # yvec) = yvec1 @ x # yvec2\ + have "remove1 x' yvec = yvec1a @ x # yvec2" unfolding Cons \y1=y\ using False + by simp + then have "\yvec1 yvec2. yvec = yvec1 @ x # yvec2" + by(rule Cons2) + then obtain yvec1 yvec2 where "yvec = yvec1 @ x # yvec2" + by blast + then show ?thesis + by (metis append_Cons) + qed + qed + qed + qed + } + then show ?thesis using 1 + by blast + qed + from \(x # xvec) \* \\ have "x \ \" and "xvec \* \" by auto + have "mset (x # xvec) = mset(yvec1 @ x # yvec2)" + unfolding yveceq[symmetric] by fact + then have "mset (xvec) = mset(yvec1 @ yvec2)" + by(subst add_right_cancel[symmetric,where a="{#x#}"]) (simp add: add.assoc) + then have "(\, \\*xvec\P, \\*(yvec1@yvec2)\P) \ resCommRel\<^sup>\" using \xvec \* \\ + by(intro Cons) auto + moreover have "(\, \\*(yvec1@yvec2)\P, \\*(yvec2@yvec1)\P) \ resCommRel\<^sup>\" + proof - + have e: "yvec2 @ yvec1 = rotate (length yvec1) (yvec1@yvec2)" + apply(cases yvec2) + apply simp + apply(simp add: rotate_drop_take) + done + have "(yvec1@yvec2) \* \" using \mset (xvec) = mset(yvec1 @ yvec2)\ \xvec \* \\ + by(rule fresh_same_multiset) + then show ?thesis unfolding e + by(rule nStepPerm_in_rel) + qed + ultimately have "(\, \\*xvec\P, \\*(yvec2 @ yvec1)\P) \ resCommRel\<^sup>\" + by(rule rel_trancl_transitive) + then have "(\, \\*(x#xvec)\P, \\*(x # yvec2 @ yvec1)\P) \ resCommRel\<^sup>\" + unfolding resChain.simps using \x \ \\ + by(rule resCommRelStarRes) + moreover have "(\, \\*(x # yvec2 @ yvec1)\P, \\*(yvec1 @ x # yvec2)\P) \ resCommRel\<^sup>\" + proof - + have e: "yvec1 @ x # yvec2 = rotate (1+length yvec2) (x#yvec2@yvec1)" + apply(simp add: rotate_drop_take) + apply(cases yvec1) + by simp+ + have "(yvec1 @ x # yvec2) \* \" using \mset (x # xvec) = mset(yvec1 @ x # yvec2)\ \(x#xvec) \* \\ + by(rule fresh_same_multiset) + then have "(x # yvec2 @ yvec1) \* \" by simp + then show ?thesis unfolding e + by(rule nStepPerm_in_rel) + qed + ultimately show ?case unfolding yveceq + by(rule rel_trancl_transitive) +qed + +lemma bisimResComm: + fixes x :: name + and \ :: 'b + and y :: name + and P :: "('a, 'b, 'c) psi" + +shows "\ \ \\x\(\\y\P) \ \\y\(\\x\P)" +proof(cases "x=y") + case True + then show ?thesis by(blast intro: bisimReflexive) +next + case False + { + fix x::name and y::name and P::"('a, 'b, 'c) psi" + assume "x \ \" and "y \ \" + let ?X = resCommRel + from \x \ \\ \y \ \\ have "(\, \\x\(\\y\P), \\y\(\\x\P)) \ ?X" + by(rule resCommRel_swap) + then have "\ \ \\x\(\\y\P) \ \\y\(\\x\P)" using eqvtResCommRel + proof(coinduct rule: bisimStarWeakCoinduct) + case(cStatEq \ R S) + then show ?case + proof(induct rule: resCommRel.induct) + case resCommRel_refl then show ?case by(rule FrameStatEqRefl) + next + case(resCommRel_swap x \ y P) + moreover obtain A\<^sub>P \\<^sub>P where "extractFrame P = \A\<^sub>P, \\<^sub>P\" and "A\<^sub>P \* \" and "x \ A\<^sub>P" and "y \ A\<^sub>P" + by(rule freshFrame[where C="(x, y, \)"]) auto + ultimately show ?case by(force intro: frameResComm FrameStatEqTrans) + next + case(resCommRel_res \ P Q a) + then show ?case by(auto intro: frameResPres) + qed + next + case(cSim \ R S) + have "eqvt(resCommRel\<^sup>\)" + by(rule rel_trancl_eqvt) (rule eqvtResCommRel) + from cSim show ?case + proof(induct rule: resCommRel.induct) + case(resCommRel_refl \ P) + then show ?case + by(rule simI[OF \eqvt(resCommRel\<^sup>\)\]) (blast intro: resCommRel.intros) + next + case(resCommRel_swap a \ b P) + show ?case + by(rule resComm) (fact|auto intro: resCommRel.intros any_perm_in_rel)+ + next + case(resCommRel_res \ P Q a) + show ?case + by(rule resPres[where Rel="resCommRel\<^sup>\"]) (fact|simp add: resCommRelStarResChain)+ + qed + next + case(cExt \ R S \') then show ?case + proof(induct arbitrary: \' rule: resCommRel.induct) + case resCommRel_refl then show ?case by(rule resCommRel_refl) + next + case(resCommRel_swap a \ b P) + then show ?case + proof(cases "a=b") + case True show ?thesis unfolding \a=b\ by(rule resCommRel_refl) + next + case False + obtain x::name and y::name where "x\y" and "x \ \" and "x \ \'" and "x \ P" and "x \ a" and "x\b" and "y \ \" and "y \ \'" and "y \ P" and "y \ a" and "y\b" + apply(generate_fresh "name") + apply(generate_fresh "name") + by force + then show ?thesis using False + apply(subst (1 2) alphaRes[where x=a and y=x]) + apply assumption + apply(simp add: fresh_abs_fun_iff[OF pt_name_inst, OF at_name_inst, OF fin_supp]) + apply(subst (1 2) alphaRes[where x=b and y=y]) + apply(simp add: fresh_abs_fun_iff[OF pt_name_inst, OF at_name_inst, OF fin_supp] fresh_left) + apply assumption + unfolding eqvts + apply(subst (1) cp1[OF cp_pt_inst, OF pt_name_inst, OF at_name_inst]) + by(auto simp add: eqvts swap_simps intro: resCommRel.intros) + qed + next + case(resCommRel_res \ P Q a) + obtain b::name where "b \ \" and "b \ a" and "b \ P" and "b \ Q" and "b \ \'" + by(generate_fresh "name") auto + have "(\ \ ([(a,b)]\\'), P, Q) \ resCommRel" by fact + moreover from \b \ \'\ have "a \ ([(a, b)] \ \')" by(simp add: fresh_left swap_simps) + with \a \ \\ have "a \ \ \ ([(a,b)]\\')" by force + ultimately have "(\ \ ([(a,b)]\\'), \\a\P, \\a\Q) \ resCommRel" + by(rule resCommRel.intros) + then have "([(a,b)]\(\ \ ([(a,b)]\\'), \\a\P, \\a\Q)) \ resCommRel" using eqvtResCommRel + by(intro eqvtI) + then have "(\ \ \', \\b\([(a,b)]\P), \\b\([(a,b)]\Q)) \ resCommRel" using \a \ \\ \b \ \\ + by(simp add: eqvts swap_simps) + then show ?case using \b \ Q\ \b \ P\ + apply(subst (1 2) alphaRes[where x=a and y=b]) + by(assumption|simp only: eqvts)+ + qed + next + case(cSym \ R S) then show ?case + by(induct rule: resCommRel.inducts) (auto intro: resCommRel.intros) + qed + } + moreover obtain x'::name where "x' \ \" and "x' \ P" and "x' \ x" and "x' \ y" + by(generate_fresh "name") auto + moreover obtain y'::name where "y' \ \" and "y' \ P" and "y' \ x" and "y' \ y" and "y' \ x'" + by(generate_fresh "name") auto + ultimately have "\ \ \\x'\(\\y'\([(y, y'), (x, x')] \ P)) \ \\y'\(\\x'\([(y, y'), (x, x')] \ P))" by auto + then show ?thesis using \x' \ P\ \x' \ x\ \x' \ y\ \y' \ P\ \y' \ x\ \y' \ y\ \y' \ x'\ \x \ y\ + apply(subst alphaRes[where x=x and y=x' and P=P], auto) + apply(subst alphaRes[where x=y and y=y' and P=P], auto) + apply(subst alphaRes[where x=x and y=x' and P="\\y'\([(y, y')] \ P)"], auto simp add: abs_fresh fresh_left) + apply(subst alphaRes[where x=y and y=y' and P="\\x'\([(x, x')] \ P)"], auto simp add: abs_fresh fresh_left) + by(subst perm_compose) (simp add: eqvts calc_atm) +qed + +lemma bisimResComm': + fixes x :: name + and \ :: 'b + and xvec :: "name list" + and P :: "('a, 'b, 'c) psi" + +assumes "x \ \" + and "xvec \* \" + +shows "\ \ \\x\(\\*xvec\P) \ \\*xvec\(\\x\P)" + using assms + by(induct xvec) (auto intro: bisimResComm bisimReflexive bisimResPres bisimTransitive) + +lemma bisimParPresSym: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and Q :: "('a, 'b, 'c) psi" + and R :: "('a, 'b, 'c) psi" + +assumes "\ \ P \ Q" + +shows "\ \ R \ P \ R \ Q" + using assms + by(metis bisimParComm bisimParPres bisimTransitive) + +lemma bisimScopeExt: + fixes x :: name + and \ :: 'b + and P :: "('a, 'b, 'c) psi" + and Q :: "('a, 'b, 'c) psi" + +assumes "x \ P" + +shows "\ \ \\x\(P \ Q) \ P \ \\x\Q" +proof - + { + fix x::name and Q :: "('a, 'b, 'c) psi" + assume "x \ \" and "x \ P" + let ?X1 = "{((\::'b), \\*xvec\(\\*yvec\((P::('a, 'b, 'c) psi) \ Q)), \\*xvec\(P \ (\\*yvec\Q))) | \ xvec yvec P Q. yvec \* \ \ yvec \* P \ xvec \* \}" + let ?X2 = "{((\::'b), \\*xvec\((P::('a, 'b, 'c) psi) \ (\\*yvec\Q)), \\*xvec\(\\*yvec\(P \ Q))) | \ xvec yvec P Q. yvec \* \ \ yvec \* P \ xvec \* \}" + let ?X = "?X1 \ ?X2" + from \x \ \\ \x \ P\ have "(\, \\x\(P \ Q), P \ \\x\Q) \ ?X" + proof - + from \x \ \\ \x \ P\ have "(\, \\x\(P \ Q), P \ \\x\Q) \ ?X1" + by (smt (verit, del_insts) mem_Collect_eq freshSets(1) freshSets(5) resChain.base resChain.step) + then show ?thesis by auto + qed + moreover have "eqvt ?X" + by (rule eqvtUnion) (clarsimp simp add: eqvt_def eqvts, metis fresh_star_bij(2))+ + ultimately have "\ \ \\x\(P \ Q) \ P \ \\x\Q" + proof(coinduct rule: transitiveStarCoinduct) + case(cStatEq \ R T) + then have "(\, R, T) \ ?X1 \ (\, R, T) \ ?X2" + by blast + then show ?case + proof(rule disjE) + assume "(\, R, T) \ ?X1" + then obtain xvec yvec P Q where "R = \\*xvec\(\\*yvec\(P \ Q))" and "T = \\*xvec\(P \ (\\*yvec\Q))" and "xvec \* \" and "yvec \* P" and "yvec \* \" + by auto + moreover obtain A\<^sub>P \\<^sub>P where FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" and "A\<^sub>P \* \" and "yvec \* A\<^sub>P" and "A\<^sub>P \* Q" + by(rule freshFrame[where C="(\, yvec, Q)"]) auto + moreover obtain A\<^sub>Q \\<^sub>Q where FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" and "A\<^sub>Q \* \" and "yvec \* A\<^sub>Q" and "A\<^sub>Q \* A\<^sub>P" and "A\<^sub>Q \* \\<^sub>P" + by(rule freshFrame[where C="(\, yvec, A\<^sub>P, \\<^sub>P)"]) auto + moreover from FrQ \A\<^sub>P \* Q\ \A\<^sub>Q \* A\<^sub>P\ have "A\<^sub>P \* \\<^sub>Q" + by(auto dest: extractFrameFreshChain) + moreover from \yvec \* P\ \yvec \* A\<^sub>P\ FrP have "yvec \* \\<^sub>P" + by(drule_tac extractFrameFreshChain) auto + ultimately show ?case + by(auto simp add: frameChainAppend[symmetric] freshChainAppend) (auto simp add: frameChainAppend intro: frameResChainPres frameResChainComm) + next + assume "(\, R, T) \ ?X2" + then obtain xvec yvec P Q where "T = \\*xvec\(\\*yvec\(P \ Q))" and "R = \\*xvec\(P \ (\\*yvec\Q))" and "xvec \* \" and "yvec \* P" and "yvec \* \" + by auto + moreover obtain A\<^sub>P \\<^sub>P where FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" and "A\<^sub>P \* \" and "yvec \* A\<^sub>P" and "A\<^sub>P \* Q" + by(rule freshFrame[where C="(\, yvec, Q)"]) auto + moreover obtain A\<^sub>Q \\<^sub>Q where FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" and "A\<^sub>Q \* \" and "yvec \* A\<^sub>Q" and "A\<^sub>Q \* A\<^sub>P" and "A\<^sub>Q \* \\<^sub>P" + by(rule freshFrame[where C="(\, yvec, A\<^sub>P, \\<^sub>P)"]) auto + moreover from FrQ \A\<^sub>P \* Q\ \A\<^sub>Q \* A\<^sub>P\ have "A\<^sub>P \* \\<^sub>Q" + by(auto dest: extractFrameFreshChain) + moreover from \yvec \* P\ \yvec \* A\<^sub>P\ FrP have "yvec \* \\<^sub>P" + by(drule_tac extractFrameFreshChain) auto + ultimately show ?case + by(auto simp add: frameChainAppend[symmetric] freshChainAppend) (auto simp add: frameChainAppend intro: frameResChainPres frameResChainComm) + qed + next + case(cSim \ R T) + let ?Y = "{(\, P, Q) | \ P P' Q' Q. \ \ P \ P' \ ((\, P', Q') \ ?X \ \ \ P' \ Q') \ \ \ Q' \ Q}" + from \eqvt ?X\ have "eqvt ?Y" by blast + have C1: "\\ R T y. \(\, R, T) \ ?Y; (y::name) \ \\ \ (\, \\y\R, \\y\T) \ ?Y" + proof - + fix \ R T y + assume "(\, R, T) \ ?Y" + then obtain R' T' where "\ \ R \ R'" and "(\, R', T') \ (?X \ bisim)" and "\ \ T' \ T" by force + assume "(y::name) \ \" + show "(\, \\y\R, \\y\T) \ ?Y" + proof(cases "(\, R', T') \ ?X") + assume "(\, R', T') \ ?X" + show ?thesis + proof(cases "(\, R', T') \ ?X1") + assume "(\, R', T') \ ?X1" + then obtain xvec yvec P Q where R'eq: "R' = \\*xvec\(\\*yvec\(P \ Q))" and T'eq: "T' = \\*xvec\(P \ (\\*yvec\Q))" + and "xvec \* \" and "yvec \* P" and "yvec \* \" + by auto + from \\ \ R \ R'\ \y \ \\ have "\ \ \\y\R \ \\y\R'" by(rule bisimResPres) + moreover from \xvec \* \\ \y \ \\ \yvec \* P\ \yvec \* \\ have "(\, \\*(y#xvec)\\\*yvec\(P \ Q), \\*(y#xvec)\(P \ (\\*yvec\Q))) \ ?X1" + by(force simp del: resChain.simps) + with R'eq T'eq have "(\, \\y\R', \\y\T') \ ?X \ bisim" by simp + moreover from \\ \ T' \ T\ \y \ \\ have "\ \ \\y\T' \ \\y\T" by(rule bisimResPres) + ultimately show ?thesis by blast + next + assume "(\, R', T') \ ?X1" + with \(\, R', T') \ ?X\ have "(\, R', T') \ ?X2" by blast + then obtain xvec yvec P Q where T'eq: "T' = \\*xvec\(\\*yvec\(P \ Q))" and R'eq: "R' = \\*xvec\(P \ (\\*yvec\Q))" and "xvec \* \" and "yvec \* P" and "yvec \* \" + by auto + from \\ \ R \ R'\ \y \ \\ have "\ \ \\y\R \ \\y\R'" by(rule bisimResPres) + moreover from \xvec \* \\ \y \ \\ \yvec \* P\ \yvec \* \\ have "(\, \\*(y#xvec)\(P \ (\\*yvec\Q)), \\*(y#xvec)\(\\*yvec\(P \ Q))) \ ?X2" + by(force simp del: resChain.simps) + with R'eq T'eq have "(\, \\y\R', \\y\T') \ ?X \ bisim" by simp + moreover from \\ \ T' \ T\ \y \ \\ have "\ \ \\y\T' \ \\y\T" by(rule bisimResPres) + ultimately show ?thesis by blast + qed + next + assume "(\, R', T') \ ?X" + with \(\, R', T') \ ?X \ bisim\ have "\ \ R' \ T'" by blast + with \\ \ R \ R'\ \\ \ T' \ T\ \y \ \\ show ?thesis + by(blast dest: bisimResPres) + qed + qed + have C1': "\\ R T y. \(\, R, T) \ ?Y\<^sup>\; (y::name) \ \\ \ (\, \\y\R, \\y\T) \ ?Y\<^sup>\" + proof - + fix \ R + and T::"('a,'b,'c) psi" + and y::name + assume "(\, R, T) \ ?Y\<^sup>\" + and "y \ \" + then show "(\, \\y\R, \\y\T) \ ?Y\<^sup>\" + proof(induct rule: rel_trancl.induct) + case(r_into_rel_trancl \ P Q) + then show ?case + by (intro rel_trancl.intros(1)) (rule C1) + next + case(rel_trancl_into_rel_trancl \ P Q R) + then show ?case + using rel_trancl.intros(2) C1 by meson + qed + qed + have C1'': "\\ R T yvec. \(\, R, T) \ ?Y\<^sup>\; (yvec::name list) \* \\ \ (\, \\*yvec\R, \\*yvec\T) \ ?Y\<^sup>\" + proof - + fix \ R T + and yvec::"name list" + assume "(\, R, T) \ ?Y\<^sup>\" + and "yvec \* \" + then show "(\, \\*yvec\R, \\*yvec\T) \ ?Y\<^sup>\" + apply(induct yvec) + apply simp + unfolding resChain.simps + by(rule C1') simp+ + qed + have C2: "\y \' R S zvec. \y \ \'; y \ R; zvec \* \'\ \ (\', \\y\(\\*zvec\(R \ S)), \\*zvec\(R \ \\y\S)) \ ?Y\<^sup>\" + proof - + fix y::name + and \::'b + and R::"('a,'b,'c) psi" + and S::"('a,'b,'c) psi" + and zvec::"name list" + assume "y \ \" + and "y \ R" + and "zvec \* \" + have "\ \ \\y\(\\*zvec\(R \ S)) \ (\\*zvec\(\\y\(R \ S)))" + by(rule bisimResComm') fact+ + moreover have "(\, (\\*zvec\(\\y\(R \ S))), \\*zvec\(R \ \\y\S)) \ ?X" using \y \ \\ \y \ R\ \zvec \* \\ + apply clarsimp + apply(rule exI[where x=zvec]) + apply(rule exI[where x="[y]"]) + by auto + moreover have "\ \ \\*zvec\(R \ \\y\S) \ \\*zvec\(R \ \\y\S)" + by(rule bisimReflexive) + ultimately show "(\, \\y\(\\*zvec\(R \ S)), \\*zvec\(R \ \\y\S)) \ ?Y\<^sup>\" + by blast + qed + have C2': "\y \' R S zvec. \y \ \'; y \ R; zvec \* \'\ \ (\', \\*zvec\(R \ \\y\S), \\y\(\\*zvec\(R \ S))) \ ?Y\<^sup>\" + proof - + fix y::name + and \::'b + and R::"('a,'b,'c) psi" + and S::"('a,'b,'c) psi" + and zvec::"name list" + assume "y \ \" + and "y \ R" + and "zvec \* \" + have "\ \ (\\*zvec\(\\y\(R \ S))) \ \\y\(\\*zvec\(R \ S))" + by(rule bisimSymmetric[OF bisimResComm']) fact+ + moreover have "(\, \\*zvec\(R \ \\y\S), (\\*zvec\(\\y\(R \ S)))) \ ?X" using \y \ \\ \y \ R\ \zvec \* \\ + apply(intro UnI2) + apply clarsimp + apply(rule exI[where x=zvec]) + apply(rule exI[where x="[y]"]) + by auto + moreover have "\ \ \\*zvec\(R \ \\y\S) \ \\*zvec\(R \ \\y\S)" + by(rule bisimReflexive) + ultimately show "(\, \\*zvec\(R \ \\y\S), \\y\(\\*zvec\(R \ S))) \ ?Y\<^sup>\" + by blast + qed + have C3: "\\' zvec R y. \y \ \'; zvec \* \'\ \ (\', \\y\(\\*zvec\R), \\*zvec\(\\y\R)) \ ?Y\<^sup>\" + proof - + fix y::name + and \::'b + and R::"('a,'b,'c) psi" + and zvec::"name list" + assume "y \ \" + and "zvec \* \" + then have "\ \ \\y\(\\*zvec\R) \ \\*zvec\(\\y\R)" + by(rule bisimResComm') + then show "(\, \\y\(\\*zvec\R), \\*zvec\(\\y\R)) \ ?Y\<^sup>\" + by(blast intro: bisimReflexive) + qed + have C4: "\\' R S zvec. \zvec \* R; zvec \* \'\ \ (\', (\\*zvec\(R \ S)), (R \ (\\*zvec\S))) \ ?Y\<^sup>\" + proof - + fix \::'b + and R::"('a,'b,'c) psi" + and S::"('a,'b,'c) psi" + and zvec::"name list" + assume "zvec \* R" + and "zvec \* \" + then have "(\, (\\*zvec\(R \ S)), (R \ (\\*zvec\S))) \ ?X" + apply clarsimp + apply(rule exI[where x="[]"]) + apply(rule exI[where x=zvec]) + by auto + then show "(\, (\\*zvec\(R \ S)), (R \ (\\*zvec\S))) \ ?Y\<^sup>\" + by(blast intro: bisimReflexive) + qed + have C4': "\\' R S zvec. \zvec \* R; zvec \* \'\ \ (\', (R \ (\\*zvec\S)), (\\*zvec\(R \ S))) \ ?Y\<^sup>\" + proof - + fix \::'b + and R::"('a,'b,'c) psi" + and S::"('a,'b,'c) psi" + and zvec::"name list" + assume "zvec \* R" + and "zvec \* \" + then have "(\, (R \ (\\*zvec\S)), (\\*zvec\(R \ S))) \ ?X" + apply(intro UnI2) + apply clarsimp + apply(rule exI[where x="[]"]) + apply(rule exI[where x=zvec]) + by auto + then show "(\, (R \ (\\*zvec\S)), (\\*zvec\(R \ S))) \ ?Y\<^sup>\" + by(blast intro: bisimReflexive) + qed + { + fix \ P Q R + assume "\ \ P \[?Y\<^sup>\] Q" + and "\ \ Q \[?Y\<^sup>\] R" + moreover note rel_trancl_eqvt[OF \eqvt ?Y\] + moreover have "{(\, P, R) | \ P R. \Q. (\, P, Q) \ ?Y\<^sup>\ \ (\, Q, R) \ ?Y\<^sup>\} \ ?Y\<^sup>\" + by(auto intro: rel_trancl_transitive) + ultimately have "\ \ P \[?Y\<^sup>\] R" + by(rule transitive) + } + note trans = this + + show ?case + proof(cases "(\, R, T) \ ?X1") + assume "(\, R, T) \ ?X1" + then obtain xvec yvec P Q where Req: "R = \\*xvec\(\\*yvec\(P \ Q))" and Teq: "T = \\*xvec\(P \ (\\*yvec\Q))" and "xvec \* \" and "yvec \* P" and "yvec \* \" + by auto + have "\ \ \\*xvec\(\\*yvec\(P \ Q)) \[?Y\<^sup>\] \\*xvec\(P \ (\\*yvec\Q))" + proof - + have "\ \ \\*yvec\(P \ Q) \[?Y\<^sup>\] P \ (\\*yvec\Q)" using \yvec \* \\ \yvec \* P\ + proof(induct yvec arbitrary: Q) + case Nil show ?case + unfolding resChain.simps + by(rule monotonic[where A="?Y"]) (blast intro: reflexive bisimReflexive)+ + next + case(Cons y yvec) + then have "yvec \* P" and "yvec \* \" and "y \ \" and "y \ P" + by simp+ + have "\ \ \\*(y#yvec)\P \ Q \[?Y\<^sup>\] \\*yvec\(\\y\(P \ Q))" + proof - + have "\ \ \\*(y#yvec)\P \ Q \[bisim] \\*yvec\(\\y\(P \ Q))" + unfolding resChain.simps + apply(rule bisimE) + apply(rule bisimResComm') + by fact+ + then have "\ \ \\*(y#yvec)\P \ Q \[?Y] \\*yvec\(\\y\(P \ Q))" + apply - + apply(drule monotonic[where B="?Y"]) + by auto (blast intro: bisimTransitive bisimReflexive) + then show ?thesis + apply - + apply(drule monotonic[where B="?Y\<^sup>\"]) + by blast + qed + moreover have "\ \ \\*yvec\(\\y\(P \ Q)) \[?Y\<^sup>\] \\*yvec\(P \ \\y\Q)" + proof - + have "\ \ \\y\(P \ Q) \[?Y\<^sup>\] P \ \\y\Q" + apply(rule scopeExtLeft) + apply fact + apply fact + apply(rule rel_trancl_eqvt) + apply fact + apply(blast intro: bisimReflexive) + by fact+ + then show ?thesis using \yvec \* \\ + proof(induct yvec) + case Nil then show ?case by simp + next + case(Cons y' yvec) + then show ?case + unfolding resChain.simps + apply - + apply(rule resPres[where Rel="?Y\<^sup>\" and Rel'="?Y\<^sup>\"]) + apply simp + apply(rule rel_trancl_eqvt[OF \eqvt ?Y\]) + apply simp + apply simp + by(rule C1'') + qed + qed + moreover have "\ \ \\*yvec\(P \ \\y\Q) \[?Y\<^sup>\] P \ (\\*yvec\(\\y\Q))" + by(rule Cons) fact+ + moreover have "\ \ P \ (\\*yvec\(\\y\Q)) \[?Y\<^sup>\] P \ (\\*(y#yvec)\Q)" + proof - + have "\ \ P \ (\\*yvec\(\\y\Q)) \[bisim] P \ (\\*(y#yvec)\Q)" + unfolding resChain.simps + apply(rule bisimE) + apply(rule bisimParPresSym) + apply(rule bisimSymmetric[OF bisimResComm']) + by fact+ + then have "\ \ P \ (\\*yvec\(\\y\Q)) \[?Y] P \ (\\*(y#yvec)\Q)" + apply - + apply(drule monotonic[where B="?Y"]) + by auto (blast intro: bisimTransitive bisimReflexive) + then show ?thesis + apply - + apply(drule monotonic[where B="?Y\<^sup>\"]) + by blast + qed + moreover have "\ \ \\*yvec\P \ \\y\Q \[?Y\<^sup>\] P \ (\\*yvec\(\\y\Q))" + by(rule Cons) fact+ + moreover have "\ \ P \ (\\*yvec\(\\y\Q)) \[?Y\<^sup>\] P \ (\\*(y#yvec)\Q)" + proof - + have "\ \ P \ (\\*yvec\(\\y\Q)) \[bisim] P \ (\\y\(\\*yvec\Q))" + apply(rule bisimE) + apply(rule bisimParPresSym) + apply(rule bisimSymmetric[OF bisimResComm']) + by fact+ + then have "\ \ P \ (\\*yvec\(\\y\Q)) \[?Y] P \ (\\y\(\\*yvec\Q))" + apply - + apply(drule monotonic[where B="?Y"]) + by auto (blast intro: bisimTransitive bisimReflexive) + then show ?thesis + unfolding resChain.simps + apply - + apply(drule monotonic[where B="?Y\<^sup>\"]) + by blast + qed + ultimately show ?case + by(blast dest: trans) + qed + then show ?thesis using \xvec \* \\ + apply(induct xvec) + apply simp + unfolding resChain.simps + apply(rule resPres) + apply simp + apply(rule rel_trancl_eqvt[OF \eqvt ?Y\]) + apply simp + apply simp + apply(rule C1'') + apply simp + by assumption + qed + then show ?case unfolding Req Teq + by - + next + assume "(\, R, T) \ ?X1" + then have "(\, R, T) \ ?X2" using \(\, R, T) \ ?X\ by blast + then obtain xvec yvec P Q where Teq: "T = \\*xvec\(\\*yvec\(P \ Q))" and Req: "R = \\*xvec\(P \ (\\*yvec\Q))" and "xvec \* \" and "yvec \* P" and "yvec \* \" + by auto + have "\ \ \\*xvec\(P \ (\\*yvec\Q)) \[?Y\<^sup>\] \\*xvec\(\\*yvec\(P \ Q))" + proof - + have "\ \ P \ (\\*yvec\Q) \[?Y\<^sup>\] \\*yvec\(P \ Q)" using \yvec \* P\ \yvec \* \\ + proof(induct yvec arbitrary: Q) + case Nil show ?case + unfolding resChain.simps + by(rule monotonic[where A="?Y"]) (blast intro: reflexive bisimReflexive)+ + next + case(Cons y yvec) + then have "yvec \* P" and "yvec \* \" and "y \ \" and "y \ P" + by simp+ + have "\ \ \\*yvec\(\\y\(P \ Q)) \[?Y\<^sup>\] \\*(y#yvec)\P \ Q" + proof - + have "\ \ \\*yvec\(\\y\(P \ Q)) \[bisim] \\*(y#yvec)\P \ Q" + unfolding resChain.simps + apply(rule bisimE) + apply(rule bisimSymmetric[OF bisimResComm']) + by fact+ + then have "\ \ \\*yvec\(\\y\(P \ Q)) \[?Y] \\*(y#yvec)\P \ Q" + apply - + apply(drule monotonic[where B="?Y"]) + by auto (blast intro: bisimTransitive bisimReflexive) + then show ?thesis + apply - + apply(drule monotonic[where B="?Y\<^sup>\"]) + by blast + qed + moreover have "\ \ \\*yvec\(P \ \\y\Q) \[?Y\<^sup>\] \\*yvec\(\\y\(P \ Q))" + proof - + have "\ \ P \ \\y\Q \[?Y\<^sup>\] \\y\(P \ Q)" + apply(rule scopeExtRight) + apply fact + apply fact + apply(rule rel_trancl_eqvt) + apply fact + apply(blast intro: bisimReflexive) + by fact+ + then show ?thesis using \yvec \* \\ + proof(induct yvec) + case Nil then show ?case by simp + next + case(Cons y' yvec) + then show ?case + unfolding resChain.simps + apply - + apply(rule resPres[where Rel="?Y\<^sup>\" and Rel'="?Y\<^sup>\"]) + apply simp + apply(rule rel_trancl_eqvt[OF \eqvt ?Y\]) + apply simp + apply simp + by(rule C1'') + qed + moreover have "\ \ P \ (\\*yvec\(\\y\Q)) \[?Y\<^sup>\] \\*yvec\(P \ \\y\Q)" + by(rule Cons) fact+ + moreover have "\ \ P \ (\\*(y#yvec)\Q) \[?Y\<^sup>\] P \ (\\*yvec\(\\y\Q))" + proof - + have "\ \ P \ (\\*(y#yvec)\Q) \[bisim] P \ (\\*yvec\(\\y\Q))" + unfolding resChain.simps + apply(rule bisimE) + apply(rule bisimParPresSym) + apply(rule bisimResComm') + by fact+ + then have "\ \ P \ (\\*(y#yvec)\Q) \[?Y] P \ (\\*yvec\(\\y\Q))" + apply - + apply(drule monotonic[where B="?Y"]) + by auto (blast intro: bisimTransitive bisimReflexive) + then show ?thesis + apply - + apply(drule monotonic[where B="?Y\<^sup>\"]) + by blast + qed + qed + moreover have "\ \ P \ (\\*yvec\(\\y\Q)) \[?Y\<^sup>\] \\*yvec\P \ \\y\Q" + by(rule Cons) fact+ + moreover have "\ \ P \ (\\*(y#yvec)\Q) \[?Y\<^sup>\] P \ (\\*yvec\(\\y\Q))" + proof - + have "\ \ P \ (\\y\(\\*yvec\Q)) \[bisim] P \ (\\*yvec\(\\y\Q))" + apply(rule bisimE) + apply(rule bisimParPresSym) + apply(rule bisimResComm') + by fact+ + then have "\ \ P \ (\\y\(\\*yvec\Q)) \[?Y] P \ (\\*yvec\(\\y\Q))" + apply - + apply(drule monotonic[where B="?Y"]) + by auto (blast intro: bisimTransitive bisimReflexive) + then show ?thesis + unfolding resChain.simps + apply - + apply(drule monotonic[where B="?Y\<^sup>\"]) + by blast + qed + ultimately show ?case + by(blast dest: trans) + qed + then show ?thesis using \xvec \* \\ + apply(induct xvec) + apply simp + unfolding resChain.simps + apply(rule resPres) + apply simp + apply(rule rel_trancl_eqvt[OF \eqvt ?Y\]) + apply simp + apply simp + apply(rule C1'') + apply simp + by assumption + qed + then show ?case unfolding Req Teq + by - + qed + next + case(cExt \ R T \') + show ?case + proof(cases "(\, R, T) \ ?X1") + assume "(\, R, T) \ ?X1" + then obtain xvec yvec P Q where Req: "R = \\*xvec\(\\*yvec\(P \ Q))" and Teq: "T = \\*xvec\(P \ (\\*yvec\Q))" and "xvec \* \" and "yvec \* P" and "yvec \* \" + by auto + obtain p where "(p \ yvec) \* \" and "(p \ yvec) \* P" and "(p \ yvec) \* Q" and "(p \ yvec) \* \'" and "(p \ yvec) \* xvec" + and S: "(set p) \ (set yvec) \ (set(p \ yvec))" and "distinctPerm p" + by(rule name_list_avoiding[where c="(\, P, Q, xvec, \')"]) auto + obtain q where "(q \ xvec) \* \" and "(q \ xvec) \* \'" and "(q \ xvec) \* P" and "(q \ xvec) \* yvec" and "(q \ xvec) \* Q" and "(q \ xvec) \* (p\P)" and "(q \ xvec) \* (p\Q)" and "distinctPerm q" and T: "(set q) \ (set xvec) \ (set(q\xvec))" and "(q \ xvec) \* (p\yvec)" + by(rule name_list_avoiding[where c="(\, P, Q, yvec, p\yvec, \',p\P,p\Q)"]) auto + note \(p \ yvec) \* Q\ \set p \ set yvec \ set (p \ yvec)\ + moreover have "(p \ yvec) \* (P \ Q)" using \(p \ yvec) \* Q\ \(p \ yvec) \* P\ + by simp + moreover have "(\ \ \', \\*xvec\\\*p \ yvec\p \ P \ Q, \\*xvec\P \ (\\*p \ yvec\p \ Q)) \ ?X" + proof - + have Pdef: "(p \ P) = P" using \set p \ set yvec \ set(p\yvec)\ \yvec \* P\ \(p\yvec) \* P\ + by simp + have yvecdef: "(q \ p \ yvec) = (p\ yvec)" using \set q \ set xvec \ set(q\xvec)\ \(p\yvec) \* xvec\ \(q\xvec) \* (p\yvec)\ + by simp + have "(q \ xvec) \* (P \ (\\*p \ yvec\p \ Q))" using \(q \ xvec) \* P\ \(q \ xvec) \* (p\Q)\ \(q \ xvec) \* (p\yvec)\ + by simp + moreover have "(q \ xvec) \* (\\*p \ yvec\p \ P \ Q)" + unfolding eqvts Pdef + using \(q \ xvec) \* P\ \(q \ xvec) \* (p\Q)\ \(q \ xvec) \* (p\yvec)\ + by simp + moreover note \set q \ set xvec \ set (q \ xvec)\ + moreover have "(\ \ \', \\*q \ xvec\q \ \\*p \ yvec\p \ P \ Q, + \\*q \ xvec\q \ P \ (\\*p \ yvec\p \ Q)) \ ?X" + proof - + have "(p\yvec) \* (\\\')" using \(p\yvec) \* \\ \(p\yvec) \* \'\ + by auto + moreover have "(p\yvec) \* (q\P)" using \(p\yvec) \* P\ + apply(subst yvecdef[symmetric]) + by(subst fresh_star_bij) + moreover have "(q\xvec) \* (\\\')" using \(q\xvec) \* \\ \(q\xvec) \* \'\ + by auto + ultimately show ?thesis + unfolding Pdef eqvts yvecdef + by blast + qed + ultimately show ?thesis + by(subst (1 2) resChainAlpha[where p=q and xvec=xvec]) + qed + ultimately show ?case unfolding Req Teq + apply(intro disjI1) + by(subst (1 2) resChainAlpha[where p=p and xvec=yvec]) + next + assume "(\, R, T) \ ?X1" + then have "(\, R, T) \ ?X2" using \(\,R,T) \ ?X\ + by blast + then obtain xvec yvec P Q where Teq: "T = \\*xvec\(\\*yvec\(P \ Q))" and Req: "R = \\*xvec\(P \ (\\*yvec\Q))" and "xvec \* \" and "yvec \* P" and "yvec \* \" + by auto + obtain p where "(p \ yvec) \* \" and "(p \ yvec) \* P" and "(p \ yvec) \* Q" and "(p \ yvec) \* \'" and "(p \ yvec) \* xvec" + and S: "(set p) \ (set yvec) \ (set(p \ yvec))" and "distinctPerm p" + by(rule name_list_avoiding[where c="(\, P, Q, xvec, \')"]) auto + obtain q where "(q \ xvec) \* \" and "(q \ xvec) \* \'" and "(q \ xvec) \* P" and "(q \ xvec) \* yvec" and "(q \ xvec) \* Q" and "(q \ xvec) \* (p\P)" and "(q \ xvec) \* (p\Q)" and "distinctPerm q" and T: "(set q) \ (set xvec) \ (set(q\xvec))" and "(q \ xvec) \* (p\yvec)" + by(rule name_list_avoiding[where c="(\, P, Q, yvec, p\yvec, \',p\P,p\Q)"]) auto + note \(p \ yvec) \* Q\ \set p \ set yvec \ set (p \ yvec)\ + moreover have "(p \ yvec) \* (P \ Q)" using \(p \ yvec) \* Q\ \(p \ yvec) \* P\ + by simp + moreover have "(\ \ \', \\*xvec\P \ (\\*p \ yvec\p \ Q), \\*xvec\\\*p \ yvec\p \ P \ Q) \ ?X" + proof - + have Pdef: "(p \ P) = P" using \set p \ set yvec \ set(p\yvec)\ \yvec \* P\ \(p\yvec) \* P\ + by simp + have yvecdef: "(q \ p \ yvec) = (p\ yvec)" using \set q \ set xvec \ set(q\xvec)\ \(p\yvec) \* xvec\ \(q\xvec) \* (p\yvec)\ + by simp + have "(q \ xvec) \* (P \ (\\*p \ yvec\p \ Q))" using \(q \ xvec) \* P\ \(q \ xvec) \* (p\Q)\ \(q \ xvec) \* (p\yvec)\ + by simp + moreover have "(q \ xvec) \* (\\*p \ yvec\p \ P \ Q)" + unfolding eqvts Pdef + using \(q \ xvec) \* P\ \(q \ xvec) \* (p\Q)\ \(q \ xvec) \* (p\yvec)\ + by simp + moreover note \set q \ set xvec \ set (q \ xvec)\ + moreover have "(\ \ \', \\*q \ xvec\q \ P \ (\\*p \ yvec\p \ Q), \\*q \ xvec\q \ \\*p \ yvec\p \ P \ Q) \ ?X" + proof - + have "(p\yvec) \* (\\\')" using \(p\yvec) \* \\ \(p\yvec) \* \'\ + by auto + moreover have "(p\yvec) \* (q\P)" using \(p\yvec) \* P\ + apply(subst yvecdef[symmetric]) + by(subst fresh_star_bij) + moreover have "(q\xvec) \* (\\\')" using \(q\xvec) \* \\ \(q\xvec) \* \'\ + by auto + ultimately show ?thesis + unfolding Pdef eqvts yvecdef + by blast + qed + ultimately show ?thesis + by(subst (1 2) resChainAlpha[where p=q and xvec=xvec]) + qed + ultimately show ?case unfolding Req Teq + apply(intro disjI1) + by(subst (1 2) resChainAlpha[where p=p and xvec=yvec]) + qed + next + case(cSym \ P Q) + then show ?case + by(blast dest: bisimE) + qed + } + moreover obtain y::name where "y \ \" and "y \ P" "y \ Q" + by(generate_fresh "name") auto + ultimately have "\ \ \\y\(P \ ([(x, y)] \ Q)) \ P \ \\y\([(x, y)] \ Q)" by auto + then show ?thesis using assms \y \ P\ \y \ Q\ + apply(subst alphaRes[where x=x and y=y and P=Q], auto) + by(subst alphaRes[where x=x and y=y and P="P \ Q"]) auto +qed + +lemma bisimScopeExtChain: + fixes xvec :: "name list" + and \ :: 'b + and P :: "('a, 'b, 'c) psi" + and Q :: "('a, 'b, 'c) psi" + +assumes "xvec \* \" + and "xvec \* P" + +shows "\ \ \\*xvec\(P \ Q) \ P \ (\\*xvec\Q)" + using assms + by(induct xvec) (auto intro: bisimScopeExt bisimReflexive bisimTransitive bisimResPres) + +lemma bisimParAssoc: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and Q :: "('a, 'b, 'c) psi" + and R :: "('a, 'b, 'c) psi" + +shows "\ \ (P \ Q) \ R \ P \ (Q \ R)" +proof - + let ?X = "{(\, \\*xvec\((P \ Q) \ R), \\*xvec\(P \ (Q \ R))) | \ xvec P Q R. xvec \* \}" + let ?Y = "{(\, P, Q) | \ P P' Q' Q. \ \ P \ P' \ (\, P', Q') \ ?X \ \ \ Q' \ Q}" + + have "(\, (P \ Q) \ R, P \ (Q \ R)) \ ?X" + by(clarsimp, rule exI[where x="[]"]) auto + moreover have "eqvt ?X" by(force simp add: eqvt_def simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst] eqvts) + ultimately show ?thesis + proof(coinduct rule: weakTransitiveCoinduct') + case(cStatEq \ PQR PQR') + from \(\, PQR, PQR') \ ?X\ obtain xvec P Q R where "xvec \* \" and "PQR = \\*xvec\((P \ Q) \ R)" and "PQR' = \\*xvec\(P \ (Q \ R))" + by auto + moreover obtain A\<^sub>P \\<^sub>P where FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" and "A\<^sub>P \* \" and "A\<^sub>P \* Q" and "A\<^sub>P \* R" + by(rule freshFrame[where C="(\, Q, R)"]) auto + moreover obtain A\<^sub>Q \\<^sub>Q where FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" and "A\<^sub>Q \* \" and "A\<^sub>Q \* A\<^sub>P" and "A\<^sub>Q \* \\<^sub>P" and "A\<^sub>Q \* R" + by(rule freshFrame[where C="(\, A\<^sub>P, \\<^sub>P, R)"]) auto + moreover obtain A\<^sub>R \\<^sub>R where FrR: "extractFrame R = \A\<^sub>R, \\<^sub>R\" and "A\<^sub>R \* \" and "A\<^sub>R \* A\<^sub>P" and "A\<^sub>R \* \\<^sub>P" and "A\<^sub>R \* A\<^sub>Q" and "A\<^sub>R \* \\<^sub>Q" + by(rule freshFrame[where C="(\, A\<^sub>P, \\<^sub>P, A\<^sub>Q, \\<^sub>Q)"]) auto + moreover from FrQ \A\<^sub>P \* Q\ \A\<^sub>Q \* A\<^sub>P\ have "A\<^sub>P \* \\<^sub>Q" + by(auto dest: extractFrameFreshChain) + moreover from FrR \A\<^sub>P \* R\ \A\<^sub>R \* A\<^sub>P\ have "A\<^sub>P \* \\<^sub>R" + by(auto dest: extractFrameFreshChain) + moreover from FrR \A\<^sub>Q \* R\ \A\<^sub>R \* A\<^sub>Q\ have "A\<^sub>Q \* \\<^sub>R" + by(auto dest: extractFrameFreshChain) + ultimately show ?case using freshCompChain + by auto (metis frameChainAppend compositionSym Associativity frameNilStatEq frameResChainPres) + next + case(cSim \ T S) + from \(\, T, S) \ ?X\ obtain xvec P Q R where "xvec \* \" and TEq: "T = \\*xvec\((P \ Q) \ R)" + and SEq: "S = \\*xvec\(P \ (Q \ R))" + by auto + from \eqvt ?X\have "eqvt ?Y" by blast + have C1: "\\ T S yvec. \(\, T, S) \ ?Y; yvec \* \\ \ (\, \\*yvec\T, \\*yvec\S) \ ?Y" + proof - + fix \ T S yvec + assume "(\, T, S) \ ?Y" + then obtain T' S' where "\ \ T \ T'" and "(\, T', S') \ ?X" and "\ \ S' \ S" by force + assume "(yvec::name list) \* \" + from \(\, T', S') \ ?X\ obtain xvec P Q R where T'eq: "T' = \\*xvec\((P \ Q) \ R)" and S'eq: "S' = \\*xvec\(P \ (Q \ R))" + and "xvec \* \" + by auto + from \\ \ T \ T'\ \yvec \* \\ have "\ \ \\*yvec\T \ \\*yvec\T'" by(rule bisimResChainPres) + moreover from \xvec \* \\ \yvec \* \\ have "(\, \\*(yvec@xvec)\((P \ Q) \ R), \\*(yvec@xvec)\(P \ (Q \ R))) \ ?X" + by force + with T'eq S'eq have "(\, \\*yvec\T', \\*yvec\S') \ ?X" by(simp add: resChainAppend) + moreover from \\ \ S' \ S\ \yvec \* \\ have "\ \ \\*yvec\S' \ \\*yvec\S" by(rule bisimResChainPres) + ultimately show "(\, \\*yvec\T, \\*yvec\S) \ ?Y" by blast + qed + + { + fix y + have "\\ T S. \(\, T, S) \ ?Y; y \ \\ \ (\, \\y\T, \\y\S) \ ?Y" + by(drule C1[where yvec2="[y]"]) auto + } + note C2 = this + + have "\ \ \\*xvec\((P \ Q) \ R) \[?Y] \\*xvec\(P \ (Q \ R))" + proof - + have "\ \ (P \ Q) \ R \[?Y] P \ (Q \ R)" + proof - + note \eqvt ?Y\ + moreover have "\\ P Q R. (\, (P \ Q) \ R, P \ (Q \ R)) \ ?Y" + proof - + fix \ P Q R + have "(\::'b, ((P::('a, 'b, 'c) psi) \ Q) \ R, P \ (Q \ R)) \ ?X" + by(clarsimp, rule exI[where x="[]"]) auto + then show "(\, (P \ Q) \ R, P \ (Q \ R)) \ ?Y" + by(blast intro: bisimReflexive) + qed + moreover have "\xvec \ P Q R. \xvec \* \; xvec \* P\ \ (\, \\*xvec\((P \ Q) \ R), P \ (\\*xvec\(Q \ R))) \ ?Y" + proof - + fix xvec \ P Q R + assume "(xvec::name list) \* (\::'b)" and "xvec \* (P::('a, 'b, 'c) psi)" + from \xvec \* \\ have "(\, \\*xvec\((P \ Q) \ R), \\*xvec\(P \ (Q \ R))) \ ?X" by blast + moreover from \xvec \* \\ \xvec \* P\ have "\ \ \\*xvec\(P \ (Q \ R)) \ P \ (\\*xvec\(Q \ R))" + by(rule bisimScopeExtChain) + ultimately show "(\, \\*xvec\((P \ Q) \ R), P \ (\\*xvec\(Q \ R))) \ ?Y" + by(blast intro: bisimReflexive) + qed + moreover have "\xvec \ P Q R. \xvec \* \; xvec \* R\ \ (\, (\\*xvec\(P \ Q)) \ R, \\*xvec\(P \ (Q \ R))) \ ?Y" + proof - + fix xvec \ P Q R + assume "(xvec::name list) \* (\::'b)" and "xvec \* (R::('a, 'b, 'c) psi)" + have "\ \ (\\*xvec\(P \ Q)) \ R \ R \ (\\*xvec\(P \ Q))" by(rule bisimParComm) + moreover from \xvec \* \\ \xvec \* R\ have "\ \ \\*xvec\(R \ (P \ Q)) \ R \ (\\*xvec\(P \ Q))" by(rule bisimScopeExtChain) + then have "\ \ R \ (\\*xvec\(P \ Q)) \ \\*xvec\(R \ (P \ Q))" by(rule bisimE) + moreover from \xvec \* \\ have "\ \ \\*xvec\(R \ (P \ Q)) \ \\*xvec\((P \ Q) \ R)" + by(metis bisimResChainPres bisimParComm) + moreover from \xvec \* \\ have "(\, \\*xvec\((P \ Q) \ R), \\*xvec\(P \ (Q \ R))) \ ?X" by blast + ultimately show "(\, (\\*xvec\(P \ Q)) \ R, \\*xvec\(P \ (Q \ R))) \ ?Y" by(blast dest: bisimTransitive intro: bisimReflexive) + qed + ultimately show ?thesis using C1 + by(rule parAssocLeft) + qed + then show ?thesis using \eqvt ?Y\ \xvec \* \\ C1 + by(rule resChainPres) + qed + with TEq SEq show ?case by simp + next + case(cExt \ T S \') + from \(\, T, S) \ ?X\ obtain xvec P Q R where "xvec \* \" and TEq: "T = \\*xvec\((P \ Q) \ R)" + and SEq: "S = \\*xvec\(P \ (Q \ R))" + by auto + obtain p where "(p \ xvec) \* \" and "(p \ xvec) \* P" and "(p \ xvec) \* Q" and "(p \ xvec) \* R" and "(p \ xvec) \* \'" + and S: "(set p) \ (set xvec) \ (set(p \ xvec))" and "distinctPerm p" + by(rule name_list_avoiding[where c="(\, P, Q, R, \')"]) auto + + from \(p \ xvec) \* \\ \(p \ xvec) \* \'\ have "(\ \ \', \\*(p \ xvec)\(((p \ P) \ (p \ Q)) \ (p \ R)), \\*(p \ xvec)\((p \ P) \ ((p \ Q) \ (p \ R)))) \ ?X" + by auto + moreover from TEq \(p \ xvec) \* P\ \(p \ xvec) \* Q\ \(p \ xvec) \* R\ S have "T = \\*(p \ xvec)\(((p \ P) \ (p \ Q)) \ (p \ R))" + apply clarsimp by(subst resChainAlpha[of p]) auto + moreover from SEq \(p \ xvec) \* P\ \(p \ xvec) \* Q\ \(p \ xvec) \* R\ S have "S = \\*(p \ xvec)\((p \ P) \ ((p \ Q) \ (p \ R)))" + apply clarsimp by(subst resChainAlpha[of p]) auto + ultimately show ?case by simp + next + case(cSym \ T S) + from \(\, T, S) \ ?X\ obtain xvec P Q R where "xvec \* \" and TEq: "T = \\*xvec\((P \ Q) \ R)" + and SEq: "\\*xvec\(P \ (Q \ R)) = S" + by auto + + from \xvec \* \\ have "\ \ \\*xvec\(P \ (Q \ R)) \ \\*xvec\((R \ Q) \ P)" + by(metis bisimParComm bisimParPres bisimTransitive bisimResChainPres) + moreover from \xvec \* \\ have "(\, \\*xvec\((R \ Q) \ P), \\*xvec\(R \ (Q \ P))) \ ?X" by blast + moreover from \xvec \* \\ have "\ \ \\*xvec\(R \ (Q \ P)) \ \\*xvec\((P \ Q) \ R)" + by(metis bisimParComm bisimParPres bisimTransitive bisimResChainPres) + ultimately show ?case using TEq SEq by(blast dest: bisimTransitive) + qed +qed + +lemma bisimParNil: + fixes P :: "('a, 'b, 'c) psi" + +shows "\ \ P \ \ \ P" +proof - + let ?X1 = "{(\, P \ \, P) | \ P. True}" + let ?X2 = "{(\, P, P \ \) | \ P. True}" + let ?X = "?X1 \ ?X2" + have "eqvt ?X" by(auto simp add: eqvt_def) + have "(\, P \ \, P) \ ?X" by simp + then show ?thesis + proof(coinduct rule: bisimWeakCoinduct) + case(cStatEq \ Q R) + show ?case + proof(cases "(\, Q, R) \ ?X1") + assume "(\, Q, R) \ ?X1" + then obtain P where "Q = P \ \" and "R = P" by auto + moreover obtain A\<^sub>P \\<^sub>P where "extractFrame P = \A\<^sub>P, \\<^sub>P\" and "A\<^sub>P \* \" + by(rule freshFrame) + ultimately show ?case + apply clarsimp by(metis frameResChainPres frameNilStatEq Identity Associativity AssertionStatEqTrans Commutativity) + next + assume "(\, Q, R) \ ?X1" + with \(\, Q, R) \ ?X\ have "(\, Q, R) \ ?X2" by blast + then obtain P where "Q = P" and "R = P \ \" by auto + moreover obtain A\<^sub>P \\<^sub>P where "extractFrame P = \A\<^sub>P, \\<^sub>P\" and "A\<^sub>P \* \" + by(rule freshFrame) + ultimately show ?case + apply clarsimp by(metis frameResChainPres frameNilStatEq Identity Associativity AssertionStatEqTrans AssertionStatEqSym Commutativity) + qed + next + case(cSim \ Q R) + then show ?case using \eqvt ?X\ + by(auto intro: parNilLeft parNilRight) + next + case(cExt \ Q R \') + then show ?case by auto + next + case(cSym \ Q R) + then show ?case by auto + qed +qed + +lemma bisimResNil: + fixes x :: name + and \ :: 'b + +shows "\ \ \\x\\ \ \" +proof - + { + fix x::name + assume "x \ \" + have "\ \ \\x\\ \ \" + proof - + let ?X1 = "{(\, \\x\\, \) | \ x. x \ \}" + let ?X2 = "{(\, \, \\x\\) | \ x. x \ \}" + let ?X = "?X1 \ ?X2" + + from \x \ \\ have "(\, \\x\\, \) \ ?X" by auto + then show ?thesis + proof(coinduct rule: bisimWeakCoinduct) + case(cStatEq \ P Q) + then show ?case using freshComp by(force intro: frameResFresh FrameStatEqSym) + next + case(cSim \ P Q) + then show ?case + by(force intro: resNilLeft resNilRight) + next + case(cExt \ P Q \') + obtain y where "y \ \" and "y \ \'" and "y \ x" + by(generate_fresh "name") (auto simp add: fresh_prod) + show ?case + proof(cases "(\, P, Q) \ ?X1") + assume "(\, P, Q) \ ?X1" + then obtain x where "P = \\x\\" and "Q = \" by auto + moreover have "\\x\\ = \\y\ \" by(subst alphaRes) auto + ultimately show ?case using \y \ \\ \y \ \'\ by auto + next + assume "(\, P, Q) \ ?X1" + with \(\, P, Q) \ ?X\ have "(\, P, Q) \ ?X2" by auto + then obtain x where "Q = \\x\\" and "P = \" by auto + moreover have "\\x\\ = \\y\ \" by(subst alphaRes) auto + ultimately show ?case using \y \ \\ \y \ \'\ by auto + qed + next + case(cSym \ P Q) + then show ?case by auto + qed + qed + } + moreover obtain y::name where "y \ \" by(generate_fresh "name") auto + ultimately have "\ \ \\y\\ \ \" by auto + then show ?thesis by(subst alphaRes[where x=x and y=y]) auto +qed + +lemma bisimOutputPushRes: + fixes x :: name + and \ :: 'b + and M :: 'a + and N :: 'a + and P :: "('a, 'b, 'c) psi" + +assumes "x \ M" + and "x \ N" + +shows "\ \ \\x\(M\N\.P) \ M\N\.\\x\P" +proof - + { + fix x::name and P::"('a, 'b, 'c) psi" + assume "x \ \" and "x \ M" and "x \ N" + let ?X1 = "{(\, \\x\(M\N\.P), M\N\.\\x\P) | \ x M N P. x \ \ \ x \ M \ x \ N}" + let ?X2 = "{(\, M\N\.\\x\P, \\x\(M\N\.P)) | \ x M N P. x \ \ \ x \ M \ x \ N}" + let ?X = "?X1 \ ?X2" + + have "eqvt ?X" + by(rule eqvtUnion) (force simp add: eqvt_def pt_fresh_bij[OF pt_name_inst, OF at_name_inst] eqvts)+ + from \x \ \\ \x \ M\ \x \ N\ have "(\, \\x\(M\N\.P), M\N\.\\x\P) \ ?X" + by auto + then have "\ \ \\x\(M\N\.P) \ M\N\.\\x\P" + proof(coinduct rule: bisimCoinduct) + case(cStatEq \ Q R) + then show ?case using freshComp by(force intro: frameResFresh FrameStatEqSym) + next + case(cSim \ Q R) + then show ?case using \eqvt ?X\ + by(auto intro!: outputPushResLeft outputPushResRight bisimReflexive) + next + case(cExt \ Q R \') + show ?case + proof(cases "(\, Q, R) \ ?X1") + assume "(\, Q, R) \ ?X1" + then obtain x M N P where Qeq: "Q = \\x\(M\N\.P)" and Req: "R = M\N\.\\x\P" and "x \ \" and "x \ M" and "x \ N" by auto + obtain y::name where "y \ \" and "y \ \'" and "y \ M" and "y \ N" and "y \ P" + by(generate_fresh "name") (auto simp add: fresh_prod) + + moreover then have "(\ \ \', \\y\(M\N\.([(x, y)] \ P)), M\N\.\\y\([(x, y)] \ P)) \ ?X" by auto + moreover from Qeq \x \ M\ \y \ M\ \x \ N\ \y \ N\ \y \ P\ have "Q = \\y\(M\N\.([(x, y)] \ P))" + apply clarsimp by(subst alphaRes[of y]) (auto simp add: eqvts) + moreover from Req \y \ P\ have "R = M\N\.\\y\([(x, y)] \ P)" + apply clarsimp by(subst alphaRes[of y]) (auto simp add: eqvts) + ultimately show ?case by blast + next + assume "(\, Q, R) \ ?X1" + with \(\, Q, R) \ ?X\ have "(\, Q, R) \ ?X2" by blast + then obtain x M N P where Req: "R = \\x\(M\N\.P)" and Qeq: "Q = M\N\.\\x\P" and "x \ \" and "x \ M" and "x \ N" by auto + obtain y::name where "y \ \" and "y \ \'" and "y \ M" and "y \ N" and "y \ P" + by(generate_fresh "name") (auto simp add: fresh_prod) + + moreover then have "(\ \ \', \\y\(M\N\.([(x, y)] \ P)), M\N\.\\y\([(x, y)] \ P)) \ ?X" by auto + moreover from Req \x \ M\ \y \ M\ \x \ N\ \y \ N\ \y \ P\ have "R = \\y\(M\N\.([(x, y)] \ P))" + apply clarsimp by(subst alphaRes[of y]) (auto simp add: eqvts) + moreover from Qeq \y \ P\ have "Q = M\N\.\\y\([(x, y)] \ P)" + apply clarsimp by(subst alphaRes[of y]) (auto simp add: eqvts) + ultimately show ?case by blast + qed + next + case(cSym \ R Q) + then show ?case by blast + qed + } + moreover obtain y::name where "y \ \" and "y \ M" and "y \ N" "y \ P" + by(generate_fresh "name") auto + ultimately have "\ \ \\y\(M\N\.([(x, y)] \ P)) \ M\N\.\\y\([(x, y)] \ P)" by auto + then show ?thesis using assms \y \ P\ \y \ M\ \y \ N\ + apply(subst alphaRes[where x=x and y=y and P=P], auto) + by(subst alphaRes[where x=x and y=y and P="M\N\.P"]) auto +qed + +lemma bisimInputPushRes: + fixes x :: name + and \ :: 'b + and M :: 'a + and xvec :: "name list" + and N :: 'a + and P :: "('a, 'b, 'c) psi" + +assumes "x \ M" + and "x \ xvec" + and "x \ N" + +shows "\ \ \\x\(M\\*xvec N\.P) \ M\\*xvec N\.\\x\P" +proof - + { + fix x::name and P::"('a, 'b, 'c) psi" + assume "x \ \" and "x \ M" and "x \ N" and "x \ xvec" + let ?X1 = "{(\, \\x\(M\\*xvec N\.P), M\\*xvec N\.\\x\P) | \ x M xvec N P. x \ \ \ x \ M \ x \ xvec \ x \ N}" + let ?X2 = "{(\, M\\*xvec N\.\\x\P, \\x\(M\\*xvec N\.P)) | \ x M xvec N P. x \ \ \ x \ M \ x \ xvec \ x \ N}" + let ?X = "?X1 \ ?X2" + + have "eqvt ?X" + by(rule eqvtUnion) (force simp add: eqvt_def pt_fresh_bij[OF pt_name_inst, OF at_name_inst] eqvts)+ + + from \x \ \\ \x \ M\ \x \ xvec\ \x \ N\ have "(\, \\x\(M\\*xvec N\.P), M\\*xvec N\.\\x\P) \ ?X" + by blast + then have "\ \ \\x\(M\\*xvec N\.P) \ M\\*xvec N\.\\x\P" + proof(coinduct rule: bisimCoinduct) + case(cStatEq \ Q R) + then show ?case using freshComp by(force intro: frameResFresh FrameStatEqSym) + next + case(cSim \ Q R) + then show ?case using \eqvt ?X\ + by(auto intro!: inputPushResLeft inputPushResRight bisimReflexive) + next + case(cExt \ Q R \') + show ?case + proof(cases "(\, Q, R) \ ?X1") + assume "(\, Q, R) \ ?X1" + then obtain x M xvec N P where Qeq: "Q = \\x\(M\\*xvec N\.P)" and Req: "R = M\\*xvec N\.\\x\P" and "x \ \" + and "x \ M" and "x \ xvec" and "x \ N" by auto + obtain y::name where "y \ \" and "y \ \'" and "y \ M" and "y \ N" and "y \ P" and "y \ xvec" + by(generate_fresh "name") (auto simp add: fresh_prod) + + moreover then have "(\ \ \', \\y\(M\\*xvec N\.([(x, y)] \ P)), M\\*xvec N\.\\y\([(x, y)] \ P)) \ ?X" by force + moreover from Qeq \x \ M\ \y \ M\ \x \ xvec\ \y \ xvec\ \x \ N\ \y \ N\ \y \ P\ have "Q = \\y\(M\\*xvec N\.([(x, y)] \ P))" + apply clarsimp by(subst alphaRes[of y]) (auto simp add: eqvts inputChainFresh) + moreover from Req \y \ P\ have "R = M\\*xvec N \.\\y\([(x, y)] \ P)" + apply clarsimp by(subst alphaRes[of y]) (auto simp add: eqvts) + ultimately show ?case by blast + next + assume "(\, Q, R) \ ?X1" + with \(\, Q, R) \ ?X\ have "(\, Q, R) \ ?X2" by blast + then obtain x M xvec N P where Req: "R = \\x\(M\\*xvec N\.P)" and Qeq: "Q = M\\*xvec N\.\\x\P" and "x \ \" + and "x \ M" and "x \ xvec" and "x \ N" by auto + obtain y::name where "y \ \" and "y \ \'" and "y \ M" and "y \ N" and "y \ P" and "y \ xvec" + by(generate_fresh "name") (auto simp add: fresh_prod) + + moreover then have "(\ \ \', \\y\(M\\*xvec N\.([(x, y)] \ P)), M\\*xvec N\.\\y\([(x, y)] \ P)) \ ?X" by force + moreover from Req \x \ M\ \y \ M\ \x \ xvec\ \y \ xvec\ \x \ N\ \y \ N\ \y \ P\ have "R = \\y\(M\\*xvec N\.([(x, y)] \ P))" + apply clarsimp by(subst alphaRes[of y]) (auto simp add: eqvts inputChainFresh) + moreover from Qeq \y \ P\ have "Q = M\\*xvec N \.\\y\([(x, y)] \ P)" + apply clarsimp by(subst alphaRes[of y]) (auto simp add: eqvts) + ultimately show ?case by blast + qed + next + case(cSym \ R Q) + then show ?case by blast + qed + } + moreover obtain y::name where "y \ \" and "y \ M" and "y \ N" and "y \ P" and "y \ xvec" + by(generate_fresh "name") auto + ultimately have "\ \ \\y\(M\\*xvec N\.([(x, y)] \ P)) \ M\\*xvec N\.\\y\([(x, y)] \ P)" by auto + then show ?thesis using assms \y \ P\ \y \ M\ \y \ N\ \y \ xvec\ + apply(subst alphaRes[where x=x and y=y and P=P], auto) + by(subst alphaRes[where x=x and y=y and P="M\\*xvec N\.P"]) (auto simp add: inputChainFresh eqvts) +qed + +lemma bisimCasePushRes: + fixes x :: name + and \ :: 'b + and Cs :: "('c \ ('a, 'b, 'c) psi) list" + +assumes "x \ (map fst Cs)" + +shows "\ \ \\x\(Cases Cs) \ Cases(map (\(\, P). (\, \\x\P)) Cs)" +proof - + { + fix x::name and Cs::"('c \ ('a, 'b, 'c) psi) list" + assume "x \ \" and "x \ (map fst Cs)" + let ?X1 = "{(\, \\x\(Cases Cs), Cases(map (\(\, P). (\, \\x\P)) Cs)) | \ x Cs. x \ \ \ x \ (map fst Cs)}" + let ?X2 = "{(\, Cases(map (\(\, P). (\, \\x\P)) Cs), \\x\(Cases Cs)) | \ x Cs. x \ \ \ x \ (map fst Cs)}" + let ?X = "?X1 \ ?X2" + + have "eqvt ?X" + proof + show "eqvt ?X1" + apply(clarsimp simp add: eqvt_def eqvts) + apply (rule exI) + apply (rule exI) + apply (rule conjI) + apply (rule refl) + apply(perm_extend_simp) + apply(clarsimp simp add: eqvts) + apply (simp add: fresh_bij) + apply(drule pt_fresh_bij1[OF pt_name_inst, OF at_name_inst]) + apply(drule pt_fresh_bij1[OF pt_name_inst, OF at_name_inst]) + apply(simp add: eqvts) + apply(perm_extend_simp) + apply(simp add: eqvts) + done + next + show "eqvt ?X2" + apply(clarsimp simp add: eqvt_def eqvts) + apply (rule exI) + apply (rule exI) + apply (subst conj_commute) + apply (rule conjI) + apply (rule conjI) + apply (rule refl) + apply(perm_extend_simp) + apply (simp add: fresh_bij) + apply(drule pt_fresh_bij1[OF pt_name_inst, OF at_name_inst]) + apply(drule pt_fresh_bij1[OF pt_name_inst, OF at_name_inst]) + apply(simp add: eqvts) + apply(perm_extend_simp) + apply(simp add: eqvts) + apply(perm_extend_simp) + apply(clarsimp simp add: eqvts) + done + qed + + from \x \ \\ \x \ map fst Cs\ have "(\, \\x\(Cases Cs), Cases(map (\(\, P). (\, \\x\P)) Cs)) \ ?X" by auto + then have "\ \ \\x\(Cases Cs) \ Cases(map (\(\, P). (\, \\x\P)) Cs)" + proof(coinduct rule: bisimCoinduct) + case(cStatEq \ Q R) + then show ?case using freshComp by(force intro: frameResFresh FrameStatEqSym) + next + case(cSim \ Q R) + then show ?case using \eqvt ?X\ + by(auto intro!: casePushResLeft casePushResRight bisimReflexive) + next + case(cExt \ Q R \') + show ?case + proof(cases "(\, Q, R) \ ?X1") + assume "(\, Q, R) \ ?X1" + then obtain x Cs where Qeq: "Q = \\x\(Cases Cs)" and Req: "R = Cases(map (\(\, P). (\, \\x\P)) Cs)" + and "x \ \" and "x \ (map fst Cs)" by blast + obtain y::name where "y \ \" and "y \ \'" and "y \ Cs" + by(generate_fresh "name") (auto simp add: fresh_prod) + from \y \ Cs\ \x \ (map fst Cs)\ have "y \ map fst ([(x, y)] \ Cs)" by(induct Cs) (auto simp add: fresh_list_cons fresh_list_nil) + + moreover with \y \ \\ \y \ \'\ have "(\ \ \', \\y\(Cases ([(x, y)] \ Cs)), Cases(map (\(\, P). (\, \\y\P)) ([(x, y)] \ Cs))) \ ?X" + by auto + moreover from Qeq \y \ Cs\ have "Q = \\y\(Cases([(x, y)] \ Cs))" + apply clarsimp by(subst alphaRes[of y]) (auto simp add: eqvts) + moreover from Req \y \ Cs\ \x \ (map fst Cs)\ have "R = Cases(map (\(\, P). (\, \\y\P)) ([(x, y)] \ Cs))" + by(induct Cs arbitrary: R) (auto simp add: fresh_list_cons fresh_prod alphaRes) + ultimately show ?case by blast + next + assume "(\, Q, R) \ ?X1" + with \(\, Q, R) \ ?X\ have "(\, Q, R) \ ?X2" by blast + then obtain x Cs where Req: "R = \\x\(Cases Cs)" and Qeq: "Q = Cases(map (\(\, P). (\, \\x\P)) Cs)" + and "x \ \" and "x \ (map fst Cs)" by blast + obtain y::name where "y \ \" and "y \ \'" and "y \ Cs" + by(generate_fresh "name") (auto simp add: fresh_prod) + from \y \ Cs\ \x \ (map fst Cs)\ have "y \ map fst ([(x, y)] \ Cs)" by(induct Cs) (auto simp add: fresh_list_cons fresh_list_nil) + + moreover with \y \ \\ \y \ \'\ have "(\ \ \', \\y\(Cases ([(x, y)] \ Cs)), Cases(map (\(\, P). (\, \\y\P)) ([(x, y)] \ Cs))) \ ?X" + by auto + moreover from Req \y \ Cs\ have "R = \\y\(Cases([(x, y)] \ Cs))" + apply clarsimp by(subst alphaRes[of y]) (auto simp add: eqvts) + moreover from Qeq \y \ Cs\ \x \ (map fst Cs)\ have "Q = Cases(map (\(\, P). (\, \\y\P)) ([(x, y)] \ Cs))" + by(induct Cs arbitrary: Q) (auto simp add: fresh_list_cons fresh_prod alphaRes) + ultimately show ?case by blast + qed + next + case(cSym \ R Q) + then show ?case by blast + qed + } + moreover obtain y::name where "y \ \" and "y \ Cs" by(generate_fresh "name") auto + moreover from \x \ map fst Cs\ have "y \ map fst([(x, y)] \ Cs)" + by(induct Cs) (auto simp add: fresh_left calc_atm) + ultimately have "\ \ \\y\(Cases ([(x, y)] \ Cs)) \ Cases(map (\(\, P). (\, \\y\P)) ([(x, y)] \ Cs))" + by auto + moreover from \y \ Cs\ have "\\y\(Cases ([(x, y)] \ Cs)) = \\x\(Cases Cs)" + by(simp add: alphaRes eqvts) + moreover from \x \ map fst Cs\ \y \ Cs\ have "Cases(map (\(\, P). (\, \\y\P)) ([(x, y)] \ Cs)) = Cases(map (\(\, P). (\, \\x\P)) Cs)" + by(induct Cs) (auto simp add: alphaRes) + ultimately show ?thesis by auto +qed + +lemma bangExt: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + +assumes "guarded P" + +shows "\ \ !P \ P \ !P" +proof - + let ?X = "{(\, !P, P \ !P) | \ P. guarded P} \ {(\, P \ !P, !P) | \ P. guarded P}" + from \guarded P\ have "(\, !P, P \ !P) \ ?X" by auto + then show ?thesis + proof(coinduct rule: bisimCoinduct) + case(cStatEq \ Q R) + from \(\, Q, R) \ ?X\ obtain P where Eq: "(Q = !P \ R = P \ !P) \ (Q = P \ !P \ R = !P)" and "guarded P" + by auto + obtain A\<^sub>P \\<^sub>P where FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" and "A\<^sub>P \* \" by(rule freshFrame) + from FrP \guarded P\ have "\\<^sub>P \ SBottom'" by(blast dest: guardedStatEq) + from \\\<^sub>P \ SBottom'\ have "\ \ SBottom' \ \ \ \\<^sub>P \ SBottom'" by(metis Identity Composition AssertionStatEqTrans Commutativity AssertionStatEqSym) + then have "\A\<^sub>P, \ \ SBottom'\ \\<^sub>F \A\<^sub>P, \ \ \\<^sub>P \ SBottom'\" + by(force intro: frameResChainPres) + moreover from \A\<^sub>P \* \\ have "\\, \ \ SBottom'\ \\<^sub>F \A\<^sub>P, \ \ SBottom'\" + apply - + by(rule FrameStatEqSym) (simp add: frameResFreshChain freshCompChain(1)) + ultimately show ?case using Eq \A\<^sub>P \* \\ FrP + by auto (blast dest: FrameStatEqTrans FrameStatEqSym)+ + next + case(cSim \ Q R) + then show ?case by(auto intro: bangExtLeft bangExtRight bisimReflexive) + next + case(cExt \ Q R) + then show ?case by auto + next + case(cSym \ Q R) + then show ?case by auto + qed +qed + +lemma bisimScopeExtSym: + fixes x :: name + and Q :: "('a, 'b, 'c) psi" + and P :: "('a, 'b, 'c) psi" + +assumes "x \ \" + and "x \ Q" + +shows "\ \ \\x\(P \ Q) \ (\\x\P) \ Q" + using assms + by(metis bisimScopeExt bisimTransitive bisimParComm bisimSymmetric bisimResPres) + +lemma bisimScopeExtChainSym: + fixes xvec :: "name list" + and Q :: "('a, 'b, 'c) psi" + and P :: "('a, 'b, 'c) psi" + +assumes "xvec \* \" + and "xvec \* Q" + +shows "\ \ \\*xvec\(P \ Q) \ (\\*xvec\P) \ Q" + using assms + by(induct xvec) (auto intro: bisimScopeExtSym bisimReflexive bisimTransitive bisimResPres) + +lemma bisimParPresAuxSym: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and Q :: "('a, 'b, 'c) psi" + and R :: "('a, 'b, 'c) psi" + +assumes "\ \ \\<^sub>R \ P \ Q" + and "extractFrame R = \A\<^sub>R, \\<^sub>R\" + and "A\<^sub>R \* \" + and "A\<^sub>R \* P" + and "A\<^sub>R \* Q" + +shows "\ \ R \ P \ R \ Q" + using assms + by(metis bisimParComm bisimParPresAux bisimTransitive) + +lemma bangDerivative: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and \ :: "'a action" + and P' :: "('a, 'b, 'c) psi" + +assumes "\ \ !P \\ \ P'" + and "\ \ P \ Q" + and "bn \ \* \" + and "bn \ \* P" + and "bn \ \* Q" + and "bn \ \* subject \" + and "guarded Q" + +obtains Q' R T where "\ \ !Q \\ \ Q'" and "\ \ P' \ R \ !P" and "\ \ Q' \ T \ !Q" and "\ \ R \ T" + and "((supp R)::name set) \ supp P'" and "((supp T)::name set) \ supp Q'" +proof - + from \\ \ !P \\ \ P'\ have "guarded P" + apply - + by(ind_cases "\ \ !P \\ \ P'") (auto simp add: psi.inject) + assume "\Q' R T. \\ \ !Q \\ \ Q'; \ \ P' \ R \ !P; \ \ Q' \ T \ !Q; \ \ R \ T; ((supp R)::name set) \ supp P'; + ((supp T)::name set) \ supp Q'\ \ thesis" + moreover from \\ \ !P \\ \ P'\ \bn \ \* subject \\ \bn \ \* \\ \bn \ \* P\ \bn \ \* Q\ \\ \ P \ Q\ \guarded Q\ + have "\Q' T R . \ \ !Q \\ \ Q' \ \ \ P' \ R \ !P \ \ \ Q' \ T \ !Q \ \ \ R \ T \ + ((supp R)::name set) \ supp P' \ ((supp T)::name set) \ supp Q'" + proof(nominal_induct avoiding: Q rule: bangInduct') + case(cAlpha \ P' p Q) + then obtain Q' T R where QTrans: "\ \ !Q \\ \ Q'" and "\ \ P' \ R \ (P \ !P)" and "\ \ Q' \ T \ !Q" and "\ \ R \ T" + and suppR: "((supp R)::name set) \ supp P'" and suppT: "((supp T)::name set) \ supp Q'" + by blast + from QTrans have "distinct(bn \)" + by(rule boundOutputDistinct) + have S: "set p \ set(bn \) \ set(bn(p \ \))" + by fact + from QTrans \bn(p \ \) \* Q\ \bn(p \ \) \* \\ \bn \ \* subject \\ \distinct(bn \)\ have "bn(p \ \) \* Q'" + by(drule_tac freeFreshChainDerivative) simp+ + with QTrans \bn(p \ \) \* \\ S \bn \ \* subject \\ have "\ \ !Q \(p \ \) \ (p \ Q')" + by(force simp add: residualAlpha) + moreover from \\ \ P' \ R \ (P \ !P)\ have "(p \ \) \ (p \ P') \ (p \ (R \ (P \ !P)))" + by(rule bisimClosed) + with \bn \ \* \\ \bn \ \* P\ \bn(p \ \) \* \\ \bn(p \ \) \* P\ S have "\ \ (p \ P') \ (p \ R) \ (P \ !P)" + by(simp add: eqvts) + moreover from \\ \ Q' \ T \ !Q\ have "(p \ \) \ (p \ Q') \ (p \ (T \ !Q))" + by(rule bisimClosed) + with \bn \ \* \\ \bn \ \* Q\ \bn(p \ \) \* \\ \bn(p \ \) \* Q\ S have "\ \ (p \ Q') \ (p \ T) \ !Q" + by(simp add: eqvts) + moreover from \\ \ R \ T\ have "(p \ \) \ (p \ R) \ (p \ T)" + by(rule bisimClosed) + with \bn \ \* \\ \bn(p \ \) \* \\ S have "\ \ (p \ R) \ (p \ T)" + by(simp add: eqvts) + moreover from suppR have "((supp(p \ R))::name set) \ supp(p \ P')" + apply(erule_tac rev_mp) + by(subst subsetClosed[of p, symmetric]) (simp add: eqvts) + moreover from suppT have "((supp(p \ T))::name set) \ supp(p \ Q')" + apply(erule_tac rev_mp) + by(subst subsetClosed[of p, symmetric]) (simp add: eqvts) + ultimately show ?case by blast + next + case(cPar1 \ P' Q) + from \\ \ P \ Q\ \\ \ P \\ \ P'\ \bn \ \* \\ \bn \ \* Q\ + obtain Q' where QTrans: "\ \ Q \\ \ Q'" and "\ \ P' \ Q'" + by(blast dest: bisimE simE) + from QTrans have "\ \ SBottom' \ Q \\ \ Q'" + by(metis statEqTransition Identity AssertionStatEqSym) + then have "\ \ Q \ !Q \\ \ (Q' \ !Q)" + using \bn \ \* Q\ by(intro Par1) (assumption | simp)+ + then have "\ \ !Q \\ \ (Q' \ !Q)" + using \guarded Q\ by(rule Bang) + moreover from \guarded P\ have "\ \ P' \ !P \ P' \ (P \ !P)" + by(metis bangExt bisimParPresSym) + moreover have "\ \ Q' \ !Q \ Q' \ !Q" + by(rule bisimReflexive) + ultimately show ?case using \\ \ P' \ Q'\ + by(force simp add: psi.supp) + next + case(cPar2 \ P' Q) + then obtain Q' T R where QTrans: "\ \ !Q \\ \ Q'" and "\ \ P' \ R \ !P" and "\ \ Q' \ T \ !Q" + and "\ \ R \ T" and suppR: "((supp R)::name set) \ supp P'" and suppT: "((supp T)::name set) \ supp Q'" + by blast + note QTrans + from \\ \ P' \ R \ !P\ have "\ \ P \ P' \ R \ (P \ !P)" + by(metis bisimParPresSym bisimParComm bisimTransitive bisimParAssoc) + with QTrans show ?case using \\ \ Q' \ T \ !Q\ \\ \ R \ T\ suppR suppT + by(force simp add: psi.supp) + next + case(cComm1 M N P' K xvec P'' Q) + from \\ \ P \ Q\ have "\ \ Q \[bisim] P" + by(metis bisimE) + with \\ \ P \M\N\ \ P'\ obtain Q' where QTrans: "\ \ Q \M\N\ \ Q'" and "\ \ Q' \ P'" + by(force dest: simE) + from QTrans have "\ \ SBottom' \ Q \M\N\ \ Q'" + by(metis statEqTransition Identity AssertionStatEqSym) + moreover obtain A\<^sub>Q \\<^sub>Q where FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" and "A\<^sub>Q \* \" and "A\<^sub>Q \* Q" and "A\<^sub>Q \* M" + by(rule freshFrame[where C="(\, Q, M)"]) auto + note FrQ + moreover from FrQ \guarded Q\ have "\\<^sub>Q \ SBottom'" + by(blast dest: guardedStatEq) + from \\ \ !P \K\\*xvec\\N\ \ P''\ \xvec \* K\ \\ \ P \ Q\ \xvec \* \\ \xvec \* P\ \xvec \* Q\ \guarded Q\ + obtain Q'' T R where QTrans': "\ \ !Q \K\\*xvec\\N\ \ Q''" and "\ \ P'' \ R \ !P" and "\ \ Q'' \ T \ !Q" and "\ \ R \ T" + and suppR: "((supp R)::name set) \ supp P''" and suppT: "((supp T)::name set) \ supp Q''" + using cComm1 by force + from QTrans' \\\<^sub>Q \ SBottom'\ have "\ \ \\<^sub>Q \ !Q \K\\*xvec\\N\ \ Q''" + by(metis statEqTransition Identity compositionSym AssertionStatEqSym) + moreover from \\ \ M \ K\ \\\<^sub>Q \ SBottom'\ have "\ \ \\<^sub>Q \ SBottom' \ M \ K" + by(metis statEqEnt Identity compositionSym AssertionStatEqSym) + ultimately have "\ \ Q \ !Q \\ \ (\\*xvec\(Q' \ Q''))" using \A\<^sub>Q \* \\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* M\ \xvec \* Q\ + by(intro Comm1) (assumption | simp)+ + then have "\ \ !Q \\ \ (\\*xvec\(Q' \ Q''))" + using \guarded Q\ by(rule Bang) + moreover from \\ \ P'' \ R \ !P\ \guarded P\ \xvec \* \\ have "\ \ \\*xvec\(P' \ P'') \ \\*xvec\((P' \ R) \ (P \ !P))" + by(metis bisimParPresSym bangExt bisimTransitive bisimParAssoc bisimSymmetric bisimResChainPres) + with \xvec \* \\ \xvec \* P\ have "\ \ \\*xvec\(P' \ P'') \ (\\*xvec\(P' \ R)) \ (P \ !P)" + by(metis bisimScopeExtChainSym bisimTransitive psiFreshVec) + moreover from \\ \ Q'' \ T \ !Q\ \xvec \* \\ \xvec \* Q\ have "\ \ \\*xvec\(Q' \ Q'') \ (\\*xvec\(Q' \ T)) \ !Q" + by(metis bisimParPresSym bisimTransitive bisimParAssoc bisimSymmetric bisimResChainPres bisimScopeExtChainSym psiFreshVec) + moreover from \\ \ R \ T\ \\ \ Q' \ P'\ \xvec \* \\ have "\ \ \\*xvec\(P' \ R) \ \\*xvec\(Q' \ T)" + by(metis bisimParPresSym bisimTransitive bisimResChainPres bisimParComm bisimE(4)) + moreover from suppR have "((supp(\\*xvec\(P' \ R)))::name set) \ supp((\\*xvec\(P' \ P'')))" + by(auto simp add: psi.supp resChainSupp) + moreover from suppT have "((supp(\\*xvec\(Q' \ T)))::name set) \ supp((\\*xvec\(Q' \ Q'')))" + by(auto simp add: psi.supp resChainSupp) + ultimately show ?case by blast + next + case(cComm2 M xvec N P' K P'' Q) + from \\ \ P \ Q\ \\ \ P \M\\*xvec\\N\ \ P'\ \xvec \* \\ \xvec \* Q\ + obtain Q' where QTrans: "\ \ Q \M\\*xvec\\N\ \ Q'" and "\ \ P' \ Q'" + by(metis bisimE simE bn.simps) + from QTrans have "\ \ SBottom' \ Q \M\\*xvec\\N\ \ Q'" + by(metis statEqTransition Identity AssertionStatEqSym) + moreover obtain A\<^sub>Q \\<^sub>Q where FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" and "A\<^sub>Q \* \" and "A\<^sub>Q \* Q" and "A\<^sub>Q \* M" + by(rule freshFrame[where C="(\, Q, M)"]) auto + note FrQ + moreover from FrQ \guarded Q\ have "\\<^sub>Q \ SBottom'" + by(blast dest: guardedStatEq) + from \\ \ !P \K\N\ \ P''\ \\ \ P \ Q\ \guarded Q\ + obtain Q'' T R where QTrans': "\ \ !Q \K\N\ \ Q''" and "\ \ P'' \ R \ !P" and "\ \ Q'' \ T \ !Q" and "\ \ R \ T" + and suppR: "((supp R)::name set) \ supp P''" and suppT: "((supp T)::name set) \ supp Q''" using cComm2 + by force + from QTrans' \\\<^sub>Q \ SBottom'\ have "\ \ \\<^sub>Q \ !Q \K\N\ \ Q''" + by(metis statEqTransition Identity compositionSym AssertionStatEqSym) + moreover from \\ \ M \ K\ \\\<^sub>Q \ SBottom'\ have "\ \ \\<^sub>Q \ SBottom' \ M \ K" + by(metis statEqEnt Identity compositionSym AssertionStatEqSym) + ultimately have "\ \ Q \ !Q \\ \ (\\*xvec\(Q' \ Q''))" using \A\<^sub>Q \* \\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* M\ \xvec \* Q\ + by(intro Comm2) (assumption | simp)+ + then have "\ \ !Q \\ \ (\\*xvec\(Q' \ Q''))" using \guarded Q\ by(rule Bang) + moreover from \\ \ P'' \ R \ !P\ \guarded P\ \xvec \* \\ have "\ \ \\*xvec\(P' \ P'') \ \\*xvec\((P' \ R) \ (P \ !P))" + by(metis bisimParPresSym bangExt bisimTransitive bisimParAssoc bisimSymmetric bisimResChainPres) + with \xvec \* \\ \xvec \* P\ have "\ \ \\*xvec\(P' \ P'') \ (\\*xvec\(P' \ R)) \ (P \ !P)" + by(metis bisimScopeExtChainSym bisimTransitive psiFreshVec) + moreover from \\ \ Q'' \ T \ !Q\ \xvec \* \\ \xvec \* Q\ have "\ \ \\*xvec\(Q' \ Q'') \ (\\*xvec\(Q' \ T)) \ !Q" + by(metis bisimParPresSym bisimTransitive bisimParAssoc bisimSymmetric bisimResChainPres bisimScopeExtChainSym psiFreshVec) + moreover from \\ \ R \ T\ \\ \ P' \ Q'\ \xvec \* \\ have "\ \ \\*xvec\(P' \ R) \ \\*xvec\(Q' \ T)" + by(metis bisimParPresSym bisimTransitive bisimResChainPres bisimParComm) + moreover from suppR have "((supp(\\*xvec\(P' \ R)))::name set) \ supp((\\*xvec\(P' \ P'')))" + by(auto simp add: psi.supp resChainSupp) + moreover from suppT have "((supp(\\*xvec\(Q' \ T)))::name set) \ supp((\\*xvec\(Q' \ Q'')))" + by(auto simp add: psi.supp resChainSupp) + ultimately show ?case by blast + next + case(cBang \ P' Q) + then obtain Q' T R where QTrans: "\ \ !Q \\ \ Q'" and "\ \ P' \ R \ (P \ !P)" and "\ \ Q' \ T \ !Q" and "\ \ R \ T" + and suppR: "((supp R)::name set) \ supp P'" and suppT: "((supp T)::name set) \ supp Q'" + by blast + from \\ \ P' \ R \ (P \ !P)\ \guarded P\ have "\ \ P' \ R \ !P" + by(metis bangExt bisimParPresSym bisimTransitive bisimSymmetric) + with QTrans show ?case using \\ \ Q' \ T \ !Q\ \\ \ R \ T\ suppR suppT + by blast + next + case(cBrMerge M N P' P'' Q) + then obtain Q' T R where "\ \ !Q \ \M\N\ \ Q'" and + p''eq: "\ \ P'' \ R \ !P" and + "\ \ Q' \ T \ !Q" and "\ \ R \ T" and + "supp R \ (supp P''::name set)" and + "supp T \ (supp Q'::name set)" + by blast + obtain Q'' where "\ \ Q \ \M\N\ \ Q''" and "\ \ Q'' \ P'" + proof - + assume g: "(\Q''. \\ \ Q \ \M\N\ \ Q''; \ \ Q'' \ P'\ \ thesis)" + have "\Q'. \ \ Q \ \M\N\ \ Q' \ (\, Q', P') \ bisim" + apply(rule simE) + apply(rule bisimE) + apply(rule bisimSymmetric) + by fact+ + then show thesis + using g by blast + qed + have "\ \ \ \ Q \ \M\N\ \ Q''" using \\ \ Q \ \M\N\ \ Q''\ + by(rule statEqTransition) (rule AssertionStatEqSym[OF Identity]) + obtain A\<^sub>Q \\<^sub>Q where "extractFrame Q = \A\<^sub>Q,\\<^sub>Q\" and "distinct A\<^sub>Q" and "A\<^sub>Q \* \" and "A\<^sub>Q \* Q" and "A\<^sub>Q \* M" + by(rule freshFrame[where C="(\,Q,M)"]) force + then have "\\<^sub>Q \ \" using \guarded Q\ + by(auto dest: guardedStatEq) + have "\ \ \\<^sub>Q \ !Q \ \M\N\ \ Q'" using \\ \ !Q \ \M\N\ \ Q'\ + by(rule statEqTransition) (metis \\\<^sub>Q \ \\ AssertionStatEqTrans AssertionStatEqSym Identity compositionSym) + have "\ \ !Q \ \M\N\ \ Q'' \ Q'" + by(rule Bang) (rule BrMerge|fact|simp)+ + moreover have "\ \ P' \ P'' \ (P' \ R) \ (P \ !P)" + apply(rule bisimTransitive[OF bisimParPresSym, OF p''eq]) + apply(rule bisimTransitive[OF bisimSymmetric, OF bisimParAssoc]) + apply(rule bisimParPresSym) + apply(rule bangExt[OF \guarded P\]) + done + moreover have "\ \ Q'' \ Q' \ Q'' \ T \ !Q" using \\ \ Q' \ T \ !Q\ + by(metis bisimTransitive bisimParAssoc bisimParComm bisimParPres bisimSymmetric) + moreover have "\ \ (P' \ R) \ Q'' \ T" using \\ \ R \ T\ and \\ \ Q'' \ P'\ + apply - + apply(erule bisimTransitive[OF bisimParPres, OF bisimSymmetric]) + apply(erule bisimTransitive[OF bisimParPresSym]) + by(rule bisimReflexive) + moreover have "supp(P' \ R) \ (supp(P'\P'')::name set)" + using \supp R \ (supp P''::name set)\ by(auto simp add: psi.supp) + moreover have "supp(Q'' \ T) \ (supp(Q'' \ Q')::name set)" + using \supp T \ (supp Q'::name set)\ by(auto simp add: psi.supp) + ultimately show ?case + by blast + next + case(cBrComm1 M N P' xvec P'' Q) + from \\ \ P \ Q\ have "\ \ Q \[bisim] P" by(metis bisimE) + with \\ \ P \\M\N\ \ P'\ obtain Q' where QTrans: "\ \ Q \\M\N\ \ Q'" and "\ \ Q' \ P'" + by(force dest: simE) + from QTrans have "\ \ SBottom' \ Q \\M\N\ \ Q'" + by(metis statEqTransition Identity AssertionStatEqSym) + moreover obtain A\<^sub>Q \\<^sub>Q where FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" and "A\<^sub>Q \* \" and "A\<^sub>Q \* Q" and "A\<^sub>Q \* M" + by(rule freshFrame[where C="(\, Q, M)"]) auto + note FrQ + moreover from FrQ \guarded Q\ have "\\<^sub>Q \ SBottom'" + by(blast dest: guardedStatEq) + from \\ \ !P \\M\\*xvec\\N\ \ P''\ \\ \ P \ Q\ \xvec \* \\ \xvec \* P\ \xvec \* Q\ \guarded Q\ + obtain Q'' T R where QTrans': "\ \ !Q \\M\\*xvec\\N\ \ Q''" and "\ \ P'' \ R \ !P" and "\ \ Q'' \ T \ !Q" + and "\ \ R \ T" and suppR: "((supp R)::name set) \ supp P''" and suppT: "((supp T)::name set) \ supp Q''" + using cBrComm1 by force + from QTrans' \\\<^sub>Q \ SBottom'\ have "\ \ \\<^sub>Q \ !Q \\M\\*xvec\\N\ \ Q''" + by(metis statEqTransition Identity compositionSym AssertionStatEqSym) + ultimately have "\ \ Q \ !Q \\M\\*xvec\\N\ \ Q' \ Q''" + using \A\<^sub>Q \* \\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* M\ \xvec \* Q\ by(intro BrComm1) (assumption | simp)+ + then have "\ \ !Q \\M\\*xvec\\N\ \ Q' \ Q''" + using \guarded Q\ by(rule Bang) + moreover from \\ \ P'' \ R \ !P\ \guarded P\ \xvec \* \\ have "\ \ P' \ P'' \ (P' \ R) \ (P \ !P)" + by(metis bisimParPresSym bangExt bisimTransitive bisimParAssoc bisimSymmetric) + moreover from \\ \ Q'' \ T \ !Q\ \xvec \* \\ \xvec \* Q\ have "\ \ Q' \ Q'' \ (Q' \ T) \ !Q" + by(metis bisimParPresSym bisimTransitive bisimParAssoc bisimSymmetric) + moreover from \\ \ R \ T\ \\ \ Q' \ P'\ \xvec \* \\ have "\ \ P' \ R \ Q' \ T" + by(metis bisimParPresSym bisimTransitive bisimParComm bisimE(4)) + moreover from suppR have "((supp(P' \ R))::name set) \ supp((P' \ P''))" + by(auto simp add: psi.supp resChainSupp) + moreover from suppT have "((supp(Q' \ T))::name set) \ supp((Q' \ Q''))" + by(auto simp add: psi.supp resChainSupp) + ultimately show ?case by blast + next + case(cBrComm2 M N P' xvec P'' Q) + from \\ \ P \ Q\ have "\ \ Q \[bisim] P" by(metis bisimE) + with \\ \ P \\M\\*xvec\\N\ \ P'\ obtain Q' where QTrans: "\ \ Q \\M\\*xvec\\N\ \ Q'" and "\ \ Q' \ P'" + using \xvec \* \\ \xvec \* P\ \xvec \* Q\ by(auto dest: simE) + from QTrans have "\ \ SBottom' \ Q \\M\\*xvec\\N\ \ Q'" + by(metis statEqTransition Identity AssertionStatEqSym) + moreover obtain A\<^sub>Q \\<^sub>Q where FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" and "A\<^sub>Q \* \" and "A\<^sub>Q \* Q" and "A\<^sub>Q \* M" + by(rule freshFrame[where C="(\, Q, M)"]) auto + note FrQ + moreover from FrQ \guarded Q\ have "\\<^sub>Q \ SBottom'" by(blast dest: guardedStatEq) + from \\ \ !P \\M\N\ \ P''\ \\ \ P \ Q\ \xvec \* \\ \xvec \* P\ \xvec \* Q\ \guarded Q\ + obtain Q'' T R where QTrans': "\ \ !Q \\M\N\ \ Q''" and "\ \ P'' \ R \ !P" and "\ \ Q'' \ T \ !Q" + and "\ \ R \ T" and suppR: "((supp R)::name set) \ supp P''" and suppT: "((supp T)::name set) \ supp Q''" + using cBrComm2 by force + from QTrans' \\\<^sub>Q \ SBottom'\ have "\ \ \\<^sub>Q \ !Q \\M\N\ \ Q''" + by(metis statEqTransition Identity compositionSym AssertionStatEqSym) + ultimately have "\ \ Q \ !Q \\M\\*xvec\\N\ \ Q' \ Q''" + using \A\<^sub>Q \* \\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* M\ \xvec \* Q\ by(intro BrComm2) (assumption | simp)+ + then have "\ \ !Q \\M\\*xvec\\N\ \ Q' \ Q''" + using \guarded Q\ by(rule Bang) + moreover from \\ \ P'' \ R \ !P\ \guarded P\ \xvec \* \\ have "\ \ P' \ P'' \ (P' \ R) \ (P \ !P)" + by(metis bisimParPresSym bangExt bisimTransitive bisimParAssoc bisimSymmetric) + moreover from \\ \ Q'' \ T \ !Q\ \xvec \* \\ \xvec \* Q\ have "\ \ Q' \ Q'' \ (Q' \ T) \ !Q" + by(metis bisimParPresSym bisimTransitive bisimParAssoc bisimSymmetric) + moreover from \\ \ R \ T\ \\ \ Q' \ P'\ \xvec \* \\ have "\ \ P' \ R \ Q' \ T" + by(metis bisimParPresSym bisimTransitive bisimParComm bisimE(4)) + moreover from suppR have "((supp(P' \ R))::name set) \ supp((P' \ P''))" + by(auto simp add: psi.supp resChainSupp) + moreover from suppT have "((supp(Q' \ T))::name set) \ supp((Q' \ Q''))" + by(auto simp add: psi.supp resChainSupp) + ultimately show ?case by blast + qed + ultimately show ?thesis by blast +qed + +lemma structCongBisim: + fixes P :: "('a, 'b, 'c) psi" + and Q :: "('a, 'b, 'c) psi" + +assumes "P \\<^sub>s Q" + +shows "P \ Q" + using assms + by(induct rule: structCong.induct) + (auto intro: bisimReflexive bisimSymmetric bisimTransitive bisimParComm bisimParAssoc bisimParNil bisimResNil bisimResComm bisimScopeExt bisimCasePushRes bisimInputPushRes bisimOutputPushRes bangExt) + +lemma bisimBangPres: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and Q :: "('a, 'b, 'c) psi" + +assumes "\ \ P \ Q" + and "guarded P" + and "guarded Q" + +shows "\ \ !P \ !Q" +proof - + let ?X = "{(\, R \ !P, R \ !Q) | \ P Q R. \ \ P \ Q \ guarded P \ guarded Q}" + let ?Y = "{(\, P, Q) | \ P P' Q' Q. \ \ P \ P' \ (\, P', Q') \ ?X \ \ \ Q' \ Q}" + from assms have "(\, \ \ !P, \ \ !Q) \ ?X" by(blast intro: bisimReflexive) + + moreover have "eqvt ?X" + apply(clarsimp simp add: eqvt_def) + apply(drule bisimClosed) + by force + ultimately have "\ \ \ \ !P \ \ \ !Q" + proof(coinduct rule: weakTransitiveCoinduct) + case(cStatEq \ P Q) + then show ?case by auto + next + case(cSim \ RP RQ) + from \(\, RP, RQ) \ ?X\ obtain P Q R where "\ \ P \ Q" and "guarded P" and "guarded Q" + and "RP = R \ !P" and "RQ = R \ !Q" + by auto + note \\ \ P \ Q\ + moreover from \eqvt ?X\ have "eqvt ?Y" by blast + moreover note \guarded P\ \guarded Q\ bisimE(2) bisimE(3) bisimE(4) statEqBisim bisimClosed bisimParAssoc[THEN bisimSymmetric] + bisimParPres bisimParPresAuxSym bisimResChainPres bisimScopeExtChainSym bisimTransitive + moreover have "\\ P Q R T. \\ \ P \ Q; (\, Q, R) \ ?Y; \ \ R \ T\ \ (\, P, T) \ ?Y" + by auto (metis bisimTransitive) + moreover have "\\ P Q R. \\ \ P \ Q; guarded P; guarded Q\ \ (\, R \ !P, R \ !Q) \ ?Y" by(blast intro: bisimReflexive) + moreover have "\\ P \ P' Q. \\ \ !P \\ \ P'; \ \ P \ Q; bn \ \* \; bn \ \* P; bn \ \* Q; guarded Q; bn \ \* subject \\ \ + \Q' R T. \ \ !Q \\ \ Q' \ \ \ P' \ R \ !P \ \ \ Q' \ T \ !Q \ + \ \ R \ T \ ((supp R)::name set) \ supp P' \ + ((supp T)::name set) \ supp Q'" + by(blast elim: bangDerivative) + ultimately have "\ \ R \ !P \[?Y] R \ !Q" + by(rule bangPres) + with \RP = R \ !P\ \RQ = R \ !Q\ show ?case + by blast + next + case(cExt \ RP RQ \') + then show ?case by(blast dest: bisimE) + next + case(cSym \ RP RQ) + then show ?case by(blast dest: bisimE) + qed + then show ?thesis + by(metis bisimTransitive bisimParNil bisimSymmetric bisimParComm) +qed + +end + +end diff --git a/thys/Broadcast_Psi/Bisim_Subst.thy b/thys/Broadcast_Psi/Bisim_Subst.thy new file mode 100644 --- /dev/null +++ b/thys/Broadcast_Psi/Bisim_Subst.thy @@ -0,0 +1,651 @@ +theory Bisim_Subst + imports Bisim_Struct_Cong "Psi_Calculi.Close_Subst" +begin + +text \This file is a (heavily modified) variant of the theory {\it Psi\_Calculi.Bisim\_Subst} +from~\cite{DBLP:journals/afp/Bengtson12}.\ + +context env begin + +abbreviation + bisimSubstJudge ("_ \ _ \\<^sub>s _" [70, 70, 70] 65) where "\ \ P \\<^sub>s Q \ (\, P, Q) \ closeSubst bisim" +abbreviation + bisimSubstNilJudge ("_ \\<^sub>s _" [70, 70] 65) where "P \\<^sub>s Q \ SBottom' \ P \\<^sub>s Q" + +lemmas bisimSubstClosed[eqvt] = closeSubstClosed[OF bisimEqvt] +lemmas bisimSubstEqvt[simp] = closeSubstEqvt[OF bisimEqvt] + +lemma bisimSubstOutputPres: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and Q :: "('a, 'b, 'c) psi" + and M :: 'a + and N :: 'a + +assumes "\ \ P \\<^sub>s Q" + +shows "\ \ M\N\.P \\<^sub>s M\N\.Q" + using assms closeSubstI closeSubstE bisimOutputPres by force + + +lemma seqSubstInputChain[simp]: + fixes xvec :: "name list" + and N :: "'a" + and P :: "('a, 'b, 'c) psi" + and \ :: "(name list \ 'a list) list" + +assumes "xvec \* \" + +shows "seqSubs' (inputChain xvec N P) \ = inputChain xvec (substTerm.seqSubst N \) (seqSubs P \)" + using assms + by(induct xvec) auto + +lemma bisimSubstInputPres: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and Q :: "('a, 'b, 'c) psi" + and M :: 'a + and xvec :: "name list" + and N :: 'a + +assumes "\ \ P \\<^sub>s Q" + and "xvec \* \" + and "distinct xvec" + +shows "\ \ M\\*xvec N\.P \\<^sub>s M\\*xvec N\.Q" +proof(rule closeSubstI) + fix \ + assume "wellFormedSubst(\::(name list \ 'a list) list)" + obtain p where "(p \ xvec) \* \" + and "(p \ xvec) \* P" and "(p \ xvec) \* Q" and "(p \ xvec) \* \" and "(p \ xvec) \* N" + and S: "set p \ set xvec \ set (p \ xvec)" + by(rule name_list_avoiding[where c="(\, P, Q, \, N)"]) auto + + from \\ \ P \\<^sub>s Q\ have "(p \ \) \ (p \ P) \\<^sub>s (p \ Q)" + by(rule bisimSubstClosed) + with \xvec \* \\ \(p \ xvec) \* \\ S have "\ \ (p \ P) \\<^sub>s (p \ Q)" + by simp + + { + fix Tvec :: "'a list" + from \\ \ (p \ P) \\<^sub>s (p \ Q)\ \wellFormedSubst \\ have "\ \ (p \ P)[<\>] \\<^sub>s (p \ Q)[<\>]" + by(rule closeSubstUnfold) + moreover assume "length xvec = length Tvec" and "distinct xvec" + ultimately have "\ \ ((p \ P)[<\>])[(p \ xvec)::=Tvec] \ ((p \ Q)[<\>])[(p \ xvec)::=Tvec]" + apply - + by(drule closeSubstE[where \="[((p \ xvec), Tvec)]"]) auto + } + + with \(p \ xvec) \* \\ \distinct xvec\ + have "\ \ (M\\*(p \ xvec) (p \ N)\.(p \ P))[<\>] \ (M\\*(p \ xvec) (p \ N)\.(p \ Q))[<\>]" + by(force intro: bisimInputPres) + moreover from \(p \ xvec) \* N\ \(p \ xvec) \* P\ S have "M\\*(p \ xvec) (p \ N)\.(p \ P) = M\\*xvec N\.P" + apply(simp add: psi.inject) by(rule inputChainAlpha[symmetric]) auto + moreover from \(p \ xvec) \* N\ \(p \ xvec) \* Q\ S have "M\\*(p \ xvec) (p \ N)\.(p \ Q) = M\\*xvec N\.Q" + apply(simp add: psi.inject) by(rule inputChainAlpha[symmetric]) auto + ultimately show "\ \ (M\\*xvec N\.P)[<\>] \ (M\\*xvec N\.Q)[<\>]" + by force +qed + +lemma bisimSubstCasePresAux: + fixes \ :: 'b + and CsP :: "('c \ ('a, 'b, 'c) psi) list" + and CsQ :: "('c \ ('a, 'b, 'c) psi) list" + +assumes C1: "\\ P. (\, P) \ set CsP \ \Q. (\, Q) \ set CsQ \ guarded Q \ \ \ P \\<^sub>s Q" + and C2: "\\ Q. (\, Q) \ set CsQ \ \P. (\, P) \ set CsP \ guarded P \ \ \ P \\<^sub>s Q" + +shows "\ \ Cases CsP \\<^sub>s Cases CsQ" +proof - + { + fix \ :: "(name list \ 'a list) list" + + assume "wellFormedSubst \" + + have "\ \ Cases(caseListSeqSubst CsP \) \ Cases(caseListSeqSubst CsQ \)" + proof(rule bisimCasePres) + fix \ P + assume "(\, P) \ set (caseListSeqSubst CsP \)" + then obtain \' P' where "(\', P') \ set CsP" and "\ = substCond.seqSubst \' \" and PeqP': "P = (P'[<\>])" + by(induct CsP) force+ + from \(\', P') \ set CsP\ obtain Q' where "(\', Q') \ set CsQ" and "guarded Q'" and "\ \ P' \\<^sub>s Q'" by(blast dest: C1) + from \(\', Q') \ set CsQ\ \\ = substCond.seqSubst \' \\ obtain Q where "(\, Q) \ set (caseListSeqSubst CsQ \)" and "Q = Q'[<\>]" + by(induct CsQ) auto + with PeqP' \guarded Q'\ \\ \ P' \\<^sub>s Q'\ \wellFormedSubst \\ show "\Q. (\, Q) \ set (caseListSeqSubst CsQ \) \ guarded Q \ \ \ P \ Q" + by(blast dest: closeSubstE guardedSeqSubst) + next + fix \ Q + assume "(\, Q) \ set (caseListSeqSubst CsQ \)" + then obtain \' Q' where "(\', Q') \ set CsQ" and "\ = substCond.seqSubst \' \" and QeqQ': "Q = Q'[<\>]" + by(induct CsQ) force+ + from \(\', Q') \ set CsQ\ obtain P' where "(\', P') \ set CsP" and "guarded P'" and "\ \ P' \\<^sub>s Q'" by(blast dest: C2) + from \(\', P') \ set CsP\ \\ = substCond.seqSubst \' \\ obtain P where "(\, P) \ set (caseListSeqSubst CsP \)" and "P = P'[<\>]" + by(induct CsP) auto + with QeqQ' \guarded P'\ \\ \ P' \\<^sub>s Q'\ \wellFormedSubst \\ show "\P. (\, P) \ set (caseListSeqSubst CsP \) \ guarded P \ \ \ P \ Q" + by(blast dest: closeSubstE guardedSeqSubst) + qed + } + then show ?thesis + by(intro closeSubstI) auto +qed + +lemma bisimSubstReflexive: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + +shows "\ \ P \\<^sub>s P" + by(auto intro: closeSubstI bisimReflexive) + +lemma bisimSubstTransitive: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and Q :: "('a, 'b, 'c) psi" + and R :: "('a, 'b, 'c) psi" + +assumes "\ \ P \\<^sub>s Q" + and "\ \ Q \\<^sub>s R" + +shows "\ \ P \\<^sub>s R" + using assms + by(auto intro: closeSubstI closeSubstE bisimTransitive) + +lemma bisimSubstSymmetric: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and Q :: "('a, 'b, 'c) psi" + +assumes "\ \ P \\<^sub>s Q" + +shows "\ \ Q \\<^sub>s P" + using assms + by(auto intro: closeSubstI closeSubstE bisimE) + +lemma bisimSubstCasePres: + fixes \ :: 'b + and CsP :: "('c \ ('a, 'b, 'c) psi) list" + and CsQ :: "('c \ ('a, 'b, 'c) psi) list" + +assumes "length CsP = length CsQ" + and C: "\(i::nat) \ P \' Q. \i <= length CsP; (\, P) = nth CsP i; (\', Q) = nth CsQ i\ \ \ = \' \ \ \ P \\<^sub>s Q \ guarded P \ guarded Q" + +shows "\ \ Cases CsP \\<^sub>s Cases CsQ" +proof - + { + fix \ + and P + + assume "(\, P) \ set CsP" + + with \length CsP = length CsQ\ have "\Q. (\, Q) \ set CsQ \ guarded Q \ \ \ P \\<^sub>s Q" using C + proof(induct n=="length CsP" arbitrary: CsP CsQ rule: nat.induct) + case zero then show ?case by simp + next + case(Suc n) then show ?case + proof(cases CsP) + case Nil then show ?thesis using \Suc n = length CsP\ by simp + next + case(Cons P'\ CsP') + note CsPdef = this + then show ?thesis + proof(cases CsQ) + case Nil then show ?thesis using CsPdef \length CsP = length CsQ\ + by simp + next + case(Cons Q'\ CsQ') + obtain Q' \' where Q'phi: "Q'\=(\',Q')" + by(induct Q'\) auto + show ?thesis + proof(cases "P'\ = (\,P)") + case True + have "0 <= length CsP" unfolding CsPdef + by simp + moreover from True have "(\, P) = nth CsP 0" using \(\, P) \ set CsP\ unfolding CsPdef + by simp + moreover have "(\', Q') = nth CsQ 0" unfolding Cons Q'phi by simp + ultimately have "\ = \' \ \ \ P \\<^sub>s Q' \ guarded P \ guarded Q'" + by(rule Suc) + then show ?thesis unfolding Cons Q'phi + by(intro exI[where x=Q']) auto + next + case False + have "n = length CsP'" using \Suc n = length CsP\ unfolding CsPdef + by simp + moreover have "length CsP' = length CsQ'" using \length CsP = length CsQ\ unfolding CsPdef Cons by simp + moreover from \(\,P) \ set CsP\ False have "(\, P) \ set CsP'" unfolding CsPdef by simp + moreover + { + fix i::nat + and \::'c + and \'::'c + and P::"('a,'b,'c) psi" + and Q::"('a,'b,'c) psi" + assume "i \ length CsP'" + and "(\, P) = CsP' ! i" + and "(\', Q) = CsQ' ! i" + then have "(i+1) \ length CsP" + and "(\, P) = CsP ! (i+1)" + and "(\', Q) = CsQ ! (i+1)" + unfolding CsPdef Cons + by simp+ + then have "\ = \' \ \ \ P \\<^sub>s Q \ guarded P \ guarded Q" + by(rule Suc) + } + note this + ultimately have "\Q. (\, Q) \ set CsQ' \ guarded Q \ \ \ P \\<^sub>s Q" + by(rule Suc) + then show ?thesis unfolding Cons by auto + qed + qed + qed + qed + } + note this + moreover + { + fix \ + and Q + + assume "(\, Q) \ set CsQ" + + with \length CsP = length CsQ\ have "\P. (\, P) \ set CsP \ guarded P \ \ \ P \\<^sub>s Q" using C + proof(induct n=="length CsQ" arbitrary: CsP CsQ rule: nat.induct) + case zero then show ?case by simp + next + case(Suc n) then show ?case + proof(cases CsQ) + case Nil then show ?thesis using \Suc n = length CsQ\ by simp + next + case(Cons Q'\ CsQ') + note CsPdef = this + then show ?thesis + proof(cases CsP) + case Nil then show ?thesis using CsPdef \length CsP = length CsQ\ + by simp + next + case(Cons P'\ CsP') + obtain P' \' where P'phi: "P'\=(\',P')" + by(induct P'\) auto + show ?thesis + proof(cases "Q'\ = (\,Q)") + case True + have "0 <= length CsP" unfolding CsPdef + by simp + moreover have "(\', P') = nth CsP 0" unfolding Cons P'phi by simp + moreover from True have "(\, Q) = nth CsQ 0" using \(\, Q) \ set CsQ\ unfolding CsPdef + by simp + ultimately have "\' = \ \ \ \ P' \\<^sub>s Q \ guarded P' \ guarded Q" + by(rule Suc) + then show ?thesis unfolding Cons P'phi + by(intro exI[where x=P']) auto + next + case False + have "n = length CsQ'" using \Suc n = length CsQ\ unfolding CsPdef + by simp + moreover have "length CsP' = length CsQ'" using \length CsP = length CsQ\ unfolding CsPdef Cons by simp + moreover from \(\,Q) \ set CsQ\ False have "(\, Q) \ set CsQ'" unfolding CsPdef by simp + moreover + { + fix i::nat + and \::'c + and \'::'c + and P::"('a,'b,'c) psi" + and Q::"('a,'b,'c) psi" + assume "i \ length CsP'" + and "(\, P) = CsP' ! i" + and "(\', Q) = CsQ' ! i" + then have "(i+1) \ length CsP" + and "(\, P) = CsP ! (i+1)" + and "(\', Q) = CsQ ! (i+1)" + unfolding CsPdef Cons + by simp+ + then have "\ = \' \ \ \ P \\<^sub>s Q \ guarded P \ guarded Q" + by(rule Suc) + } + note this + ultimately have "\P. (\, P) \ set CsP' \ guarded P \ \ \ P \\<^sub>s Q" + by(rule Suc) + then show ?thesis unfolding Cons by auto + qed + qed + qed + qed + } + ultimately show ?thesis + by(rule bisimSubstCasePresAux) +qed + +lemma bisimSubstParPres: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and Q :: "('a, 'b, 'c) psi" + and R :: "('a, 'b, 'c) psi" + +assumes "\ \ P \\<^sub>s Q" + +shows "\ \ P \ R \\<^sub>s Q \ R" + using assms closeSubstI closeSubstE bisimParPres by force + +lemma bisimSubstResPres: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and Q :: "('a, 'b, 'c) psi" + and x :: name + +assumes "\ \ P \\<^sub>s Q" + and "x \ \" + +shows "\ \ \\x\P \\<^sub>s \\x\Q" +proof(rule closeSubstI) + fix \ :: "(name list \ 'a list) list" + + assume "wellFormedSubst \" + obtain y::name where "y \ \" and "y \ P" and "y \ Q" and "y \ \" + by(generate_fresh "name") (auto simp add: fresh_prod) + + from \\ \ P \\<^sub>s Q\ have "([(x, y)] \ \) \ ([(x, y)] \ P) \\<^sub>s ([(x, y)] \ Q)" + by(rule bisimSubstClosed) + with \x \ \\ \y \ \\ have "\ \ ([(x, y)] \ P) \\<^sub>s ([(x, y)] \ Q)" + by simp + then have "\ \ ([(x, y)] \ P)[<\>] \ ([(x, y)] \ Q)[<\>]" using \wellFormedSubst \\ + by(rule closeSubstE) + then have "\ \ \\y\(([(x, y)] \ P)[<\>]) \ \\y\(([(x, y)] \ Q)[<\>])" using \y \ \\ + by(rule bisimResPres) + with \y \ P\ \y \ Q\ \y \ \\ + show "\ \ (\\x\P)[<\>] \ (\\x\Q)[<\>]" + by(simp add: alphaRes) +qed + +lemma bisimSubstBangPres: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and Q :: "('a, 'b, 'c) psi" + +assumes "\ \ P \\<^sub>s Q" + and "guarded P" + and "guarded Q" + +shows "\ \ !P \\<^sub>s !Q" + using assms closeSubstI closeSubstE bisimBangPres guardedSeqSubst by force + +lemma substNil[simp]: + fixes xvec :: "name list" + and Tvec :: "'a list" + +assumes "wellFormedSubst \" + and "distinct xvec" + +shows "(\[<\>]) = \" + using assms + by simp + +lemma bisimSubstParNil: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + +shows "\ \ P \ \ \\<^sub>s P" + by(auto intro!: closeSubstI bisimParNil) + +lemma bisimSubstParComm: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and Q :: "('a, 'b, 'c) psi" + +shows "\ \ P \ Q \\<^sub>s Q \ P" + by(auto intro!: closeSubstI bisimParComm) + +lemma bisimSubstParAssoc: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and Q :: "('a, 'b, 'c) psi" + and R :: "('a, 'b, 'c) psi" + +shows "\ \ (P \ Q) \ R \\<^sub>s P \ (Q \ R)" + by(auto intro!: closeSubstI bisimParAssoc) + +lemma bisimSubstResNil: + fixes \ :: 'b + and x :: name + +shows "\ \ \\x\\ \\<^sub>s \" +proof(rule closeSubstI) + fix \:: "(name list \ 'a list) list" + + assume "wellFormedSubst \" + obtain y::name where "y \ \" and "y \ \" + by(generate_fresh "name") (auto simp add: fresh_prod) + have "\ \ \\y\\ \ \" by(rule bisimResNil) + with \y \ \\ \wellFormedSubst \\ show "\ \ (\\x\\)[<\>] \ \[<\>]" + by(subst alphaRes[of y]) auto +qed + +lemma seqSubst2: + fixes x :: name + and P :: "('a, 'b, 'c) psi" + +assumes "wellFormedSubst \" + and "x \ \" + and "x \ P" + +shows "x \ P[<\>]" + using assms + by(induct \ arbitrary: P, auto) (blast dest: subst2) + +notation substTerm.seqSubst ("_[<_>]" [100, 100] 100) + +lemma bisimSubstScopeExt: + fixes \ :: 'b + and x :: name + and P :: "('a, 'b, 'c) psi" + and Q :: "('a, 'b, 'c) psi" + +assumes "x \ P" + +shows "\ \ \\x\(P \ Q) \\<^sub>s P \ \\x\Q" +proof(rule closeSubstI) + fix \:: "(name list \ 'a list) list" + + assume "wellFormedSubst \" + obtain y::name where "y \ \" and "y \ \" and "y \ P" and "y \ Q" + by(generate_fresh "name") (auto simp add: fresh_prod) + moreover from \wellFormedSubst \\ \y \ \\ \y \ P\ have "y \ P[<\>]" + by(rule seqSubst2) + then have "\ \ \\y\((P[<\>]) \ (([(x, y)] \ Q)[<\>])) \ (P[<\>]) \ \\y\(([(x, y)] \ Q)[<\>])" + by(rule bisimScopeExt) + with \x \ P\ \y \ P\ \y \ Q\ \y \ \\ show "\ \ (\\x\(P \ Q))[<\>] \ (P \ \\x\Q)[<\>]" + apply(subst alphaRes[of y], simp) + apply(subst alphaRes[of y Q], simp) + by(simp add: eqvts) +qed + +lemma bisimSubstCasePushRes: + fixes x :: name + and \ :: 'b + and Cs :: "('c \ ('a, 'b, 'c) psi) list" + +assumes "x \ map fst Cs" + +shows "\ \ \\x\(Cases Cs) \\<^sub>s Cases map (\(\, P). (\, \\x\P)) Cs" +proof(rule closeSubstI) + fix \:: "(name list \ 'a list) list" + + assume "wellFormedSubst \" + obtain y::name where "y \ \" and "y \ \" and "y \ Cs" + by(generate_fresh "name") (auto simp add: fresh_prod) + + { + fix x :: name + and Cs :: "('c \ ('a, 'b, 'c) psi) list" + and \ :: "(name list \ 'a list) list" + + assume "x \ \" + + then have "(Cases map (\(\, P). (\, \\x\P)) Cs)[<\>] = Cases map (\(\, P). (\, \\x\P)) (caseListSeqSubst Cs \)" + by(induct Cs) auto + } + note C1 = this + + { + fix x :: name + and y :: name + and Cs :: "('c \ ('a, 'b, 'c) psi) list" + + assume "x \ map fst Cs" + and "y \ map fst Cs" + and "y \ Cs" + + then have "(Cases map (\(\, P). (\, \\x\P)) Cs) = Cases map (\(\, P). (\, \\y\P)) ([(x, y)] \ Cs)" + by(induct Cs) (auto simp add: fresh_list_cons alphaRes) + } + note C2 = this + + from \y \ Cs\ have "y \ map fst Cs" by(induct Cs) (auto simp add: fresh_list_cons fresh_list_nil) + from \y \ Cs\ \y \ \\ \x \ map fst Cs\ \wellFormedSubst \\ have "y \ map fst (caseListSeqSubst ([(x, y)] \ Cs) \)" + by(induct Cs) (auto intro: substCond.seqSubst2 simp add: fresh_list_cons fresh_list_nil fresh_prod) + then have "\ \ \\y\(Cases(caseListSeqSubst ([(x, y)] \ Cs) \)) \ Cases map (\(\, P). (\, \\y\P)) (caseListSeqSubst ([(x, y)] \ Cs) \)" + by(rule bisimCasePushRes) + + with \y \ Cs\ \x \ map fst Cs\ \y \ map fst Cs\ \y \ \\ \wellFormedSubst \\ + show "\ \ (\\x\(Cases Cs))[<\>] \ (Cases map (\(\, P). (\, \\x\P)) Cs)[<\>]" + apply(subst C2[of x Cs y]) + apply assumption+ + apply(subst C1) + apply assumption+ + apply(subst alphaRes[of y], simp) + by(simp add: eqvts) +qed + +lemma bisimSubstOutputPushRes: + fixes x :: name + and \ :: 'b + and M :: 'a + and N :: 'a + and P :: "('a, 'b, 'c) psi" + +assumes "x \ M" + and "x \ N" + +shows "\ \ \\x\(M\N\.P) \\<^sub>s M\N\.\\x\P" +proof(rule closeSubstI) + fix \:: "(name list \ 'a list) list" + + assume "wellFormedSubst \" + obtain y::name where "y \ \" and "y \ \" and "y \ P" and "y \ M" and "y \ N" + by(generate_fresh "name") (auto simp add: fresh_prod) + from \wellFormedSubst \\ \y \ M\ \y \ \\ have "y \ M[<\>]" by auto + moreover from \wellFormedSubst \\ \y \ N\ \y \ \\ have "y \ N[<\>]" by auto + ultimately have "\ \ \\y\((M[<\>])\(N[<\>])\.(([(x, y)] \ P)[<\>])) \ (M[<\>])\(N[<\>])\.(\\y\(([(x, y)] \ P)[<\>]))" + by(rule bisimOutputPushRes) + with \y \ M\ \y \ N\ \y \ P\ \x \ M\ \x \ N\ \y \ \\ \wellFormedSubst \\ + show "\ \ (\\x\(M\N\.P))[<\>] \ (M\N\.\\x\P)[<\>]" + apply(subst alphaRes[of y], simp) + apply(subst alphaRes[of y P], simp) + by(simp add: eqvts) +qed + +lemma bisimSubstInputPushRes: + fixes x :: name + and \ :: 'b + and M :: 'a + and xvec :: "name list" + and N :: 'a + +assumes "x \ M" + and "x \ xvec" + and "x \ N" + +shows "\ \ \\x\(M\\*xvec N\.P) \\<^sub>s M\\*xvec N\.\\x\P" +proof(rule closeSubstI) + fix \:: "(name list \ 'a list) list" + + assume "wellFormedSubst \" + obtain y::name where "y \ \" and "y \ \" and "y \ P" and "y \ M" and "y \ xvec" and "y \ N" + by(generate_fresh "name") (auto simp add: fresh_prod) + obtain p::"name prm" where "(p \ xvec) \* N" and "(p \ xvec) \* P" and "x \ (p \ xvec)" and "y \ (p \ xvec)" and "(p \ xvec) \* \" + and S: "set p \ set xvec \ set(p \ xvec)" + by(rule name_list_avoiding[where c="(N, P, x, y, \)"]) auto + + from \wellFormedSubst \\ \y \ M\ \y \ \ \ have "y \ M[<\>]" by auto + moreover note \y \ (p \ xvec)\ + moreover from \y \ N\ have "(p \ y) \ (p \ N)" by(simp add: pt_fresh_bij[OF pt_name_inst, OF at_name_inst]) + with \y \ xvec\ \y \ (p \ xvec)\ S have "y \ p \ N" by simp + with \wellFormedSubst \\ have "y \ (p \ N)[<\>]" using \y \ \\ by auto + ultimately have "\ \ \\y\((M[<\>])\\*(p \ xvec) ((p \ N)[<\>])\.(([(x, y)] \ (p \ P))[<\>])) \ (M[<\>])\\*(p \ xvec) ((p \ N)[<\>])\.(\\y\(([(x, y)] \ p \ P)[<\>]))" + by(rule bisimInputPushRes) + with \y \ M\ \y \ N\ \y \ P\ \x \ M\ \x \ N\ \y \ xvec\ \x \ xvec\ \(p \ xvec) \* N\ \(p \ xvec) \* P\ + \x \ (p \ xvec)\ \y \ (p \ xvec)\ \y \ \\ \(p \ xvec) \* \\ S \wellFormedSubst \\ + show "\ \ (\\x\(M\\*xvec N\.P))[<\>] \ (M\\*xvec N\.\\x\P)[<\>]" + apply(subst inputChainAlpha') + apply assumption+ + apply(subst inputChainAlpha'[of p xvec]) + apply(simp add: abs_fresh_star) + apply assumption+ + apply(simp add: eqvts) + apply(subst alphaRes[of y], simp) + apply(simp add: inputChainFresh) + apply(simp add: freshChainSimps) + apply(subst alphaRes[of y "(p \ P)"]) + apply(simp add: freshChainSimps) + by(simp add: freshChainSimps eqvts) +qed + +lemma bisimSubstResComm: + fixes x :: name + and y :: name + +shows "\ \ \\x\(\\y\P) \\<^sub>s \\y\(\\x\P)" +proof(cases "x = y") + assume "x = y" + then show ?thesis by(force intro: bisimSubstReflexive) +next + assume "x \ y" + show ?thesis + proof(rule closeSubstI) + fix \:: "(name list \ 'a list) list" + assume "wellFormedSubst \" + + + obtain x'::name where "x' \ \" and "x' \ \" and "x' \ P" and "x \ x'" and "y \ x'" + by(generate_fresh "name") (auto simp add: fresh_prod) + obtain y'::name where "y' \ \" and "y' \ \" and "y' \ P" and "x \ y'" and "y \ y'" and "x' \ y'" + by(generate_fresh "name") (auto simp add: fresh_prod) + + have "\ \ \\x'\(\\y'\(([(x, x')] \ [(y, y')] \ P)[<\>])) \ \\y'\(\\x'\(([(x, x')] \ [(y, y')] \ P)[<\>]))" + by(rule bisimResComm) + moreover from \x' \ P\ \y' \ P\ \x \ y'\ \x' \ y'\ have "\\x\(\\y\P) = \\x'\(\\y'\(([(x, x')] \ [(y, y')] \ P)))" + apply(subst alphaRes[of y' P], simp) + by(subst alphaRes[of x']) (auto simp add: abs_fresh fresh_left calc_atm eqvts) + moreover from \x' \ P\ \y' \ P\ \y \ x'\ \x \ y'\ \x' \ y'\ \x \ x'\ \x \ y\ have "\\y\(\\x\P) = \\y'\(\\x'\(([(x, x')] \ [(y, y')] \ P)))" + apply(subst alphaRes[of x' P], simp) + apply(subst alphaRes[of y'], simp add: abs_fresh fresh_left calc_atm) + apply(simp add: eqvts calc_atm) + by(subst perm_compose) (simp add: calc_atm) + + ultimately show "\ \ (\\x\(\\y\P))[<\>] \ (\\y\(\\x\P))[<\>]" + using \wellFormedSubst \\ \x' \ \\ \y' \ \\ + by simp + qed +qed + +lemma bisimSubstExtBang: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + +assumes "guarded P" + +shows "\ \ !P \\<^sub>s P \ !P" + using assms closeSubstI bangExt guardedSeqSubst by force + +lemma structCongBisimSubst: + fixes P :: "('a, 'b, 'c) psi" + and Q :: "('a, 'b, 'c) psi" + +assumes "P \\<^sub>s Q" + +shows "P \\<^sub>s Q" + using assms + by(induct rule: structCong.induct) + (auto intro: bisimSubstReflexive bisimSubstSymmetric bisimSubstTransitive bisimSubstParComm bisimSubstParAssoc bisimSubstParNil bisimSubstResNil bisimSubstResComm bisimSubstScopeExt bisimSubstCasePushRes bisimSubstInputPushRes bisimSubstOutputPushRes bisimSubstExtBang) + +end + +end diff --git a/thys/Broadcast_Psi/Bisimulation.thy b/thys/Broadcast_Psi/Bisimulation.thy new file mode 100644 --- /dev/null +++ b/thys/Broadcast_Psi/Bisimulation.thy @@ -0,0 +1,993 @@ +theory Bisimulation + imports Simulation +begin + +text \This file is a (heavily modified) variant of the theory {\it Psi\_Calculi.Bisimulation} +from~\cite{DBLP:journals/afp/Bengtson12}.\ + +context env begin + +lemma monoCoinduct: "\x y xa xb xc P Q \. + x \ y \ + (\ \ Q \[{(xc, xb, xa). x xc xb xa}] P) \ + (\ \ Q \[{(xb, xa, xc). y xb xa xc}] P)" + apply(rule impI) + apply(rule monotonic) + by(auto dest: le_funE) + +coinductive_set bisim :: "('b \ ('a, 'b, 'c) psi \ ('a, 'b, 'c) psi) set" + where + step: "\(insertAssertion (extractFrame P)) \ \\<^sub>F (insertAssertion (extractFrame Q) \); + \ \ P \[bisim] Q; + \\'. (\ \ \', P, Q) \ bisim; (\, Q, P) \ bisim\ \ (\, P, Q) \ bisim" + monos monoCoinduct + +abbreviation + bisimJudge ("_ \ _ \ _" [70, 70, 70] 65) where "\ \ P \ Q \ (\, P, Q) \ bisim" +abbreviation + bisimNilJudge ("_ \ _" [70, 70] 65) where "P \ Q \ SBottom' \ P \ Q" + +lemma bisimCoinductAux[consumes 1]: + fixes F :: 'b + and P :: "('a, 'b, 'c) psi" + and Q :: "('a, 'b, 'c) psi" + and X :: "('b \ ('a, 'b, 'c) psi \ ('a, 'b, 'c) psi) set" + +assumes "(\, P, Q) \ X" + and "\\ P Q. (\, P, Q) \ X \ insertAssertion (extractFrame P) \ \\<^sub>F insertAssertion (extractFrame Q) \ \ + (\ \ P \[(X \ bisim)] Q) \ + (\\'. (\ \ \', P, Q) \ X \ (\ \ \', P, Q) \ bisim) \ + ((\, Q, P) \ X \ (\, Q, P) \ bisim)" + +shows "(\, P, Q) \ bisim" +proof - + have "X \ bisim = {(\, P, Q). (\, P, Q) \ X \ (\, P, Q) \ bisim}" by auto + with assms show ?thesis + by coinduct simp +qed + +lemma bisimCoinduct[consumes 1, case_names cStatEq cSim cExt cSym]: + fixes F :: 'b + and P :: "('a, 'b, 'c) psi" + and Q :: "('a, 'b, 'c) psi" + and X :: "('b \ ('a, 'b, 'c) psi \ ('a, 'b, 'c) psi) set" + +assumes "(\, P, Q) \ X" + and "\\' R S. (\', R, S) \ X \ insertAssertion (extractFrame R) \' \\<^sub>F insertAssertion (extractFrame S) \'" + and "\\' R S. (\', R, S) \ X \ \' \ R \[(X \ bisim)] S" + and "\\' R S \''. (\', R, S) \ X \ (\' \ \'', R, S) \ X \ (\' \ \'', R, S) \ bisim" + and "\\' R S. (\', R, S) \ X \ (\', S, R) \ X \ (\', S, R) \ bisim" + +shows "(\, P, Q) \ bisim" +proof - + have "X \ bisim = {(\, P, Q). (\, P, Q) \ X \ (\, P, Q) \ bisim}" by auto + with assms show ?thesis + by coinduct simp +qed + +lemma bisimWeakCoinductAux[consumes 1]: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and Q :: "('a, 'b, 'c) psi" + and X :: "('b \ ('a, 'b, 'c) psi \ ('a, 'b, 'c) psi) set" + +assumes "(\, P, Q) \ X" + and "\\ P Q. (\, P, Q) \ X \ insertAssertion (extractFrame P) \ \\<^sub>F insertAssertion (extractFrame Q) \ \ + \ \ P \[X] Q \ + (\\'. (\ \ \', P, Q) \ X) \ (\, Q, P) \ X" + +shows "(\, P, Q) \ bisim" + using assms + by(coinduct rule: bisimCoinductAux) (blast intro: monotonic) + +lemma bisimWeakCoinduct[consumes 1, case_names cStatEq cSim cExt cSym]: + fixes F :: 'b + and P :: "('a, 'b, 'c) psi" + and Q :: "('a, 'b, 'c) psi" + and X :: "('b \ ('a, 'b, 'c) psi \ ('a, 'b, 'c) psi) set" + +assumes "(\, P, Q) \ X" + and "\\ P Q. (\, P, Q) \ X \ insertAssertion (extractFrame P) \ \\<^sub>F insertAssertion (extractFrame Q) \" + and "\\ P Q. (\, P, Q) \ X \ \ \ P \[X] Q" + and "\\ P Q \'. (\, P, Q) \ X \ (\ \ \', P, Q) \ X" + and "\\ P Q. (\, P, Q) \ X \ (\, Q, P) \ X" + +shows "(\, P, Q) \ bisim" +proof - + have "X \ bisim = {(\, P, Q). (\, P, Q) \ X \ (\, P, Q) \ bisim}" by auto + with assms show ?thesis + by(coinduct rule: bisimCoinduct) (blast intro: monotonic)+ +qed + +lemma bisimE: + fixes P :: "('a, 'b, 'c) psi" + and Q :: "('a, 'b, 'c) psi" + and \ :: 'b + and \' :: 'b + +assumes "(\, P, Q) \ bisim" + +shows "insertAssertion (extractFrame P) \ \\<^sub>F insertAssertion (extractFrame Q) \" + and "\ \ P \[bisim] Q" + and "(\ \ \', P, Q) \ bisim" + and "(\, Q, P) \ bisim" + using assms + by(auto simp add: intro: bisim.cases) + +lemma bisimI: + fixes P :: "('a, 'b, 'c) psi" + and Q :: "('a, 'b, 'c) psi" + and \ :: 'b + +assumes "insertAssertion (extractFrame P) \ \\<^sub>F insertAssertion (extractFrame Q) \" + and "\ \ P \[bisim] Q" + and "\\'. (\ \ \', P, Q) \ bisim" + and "(\, Q, P) \ bisim" + +shows "(\, P, Q) \ bisim" + using assms + by(auto intro: bisim.step) + +lemma bisimReflexive: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + + +shows "\ \ P \ P" +proof - + let ?X = "{(\, P, P) | \ P. True}" + have "(\, P, P) \ ?X" by simp + then show ?thesis + by(coinduct rule: bisimWeakCoinduct, auto intro: reflexive) +qed + +lemma bisimClosed: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and Q :: "('a, 'b, 'c) psi" + and p :: "name prm" + +assumes PBisimQ: "\ \ P \ Q" + +shows "(p \ \) \ (p \ P) \ (p \ Q)" +proof - + let ?X = "{(p \ \, p \ P, p \ Q) | (p::name prm) \ P Q. \ \ P \ Q}" + from PBisimQ have "(p \ \, p \ P, p \ Q) \ ?X" by blast + then show ?thesis + proof(coinduct rule: bisimWeakCoinduct) + case(cStatEq \ P Q) + have "\\ P Q (p::name prm). insertAssertion (extractFrame P) \ \\<^sub>F insertAssertion (extractFrame Q) \ \ + insertAssertion (extractFrame(p \ P)) (p \ \) \\<^sub>F insertAssertion (extractFrame(p \ Q)) (p \ \)" + by(drule FrameStatEqClosed) (simp add: eqvts) + + with \(\, P, Q) \ ?X\ show ?case by(blast dest: bisimE) + next + case(cSim \ P Q) + { + fix p :: "name prm" + fix \ P Q + have "eqvt ?X" + apply(clarsimp simp add: eqvt_def) + by (metis (no_types, opaque_lifting) pt_name2) + moreover assume "\ \ P \[bisim] Q" + then have "\ \ P \[?X] Q" + apply(rule monotonic[where A=bisim], auto) + by(rule exI[where x="[]::name prm"]) auto + ultimately have "((p::name prm) \ \) \ (p \ P) \[?X] (p \ Q)" + by(rule simClosed) + } + with \(\, P, Q) \ ?X\ show ?case + by(blast dest: bisimE) + next + case(cExt \ P Q \') + { + fix p :: "name prm" + fix \ P Q \' + assume "\\'. (\ \ \', P, Q) \ bisim" + then have "((p \ \) \ \', p \ P, p \ Q) \ ?X" + apply(clarsimp) + apply(rule exI[where x=p]) + apply(rule exI[where x="\ \ (rev p \ \')"]) + by(auto simp add: eqvts) + } + with \(\, P, Q) \ ?X\ show ?case + by(blast dest: bisimE) + next + case(cSym \ P Q) + then show ?case + by(blast dest: bisimE) + qed +qed + +lemma bisimEqvt[simp]: + shows "eqvt bisim" + by(auto simp add: eqvt_def bisimClosed) + +lemma statEqBisim: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and Q :: "('a, 'b, 'c) psi" + and \' :: 'b + +assumes "\ \ P \ Q" + and "\ \ \'" + +shows "\' \ P \ Q" +proof - + let ?X = "{(\', P, Q) | \ P Q \'. \ \ P \ Q \ \ \ \'}" + from \\ \ P \ Q\ \\ \ \'\ have "(\', P, Q) \ ?X" by auto + then show ?thesis + proof(coinduct rule: bisimCoinduct) + case(cStatEq \' P Q) + from \(\', P, Q) \ ?X\ obtain \ where "\ \ P \ Q" and "\ \ \'" + by auto + from \\ \ P \ Q\ have PeqQ: "insertAssertion (extractFrame P) \ \\<^sub>F insertAssertion (extractFrame Q) \" + by(rule bisimE) + + obtain A\<^sub>P \\<^sub>P where FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" and "A\<^sub>P \* \" and "A\<^sub>P \* \'" + by(rule freshFrame[where C="(\, \')"]) auto + obtain A\<^sub>Q \\<^sub>Q where FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" and "A\<^sub>Q \* \" and "A\<^sub>Q \* \'" + by(rule freshFrame[where C="(\, \')"]) auto + + from PeqQ FrP FrQ \A\<^sub>P \* \\ \A\<^sub>Q \* \\ \\ \ \'\ + have "\A\<^sub>P, \' \ \\<^sub>P\ \\<^sub>F \A\<^sub>Q, \' \ \\<^sub>Q\" + by simp (metis frameIntComposition FrameStatEqTrans FrameStatEqSym) + with FrP FrQ \A\<^sub>P \* \'\ \A\<^sub>Q \* \'\ show ?case by simp + next + case(cSim \' P Q) + from \(\', P, Q) \ ?X\ obtain \ where "\ \ P \ Q" and "\ \ \'" + by auto + from \\ \ P \ Q\ have "\ \ P \[bisim] Q" by(blast dest: bisimE) + moreover have "eqvt ?X" + by(auto simp add: eqvt_def) (metis bisimClosed AssertionStatEqClosed) + then have "eqvt(?X \ bisim)" by auto + moreover note \\ \ \'\ + moreover have "\\ P Q \'. \\ \ P \ Q; \ \ \'\ \ (\', P, Q) \ ?X \ bisim" + by auto + ultimately show ?case + by(rule statEqSim) + next + case(cExt \' P Q \'') + from \(\', P, Q) \ ?X\ obtain \ where "\ \ P \ Q" and "\ \ \'" + by auto + from \\ \ P \ Q\ have "\ \ \'' \ P \ Q" by(rule bisimE) + moreover from \\ \ \'\ have "\ \ \'' \ \' \ \''" by(rule Composition) + ultimately show ?case by blast + next + case(cSym \' P Q) + from \(\', P, Q) \ ?X\ obtain \ where "\ \ P \ Q" and "\ \ \'" + by auto + from \\ \ P \ Q\ have "\ \ Q \ P" by(rule bisimE) + then show ?case using \\ \ \'\ by auto + qed +qed + +lemma bisimTransitive: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and Q :: "('a, 'b, 'c) psi" + and R :: "('a, 'b, 'c) psi" + +assumes PQ: "\ \ P \ Q" + and QR: "\ \ Q \ R" + +shows "\ \ P \ R" +proof - + let ?X = "{(\, P, R) | \ P Q R. \ \ P \ Q \ \ \ Q \ R}" + from PQ QR have "(\, P, R) \ ?X" by auto + then show ?thesis + proof(coinduct rule: bisimCoinduct) + case(cStatEq \ P R) + then show ?case by(blast dest: bisimE FrameStatEqTrans) + next + case(cSim \ P R) + { + fix \ P Q R + assume "\ \ P \[bisim] Q" and "\ \ Q \[bisim] R" + moreover have "eqvt ?X" + by(force simp add: eqvt_def dest: bisimClosed) + with bisimEqvt have "eqvt (?X \ bisim)" by blast + moreover have "?X \ ?X \ bisim" by auto + ultimately have "\ \ P \[(?X \ bisim)] R" + by(force intro: transitive) + } + with \(\, P, R) \ ?X\ show ?case + by(blast dest: bisimE) + next + case(cExt \ P R \') + then show ?case by(blast dest: bisimE) + next + case(cSym \ P R) + then show ?case by(blast dest: bisimE) + qed +qed + +lemma weakTransitiveCoinduct[case_names cStatEq cSim cExt cSym, case_conclusion bisim step, consumes 2]: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and Q :: "('a, 'b, 'c) psi" + and X :: "('b \ ('a, 'b, 'c) psi \ ('a, 'b, 'c) psi) set" + +assumes p: "(\, P, Q) \ X" + and Eqvt: "eqvt X" + and rStatEq: "\\ P Q. (\, P, Q) \ X \ insertAssertion (extractFrame P) \ \\<^sub>F insertAssertion (extractFrame Q) \" + and rSim: "\\ P Q. (\, P, Q) \ X \ \ \ P \[({(\, P, Q) | \ P P' Q' Q. \ \ P \ P' \ + (\, P', Q') \ X \ + \ \ Q' \ Q})] Q" + and rExt: "\\ P Q \'. (\, P, Q) \ X \ (\ \ \', P, Q) \ X" + and rSym: "\\ P Q. (\, P, Q) \ X \ (\, Q, P) \ X" + +shows "\ \ P \ Q" +proof - + let ?X = "{(\, P, Q) | \ P P' Q' Q. \ \ P \ P' \ (\, P', Q') \ X \ \ \ Q' \ Q}" + from p have "(\, P, Q) \ ?X" + by(blast intro: bisimReflexive) + then show ?thesis + proof(coinduct rule: bisimWeakCoinduct) + case(cStatEq \ P Q) + then show ?case + by(blast dest: rStatEq bisimE FrameStatEqTrans) + next + case(cSim \ P Q) + { + fix \ P P' Q' Q + assume "\ \ P \[bisim] P'" + moreover assume P'RelQ': "(\, P', Q') \ X" + then have "\ \ P' \[?X] Q'" by(rule rSim) + moreover from \eqvt X\ P'RelQ' have "eqvt ?X" + apply(clarsimp simp add: eqvt_def) + apply(drule bisimClosed) + apply(drule bisimClosed) + apply(rule exI) + apply(rule conjI) + apply assumption + apply(rule exI) + by auto + ultimately have "\ \ P \[?X] Q'" + by(force intro: transitive dest: bisimTransitive) + moreover assume "\ \ Q' \[bisim] Q" + ultimately have "\ \ P \[?X] Q" using \eqvt ?X\ + by(force intro: transitive dest: bisimTransitive) + } + with \(\, P, Q) \ ?X\ show ?case + by(blast dest: bisimE) + next + case(cExt \ P Q \') + then show ?case by(blast dest: bisimE intro: rExt) + next + case(cSym \ P Q) + then show ?case by(blast dest: bisimE intro: rSym) + qed +qed + +lemma weakTransitiveCoinduct'[case_names cStatEq cSim cExt cSym, case_conclusion bisim step, consumes 2]: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and Q :: "('a, 'b, 'c) psi" + and X :: "('b \ ('a, 'b, 'c) psi \ ('a, 'b, 'c) psi) set" + +assumes p: "(\, P, Q) \ X" + and Eqvt: "eqvt X" + and rStatEq: "\\ P Q. (\, P, Q) \ X \ insertAssertion (extractFrame P) \ \\<^sub>F insertAssertion (extractFrame Q) \" + and rSim: "\\ P Q. (\, P, Q) \ X \ \ \ P \[({(\, P, Q) | \ P P' Q' Q. \ \ P \ P' \ + (\, P', Q') \ X \ + \ \ Q' \ Q})] Q" + and rExt: "\\ P Q \'. (\, P, Q) \ X \ (\ \ \', P, Q) \ X" + and rSym: "\\ P Q. (\, P, Q) \ X \ + (\, Q, P) \ {(\, P, Q) | \ P P' Q' Q. \ \ P \ P' \ (\, P', Q') \ X \ \ \ Q' \ Q}" + +shows "\ \ P \ Q" +proof - + let ?X = "{(\, P, Q) | \ P P' Q' Q. \ \ P \ P' \ (\, P', Q') \ X \ \ \ Q' \ Q}" + from p have "(\, P, Q) \ ?X" + by(blast intro: bisimReflexive) + then show ?thesis + proof(coinduct rule: bisimWeakCoinduct) + case(cStatEq \ P Q) + then show ?case + by(blast dest: rStatEq bisimE FrameStatEqTrans) + next + case(cSim \ P Q) + { + fix \ P P' Q' Q + assume "\ \ P \[bisim] P'" + moreover assume P'RelQ': "(\, P', Q') \ X" + then have "\ \ P' \[?X] Q'" by(rule rSim) + moreover from \eqvt X\ P'RelQ' have "eqvt ?X" + apply(clarsimp simp add: eqvt_def) + apply(drule bisimClosed) + apply(drule bisimClosed) + apply(rule exI) + apply(rule conjI) + apply assumption + apply(rule exI) + by auto + ultimately have "\ \ P \[?X] Q'" + by(force intro: transitive dest: bisimTransitive) + moreover assume "\ \ Q' \[bisim] Q" + ultimately have "\ \ P \[?X] Q" using \eqvt ?X\ + by(force intro: transitive dest: bisimTransitive) + } + with \(\, P, Q) \ ?X\ show ?case + by(blast dest: bisimE) + next + case(cExt \ P Q \') + then show ?case by(blast dest: bisimE intro: rExt) + next + case(cSym \ P Q) + then show ?case + apply clarsimp + apply(drule rSym) + apply clarsimp + by(metis bisimTransitive bisimE(4)) + qed +qed + +lemma weakTransitiveCoinduct''[case_names cStatEq cSim cExt cSym, case_conclusion bisim step, consumes 2]: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and Q :: "('a, 'b, 'c) psi" + and X :: "('b \ ('a, 'b, 'c) psi \ ('a, 'b, 'c) psi) set" + +assumes p: "(\, P, Q) \ X" + and Eqvt: "eqvt X" + and rStatEq: "\\ P Q. (\, P, Q) \ X \ insertAssertion (extractFrame P) \ \\<^sub>F insertAssertion (extractFrame Q) \" + and rSim: "\\ P Q. (\, P, Q) \ X \ \ \ P \[({(\, P, Q) | \ P P' Q' Q. \ \ P \ P' \ + (\, P', Q') \ X \ + \ \ Q' \ Q})] Q" + and rExt: "\\ P Q \'. (\, P, Q) \ {(\, P, Q) | \ P P' Q' Q. \ \ P \ P' \ (\, P', Q') \ X \ \ \ Q' \ Q} \ + (\ \ \', P, Q) \ {(\, P, Q) | \ P P' Q' Q. \ \ P \ P' \ (\, P', Q') \ X \ \ \ Q' \ Q}" + and rSym: "\\ P Q. (\, P, Q) \ {(\, P, Q) | \ P P' Q' Q. \ \ P \ P' \ (\, P', Q') \ X \ \ \ Q' \ Q} \ + (\, Q, P) \ {(\, P, Q) | \ P P' Q' Q. \ \ P \ P' \ (\, P', Q') \ X \ \ \ Q' \ Q}" + +shows "\ \ P \ Q" +proof - + let ?X = "{(\, P, Q) | \ P P' Q' Q. \ \ P \ P' \ (\, P', Q') \ X \ \ \ Q' \ Q}" + from p have "(\, P, Q) \ ?X" + by(blast intro: bisimReflexive) + then show ?thesis + proof(coinduct rule: bisimWeakCoinduct) + case(cStatEq \ P Q) + then show ?case + by(blast dest: rStatEq bisimE FrameStatEqTrans) + next + case(cSim \ P Q) + { + fix \ P P' Q' Q + assume "\ \ P \[bisim] P'" + moreover assume P'RelQ': "(\, P', Q') \ X" + then have "\ \ P' \[?X] Q'" by(rule rSim) + moreover from \eqvt X\ P'RelQ' have "eqvt ?X" + apply(clarsimp simp add: eqvt_def) + apply(drule bisimClosed) + apply(drule bisimClosed) + apply(rule exI) + apply(rule conjI) + apply assumption + apply(rule exI) + by auto + ultimately have "\ \ P \[?X] Q'" + by(force intro: transitive dest: bisimTransitive) + moreover assume "\ \ Q' \[bisim] Q" + ultimately have "\ \ P \[?X] Q" using \eqvt ?X\ + by(force intro: transitive dest: bisimTransitive) + } + with \(\, P, Q) \ ?X\ show ?case + by(blast dest: bisimE) + next + case(cExt \ P Q \') + then show ?case by(rule rExt) + next + case(cSym \ P Q) + then show ?case by(rule rSym) + qed +qed + +lemma transitiveCoinduct[case_names cStatEq cSim cExt cSym, case_conclusion bisim step, consumes 2]: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and Q :: "('a, 'b, 'c) psi" + and X :: "('b \ ('a, 'b, 'c) psi \ ('a, 'b, 'c) psi) set" + +assumes p: "(\, P, Q) \ X" + and Eqvt: "eqvt X" + and rStatEq: "\\' R S. (\', R, S) \ X \ insertAssertion (extractFrame R) \' \\<^sub>F insertAssertion (extractFrame S) \'" + and rSim: "\\' R S. (\', R, S) \ X \ \' \ R \[({(\', R, S) | \' R R' S' S. \' \ R \ R' \ + ((\', R', S') \ X \ \' \ R' \ S') \ + \' \ S' \ S})] S" + and rExt: "\\' R S \''. (\', R, S) \ X \ (\' \ \'', R, S) \ X \ \' \ \'' \ R \ S" + and rSym: "\\' R S. (\', R, S) \ X \ (\', S, R) \ X \ \' \ S \ R" + + +shows "\ \ P \ Q" +proof - + from p have "(\, P, Q) \ (X \ bisim)" + by blast + moreover from \eqvt X\ bisimEqvt have "eqvt (X \ bisim)" + by auto + ultimately show ?thesis + proof(coinduct rule: weakTransitiveCoinduct') + case(cStatEq \ P Q) + then show ?case + by(blast intro: rStatEq dest: bisimE) + next + case(cSim \ P Q) + then show ?case + apply clarsimp + apply (erule disjE) + apply(blast intro: rSim) + apply(drule bisimE(2)) + apply(rule monotonic, simp) + by(force intro: bisimReflexive) + next + case(cExt \ P Q \') + then show ?case + by(blast dest: bisimE rExt) + next + case(cSym \ P Q) + then show ?case by(blast dest: bisimE rSym intro: bisimReflexive) + qed +qed + +lemma transitiveCoinduct'[case_names cStatEq cSim cExt cSym, case_conclusion bisim step, consumes 2]: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and Q :: "('a, 'b, 'c) psi" + and X :: "('b \ ('a, 'b, 'c) psi \ ('a, 'b, 'c) psi) set" + +assumes p: "(\, P, Q) \ X" + and Eqvt: "eqvt X" + and rStatEq: "\\ P Q. (\, P, Q) \ X \ insertAssertion (extractFrame P) \ \\<^sub>F insertAssertion (extractFrame Q) \" + and rSim: "\\ P Q. (\, P, Q) \ X \ \ \ P \[({(\, P, Q) | \ P P' Q' Q. \ \ P \ P' \ + (\, P', Q') \ (X \ bisim) \ + \ \ Q' \ Q})] Q" + and rExt: "\\ P Q \'. (\, P, Q) \ X \ (\ \ \', P, Q) \ X \ \ \ \' \ P \ Q" + and rSym: "\\ P Q. (\, P, Q) \ X \ + (\, Q, P) \ {(\, P, Q) | \ P P' Q' Q. \ \ P \ P' \ ((\, P', Q') \ (X \ bisim)) \ \ \ Q' \ Q}" + +shows "\ \ P \ Q" +proof - + from p have "(\, P, Q) \ (X \ bisim)" + by blast + moreover from \eqvt X\ bisimEqvt have "eqvt (X \ bisim)" + by auto + ultimately show ?thesis + proof(coinduct rule: weakTransitiveCoinduct') + case(cStatEq \ P Q) + then show ?case + by(blast intro: rStatEq dest: bisimE) + next + case(cSim \ P Q) + then show ?case + apply - + apply(cases "(\, P, Q) \ X") + apply(rule rSim) + apply simp + apply(clarify) + apply(drule bisimE(2)) + apply(rule monotonic, simp) + by(force intro: bisimReflexive) + next + case(cExt \ P Q \') + then show ?case + by(blast dest: bisimE rExt) + next + case(cSym \ P Q) + then show ?case + apply clarsimp + apply (erule disjE) + apply(drule rSym) + apply clarsimp + apply(rule exI[where x=Q]) + apply(simp add: bisimReflexive) + apply(rule exI) + by(auto intro: bisimReflexive dest: bisimE(4)) + qed +qed + +lemma bisimSymmetric: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and Q :: "('a, 'b, 'c) psi" + +assumes "\ \ P \ Q" + +shows "\ \ Q \ P" + using assms + by(rule bisimE) + +lemma eqvtTrans[intro]: + assumes "eqvt X" + +shows "eqvt {(\, P, Q) | \ P P' Q' Q. \ \ P \ P' \ ((\, P', Q') \ X \ \ \ P' \ Q') \ \ \ Q' \ Q}" + using assms + apply(clarsimp simp add: eqvt_def) + apply(erule disjE) + using ballE bisimClosed apply fastforce + by(blast dest: bisimClosed) + +lemma eqvtWeakTrans[intro]: + assumes "eqvt X" + +shows "eqvt {(\, P, Q) | \ P P' Q' Q. \ \ P \ P' \ (\, P', Q') \ X \ \ \ Q' \ Q}" + using assms + apply(clarsimp simp add: eqvt_def) + using ballE bisimClosed by fastforce + +inductive_set + rel_trancl :: "('b \ ('a,'b,'c) psi \ ('a,'b,'c) psi) set \ ('b \ ('a,'b,'c) psi \ ('a,'b,'c) psi) set" ("(_\<^sup>\)" [1000] 999) + for r :: "('b \ ('a,'b,'c) psi \ ('a,'b,'c) psi) set" + where + r_into_rel_trancl [intro, Pure.intro]: "(\, P,Q) : r ==> (\,P,Q) : r\<^sup>\" + | rel_trancl_into_rel_trancl [Pure.intro]: "(\,P,Q) : r\<^sup>\ ==> (\,Q,R) : r ==> (\,P,R) : r\<^sup>\" + +lemma rel_trancl_transitive: + assumes "(\,P,Q) \ Rel\<^sup>\" + and "(\,Q,R) \ Rel\<^sup>\" + shows "(\,P,R) \ Rel\<^sup>\" + using \(\,Q,R) \ Rel\<^sup>\\ \(\,P,Q) \ Rel\<^sup>\\ + by(induct rule: rel_trancl.induct) (auto intro: rel_trancl_into_rel_trancl) + +lemma rel_trancl_eqvt: + assumes "eqvt X" + shows "eqvt(X\<^sup>\)" +proof - + { + fix p::"name prm" + and \ P Q + assume "(\,P,Q) \ X\<^sup>\" + then have "(p\(\,P,Q)) \ X\<^sup>\" + proof(induct rule: rel_trancl.induct) + case(r_into_rel_trancl \ P Q) + with \eqvt X\ have "(p\(\, P, Q)) \ X" + unfolding eqvt_def by auto + then show ?case by auto + next + case(rel_trancl_into_rel_trancl \ P Q R) + from \(\, Q, R) \ X\ \eqvt X\ have "(p\(\, Q, R)) \ X" + unfolding eqvt_def by auto + then have "(p\(\, Q, R)) \ X\<^sup>\" + by auto + with \p \ (\, P, Q) \ X\<^sup>\\ show ?case by(auto dest: rel_trancl_transitive) + qed + } + then show ?thesis unfolding eqvt_def + by auto +qed + +lemma bisimStarWeakCoinduct[consumes 2, case_names cStatEq cSim cExt cSym]: + fixes F :: 'b + and P :: "('a, 'b, 'c) psi" + and Q :: "('a, 'b, 'c) psi" + and X :: "('b \ ('a, 'b, 'c) psi \ ('a, 'b, 'c) psi) set" + +assumes "(\, P, Q) \ X" + and "eqvt X" + and rStatEq: "\\' R S. (\', R, S) \ X \ insertAssertion (extractFrame R) \' \\<^sub>F insertAssertion (extractFrame S) \'" + and rSim: "\\' R S. (\', R, S) \ X \ \' \ R \[X\<^sup>\] S" + and rExt: "\\' R S \''. (\', R, S) \ X \ (\' \ \'', R, S) \ X" + and rSym: "\\' R S. (\', R, S) \ X \ (\', S, R) \ X" + +shows "(\, P, Q) \ bisim" +proof - + have "eqvt(X\<^sup>\)" using \eqvt X\ + by(rule rel_trancl_eqvt) + from \(\, P, Q) \ X\ + have "(\, P, Q) \ X\<^sup>\" + by force + then show ?thesis + proof(coinduct rule: bisimWeakCoinduct) + case(cSim \' R S) + then show ?case + proof(induct rule: rel_trancl.induct) + case(r_into_rel_trancl \ P Q) + then show ?case + by(rule rSim) + next + case(rel_trancl_into_rel_trancl \ P Q R) + note \\ \ P \[X\<^sup>\] Q\ + moreover from \(\, Q, R) \ X\ have "\ \ Q \[X\<^sup>\] R" + by(rule rSim) + moreover note \eqvt(X\<^sup>\)\ + moreover have "{(\, P, R) |\ P R. \Q. (\, P, Q) \ X\<^sup>\ \ (\, Q, R) \ X\<^sup>\} \ X\<^sup>\" + by(blast intro: rel_trancl_transitive) + ultimately show ?case + using transitive by meson + qed + next + case (cStatEq \ P Q) + then show ?case + proof(induct rule: rel_trancl.induct) + case(r_into_rel_trancl \ P Q) + then show ?case + by(rule rStatEq) + next + case(rel_trancl_into_rel_trancl \ P Q R) + from \(\, Q, R) \ X\ have "insertAssertion (extractFrame Q) \ \\<^sub>F insertAssertion (extractFrame R) \" + by(rule rStatEq) + with \insertAssertion (extractFrame P) \ \\<^sub>F insertAssertion (extractFrame Q) \\ + show ?case + by(rule FrameStatEqTrans) + qed + next + case(cExt \ P Q \') + then show ?case + proof(induct rule: rel_trancl.induct) + case(r_into_rel_trancl \ P Q) + then have "(\ \ \', P, Q) \ X" by(rule rExt) + then show ?case + by force + next + case(rel_trancl_into_rel_trancl \ P Q R) + from \(\, Q, R) \ X\ have "(\ \ \', Q, R) \ X" by(rule rExt) + then have "(\ \ \', Q, R) \ X\<^sup>\" by force + with \(\ \ \', P, Q) \ X\<^sup>\\ + show ?case + by(rule rel_trancl_transitive) + qed + next + case(cSym \ P Q) + then show ?case + proof(induct rule: rel_trancl.induct) + case(r_into_rel_trancl \ P Q) + then have "(\, Q, P) \ X" by(rule rSym) + then show ?case + by force + next + case(rel_trancl_into_rel_trancl \ P Q R) + from \(\, Q, R) \ X\ have "(\, R, Q) \ X" by(rule rSym) + then have "(\, R, Q) \ X\<^sup>\" by force + then show ?case using \(\, Q, P) \ X\<^sup>\\ + by(rule rel_trancl_transitive) + qed + qed +qed + +lemma weakTransitiveStarCoinduct[case_names cStatEq cSim cExt cSym, case_conclusion bisim step, consumes 2]: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and Q :: "('a, 'b, 'c) psi" + and X :: "('b \ ('a, 'b, 'c) psi \ ('a, 'b, 'c) psi) set" + +assumes p: "(\, P, Q) \ X" + and Eqvt: "eqvt X" + and rStatEq: "\\ P Q. (\, P, Q) \ X \ insertAssertion (extractFrame P) \ \\<^sub>F insertAssertion (extractFrame Q) \" + and rSim: "\\ P Q. (\, P, Q) \ X \ \ \ P \[({(\, P, Q) | \ P P' Q' Q. \ \ P \ P' \ + (\, P', Q') \ X \ + \ \ Q' \ Q})\<^sup>\] Q" + and rExt: "\\ P Q \'. (\, P, Q) \ X \ (\ \ \', P, Q) \ X" + and rSym: "\\ P Q. (\, P, Q) \ X \ (\, Q, P) \ X" + +shows "\ \ P \ Q" +proof - + let ?X = "{(\, P, Q) | \ P P' Q' Q. \ \ P \ P' \ (\, P', Q') \ X \ \ \ Q' \ Q}" + from p have "(\, P, Q) \ ?X" + by(blast intro: bisimReflexive) + moreover from \eqvt X\ have "eqvt ?X" by auto + ultimately show ?thesis + proof(coinduct rule: bisimStarWeakCoinduct) + case(cStatEq \ P Q) + then show ?case + by(blast dest: rStatEq bisimE FrameStatEqTrans) + next + case(cSim \ P Q) + then obtain P' Q' where "\ \ P \ P'" and "(\, P', Q') \ X" and "\ \ Q' \ Q" + by blast + then have "\ \ P \[bisim] P'" and "\ \ Q' \[bisim] Q" + by(auto dest: bisimE) + { + fix \ P Q + and Q'::"('a,'b,'c) psi" + assume "\ \ P \ Q" + and "(\,Q,Q') \ ?X\<^sup>\" + note \(\,Q,Q') \ ?X\<^sup>\\ \\ \ P \ Q\ + then have "(\,P,Q') \ ?X\<^sup>\" + proof(induct rule: rel_trancl.inducts) + case(r_into_rel_trancl \ Q Q') then show ?case + by(blast dest: bisimTransitive) + next + case(rel_trancl_into_rel_trancl \ P' Q Q') + then show ?case + by(blast dest: rel_trancl_transitive) + qed + } + note tonic = this + { + fix \ P Q + and Q'::"('a,'b,'c) psi" + assume "(\,P,Q) \ ?X\<^sup>\" + and "\ \ Q \ Q'" + then have "(\,P,Q') \ ?X\<^sup>\" + proof(induct arbitrary: Q' rule: rel_trancl.inducts) + case(r_into_rel_trancl \ P Q) then show ?case + by(blast dest: bisimTransitive) + next + case(rel_trancl_into_rel_trancl \ P P' Q) + from \(\, P', Q) \ ?X\ \\ \ Q \ Q'\ have "(\, P', Q') \ ?X" + by(blast dest: bisimTransitive) + then have "(\, P', Q') \ ?X\<^sup>\" + by blast + with \(\, P, P') \ ?X\<^sup>\\ + show ?case + by(rule rel_trancl_transitive) + qed + } + note tonic2 = this + from \(\, P', Q') \ X\ have "\ \ P' \[?X\<^sup>\] Q'" + by(rule rSim) + with \\ \ P \[bisim] P'\ have "\ \ P \[?X\<^sup>\] Q'" + apply - + apply(rule transitive) + apply assumption + apply assumption + apply(rule rel_trancl_eqvt) + apply(rule \eqvt ?X\) + by(blast intro: tonic) + with \\ \ Q' \[bisim] Q\ show ?case + apply - + apply(rule transitive) + apply assumption + apply assumption + apply(rule rel_trancl_eqvt) + apply(rule \eqvt ?X\) + by(blast intro: tonic2) + next + case(cExt \ P Q \') + then show ?case by(blast dest: bisimE intro: rExt) + next + case(cSym \ P Q) + then show ?case by(blast dest: bisimE intro: rSym) + qed +qed + +lemma weakTransitiveStarCoinduct'[case_names cStatEq cSim cExt cSym, case_conclusion bisim step, consumes 2]: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and Q :: "('a, 'b, 'c) psi" + and X :: "('b \ ('a, 'b, 'c) psi \ ('a, 'b, 'c) psi) set" + +assumes p: "(\, P, Q) \ X" + and Eqvt: "eqvt X" + and rStatEq: "\\ P Q. (\, P, Q) \ X \ insertAssertion (extractFrame P) \ \\<^sub>F insertAssertion (extractFrame Q) \" + and rSim: "\\ P Q. (\, P, Q) \ X \ \ \ P \[({(\, P, Q) | \ P P' Q' Q. \ \ P \ P' \ + (\, P', Q') \ X \ + \ \ Q' \ Q})\<^sup>\] Q" + and rExt: "\\ P Q \'. (\, P, Q) \ X \ (\ \ \', P, Q) \ X" + and rSym: "\\ P Q. (\, P, Q) \ X \ + (\, Q, P) \ {(\, P, Q) | \ P P' Q' Q. \ \ P \ P' \ (\, P', Q') \ X \ \ \ Q' \ Q}" + +shows "\ \ P \ Q" +proof - + let ?X = "{(\, P, Q) | \ P P' Q' Q. \ \ P \ P' \ (\, P', Q') \ X \ \ \ Q' \ Q}" + from p have "(\, P, Q) \ ?X" + by(blast intro: bisimReflexive) + moreover from \eqvt X\ have "eqvt ?X" by auto + ultimately show ?thesis + proof(coinduct rule: bisimStarWeakCoinduct) + case(cStatEq \ P Q) + then show ?case + by(blast dest: rStatEq bisimE FrameStatEqTrans) + next + case(cSim \ P Q) + then obtain P' Q' where "\ \ P \ P'" and "(\, P', Q') \ X" and "\ \ Q' \ Q" + by blast + then have "\ \ P \[bisim] P'" and "\ \ Q' \[bisim] Q" + by(auto dest: bisimE) + { + fix \ P Q + and Q'::"('a,'b,'c) psi" + assume "\ \ P \ Q" + and "(\,Q,Q') \ ?X\<^sup>\" + note \(\,Q,Q') \ ?X\<^sup>\\ \\ \ P \ Q\ + then have "(\,P,Q') \ ?X\<^sup>\" + proof(induct rule: rel_trancl.inducts) + case(r_into_rel_trancl \ Q Q') then show ?case + by(blast dest: bisimTransitive) + next + case(rel_trancl_into_rel_trancl \ P' Q Q') + then show ?case + by(blast dest: rel_trancl_transitive) + qed + } + note tonic = this + { + fix \ P Q + and Q'::"('a,'b,'c) psi" + assume "(\,P,Q) \ ?X\<^sup>\" + and "\ \ Q \ Q'" + then have "(\,P,Q') \ ?X\<^sup>\" + proof(induct arbitrary: Q' rule: rel_trancl.inducts) + case(r_into_rel_trancl \ P Q) then show ?case + by(blast dest: bisimTransitive) + next + case(rel_trancl_into_rel_trancl \ P P' Q) + from \(\, P', Q) \ ?X\ \\ \ Q \ Q'\ have "(\, P', Q') \ ?X" + by(blast dest: bisimTransitive) + then have "(\, P', Q') \ ?X\<^sup>\" + by blast + with \(\, P, P') \ ?X\<^sup>\\ + show ?case + by(rule rel_trancl_transitive) + qed + } + note tonic2 = this + from \(\, P', Q') \ X\ have "\ \ P' \[?X\<^sup>\] Q'" + by(rule rSim) + with \\ \ P \[bisim] P'\ have "\ \ P \[?X\<^sup>\] Q'" + apply - + apply(rule transitive) + apply assumption + apply assumption + apply(rule rel_trancl_eqvt) + apply(rule \eqvt ?X\) + by(blast intro: tonic) + with \\ \ Q' \[bisim] Q\ show ?case + apply - + apply(rule transitive) + apply assumption + apply assumption + apply(rule rel_trancl_eqvt) + apply(rule \eqvt ?X\) + by(blast intro: tonic2) + next + case(cExt \ P Q \') + then show ?case by(blast dest: bisimE intro: rExt) + next + case(cSym \ P Q) + then show ?case + apply clarsimp + apply(drule rSym) + apply clarsimp + by(metis bisimTransitive bisimE(4)) + qed +qed + +lemma transitiveStarCoinduct[case_names cStatEq cSim cExt cSym, case_conclusion bisim step, consumes 2]: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and Q :: "('a, 'b, 'c) psi" + and X :: "('b \ ('a, 'b, 'c) psi \ ('a, 'b, 'c) psi) set" + +assumes p: "(\, P, Q) \ X" + and Eqvt: "eqvt X" + and rStatEq: "\\' R S. (\', R, S) \ X \ insertAssertion (extractFrame R) \' \\<^sub>F insertAssertion (extractFrame S) \'" + and rSim: "\\' R S. (\', R, S) \ X \ \' \ R \[({(\', R, S) | \' R R' S' S. \' \ R \ R' \ + ((\', R', S') \ X \ \' \ R' \ S') \ + \' \ S' \ S})\<^sup>\] S" + and rExt: "\\' R S \''. (\', R, S) \ X \ (\' \ \'', R, S) \ X \ \' \ \'' \ R \ S" + and rSym: "\\' R S. (\', R, S) \ X \ (\', S, R) \ X \ \' \ S \ R" + + +shows "\ \ P \ Q" +proof - + from p have "(\, P, Q) \ (X \ bisim)" + by blast + moreover from \eqvt X\ bisimEqvt have "eqvt (X \ bisim)" + by auto + ultimately show ?thesis + proof(coinduct rule: weakTransitiveStarCoinduct') + case(cStatEq \ P Q) + then show ?case + by(blast intro: rStatEq dest: bisimE) + next + case(cSim \ P Q) + then show ?case + apply clarsimp + apply (erule disjE) + apply(blast intro: rSim) + apply(drule bisimE(2)) + apply(rule monotonic, simp) + by(force intro: bisimReflexive) + next + case(cExt \ P Q \') + then show ?case + by(blast dest: bisimE rExt) + next + case(cSym \ P Q) + then show ?case by(blast dest: bisimE rSym intro: bisimReflexive) + qed +qed + +end + +end diff --git a/thys/Broadcast_Psi/Broadcast_Chain.thy b/thys/Broadcast_Psi/Broadcast_Chain.thy new file mode 100644 --- /dev/null +++ b/thys/Broadcast_Psi/Broadcast_Chain.thy @@ -0,0 +1,249 @@ +theory Broadcast_Chain + imports "Psi_Calculi.Chain" +begin + +lemma pair_perm_fresh_contr: + fixes a::'a and b::'a + assumes + at: "at TYPE('a)" + and + prems: "b \ pi" "(a, b) \ set pi" + shows False + using prems + by(induct pi) + (auto simp add: supp_list_cons fresh_list_nil supp_prod at_supp[OF at] + fresh_list_cons fresh_prod at_fresh[OF at]) + +lemma pair_perm_fresh_contr': + fixes a::'a and b::'a + assumes + at: "at TYPE('a)" + and + prems: "a \ pi" "(a, b) \ set pi" + shows False + using prems + by(induct pi) + (auto simp add: supp_list_cons fresh_list_nil supp_prod at_supp[OF at] + fresh_list_cons fresh_prod at_fresh[OF at]) + +lemma list_set_supp: + fixes l :: "('d::fs_name) list" + shows "supp (set l) = (supp l :: name set)" +proof(induct l) + case Nil then show ?case + by (simp add: supp_list_nil supp_set_empty) +next + case (Cons x xs) then show ?case + using at_name_inst fs_name_inst pt_list_set_supp pt_name_inst by blast +qed + +lemma name_set_supp: + assumes "finite a" + shows "supp a = (a::name set)" + using assms + by(rule at_fin_set_supp[OF at_name_inst]) + +lemma supp_idem: + fixes l :: "('d::fs_name)" + shows "supp((supp l)::name set) = (supp(l)::name set)" +proof - + have f: "finite((supp l)::name set)" + by finite_guess + show ?thesis + by(rule name_set_supp[OF f]) +qed + +lemma fresh_supp: + fixes a :: name + and X :: "('d::fs_name)" + shows "a \ ((supp X)::name set) = a \ X" + by(simp add: fresh_def supp_idem) + +lemma fresh_chain_supp: + fixes A :: "name list" + and X :: "('d::fs_name)" + shows "A \* ((supp X)::name set) = A \* X" + unfolding fresh_star_def + by(simp add: fresh_supp) + +lemma fresh_chain_fin_union: + fixes X::"('d::fs_name set)" + and Y::"('d::fs_name set)" + and A::"name list" + assumes f1: "finite X" + and f2: "finite Y" + shows "A\*(X\Y) = (A\*X \ A\*Y)" + unfolding fresh_star_def + apply(subst fresh_fin_union[OF pt_name_inst, OF at_name_inst, OF fs_name_inst, OF f1, OF f2]) + by blast + +lemma fresh_subset: + fixes S :: "name set" + and S' :: "name set" + and a :: name + assumes "a \ S" + and "S' \ S" + and "finite S" + +shows "a \ S'" +proof - + have "finite S'" using \S' \ S\ \finite S\ + by(rule Finite_Set.finite_subset) + with assms show ?thesis + by(auto simp add: fresh_def supp_subset Chain.name_list_supp name_set_supp) +qed + +lemma fresh_subset': + fixes S :: "'d::fs_name set" + and S' :: "'d::fs_name set" + and a :: name + assumes "a \ S" + and "S' \ S" + and "finite S" + +shows "a \ S'" +proof - + have "finite S'" using \S' \ S\ \finite S\ + by(rule Finite_Set.finite_subset) + have "supp S' \ ((supp S)::name set)" + apply(rule Chain.supp_subset) + by fact+ + with assms show ?thesis + unfolding fresh_def + by auto +qed + +lemma fresh_star_subset': + fixes S :: "'d::fs_name set" + and S' :: "'d::fs_name set" + and A :: "name list" + assumes "A \* S" + and "S' \ S" + and "finite S" + +shows "A \* S'" + using assms + unfolding fresh_star_def + by(auto simp add: fresh_subset') + +lemma fresh_star_subset: + fixes S :: "name set" + and S' :: "name set" + and A :: "name list" + assumes "A \* S" + and "S' \ S" + and "finite S" + +shows "A \* S'" + using assms + unfolding fresh_star_def + by(auto simp add: fresh_subset) + +lemma times_set_fresh: + fixes a :: name + and S :: "name list" + and S' :: "name list" + assumes "a \ set S" + and "a \ set S'" + shows "a \ set S \ set S'" + using assms +proof(cases S) + case Nil then show ?thesis by(simp add: fresh_set_empty) +next + case (Cons s Svec) then show ?thesis + proof(cases S') + case Nil then show ?thesis by(simp add: fresh_set_empty) + next + case (Cons s' S') then show ?thesis + using \S = s # Svec\ assms + apply(subst fresh_cart_prod[OF pt_name_inst, OF pt_name_inst, OF fs_name_inst, OF fs_name_inst, OF at_name_inst]) + by(simp+) + qed +qed + +lemma times_set_fresh_star: + fixes A :: "name list" + and S :: "name list" + and S' :: "name list" + assumes "A \* set S" + and "A \* set S'" + shows "A \* (set S \ set S')" + using assms + unfolding fresh_star_def +proof(induct A) + case Nil show ?case by simp +next + case(Cons a A) + have "a \ set S" and "a \ set S'" using Cons by simp+ + then have "a \ (set S \ set S')" by(rule times_set_fresh) + have "\ x\set A. (x \ set S)" and "\ x\set A. (x \ set S')" using Cons + by simp+ + then have "\x\set A. x \ set S \ set S'" using Cons by simp + then show ?case using \a \ (set S \ set S')\ + by simp +qed + +lemma supp_list_set: + fixes M::"'d::fs_name list" + shows "(supp M) = ((supp(set M))::name set)" +proof(induct M) + case Nil then show ?case by(simp add: supp_set_empty supp_list_nil) +next + case (Cons m M) + have lhs: "(supp (m # M)::name set) = supp m \ supp M" by(simp add: supp_list_cons) + have rhs: "(supp (set (m # M))::name set) = supp m \ supp M" + proof - + have "supp (set (m # M)) = (supp (set [m] \ set M)::name set)" + by simp + moreover have "\ = supp (set [m]) \ supp(set M)" + apply(rule supp_fin_union[OF pt_name_inst, OF at_name_inst, OF fs_name_inst]) + by simp+ + moreover have "\ = supp m \ supp M" + using calculation(1) calculation(2) lhs list_set_supp by blast + ultimately show ?thesis + by simp + qed + show ?case + unfolding lhs rhs + by(rule refl) +qed + +lemma fresh_list_set: + fixes M::"'d::fs_name list" + and A::"name list" + shows "A \* set M = A \* M" + unfolding fresh_star_def fresh_def supp_list_set + by(rule refl) + +lemma permSupp: + fixes \ :: "name prm" + and \' :: "'d::fs_name" + +shows "(supp(\ \ \')::name set) \ ((supp \) \ (supp \'))" +proof - + { + fix x::name + let ?P = "\y. ([(x, y)] \ \) \ ([(x, y)] \ \') \ \ \ \'" + let ?Q = "\y \. ([(x, y)] \ \) \ \" + assume "finite {y. ?Q y \'}" + moreover assume "finite {y. ?Q y \}" + and "infinite {b. [(x, b)] \ \ \ \' \ \ \ \'}" + from \infinite {b. [(x, b)] \ \ \ \' \ \ \ \'}\ have "infinite {y. ?P(y)}" + by(subst cp1[symmetric, OF cp_pt_inst, OF pt_name_inst, OF at_name_inst]) + then have "infinite({y. ?P(y)} - {y. ?Q y \})" + using \finite {y. ?Q y \'}\ \finite {y. ?Q y \}\ + by - (rule Diff_infinite_finite) + ultimately have "infinite(({y. ?P(y)} - {y. ?Q y \}) - {y. ?Q y \'})" + by(rule Diff_infinite_finite) + then have "infinite({y. ?P(y) \ \(?Q y \) \ \ (?Q y \')})" by(simp add: set_diff_eq) + moreover have "{y. ?P(y) \ \(?Q y \) \ \ (?Q y \')} = {}" + by auto + ultimately have "infinite {}" + by - (drule Infinite_cong, auto) + then have False by simp + } + from this show ?thesis + by(force simp add: supp_def) +qed + +end diff --git a/thys/Broadcast_Psi/Broadcast_Frame.thy b/thys/Broadcast_Psi/Broadcast_Frame.thy new file mode 100644 --- /dev/null +++ b/thys/Broadcast_Psi/Broadcast_Frame.thy @@ -0,0 +1,157 @@ +theory Broadcast_Frame + imports "Psi_Calculi.Frame" +begin + +locale assertionAux = Frame.assertionAux SCompose SImp SBottom SChanEq + for SCompose :: "'b::fs_name \ 'b \ 'b" (infixr "\" 80) + and SImp :: "'b \ 'c::fs_name \ bool" ("_ \ _" [70, 70] 70) + and SBottom :: 'b ("\" 90) + and SChanEq :: "('a::fs_name \ 'a \ 'c)" ("_ \ _" [80, 80] 80) + + + fixes SOutCon :: "'a::fs_name \ 'a \ 'c" ("_ \ _" [80, 80] 80) + and SInCon :: "'a::fs_name \ 'a \ 'c" ("_ \ _" [80, 80] 80) + +assumes statEqvt'''[eqvt]: "\p::name prm. p \ (M \ N) = (p \ M) \ (p \ N)" + and statEqvt''''[eqvt]: "\p::name prm. p \ (M \ N) = (p \ M) \ (p \ N)" + +begin + +lemma chanInConSupp: + fixes M :: 'a + and N :: 'a + +shows "(supp(M \ N)::name set) \ ((supp M) \ (supp N))" +proof - + { + fix x::name + let ?P = "\y. ([(x, y)] \ M) \ [(x, y)] \ N \ M \ N" + let ?Q = "\y M. ([(x, y)] \ M) \ M" + assume "finite {y. ?Q y N}" + moreover assume "finite {y. ?Q y M}" and "infinite {y. ?P(y)}" + then have "infinite({y. ?P(y)} - {y. ?Q y M})" by(rule Diff_infinite_finite) + ultimately have "infinite(({y. ?P(y)} - {y. ?Q y M}) - {y. ?Q y N})" by(rule Diff_infinite_finite) + then have "infinite({y. ?P(y) \ \(?Q y M) \ \ (?Q y N)})" by(simp add: set_diff_eq) + moreover have "{y. ?P(y) \ \(?Q y M) \ \ (?Q y N)} = {}" by auto + ultimately have "infinite {}" by(blast dest: Infinite_cong) + then have False by simp + } + then show ?thesis by(auto simp add: eqvts supp_def) +qed + +lemma chanOutConSupp: + fixes M :: 'a + and N :: 'a + +shows "(supp(M \ N)::name set) \ ((supp M) \ (supp N))" +proof - + { + fix x::name + let ?P = "\y. ([(x, y)] \ M) \ [(x, y)] \ N \ M \ N" + let ?Q = "\y M. ([(x, y)] \ M) \ M" + assume "finite {y. ?Q y N}" + moreover assume "finite {y. ?Q y M}" and "infinite {y. ?P(y)}" + then have "infinite({y. ?P(y)} - {y. ?Q y M})" by(rule Diff_infinite_finite) + ultimately have "infinite(({y. ?P(y)} - {y. ?Q y M}) - {y. ?Q y N})" by(rule Diff_infinite_finite) + then have "infinite({y. ?P(y) \ \(?Q y M) \ \ (?Q y N)})" by(simp add: set_diff_eq) + moreover have "{y. ?P(y) \ \(?Q y M) \ \ (?Q y N)} = {}" by auto + ultimately have "infinite {}" by(blast dest: Infinite_cong) + then have False by simp + } + then show ?thesis by (auto simp add: eqvts supp_def) +qed + +lemma freshInCon[intro]: + fixes x :: name + and M :: 'a + and N :: 'a + +assumes "x \ M" + and "x \ N" + +shows "x \ M \ N" + using assms chanInConSupp + by(auto simp add: fresh_def) + +lemma freshInConChain[intro]: + fixes xvec :: "name list" + and Xs :: "name set" + and M :: 'a + and N :: 'a + +shows "\xvec \* M; xvec \* N\ \ xvec \* (M \ N)" + and "\Xs \* M; Xs \* N\ \ Xs \* (M \ N)" + by(auto simp add: fresh_star_def) + +lemma freshOutCon[intro]: + fixes x :: name + and M :: 'a + and N :: 'a + +assumes "x \ M" + and "x \ N" + +shows "x \ M \ N" + using assms chanOutConSupp + by(auto simp add: fresh_def) + +lemma freshOutConChain[intro]: + fixes xvec :: "name list" + and Xs :: "name set" + and M :: 'a + and N :: 'a + +shows "\xvec \* M; xvec \* N\ \ xvec \* (M \ N)" + and "\Xs \* M; Xs \* N\ \ Xs \* (M \ N)" + by(auto simp add: fresh_star_def) + +lemma chanOutConClosed: + fixes \ :: 'b + and M :: 'a + and N :: 'a + and p :: "name prm" + +assumes "\ \ M \ N" + +shows "(p \ \) \ (p \ M) \ (p \ N)" +proof - + from \\ \ M \ N\ have "(p \ \) \ p \ (M \ N)" + by(rule statClosed) + then show ?thesis by(auto simp add: eqvts) +qed + +lemma chanInConClosed: + fixes \ :: 'b + and M :: 'a + and N :: 'a + and p :: "name prm" + +assumes "\ \ M \ N" + +shows "(p \ \) \ (p \ M) \ (p \ N)" +proof - + from \\ \ M \ N\ have "(p \ \) \ p \ (M \ N)" + by(rule statClosed) + then show ?thesis by(auto simp add: eqvts) +qed + +end + +locale assertion = assertionAux SCompose SImp SBottom SChanEq SOutCon SInCon + assertion SCompose SImp SBottom SChanEq + for SCompose :: "'b::fs_name \ 'b \ 'b" + and SImp :: "'b \ 'c::fs_name \ bool" + and SBottom :: 'b + and SChanEq :: "'a::fs_name \ 'a \ 'c" + and SOutCon :: "'a::fs_name \ 'a \ 'c" + and SInCon :: "'a::fs_name \ 'a \ 'c" + + + assumes chanOutConSupp: "SImp \ (SOutCon M N) \ (((supp N)::name set) \ ((supp M)::name set))" + and chanInConSupp: "SImp \ (SInCon N M) \ (((supp N)::name set) \ ((supp M)::name set))" + +begin + +notation SOutCon ("_ \ _" [90, 90] 90) +notation SInCon ("_ \ _" [90, 90] 90) + +end + +end diff --git a/thys/Broadcast_Psi/Broadcast_Thms.thy b/thys/Broadcast_Psi/Broadcast_Thms.thy new file mode 100644 --- /dev/null +++ b/thys/Broadcast_Psi/Broadcast_Thms.thy @@ -0,0 +1,202 @@ +theory Broadcast_Thms + imports Broadcast_Chain Broadcast_Frame Semantics Simulation Bisimulation Sim_Pres + Sim_Struct_Cong Bisim_Pres Bisim_Struct_Cong Bisim_Subst +begin + +context env +begin + + subsection \Syntax\ + + subsubsection \Psi calculus agents~-- Definition~2\ + + text \ + @{term M}, @{term N} range over message terms. @{term P}, @{term Q} range over processes. + @{term C} ranges over cases. + + \<^item> Output: @{term "M\N\.P"} + \<^item> Input: @{term "M\\*xvec N\.P"} + \<^item> Case: @{term "Case C"} + \<^item> Par: @{term "(P \ Q)"} + \<^item> Res: @{term "\\x\(P::('a,'b,'c) psi)"} + \<^item> Assert: @{term "\\\"} + \<^item> Bang: @{term "!P"} + \ + + text \ + \<^item> Cases: @{term "\ \ \ P C"} + \ + + subsubsection \Parameters~-- Definition~1\ + + text \ + \<^item> Channel equivalence: @{term "M \ N"} + \<^item> Composition: @{term "\\<^sub>P \ \\<^sub>Q"} + \<^item> Unit: @{term "\"} + \<^item> Entailment: @{term "\ \ \"} + \ + + subsubsection \Extra predicates for broadcast~-- Definition~5\ + + text \ + \<^item> Output connectivity: @{term "M \ N"} + \<^item> Input connectivity: @{term "M \ N"} + \ + + subsubsection \Transitions~-- Definition~3\ + + text \ + \<^item> @{term "\ \ P \\ P'"} + \ + + subsubsection \Actions (\texorpdfstring{@{term \}}{alpha})~-- Definition~7\ + + text \ + \<^item> Input: @{term "M\N\"} + \<^item> Output: @{term "M\\*xvec\\N\"} + \<^item> Broadcast input: @{term "\M\N\"} + \<^item> Broadcast output: @{term "\M\\*xvec\\N\"} + \<^item> Silent action: @{term "\"} + \ + + subsection \Semantics\ + + subsubsection \Basic Psi semantics~-- Table~1\ + + text \ + \<^item> Theorem {\it Input}: + @{thm [display] Input[no_vars]} + \<^item> Theorem {\it Output}: + @{thm [display] Output[no_vars]} + \<^item> Theorem {\it Case}: + @{thm [display] Case[no_vars]} + \<^item> Theorems {\it Par1} and {\it Par2}: + @{thm [display] Par1[no_vars]} + @{thm [display] Par2[no_vars]} + \<^item> Theorems {\it Comm1} and {\it Comm2}: + @{thm [display] Comm1[no_vars]} + @{thm [display] Comm2[no_vars]} + \<^item> Theorem {\it Open}: + @{thm [display] Open[no_vars]} + \<^item> Theorem {\it Scope}: + @{thm [display] Scope[no_vars]} + \<^item> Theorem {\it Bang}: + @{thm [display] Bang[no_vars]} + \ + + subsubsection \Broadcast rules~-- Table~2\ + + text \ + \<^item> Theorem {\it BrInput}: + @{thm [display] BrInput[no_vars]} + \<^item> Theorem {\it BrOutput}: + @{thm [display] BrOutput[no_vars]} + \<^item> Theorem {\it BrMerge}: + @{thm [display] BrMerge[no_vars]} + \<^item> Theorems {\it BrComm1} and {\it BrComm2}: + @{thm [display] BrComm1[no_vars]} + @{thm [display] BrComm2[no_vars]} + \<^item> Theorem {\it BrClose}: + @{thm [display] BrClose[no_vars]} + \<^item> Theorem {\it BrOpen}: + @{thm [display] BrOpen[no_vars]} + \ + + subsubsection \Requirements for broadcast~-- Definition~6\ + + text \ + \<^item> Theorem {\it chanOutConSupp}: + @{thm [display] chanOutConSupp[no_vars]} + \<^item> Theorem {\it chanInConSupp}: + @{thm [display] chanInConSupp[no_vars]} + \ + + subsubsection \Strong bisimulation~-- Definition~4\ + + text \ + \<^item> Theorem {\it bisim.step}: + @{thm [display] bisim.step[no_vars]} + \ + + subsection \Theorems\ + + subsubsection \Congruence properties of strong bisimulation~-- Theorem~8\ + + text \ + \<^item> Theorem {\it bisimOutputPres}: + @{thm [display] bisimOutputPres[no_vars]} + \<^item> Theorem {\it bisimInputPres}: + @{thm [display] bisimInputPres[no_vars]} + \<^item> Theorem {\it bisimCasePres}: + @{thm [display] bisimCasePres[no_vars]} + \<^item> Theorems {\it bisimParPres} and {\it bisimParPresSym}: + @{thm [display] bisimParPres[no_vars]} + @{thm [display] bisimParPresSym[no_vars]} + \<^item> Theorem {\it bisimResPres}: + @{thm [display] bisimResPres[no_vars]} + \<^item> Theorem {\it bisimBangPres}: + @{thm [display] bisimBangPres[no_vars]} + \ + + subsubsection \Strong congruence, bisimulation closed under substitution~-- Definition~9\ + + text \ + \<^item> Theorem {\it closeSubst\_def}: + @{thm [display] closeSubst_def[no_vars]} + \<^item> @{term "\ \ P \\<^sub>s Q"} + \ + + subsubsection \Strong congruence is a process congruence for all \texorpdfstring{@{term \}}{Psi}~-- Theorem~10\ + + text \ + \<^item> Theorem {\it bisimSubstOutputPres}: + @{thm [display] bisimSubstOutputPres[no_vars]} + \<^item> Theorem {\it bisimSubstInputPres}: + @{thm [display] bisimSubstInputPres[no_vars]} + \<^item> Theorem {\it bisimSubstCasePres}: + @{thm [display] bisimSubstCasePres[no_vars]} + \<^item> Theorem {\it bisimSubstParPres}: + @{thm [display] bisimSubstParPres[no_vars]} + \<^item> Theorem {\it bisimSubstResPres}: + @{thm [display] bisimSubstResPres[no_vars]} + \<^item> Theorem {\it bisimSubstBangPres}: + @{thm [display] bisimSubstBangPres[no_vars]} + \ + + subsubsection \Structural equivalence~-- Theorem~11\ + + text \ + \<^item> Theorem {\it bisimCasePushRes}: + @{thm [display] bisimCasePushRes[no_vars]} + \<^item> Theorem {\it bisimInputPushRes}: + @{thm [display] bisimInputPushRes[no_vars]} + \<^item> Theorem {\it bisimOutputPushRes}: + @{thm [display] bisimOutputPushRes[no_vars]} + \<^item> Theorem {\it bisimParAssoc}: + @{thm [display] bisimParAssoc[no_vars]} + \<^item> Theorem {\it bisimParComm}: + @{thm [display] bisimParComm[no_vars]} + \<^item> Theorem {\it bisimResNil}: + @{thm [display] bisimResNil[no_vars]} + \<^item> Theorem {\it bisimScopeExt}: + @{thm [display] bisimScopeExt[no_vars]} + \<^item> Theorem {\it bisimResComm}: + @{thm [display] bisimResComm[no_vars]} + \<^item> Theorem {\it bangExt}: + @{thm [display] bangExt[no_vars]} + \<^item> Theorem {\it bisimParNil}: + @{thm [display] bisimParNil[no_vars]} + \ + + subsubsection \Support of processes in broadcast transitions~-- Lemma~14\ + + text \ + \<^item> Theorem {\it brInputTermSupp}: + @{thm [display] brInputTermSupp[no_vars]} + \<^item> Theorem {\it brOutputTermSupp}: + @{thm [display] brOutputTermSupp[no_vars]} + \ + +end + +end diff --git a/thys/Broadcast_Psi/ROOT b/thys/Broadcast_Psi/ROOT new file mode 100644 --- /dev/null +++ b/thys/Broadcast_Psi/ROOT @@ -0,0 +1,11 @@ +chapter AFP + +session Broadcast_Psi (AFP) = "HOL-Nominal" + + options [timeout = 1800] + sessions + "Psi_Calculi" + theories + Broadcast_Thms + document_files + "root.bib" + "root.tex" diff --git a/thys/Broadcast_Psi/Semantics.thy b/thys/Broadcast_Psi/Semantics.thy new file mode 100644 --- /dev/null +++ b/thys/Broadcast_Psi/Semantics.thy @@ -0,0 +1,17137 @@ +theory Semantics + imports Broadcast_Chain Broadcast_Frame +begin + +text \This file is a (heavily modified) variant of the theory {\it Psi\_Calculi.Semantics} +from~\cite{DBLP:journals/afp/Bengtson12}. The nominal datatypes {\it ('a,'b,'c) residual} and +{\it 'a action} have been extended with constructors for broadcast input and output. This leads to +a different semantics.\ + +nominal_datatype ('a, 'b, 'c) boundOutput = + BOut "'a::fs_name" "('a, 'b::fs_name, 'c::fs_name) psi" ("_ \'' _" [110, 110] 110) + | BStep "\name\ ('a, 'b, 'c) boundOutput" ("\\_\_" [110, 110] 110) + +primrec BOresChain :: "name list \ ('a::fs_name, 'b::fs_name, 'c::fs_name) boundOutput \ + ('a, 'b, 'c) boundOutput" + where + Base: "BOresChain [] B = B" + | Step: "BOresChain (x#xs) B = \\x\(BOresChain xs B)" + +abbreviation + BOresChainJudge ("\\*_\_" [80, 80] 80) where "\\*xvec\B \ BOresChain xvec B" + +lemma BOresChainEqvt[eqvt]: + fixes perm :: "name prm" + and lst :: "name list" + and B :: "('a::fs_name, 'b::fs_name, 'c::fs_name) boundOutput" + +shows "perm \ (\\*xvec\B) = \\*(perm \ xvec)\(perm \ B)" + by(induct xvec) auto + +lemma BOresChainSimps[simp]: + fixes xvec :: "name list" + and N :: "'a::fs_name" + and P :: "('a, 'b::fs_name, 'c::fs_name) psi" + and N' :: 'a + and P' :: "('a, 'b, 'c) psi" + and B :: "('a, 'b, 'c) boundOutput" + and B' :: "('a, 'b, 'c) boundOutput" + +shows "(\\*xvec\N \' P = N' \' P') = (xvec = [] \ N = N' \ P = P')" + and "(N' \' P' = \\*xvec\N \' P) = (xvec = [] \ N = N' \ P = P')" + and "(N' \' P' = N \' P) = (N = N' \ P = P')" + and "(\\*xvec\B = \\*xvec\B') = (B = B')" + by(induct xvec) (auto simp add: boundOutput.inject alpha) + +lemma outputFresh[simp]: + fixes Xs :: "name set" + and xvec :: "name list" + and N :: "'a::fs_name" + and P :: "('a, 'b::fs_name, 'c::fs_name) psi" + +shows "(Xs \* (N \' P)) = ((Xs \* N) \ (Xs \* P))" + and "(xvec \* (N \' P)) = ((xvec \* N) \ (xvec \* P))" + by(auto simp add: fresh_star_def) + +lemma boundOutputFresh: + fixes x :: name + and xvec :: "name list" + and B :: "('a::fs_name, 'b::fs_name, 'c::fs_name) boundOutput" + +shows "(x \ (\\*xvec\B)) = (x \ set xvec \ x \ B)" + by (induct xvec) (simp_all add: abs_fresh) + +lemma boundOutputFreshSet: + fixes Xs :: "name set" + and xvec :: "name list" + and B :: "('a::fs_name, 'b::fs_name, 'c::fs_name) boundOutput" + and yvec :: "name list" + and x :: name + +shows "Xs \* (\\*xvec\B) = (\x\Xs. x \ set xvec \ x \ B)" + and "yvec \* (\\*xvec\B) = (\x\(set yvec). x \ set xvec \ x \ B)" + and "Xs \* (\\x\B) = Xs \* [x].B" + and "xvec \* (\\x\B) = xvec \* [x].B" + by(simp add: fresh_star_def boundOutputFresh)+ + +lemma BOresChainSupp: + fixes xvec :: "name list" + and B :: "('a::fs_name, 'b::fs_name, 'c::fs_name) boundOutput" + +shows "(supp(\\*xvec\B)::name set) = (supp B) - (supp xvec)" + by(induct xvec) + (auto simp add: boundOutput.supp supp_list_nil supp_list_cons abs_supp supp_atm) + +lemma boundOutputFreshSimps[simp]: + fixes Xs :: "name set" + and xvec :: "name list" + and B :: "('a::fs_name, 'b::fs_name, 'c::fs_name) boundOutput" + and yvec :: "name list" + and x :: name + +shows "Xs \* xvec \ (Xs \* (\\*xvec\B)) = (Xs \* B)" + and "yvec \* xvec \ yvec \* (\\*xvec\B) = yvec \* B" + and "xvec \* (\\*xvec\B)" + and "x \ xvec \ x \ \\*xvec\B = x \ B" + apply(simp add: boundOutputFreshSet) apply(force simp add: fresh_star_def name_list_supp fresh_def) + apply(simp add: boundOutputFreshSet) apply(force simp add: fresh_star_def name_list_supp fresh_def) + apply(simp add: boundOutputFreshSet) + by(simp add: BOresChainSupp fresh_def) + +lemma boundOutputChainAlpha: + fixes p :: "name prm" + and xvec :: "name list" + and B :: "('a::fs_name, 'b::fs_name, 'c::fs_name) boundOutput" + and yvec :: "name list" + +assumes xvecFreshB: "(p \ xvec) \* B" + and S: "set p \ set xvec \ set (p \ xvec)" + and "(set xvec) \ (set yvec)" + +shows "(\\*yvec\B) = (\\*(p \ yvec)\(p \ B))" +proof - + note pt_name_inst at_name_inst S + moreover from \(set xvec) \ (set yvec)\ have "set xvec \* (\\*yvec\B)" + by(force simp add: boundOutputFreshSet) + moreover from xvecFreshB \(set xvec) \ (set yvec)\ have "set (p \ xvec) \* (\\*yvec\B)" + by (simp add: boundOutputFreshSet) (simp add: fresh_star_def) + ultimately have "(\\*yvec\B) = p \ (\\*yvec\B)" + by (rule pt_freshs_freshs [symmetric]) + then show ?thesis by(simp add: eqvts) +qed + +lemma boundOutputChainAlpha': + fixes p :: "name prm" + and xvec :: "name list" + and B :: "('a::fs_name, 'b::fs_name, 'c::fs_name) boundOutput" + and yvec :: "name list" + and zvec :: "name list" + +assumes xvecFreshB: "xvec \* B" + and S: "set p \ set xvec \ set yvec" + and "yvec \* (\\*zvec\B)" + +shows "(\\*zvec\B) = (\\*(p \ zvec)\(p \ B))" +proof - + note pt_name_inst at_name_inst S \yvec \* (\\*zvec\B)\ + moreover from xvecFreshB have "set (xvec) \* (\\*zvec\B)" + by (simp add: boundOutputFreshSet) (simp add: fresh_star_def) + ultimately have "(\\*zvec\B) = p \ (\\*zvec\B)" + by(auto intro: pt_freshs_freshs [symmetric]) + then show ?thesis by(simp add: eqvts) +qed + +lemma boundOutputChainAlpha'': + fixes p :: "name prm" + and xvec :: "name list" + and M :: "'a::fs_name" + and P :: "('a::fs_name, 'b::fs_name, 'c::fs_name) psi" + and yvec :: "name list" + +assumes "(p \ xvec) \* M" + and "(p \ xvec) \* P" + and "set p \ set xvec \ set (p \ xvec)" + and "(set xvec) \ (set yvec)" + +shows "(\\*yvec\M \' P) = (\\*(p \ yvec)\(p \ M) \' (p \ P))" + using assms + by(subst boundOutputChainAlpha) auto + +lemma boundOutputChainSwap: + fixes x :: name + and y :: name + and N :: "'a::fs_name" + and P :: "('a, 'b::fs_name, 'c::fs_name) psi" + and xvec :: "name list" + +assumes "y \ N" + and "y \ P" + and "x \ (set xvec)" + +shows "\\*xvec\N \' P = \\*([(x, y)] \ xvec)\([(x ,y)] \ N) \' ([(x, y)] \ P)" +proof(cases "x=y") + assume "x=y" + then show ?thesis by simp +next + assume "x \ y" + with assms show ?thesis + by(auto simp add: calc_atm intro: boundOutputChainAlpha''[where xvec="[x]"]) +qed + +lemma alphaBoundOutput: + fixes x :: name + and y :: name + and B :: "('a::fs_name, 'b::fs_name, 'c::fs_name) boundOutput" + +assumes "y \ B" + +shows "\\x\B = \\y\([(x, y)] \ B)" + using assms + by(auto simp add: boundOutput.inject alpha fresh_left calc_atm) + +lemma boundOutputEqFresh: + fixes B :: "('a::fs_name, 'b::fs_name, 'c::fs_name) boundOutput" + and C :: "('a, 'b, 'c) boundOutput" + and x :: name + and y :: name + +assumes "\\x\B = \\y\C" + and "x \ B" + +shows "y \ C" + using assms + by(auto simp add: boundOutput.inject alpha fresh_left calc_atm) + +lemma boundOutputEqSupp: + fixes B :: "('a::fs_name, 'b::fs_name, 'c::fs_name) boundOutput" + and C :: "('a, 'b, 'c) boundOutput" + and x :: name + and y :: name + +assumes "\\x\B = \\y\C" + and "x \ supp B" + +shows "y \ supp C" + using assms + apply(clarsimp simp add: boundOutput.inject alpha fresh_left calc_atm) + apply(drule pt_set_bij2[where pi="[(x, y)]", OF pt_name_inst, OF at_name_inst]) + by(auto simp add: eqvts calc_atm) + +lemma boundOutputChainEq: + fixes xvec :: "name list" + and B :: "('a::fs_name, 'b::fs_name, 'c::fs_name) boundOutput" + and yvec :: "name list" + and B' :: "('a, 'b, 'c) boundOutput" + +assumes "\\*xvec\B = \\*yvec\B'" + and "xvec \* yvec" + and "length xvec = length yvec" + +shows "\p. (set p) \ (set xvec) \ set (yvec) \ distinctPerm p \ B = p \ B' \ (set (map fst p)) \ (supp B) \ xvec \* B' \ yvec \* B" +proof - + obtain n where "n = length xvec" by auto + with assms show ?thesis + proof(induct n arbitrary: xvec yvec B B') + case(0 xvec yvec B B') + have Eq: "\\*xvec\B = \\*yvec\B'" by fact + from \0 = length xvec\ have "xvec = []" by auto + moreover with \length xvec = length yvec\ have "yvec = []" + by(cases yvec) auto + ultimately show ?case using Eq + by(simp add: boundOutput.inject) + next + case(Suc n xvec yvec B B') + from \Suc n = length xvec\ + obtain x xvec' where "xvec = x#xvec'" and "length xvec' = n" + by(cases xvec) auto + from \\\*xvec\B = \\*yvec\B'\ \xvec = x # xvec'\ \length xvec = length yvec\ + obtain y yvec' where "\\*(x#xvec')\B = \\*(y#yvec')\B'" + and "yvec = y#yvec'" and "length xvec' = length yvec'" + by(cases yvec) auto + then have EQ: "\\x\(\\*xvec'\B) = \\y\(\\*yvec'\B')" + by simp + from \xvec = x#xvec'\ \yvec=y#yvec'\ \xvec \* yvec\ + have "x \ y" and "xvec' \* yvec'" and "x \ yvec'" and "y \ xvec'" + by auto + have IH: "\xvec yvec B B'. \\\*xvec\(B::('a::fs_name, 'b::fs_name, 'c::fs_name) boundOutput) = \\*yvec\B'; xvec \* yvec; length xvec = length yvec; n = length xvec\ \ \p. (set p) \ (set xvec) \ (set yvec) \ distinctPerm p \ B = p \ B' \ set(map fst p) \ supp B \ xvec \* B' \ yvec \* B" + by fact + from EQ \x \ y\ have EQ': "\\*xvec'\B = ([(x, y)] \ (\\*yvec'\B'))" + and xFreshB': "x \ (\\*yvec'\B')" + and yFreshB: "y \ (\\*xvec'\B)" + by(metis boundOutput.inject alpha)+ + from xFreshB' \x \ yvec'\ have "x \ B'" + by(auto simp add: boundOutputFresh) (simp add: fresh_def name_list_supp)+ + from yFreshB \y \ xvec'\ have "y \ B" + by(auto simp add: boundOutputFresh) (simp add: fresh_def name_list_supp)+ + show ?case + proof(cases "x \ \\*xvec'\B") + assume xFreshB: "x \ \\*xvec'\B" + with EQ have yFreshB': "y \ \\*yvec'\B'" + by(rule boundOutputEqFresh) + with xFreshB' EQ' have "\\*xvec'\B = \\*yvec'\B'" + by(simp) + with \xvec' \* yvec'\ \length xvec' = length yvec'\ \length xvec' = n\ IH + obtain p where S: "(set p) \ (set xvec') \ (set yvec')" and "distinctPerm p" and "B = p \ B'" + and "set(map fst p) \ supp B" and "xvec' \* B'" and "yvec' \* B" + by blast + from S have "(set p) \ set(x#xvec') \ set(y#yvec')" by auto + moreover note \xvec = x#xvec'\ \yvec=y#yvec'\ \distinctPerm p\ \B = p \ B'\ + \xvec' \* B'\ \x \ B'\ \x \ B'\ \yvec' \* B\ \y \ B\ \set(map fst p) \ supp B\ + + ultimately show ?case by auto + next + assume "\(x \ \\*xvec'\B)" + then have xSuppB: "x \ supp(\\*xvec'\B)" + by(simp add: fresh_def) + with EQ have ySuppB': "y \ supp (\\*yvec'\B')" + by(rule boundOutputEqSupp) + then have "y \ yvec'" + by(induct yvec') (auto simp add: boundOutput.supp abs_supp) + with \x \ yvec'\ EQ' have "\\*xvec'\B = \\*yvec'\([(x, y)] \ B')" + by(simp add: eqvts) + with \xvec' \* yvec'\ \length xvec' = length yvec'\ \length xvec' = n\ IH + obtain p where S: "(set p) \ (set xvec') \ (set yvec')" and "distinctPerm p" and "B = p \ [(x, y)] \ B'" + and "set(map fst p) \ supp B" and "xvec' \* ([(x, y)] \ B')" and "yvec' \* B" + by blast + + from xSuppB have "x \ xvec'" + by(induct xvec') (auto simp add: boundOutput.supp abs_supp) + with \x \ yvec'\ \y \ xvec'\ \y \ yvec'\ S have "x \ p" and "y \ p" + apply(induct p) + by(auto simp add: name_list_supp) (auto simp add: fresh_def) + from S have "(set ((x, y)#p)) \ (set(x#xvec')) \ (set(y#yvec'))" + by force + moreover from \x \ y\ \x \ p\ \y \ p\ S \distinctPerm p\ + have "distinctPerm((x,y)#p)" by simp + moreover from \B = p \ [(x, y)] \ B'\ \x \ p\ \y \ p\ have "B = [(x, y)] \ p \ B'" + by(subst perm_compose) simp + then have "B = ((x, y)#p) \ B'" by simp + moreover from \xvec' \* ([(x, y)] \ B')\ have "([(x, y)] \ xvec') \* ([(x, y)] \ [(x, y)] \ B')" + by(simp only: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \x \ xvec'\ \y \ xvec'\ \x \ B'\ have "(x#xvec') \* B'" by simp + moreover from \y \ B\ \yvec' \* B\ have "(y#yvec') \* B" by simp + moreover from \set(map fst p) \ supp B\ xSuppB \x \ xvec'\ + have "set(map fst ((x, y)#p)) \ supp B" + by(simp add: BOresChainSupp) + ultimately show ?case using \xvec=x#xvec'\ \yvec=y#yvec'\ + by metis + qed + qed +qed + +lemma boundOutputChainEqLength: + fixes xvec :: "name list" + and M :: "'a::fs_name" + and P :: "('a, 'b::fs_name, 'c::fs_name) psi" + and yvec :: "name list" + and N :: "'a::fs_name" + and Q :: "('a, 'b::fs_name, 'c::fs_name) psi" + +assumes "\\*xvec\M \' P = \\*yvec\N \' Q" + +shows "length xvec = length yvec" +proof - + obtain n where "n = length xvec" by auto + with assms show ?thesis + proof(induct n arbitrary: xvec yvec M P N Q) + case(0 xvec yvec M P N Q) + from \0 = length xvec\ have "xvec = []" by auto + moreover with \\\*xvec\M \' P = \\*yvec\N \' Q\ have "yvec = []" + by(cases yvec) auto + ultimately show ?case by simp + next + case(Suc n xvec yvec M P N Q) + from \Suc n = length xvec\ + obtain x xvec' where "xvec = x#xvec'" and "length xvec' = n" + by(cases xvec) auto + from \\\*xvec\M \' P = \\*yvec\N \' Q\ \xvec = x # xvec'\ + obtain y yvec' where "\\*(x#xvec')\M \' P = \\*(y#yvec')\N \' Q" + and "yvec = y#yvec'" + by(cases yvec) auto + then have EQ: "\\x\(\\*xvec'\M \' P) = \\y\(\\*yvec'\N \' Q)" + by simp + have IH: "\xvec yvec M P N Q. \\\*xvec\M \' P = \\*yvec\N \' (Q::('a, 'b, 'c) psi); n = length xvec\ \ length xvec = length yvec" + by fact + show ?case + proof(cases "x = y") + assume "x = y" + with EQ have "\\*xvec'\M \' P = \\*yvec'\N \' Q" + by(simp add: alpha boundOutput.inject) + with IH \length xvec' = n\ have "length xvec' = length yvec'" + by blast + with \xvec = x#xvec'\ \yvec=y#yvec'\ + show ?case by simp + next + assume "x \ y" + with EQ have "\\*xvec'\M \' P = [(x, y)] \ \\*yvec'\N \' Q" + by(simp add: alpha boundOutput.inject) + then have "\\*xvec'\M \' P = \\*([(x, y)] \ yvec')\([(x, y)] \ N) \' ([(x, y)] \ Q)" + by(simp add: eqvts) + with IH \length xvec' = n\ have "length xvec' = length ([(x, y)] \ yvec')" + by blast + then have "length xvec' = length yvec'" + by simp + with \xvec = x#xvec'\ \yvec=y#yvec'\ + show ?case by simp + qed + qed +qed + +lemma boundOutputChainEq': + fixes xvec :: "name list" + and M :: "'a::fs_name" + and P :: "('a, 'b::fs_name, 'c::fs_name) psi" + and yvec :: "name list" + and N :: 'a + and Q :: "('a::fs_name, 'b::fs_name, 'c::fs_name) psi" + +assumes "\\*xvec\M \' P = \\*yvec\N \' Q" + and "xvec \* yvec" + +shows "\p. (set p) \ (set xvec) \ set (yvec) \ distinctPerm p \ M = p \ N \ P = p \ Q \ xvec \* N \ xvec \* Q \ yvec \* M \ yvec \* P" + using assms boundOutputChainEq boundOutputChainEqLength by fastforce + +lemma boundOutputChainEq'': + fixes xvec :: "name list" + and M :: "'a::fs_name" + and P :: "('a, 'b::fs_name, 'c::fs_name) psi" + and yvec :: "name list" + and N :: 'a + and Q :: "('a::fs_name, 'b::fs_name, 'c::fs_name) psi" + +assumes "\\*xvec\M \' P = \\*yvec\N \' Q" + and "xvec \* yvec" + and "distinct xvec" + and "distinct yvec" + +obtains p where "(set p) \ (set xvec) \ set (p \ xvec)" and "distinctPerm p" and "yvec = p \ xvec" and "N = p \ M" and "Q = p \ P" and "xvec \* N" and "xvec \* Q" and "(p \ xvec) \* M" and "(p \ xvec) \* P" +proof - + + assume "\p. \set p \ set xvec \ set (p \ xvec); distinctPerm p; yvec = p \ xvec; N = p \ M; Q = p \ P; xvec \* N; xvec \* Q; (p \ xvec) \* M; (p \ xvec) \* P\ \ thesis" + + moreover obtain n where "n = length xvec" by auto + with assms have "\p. (set p) \ (set xvec) \ set (yvec) \ distinctPerm p \ yvec = p \ xvec \ N = p \ M \ Q = p \ P \ xvec \* N \ xvec \* Q \ (p \ xvec) \* M \ (p \ xvec) \* P" + proof(induct n arbitrary: xvec yvec M P N Q) + case(0 xvec yvec M P N Q) + have Eq: "\\*xvec\M \' P = \\*yvec\N \' Q" by fact + from \0 = length xvec\ have "xvec = []" by auto + moreover with Eq have "yvec = []" + by(cases yvec) auto + ultimately show ?case using Eq + by(simp add: boundOutput.inject) + next + case(Suc n xvec yvec M P N Q) + from \Suc n = length xvec\ + obtain x xvec' where "xvec = x#xvec'" and "length xvec' = n" + by(cases xvec) auto + from \\\*xvec\M \' P = \\*yvec\N \' Q\ \xvec = x # xvec'\ + obtain y yvec' where "\\*(x#xvec')\M \' P = \\*(y#yvec')\N \' Q" + and "yvec = y#yvec'" + by(cases yvec) auto + then have EQ: "\\x\(\\*xvec'\M \' P) = \\y\(\\*yvec'\N \' Q)" + by simp + from \xvec = x#xvec'\ \yvec=y#yvec'\ \xvec \* yvec\ + have "x \ y" and "xvec' \* yvec'" and "x \ yvec'" and "y \ xvec'" + by auto + from \distinct xvec\ \distinct yvec\ \xvec=x#xvec'\ \yvec=y#yvec'\ have "x \ xvec'" and "y \ yvec'" and "distinct xvec'" and "distinct yvec'" + by simp+ + have IH: "\xvec yvec M P N Q. \\\*xvec\(M::'a) \' (P::('a, 'b, 'c) psi) = \\*yvec\N \' Q; xvec \* yvec; distinct xvec; distinct yvec; n = length xvec\ \ \p. (set p) \ (set xvec) \ (set yvec) \ distinctPerm p \ yvec = p \ xvec \ N = p \ M \ Q = p \ P \ xvec \* N \ xvec \* Q \ (p \ xvec) \* M \ (p \ xvec) \* P" + by fact + from EQ \x \ y\ \x \ yvec'\ \y \ yvec'\ \y \ xvec'\ \x \ xvec'\ have "\\*xvec'\M \' P = \\*yvec'\([(x, y)] \ N) \' ([(x, y)] \ Q)" and "x \ N" and "x \ Q" and "y \ M" and "y \ P" + apply - + apply(simp add: boundOutput.inject alpha eqvts) + apply(simp add: boundOutput.inject alpha eqvts) + apply(simp add: boundOutput.inject alpha eqvts) + by(simp add: boundOutput.inject alpha' eqvts)+ + with \xvec' \* yvec'\ \distinct xvec'\ \distinct yvec'\ \length xvec' = n\ IH + obtain p where S: "(set p) \ (set xvec') \ (set yvec')" and "distinctPerm p" and "yvec' = p \ xvec'" and "([(x, y)] \ N) = p \ M" and "([(x, y)] \ Q) = p \ P" and "xvec' \* ([(x, y)] \ N)" and "xvec' \* ([(x, y)] \ Q)" and "yvec' \* M" and "yvec' \* P" + by metis + from S have "set((x, y)#p) \ set(x#xvec') \ set(y#yvec')" by auto + moreover from \x \ xvec'\ \x \ yvec'\ \y \ xvec'\ \y \ yvec'\ S have "x \ p" and "y \ p" + apply(induct p) + by(auto simp add: fresh_prod name_list_supp) (auto simp add: fresh_def) + + with S \distinctPerm p\ \x \ y\ have "distinctPerm((x, y)#p)" by auto + moreover from \yvec' = p \ xvec'\ \x \ p\ \y \ p\ \x \ xvec'\ \y \ xvec'\ have "(y#yvec') = ((x, y)#p) \ (x#xvec')" + by(simp add: eqvts calc_atm perm_compose freshChainSimps) + moreover from \([(x, y)] \ N) = p \ M\ + have "([(x, y)] \ [(x, y)] \ N) = [(x, y)] \ p \ M" + by(simp add: pt_bij) + then have "N = ((x, y)#p) \ M" by simp + moreover from \([(x, y)] \ Q) = p \ P\ + have "([(x, y)] \ [(x, y)] \ Q) = [(x, y)] \ p \ P" + by(simp add: pt_bij) + then have "Q = ((x, y)#p) \ P" by simp + moreover from \xvec' \* ([(x, y)] \ N)\ have "([(x, y)] \ xvec') \* ([(x, y)] \ [(x, y)] \ N)" + by(subst fresh_star_bij) + with \x \ xvec'\ \y \ xvec'\ have "xvec' \* N" by simp + with \x \ N\ have "(x#xvec') \* N" by simp + moreover from \xvec' \* ([(x, y)] \ Q)\ have "([(x, y)] \ xvec') \* ([(x, y)] \ [(x, y)] \ Q)" + by(subst fresh_star_bij) + with \x \ xvec'\ \y \ xvec'\ have "xvec' \* Q" by simp + with \x \ Q\ have "(x#xvec') \* Q" by simp + moreover from \y \ M\ \yvec' \* M\ have "(y#yvec') \* M" by simp + moreover from \y \ P\ \yvec' \* P\ have "(y#yvec') \* P" by simp + ultimately show ?case using \xvec=x#xvec'\ \yvec=y#yvec'\ + by metis + qed + ultimately show ?thesis by blast +qed + +lemma boundOutputEqSupp': + fixes x :: name + and xvec :: "name list" + and M :: "'a::fs_name" + and P :: "('a, 'b::fs_name, 'c::fs_name) psi" + and y :: name + and yvec :: "name list" + and N :: 'a + and Q :: "('a, 'b, 'c) psi" + +assumes Eq: "\\x\(\\*xvec\M \' P) = \\y\(\\*yvec\N \' Q)" + and "x \ y" + and "x \ yvec" + and "x \ xvec" + and "y \ xvec" + and "y \ yvec" + and "xvec \* yvec" + and "x \ supp M" + +shows "y \ supp N" +proof - + from Eq \x \ y\ \x \ yvec\ \y \ yvec\ have "\\*xvec\M \' P = \\*yvec\([(x, y)] \ N) \' ([(x, y)] \ Q)" + by(simp add: boundOutput.inject alpha eqvts) + then obtain p where S: "set p \ set xvec \ set yvec" and "M = p \ [(x, y)] \ N" and "distinctPerm p" using \xvec \* yvec\ + by(blast dest: boundOutputChainEq') + with \x \ supp M\ have "x \ supp(p \ [(x, y)] \ N)" by simp + then have "(p \ x) \ p \ supp(p \ [(x, y)] \ N)" + by(simp add: pt_set_bij[OF pt_name_inst, OF at_name_inst]) + with \x \ xvec\ \x \ yvec\ S \distinctPerm p\ have "x \ supp([(x, y)] \ N)" + by(simp add: eqvts) + then have "([(x, y)] \ x) \ ([(x, y)] \ (supp([(x, y)] \ N)))" + by(simp add: pt_set_bij[OF pt_name_inst, OF at_name_inst]) + with \x \ y\ show ?thesis by(simp add: calc_atm eqvts) +qed + +lemma boundOutputChainOpenIH: + fixes xvec :: "name list" + and x :: name + and B :: "('a::fs_name, 'b::fs_name, 'c::fs_name) boundOutput" + and yvec :: "name list" + and y :: name + and B' :: "('a, 'b, 'c) boundOutput" + +assumes Eq: "\\*xvec\(\\x\B) = \\*yvec\(\\y\B')" + and L: "length xvec = length yvec" + and xFreshB': "x \ B'" + and xFreshxvec: "x \ xvec" + and xFreshyvec: "x \ yvec" + +shows "\\*xvec\B = \\*yvec\([(x, y)] \ B')" + using assms +proof(induct n=="length xvec" arbitrary: xvec yvec y B' rule: nat.induct) + case(zero xvec yvec y B') + have "0 = length xvec" and "length xvec = length yvec" by fact+ + moreover have "\\*xvec\\\x\B = \\*yvec\\\y\B'" by fact + ultimately show ?case by(auto simp add: boundOutput.inject alpha) +next + case(Suc n xvec yvec y B') + have L: "length xvec = length yvec" and "Suc n = length xvec" by fact+ + then obtain x' xvec' y' yvec' where xEq: "xvec = x'#xvec'" and yEq: "yvec = y'#yvec'" + and L': "length xvec' = length yvec'" + by(cases xvec, auto, cases yvec, auto) + have xFreshB': "x \ B'" by fact + have "x \ xvec" and "x \ yvec" by fact+ + with xEq yEq have xineqx': "x \ x'" and xFreshxvec': "x \ xvec'" + and xineqy': "x \ y'" and xFreshyvec': "x \ yvec'" + by simp+ + have "\\*xvec\\\x\B = \\*yvec\\\y\B'" by fact + with xEq yEq have Eq: "\\x'\(\\*xvec'\\\x\B) = \\y'\(\\*yvec'\\\y\B')" by simp + have IH: "\xvec yvec y B'. + \n = length xvec; \\*xvec\\\x\B = \\*yvec\\\y\B'; length xvec = length yvec; x \ B'; x \ xvec; x \ yvec\ + \ \\*xvec\B = \\*yvec\([(x, y)] \ B')" by fact + have "Suc n = length xvec" by fact + with xEq have L'': "n = length xvec'" by simp + have "\\x'\(\\*xvec'\B) = \\y'\(\\*yvec'\([(x, y)] \ B'))" + proof(cases "x'=y'") + assume x'eqy': "x' = y'" + with Eq have "\\*xvec'\\\x\B = \\*yvec'\\\y\B'" by(simp add: boundOutput.inject alpha) + then have "\\*xvec'\B = \\*yvec'\([(x, y)] \ B')" using L' xFreshB' xFreshxvec' xFreshyvec' L'' by(metis IH) + with x'eqy' show ?thesis by(simp add: boundOutput.inject alpha) + next + assume x'ineqy': "x' \ y'" + with Eq have Eq': "\\*xvec'\\\x\B = \\*([(x', y')] \ yvec')\\\([(x', y')] \ y)\([(x', y')] \ B')" + and x'FreshB': "x' \ \\*yvec'\\\y\B'" + by(simp add: boundOutput.inject alpha eqvts)+ + from L' have "length xvec' = length ([(x', y')] \ yvec')" by simp + moreover from xineqx' xineqy' xFreshB' have "x \ [(x', y')] \ B'" by(simp add: fresh_left calc_atm) + moreover from xineqx' xineqy' xFreshyvec' have "x \ [(x', y')] \ yvec'" by(simp add: fresh_left calc_atm) + ultimately have "\\*xvec'\B = \\*([(x', y')] \ yvec')\([(x, ([(x', y')] \ y))] \ [(x', y')] \ B')" using Eq' xFreshxvec' L'' + by(metis IH) + moreover from x'FreshB' have "x' \ \\*yvec'\([(x, y)] \ B')" + proof(cases "x' \ yvec'") + assume "x' \ yvec'" + with x'FreshB' have x'FreshB': "x' \ \\y\B'" + by(simp add: fresh_def BOresChainSupp) + show ?thesis + proof(cases "x'=y") + assume x'eqy: "x' = y" + show ?thesis + proof(cases "x=y") + assume "x=y" + with xFreshB' x'eqy show ?thesis by(simp add: BOresChainSupp fresh_def) + next + assume "x \ y" + with \x \ B'\ have "y \ [(x, y)] \ B'" by(simp add: fresh_left calc_atm) + with x'eqy show ?thesis by(simp add: BOresChainSupp fresh_def) + qed + next + assume x'ineqy: "x' \ y" + with x'FreshB' have "x' \ B'" by(simp add: abs_fresh) + with xineqx' x'ineqy have "x' \ ([(x, y)] \ B')" by(simp add: fresh_left calc_atm) + then show ?thesis by(simp add: BOresChainSupp fresh_def) + qed + next + assume "\x' \ yvec'" + then show ?thesis by(simp add: BOresChainSupp fresh_def) + qed + ultimately show ?thesis using x'ineqy' xineqx' xineqy' + apply(simp add: boundOutput.inject alpha eqvts) + apply(subst perm_compose[of "[(x', y')]"]) + by(simp add: calc_atm) + qed + with xEq yEq show ?case by simp +qed + +lemma boundOutputPar1Dest: + fixes xvec :: "name list" + and M :: "'a::fs_name" + and P :: "('a, 'b::fs_name, 'c::fs_name) psi" + and yvec :: "name list" + and N :: 'a + and Q :: "('a, 'b, 'c) psi" + and R :: "('a, 'b, 'c) psi" + +assumes "\\*xvec\M \' P = \\*yvec\N \' (Q \ R)" + and "xvec \* R" + and "yvec \* R" + +obtains T where "P = T \ R" and "\\*xvec\M \' T = \\*yvec\N \' Q" +proof - + assume "\T. \P = T \ R; \\*xvec\M \' T = \\*yvec\N \' Q\ \ thesis" + moreover obtain n where "n = length xvec" by auto + with assms have "\T. P = T \ R \ \\*xvec\M \' T = \\*yvec\N \' Q" + proof(induct n arbitrary: xvec yvec M N P Q R) + case(0 xvec yvec M N P Q R) + have Eq: "\\*xvec\M \' P = \\*yvec\N \' (Q \ R)" by fact + from \0 = length xvec\ have "xvec = []" by auto + moreover with Eq have "yvec = []" + by(cases yvec) auto + ultimately show ?case using Eq + by(simp add: boundOutput.inject) + next + case(Suc n xvec yvec M N P Q R) + from \Suc n = length xvec\ + obtain x xvec' where "xvec = x#xvec'" and "length xvec' = n" + by(cases xvec) auto + from \\\*xvec\M \' P = \\*yvec\N \' (Q \ R)\ \xvec = x # xvec'\ + obtain y yvec' where "\\*(x#xvec')\M \' P = \\*(y#yvec')\N \' (Q \ R)" + and "yvec = y#yvec'" + by(cases yvec) auto + then have EQ: "\\x\(\\*xvec'\M \' P) = \\y\(\\*yvec'\N \' (Q \ R))" + by simp + from \xvec \* R\ \yvec \* R\ \xvec = x#xvec'\ \yvec = y#yvec'\ + have "x \ R" and "xvec' \* R" and "y \ R" and "yvec' \* R" by auto + have IH: "\xvec yvec M N P Q R. \\\*xvec\M \' (P::('a, 'b, 'c) psi) = \\*yvec\N \' (Q \ R); xvec \* R; yvec \* R; n = length xvec\ \ \T. P = T \ R \ \\*xvec\M \' T = \\*yvec\N \' Q" + by fact + show ?case + proof(cases "x = y") + assume "x = y" + with EQ have "\\*xvec'\M \' P = \\*yvec'\N \' (Q \ R)" + by(simp add: boundOutput.inject alpha) + with \xvec' \* R\ \yvec' \* R\ \length xvec' = n\ + obtain T where "P = T \ R" and "\\*xvec'\M \' T = \\*yvec'\N \' Q" + by(auto dest: IH) + with \xvec=x#xvec'\ \yvec=y#yvec'\ \x=y\ show ?case + by(force simp add: boundOutput.inject alpha) + next + assume "x \ y" + with EQ \x \ R\ \y \ R\ + have "\\*xvec'\M \' P = \\*([(x, y)] \ yvec')\([(x, y)] \ N) \' (([(x, y)] \ Q) \ R)" + and xFreshQR: "x \ \\*yvec'\N \' (Q \ R)" + by(simp add: boundOutput.inject alpha eqvts)+ + moreover from \yvec' \* R\ have "([(x, y)] \ yvec') \* ([(x, y)] \ R)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \x \ R\ \y \ R\ have "([(x, y)] \ yvec') \* R" by simp + moreover note \xvec' \* R\ \length xvec' = n\ + ultimately obtain T where "P = T \ R" and A: "\\*xvec'\M \' T = \\*([(x, y)] \ yvec')\([(x, y)] \ N) \' ([(x, y)] \ Q)" + by(auto dest: IH) + + from A have "\\x\(\\*xvec'\M \' T) = \\x\(\\*([(x, y)] \ yvec')\([(x, y)] \ N) \' ([(x, y)] \ Q))" + by(simp add: boundOutput.inject alpha) + moreover from xFreshQR have "x \ \\*yvec'\N \' Q" + by(force simp add: boundOutputFresh) + ultimately show ?thesis using \P = T \ R\ \xvec=x#xvec'\ \yvec=y#yvec'\ xFreshQR + by(force simp add: alphaBoundOutput name_swap eqvts) + qed + qed + ultimately show ?thesis + by blast +qed + +lemma boundOutputPar1Dest': + fixes xvec :: "name list" + and M :: "'a::fs_name" + and P :: "('a, 'b::fs_name, 'c::fs_name) psi" + and yvec :: "name list" + and N :: 'a + and Q :: "('a, 'b, 'c) psi" + and R :: "('a, 'b, 'c) psi" + +assumes "\\*xvec\M \' P = \\*yvec\N \' (Q \ R)" + and "xvec \* yvec" + +obtains T p where "set p \ set xvec \ set yvec" and "P = T \ (p \ R)" and "\\*xvec\M \' T = \\*yvec\N \' Q" +proof - + assume "\p T. \set p \ set xvec \ set yvec; P = T \ (p \ R); \\*xvec\M \' T = \\*yvec\N \' Q\ \ thesis" + moreover obtain n where "n = length xvec" by auto + with assms have "\p T. set p \ set xvec \ set yvec \ P = T \ (p \ R) \ \\*xvec\M \' T = \\*yvec\N \' Q" + proof(induct n arbitrary: xvec yvec M N P Q R) + case(0 xvec yvec M N P Q R) + have Eq: "\\*xvec\M \' P = \\*yvec\N \' (Q \ R)" by fact + from \0 = length xvec\ have "xvec = []" by auto + moreover with Eq have "yvec = []" + by(cases yvec) auto + ultimately show ?case using Eq + by(simp add: boundOutput.inject) + next + case(Suc n xvec yvec M N P Q R) + from \Suc n = length xvec\ + obtain x xvec' where "xvec = x#xvec'" and "length xvec' = n" + by(cases xvec) auto + from \\\*xvec\M \' P = \\*yvec\N \' (Q \ R)\ \xvec = x # xvec'\ + obtain y yvec' where "\\*(x#xvec')\M \' P = \\*(y#yvec')\N \' (Q \ R)" + and "yvec = y#yvec'" + by(cases yvec) auto + then have Eq: "\\x\(\\*xvec'\M \' P) = \\y\(\\*yvec'\N \' (Q \ R))" + by simp + from \xvec = x#xvec'\ \yvec=y#yvec'\ \xvec \* yvec\ have "x \ y" and "x \ yvec'" and "y \ xvec'" and "xvec' \* yvec'" + by auto + from Eq \x \ y\ have Eq': "\\*xvec'\M \' P = [(x, y)] \ \\*yvec'\N \' (Q \ R)" + and xFreshQR: "x \ \\*yvec'\N \' (Q \ R)" + by(simp add: boundOutput.inject alpha)+ + have IH: "\xvec yvec M N P Q R. \\\*xvec\M \' (P::('a, 'b, 'c) psi) = \\*yvec\N \' (Q \ R); xvec \* yvec; n = length xvec\ \ \p T. set p \ set xvec \ set yvec \ P = T \ (p \ R) \ \\*xvec\M \' T = \\*yvec\N \' Q" + by fact + show ?case + proof(cases "x \ \\*xvec'\M \' P") + assume "x \ \\*xvec'\M \' P" + with Eq have yFreshQR: "y \ \\*yvec'\N \' (Q \ R)" + by(rule boundOutputEqFresh) + with Eq' xFreshQR have "\\*xvec'\M \' P = \\*yvec'\N \' (Q \ R)" + by simp + with \xvec' \* yvec'\ \length xvec' = n\ + obtain p T where S: "set p \ set xvec' \ set yvec'" and "P = T \ (p \ R)" and A: "\\*xvec'\M \' T = \\*yvec'\N \' Q" + by(auto dest: IH) + from yFreshQR xFreshQR have yFreshQ: "y \ \\*yvec'\N \' Q" and xFreshQ: "x \ \\*yvec'\N \' Q" + by(force simp add: BOresChainSupp fresh_def boundOutput.supp psi.supp)+ + then have "\\x\(\\*yvec'\N \' Q) = \\y\(\\*yvec'\N \' Q)" by (subst alphaBoundOutput) simp+ + with A have "\\x\(\\*xvec'\M \' T) = \\y\(\\*yvec'\N \' Q)" by simp + with \xvec=x#xvec'\ \yvec=y#yvec'\ S \P = T \ (p \ R)\ show ?case + by auto + next + assume "\(x \ \\*xvec'\M \' P)" + then have "x \ supp(\\*xvec'\M \' P)" by(simp add: fresh_def) + with Eq have "y \ supp(\\*yvec'\N \' (Q \ R))" + by(rule boundOutputEqSupp) + then have "y \ yvec'" by(simp add: BOresChainSupp fresh_def) + with Eq' \x \ yvec'\ have "\\*xvec'\M \' P = \\*yvec'\([(x, y)] \ N) \' (([(x, y)] \ Q) \ ([(x, y)] \ R))" + by(simp add: eqvts) + moreover note \xvec' \* yvec'\ \length xvec' = n\ + ultimately obtain p T where S: "set p \ set xvec' \ set yvec'" and "P = T \ (p \ [(x, y)] \ R)" and A: "\\*xvec'\M \' T = \\*yvec'\([(x, y)] \ N) \' ([(x, y)] \ Q)" + by(auto dest: IH) + + from S have "set(p@[(x, y)]) \ set(x#xvec') \ set(y#yvec')" by auto + moreover from \P = T \ (p \ [(x, y)] \ R)\ have "P = T \ ((p @ [(x, y)]) \ R)" + by(simp add: pt2[OF pt_name_inst]) + moreover from xFreshQR have xFreshQ: "x \ \\*yvec'\N \' Q" + by(force simp add: BOresChainSupp fresh_def boundOutput.supp psi.supp)+ + with \x \ yvec'\ \y \ yvec'\ \x \ y\ have "y \ \\*yvec'\([(x, y)] \ N) \' ([(x, y)] \ Q)" + by(simp add: fresh_left calc_atm) + with \x \ yvec'\ \y \ yvec'\ have "\\x\(\\*yvec'\([(x, y)] \ N) \' ([(x, y)] \ Q)) = \\y\(\\*yvec'\N \' Q)" + by(subst alphaBoundOutput) (assumption | simp add: eqvts)+ + with A have "\\x\(\\*xvec'\M \' T) = \\y\(\\*yvec'\N \' Q)" by simp + ultimately show ?thesis using \xvec=x#xvec'\ \yvec=y#yvec'\ + by - (rule exI[where x="p@[(x, y)]"], force) + qed + qed + ultimately show ?thesis + by blast +qed + +lemma boundOutputPar2Dest: + fixes xvec :: "name list" + and M :: "'a::fs_name" + and P :: "('a, 'b::fs_name, 'c::fs_name) psi" + and yvec :: "name list" + and N :: 'a + and Q :: "('a, 'b, 'c) psi" + and R :: "('a, 'b, 'c) psi" + +assumes "\\*xvec\M \' P = \\*yvec\N \' (Q \ R)" + and "xvec \* Q" + and "yvec \* Q" + +obtains T where "P = Q \ T" and "\\*xvec\M \' T = \\*yvec\N \' R" +proof - + assume "\T. \P = Q \ T; \\*xvec\M \' T = \\*yvec\N \' R\ \ thesis" + moreover obtain n where "n = length xvec" by auto + with assms have "\T. P = Q \ T \ \\*xvec\M \' T = \\*yvec\N \' R" + proof(induct n arbitrary: xvec yvec M N P Q R) + case(0 xvec yvec M N P Q R) + have Eq: "\\*xvec\M \' P = \\*yvec\N \' (Q \ R)" by fact + from \0 = length xvec\ have "xvec = []" by auto + moreover with Eq have "yvec = []" + by(cases yvec) auto + ultimately show ?case using Eq + by(simp add: boundOutput.inject) + next + case(Suc n xvec yvec M N P Q R) + from \Suc n = length xvec\ + obtain x xvec' where "xvec = x#xvec'" and "length xvec' = n" + by(cases xvec) auto + from \\\*xvec\M \' P = \\*yvec\N \' (Q \ R)\ \xvec = x # xvec'\ + obtain y yvec' where "\\*(x#xvec')\M \' P = \\*(y#yvec')\N \' (Q \ R)" + and "yvec = y#yvec'" + by(cases yvec) auto + then have EQ: "\\x\(\\*xvec'\M \' P) = \\y\(\\*yvec'\N \' (Q \ R))" + by simp + from \xvec \* Q\ \yvec \* Q\ \xvec = x#xvec'\ \yvec = y#yvec'\ + have "x \ Q" and "xvec' \* Q" and "y \ Q" and "yvec' \* Q" by auto + have IH: "\xvec yvec M N P Q R. \\\*xvec\M \' (P::('a, 'b, 'c) psi) = \\*yvec\N \' (Q \ R); xvec \* Q; yvec \* Q; n = length xvec\ \ \T. P = Q \ T \ \\*xvec\M \' T = \\*yvec\N \' R" + by fact + show ?case + proof(cases "x = y") + assume "x = y" + with EQ have "\\*xvec'\M \' P = \\*yvec'\N \' (Q \ R)" + by(simp add: boundOutput.inject alpha) + with \xvec' \* Q\ \yvec' \* Q\ \length xvec' = n\ + obtain T where "P = Q \ T" and "\\*xvec'\M \' T = \\*yvec'\N \' R" + by(auto dest: IH) + with \xvec=x#xvec'\ \yvec=y#yvec'\ \x=y\ show ?case + by(force simp add: boundOutput.inject alpha) + next + assume "x \ y" + with EQ \x \ Q\ \y \ Q\ + have "\\*xvec'\M \' P = \\*([(x, y)] \ yvec')\([(x, y)] \ N) \' (Q \ ([(x, y)] \ R))" + and xFreshQR: "x \ \\*yvec'\N \' (Q \ R)" + by(simp add: boundOutput.inject alpha eqvts)+ + moreover from \yvec' \* Q\ have "([(x, y)] \ yvec') \* ([(x, y)] \ Q)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \x \ Q\ \y \ Q\ have "([(x, y)] \ yvec') \* Q" by simp + moreover note \xvec' \* Q\ \length xvec' = n\ + ultimately obtain T where "P = Q \ T" and A: "\\*xvec'\M \' T = \\*([(x, y)] \ yvec')\([(x, y)] \ N) \' ([(x, y)] \ R)" + by(auto dest: IH) + + from A have "\\x\(\\*xvec'\M \' T) = \\x\(\\*([(x, y)] \ yvec')\([(x, y)] \ N) \' ([(x, y)] \ R))" + by(simp add: boundOutput.inject alpha) + moreover from xFreshQR have "x \ \\*yvec'\N \' R" + by(force simp add: boundOutputFresh) + ultimately show ?thesis using \P = Q \ T\ \xvec=x#xvec'\ \yvec=y#yvec'\ xFreshQR + by(force simp add: alphaBoundOutput name_swap eqvts) + qed + qed + ultimately show ?thesis + by blast +qed + +lemma boundOutputPar2Dest': + fixes xvec :: "name list" + and M :: "'a::fs_name" + and P :: "('a, 'b::fs_name, 'c::fs_name) psi" + and yvec :: "name list" + and N :: 'a + and Q :: "('a, 'b, 'c) psi" + and R :: "('a, 'b, 'c) psi" + +assumes "\\*xvec\M \' P = \\*yvec\N \' (Q \ R)" + and "xvec \* yvec" + +obtains T p where "set p \ set xvec \ set yvec" and "P = (p \ Q) \ T" and "\\*xvec\M \' T = \\*yvec\N \' R" +proof - + assume "\p T. \set p \ set xvec \ set yvec; P = (p \ Q) \ T; \\*xvec\M \' T = \\*yvec\N \' R\ \ thesis" + moreover obtain n where "n = length xvec" by auto + with assms have "\p T. set p \ set xvec \ set yvec \ P = (p \ Q) \ T \ \\*xvec\M \' T = \\*yvec\N \' R" + proof(induct n arbitrary: xvec yvec M N P Q R) + case(0 xvec yvec M N P Q R) + have Eq: "\\*xvec\M \' P = \\*yvec\N \' (Q \ R)" by fact + from \0 = length xvec\ have "xvec = []" by auto + moreover with Eq have "yvec = []" + by(cases yvec) auto + ultimately show ?case using Eq + by(simp add: boundOutput.inject) + next + case(Suc n xvec yvec M N P Q R) + from \Suc n = length xvec\ + obtain x xvec' where "xvec = x#xvec'" and "length xvec' = n" + by(cases xvec) auto + from \\\*xvec\M \' P = \\*yvec\N \' (Q \ R)\ \xvec = x # xvec'\ + obtain y yvec' where "\\*(x#xvec')\M \' P = \\*(y#yvec')\N \' (Q \ R)" + and "yvec = y#yvec'" + by(cases yvec) auto + then have Eq: "\\x\(\\*xvec'\M \' P) = \\y\(\\*yvec'\N \' (Q \ R))" + by simp + from \xvec = x#xvec'\ \yvec=y#yvec'\ \xvec \* yvec\ have "x \ y" and "x \ yvec'" and "y \ xvec'" and "xvec' \* yvec'" + by auto + from Eq \x \ y\ have Eq': "\\*xvec'\M \' P = [(x, y)] \ \\*yvec'\N \' (Q \ R)" + and xFreshQR: "x \ \\*yvec'\N \' (Q \ R)" + by(simp add: boundOutput.inject alpha)+ + have IH: "\xvec yvec M N P Q R. \\\*xvec\M \' (P::('a, 'b, 'c) psi) = \\*yvec\N \' (Q \ R); xvec \* yvec; n = length xvec\ \ \p T. set p \ set xvec \ set yvec \ P = (p \ Q) \ T \ \\*xvec\M \' T = \\*yvec\N \' R" + by fact + show ?case + proof(cases "x \ \\*xvec'\M \' P") + assume "x \ \\*xvec'\M \' P" + with Eq have yFreshQR: "y \ \\*yvec'\N \' (Q \ R)" + by(rule boundOutputEqFresh) + with Eq' xFreshQR have "\\*xvec'\M \' P = \\*yvec'\N \' (Q \ R)" + by simp + with \xvec' \* yvec'\ \length xvec' = n\ + obtain p T where S: "set p \ set xvec' \ set yvec'" and "P = (p \ Q) \ T" and A: "\\*xvec'\M \' T = \\*yvec'\N \' R" + by(auto dest: IH) + from yFreshQR xFreshQR have yFreshR: "y \ \\*yvec'\N \' R" and xFreshQ: "x \ \\*yvec'\N \' R" + by(force simp add: BOresChainSupp fresh_def boundOutput.supp psi.supp)+ + then have "\\x\(\\*yvec'\N \' R) = \\y\(\\*yvec'\N \' R)" by (subst alphaBoundOutput) simp+ + with A have "\\x\(\\*xvec'\M \' T) = \\y\(\\*yvec'\N \' R)" by simp + with \xvec=x#xvec'\ \yvec=y#yvec'\ S \P = (p \ Q) \ T\ show ?case + by auto + next + assume "\(x \ \\*xvec'\M \' P)" + then have "x \ supp(\\*xvec'\M \' P)" by(simp add: fresh_def) + with Eq have "y \ supp(\\*yvec'\N \' (Q \ R))" + by(rule boundOutputEqSupp) + then have "y \ yvec'" by(simp add: BOresChainSupp fresh_def) + with Eq' \x \ yvec'\ have "\\*xvec'\M \' P = \\*yvec'\([(x, y)] \ N) \' (([(x, y)] \ Q) \ ([(x, y)] \ R))" + by(simp add: eqvts) + moreover note \xvec' \* yvec'\ \length xvec' = n\ + ultimately obtain p T where S: "set p \ set xvec' \ set yvec'" and "P = (p \ [(x, y)] \ Q) \ T" and A: "\\*xvec'\M \' T = \\*yvec'\([(x, y)] \ N) \' ([(x, y)] \ R)" + by(auto dest: IH) + + from S have "set(p@[(x, y)]) \ set(x#xvec') \ set(y#yvec')" by auto + moreover from \P = (p \ [(x, y)] \ Q) \ T\ have "P = ((p @ [(x, y)]) \ Q) \ T" + by(simp add: pt2[OF pt_name_inst]) + moreover from xFreshQR have xFreshR: "x \ \\*yvec'\N \' R" + by(force simp add: BOresChainSupp fresh_def boundOutput.supp psi.supp)+ + with \x \ yvec'\ \y \ yvec'\ \x \ y\ have "y \ \\*yvec'\([(x, y)] \ N) \' ([(x, y)] \ R)" + by(simp add: fresh_left calc_atm) + with \x \ yvec'\ \y \ yvec'\ have "\\x\(\\*yvec'\([(x, y)] \ N) \' ([(x, y)] \ R)) = \\y\(\\*yvec'\N \' R)" + by(subst alphaBoundOutput) (assumption | simp add: eqvts)+ + with A have "\\x\(\\*xvec'\M \' T) = \\y\(\\*yvec'\N \' R)" by simp + ultimately show ?thesis using \xvec=x#xvec'\ \yvec=y#yvec'\ + by(force intro!: exI[where x="p@[(x, y)]"]) + qed + qed + ultimately show ?thesis + by blast +qed + +lemma boundOutputApp: + fixes xvec :: "name list" + and yvec :: "name list" + and B :: "('a::fs_name, 'b::fs_name, 'c::fs_name) boundOutput" + +shows "\\*(xvec@yvec)\B = \\*xvec\(\\*yvec\B)" + by(induct xvec) auto + +lemma openInjectAux: + fixes xvec1 :: "name list" + and x :: name + and xvec2 :: "name list" + and yvec :: "name list" + +assumes "length(xvec1@x#xvec2) = length yvec" + +shows "\yvec1 y yvec2. yvec = yvec1@y#yvec2 \ length xvec1 = length yvec1 \ length xvec2 = length yvec2" + apply(rule exI[where x="take (length xvec1) yvec"]) + apply(rule exI[where x="yvec ! length xvec1"]) + apply(rule exI[where x="drop (length xvec1+1) yvec"]) + using assms by(auto simp add: id_take_nth_drop) + +lemma boundOutputOpenDest: + fixes yvec :: "name list" + and M :: "'a::fs_name" + and P :: "('a, 'b::fs_name, 'c::fs_name) psi" + and xvec1 :: "name list" + and x :: name + and xvec2 :: "name list" + and N :: 'a + and Q :: "('a, 'b, 'c) psi" + +assumes Eq: "\\*(xvec1@x#xvec2)\M \' P = \\*yvec\N \' Q" + and "x \ xvec1" + and "x \ yvec" + and "x \ N" + and "x \ Q" + and "distinct yvec" + + +obtains yvec1 y yvec2 where "yvec=yvec1@y#yvec2" and "length xvec1 = length yvec1" and "length xvec2 = length yvec2" + and "\\*(xvec1@xvec2)\M \' P = \\*(yvec1@yvec2)\([(x, y)] \ N) \' ([(x, y)] \ Q)" +proof - + assume Ass: "\yvec1 y yvec2. + \yvec = yvec1 @ y # yvec2; length xvec1 = length yvec1; length xvec2 = length yvec2; + \\*(xvec1 @ xvec2)\M \' P = \\*(yvec1 @ yvec2)\([(x, y)] \ N) \' ([(x, y)] \ Q)\ + \ thesis" + from Eq have "length(xvec1@x#xvec2) = length yvec" by(rule boundOutputChainEqLength) + then obtain yvec1 y yvec2 where A: "yvec = yvec1@y#yvec2" and "length xvec1 = length yvec1" + and "length xvec2 = length yvec2" + by(metis openInjectAux sym) + + from \distinct yvec\ A have "y \ yvec2" by simp + from A \x \ yvec\ have "x \ yvec2" and "x \ yvec1" by simp+ + with Eq \length xvec1 = length yvec1\ \x \ N\ \x \ Q\ \y \ yvec2\ \x \ xvec1\ A + have "\\*(xvec1@xvec2)\M \' P = \\*(yvec1@yvec2)\([(x, y)] \ N) \' ([(x, y)] \ Q)" + by(force dest: boundOutputChainOpenIH simp add: boundOutputApp BOresChainSupp fresh_def boundOutput.supp eqvts) + with \length xvec1 = length yvec1\ \length xvec2 = length yvec2\ A Ass show ?thesis + by blast +qed + +lemma boundOutputOpenDest': + fixes yvec :: "name list" + and M :: "'a::fs_name" + and P :: "('a, 'b::fs_name, 'c::fs_name) psi" + and xvec1 :: "name list" + and x :: name + and xvec2 :: "name list" + and N :: 'a + and Q :: "('a, 'b, 'c) psi" + +assumes Eq: "\\*(xvec1@x#xvec2)\M \' P = \\*yvec\N \' Q" + and "x \ xvec1" + and "x \ yvec" + and "x \ N" + and "x \ Q" + + +obtains yvec1 y yvec2 where "yvec=yvec1@y#yvec2" and "length xvec1 = length yvec1" and "length xvec2 = length yvec2" + and "\\*(xvec1@xvec2)\M \' P = \\*(yvec1@[(x, y)] \ yvec2)\([(x, y)] \ N) \' ([(x, y)] \ Q)" +proof - + assume Ass: "\yvec1 y yvec2. + \yvec = yvec1 @ y # yvec2; length xvec1 = length yvec1; length xvec2 = length yvec2; + \\*(xvec1 @ xvec2)\M \' P = \\*(yvec1 @ ([(x, y)] \ yvec2))\([(x, y)] \ N) \' ([(x, y)] \ Q)\ + \ thesis" + from Eq have "length(xvec1@x#xvec2) = length yvec" by(rule boundOutputChainEqLength) + then obtain yvec1 y yvec2 where A: "yvec = yvec1@y#yvec2" and "length xvec1 = length yvec1" + and "length xvec2 = length yvec2" + by(metis openInjectAux sym) + + from A \x \ yvec\ have "x \ yvec2" and "x \ yvec1" by simp+ + with Eq \length xvec1 = length yvec1\ \x \ N\ \x \ Q\ \x \ xvec1\ A + have "\\*(xvec1@xvec2)\M \' P = \\*(yvec1@([(x, y)] \ yvec2))\([(x, y)] \ N) \' ([(x, y)] \ Q)" + by(force dest: boundOutputChainOpenIH simp add: boundOutputApp BOresChainSupp fresh_def boundOutput.supp eqvts) + with \length xvec1 = length yvec1\ \length xvec2 = length yvec2\ A Ass show ?thesis + by blast +qed + +lemma boundOutputScopeDest: + fixes xvec :: "name list" + and M :: "'a::fs_name" + and P :: "('a, 'b::fs_name, 'c::fs_name) psi" + and yvec :: "name list" + and N :: 'a + and x :: name + and Q :: "('a, 'b, 'c) psi" + +assumes "\\*xvec\M \' P = \\*yvec\N \' \\z\Q" + and "z \ xvec" + and "z \ yvec" + +obtains R where "P = \\z\R" and "\\*xvec\M \' R = \\*yvec\N \' Q" +proof - + assume "\R. \P = \\z\R; \\*xvec\M \' R = \\*yvec\N \' Q\ \ thesis" + moreover obtain n where "n = length xvec" by auto + with assms have "\R. P = \\z\R \ \\*xvec\M \' R = \\*yvec\N \' Q" + proof(induct n arbitrary: xvec yvec M N P Q z) + case(0 xvec yvec M N P Q z) + have Eq: "\\*xvec\M \' P = \\*yvec\N \' \\z\Q" by fact + from \0 = length xvec\ have "xvec = []" by auto + moreover with Eq have "yvec = []" + by(cases yvec) auto + ultimately show ?case using Eq + by(simp add: boundOutput.inject) + next + case(Suc n xvec yvec M N P Q z) + from \Suc n = length xvec\ + obtain x xvec' where "xvec = x#xvec'" and "length xvec' = n" + by(cases xvec) auto + from \\\*xvec\M \' P = \\*yvec\N \' (\\z\Q)\ \xvec = x # xvec'\ + obtain y yvec' where "\\*(x#xvec')\M \' P = \\*(y#yvec')\N \' \\z\Q" + and "yvec = y#yvec'" + by(cases yvec) auto + then have EQ: "\\x\(\\*xvec'\M \' P) = \\y\(\\*yvec'\N \' \\z\Q)" + by simp + from \z \ xvec\ \z \ yvec\ \xvec = x#xvec'\ \yvec = y#yvec'\ + have "z \ x" and "z \ y" and "z \ xvec'" and "z \ yvec'" + by simp+ + have IH: "\xvec yvec M N P Q z. \\\*xvec\M \' (P::('a, 'b, 'c) psi) = \\*yvec\N \' \\z\Q; z \ xvec; z \ yvec; n = length xvec\ \ \R. P = \\z\R \ \\*xvec\M \' R = \\*yvec\N \' Q" + by fact + show ?case + proof(cases "x = y") + assume "x = y" + with EQ have "\\*xvec'\M \' P = \\*yvec'\N \' \\z\Q" + by(simp add: boundOutput.inject alpha) + with \z \ xvec'\ \z \ yvec'\ \length xvec' = n\ + obtain R where "P = \\z\R" and "\\*xvec'\M \' R = \\*yvec'\N \' Q" + by(auto dest: IH) + with \xvec=x#xvec'\ \yvec=y#yvec'\ \x=y\ show ?case + by(force simp add: boundOutput.inject alpha) + next + assume "x \ y" + with EQ \z \ x\ \z \ y\ + have "\\*xvec'\M \' P = \\*([(x, y)] \ yvec')\([(x, y)] \ N) \' \\z\([(x, y)] \ Q)" + and xFreshzQ: "x \ \\*yvec'\N \' \\z\Q" + by(simp add: boundOutput.inject alpha eqvts)+ + moreover from \z \ x\ \z \ y\ \z \ yvec'\ \x \ y\ have "z \ ([(x, y)] \ yvec')" + by(simp add: fresh_left calc_atm) + moreover note \z \ xvec'\ \length xvec' = n\ + ultimately obtain R where "P = \\z\R" and A: "\\*xvec'\M \' R = \\*([(x, y)] \ yvec')\([(x, y)] \ N) \' ([(x, y)] \ Q)" + by(auto dest: IH) + + from A have "\\x\(\\*xvec'\M \' R) = \\x\(\\*([(x, y)] \ yvec')\([(x, y)] \ N) \' ([(x, y)] \ Q))" + by(simp add: boundOutput.inject alpha) + moreover from xFreshzQ \z \ x\ have "x \ \\*yvec'\N \' Q" + by(simp add: boundOutputFresh abs_fresh) + ultimately show ?thesis using \P = \\z\R\ \xvec=x#xvec'\ \yvec=y#yvec'\ xFreshzQ + by(force simp add: alphaBoundOutput name_swap eqvts) + qed + qed + ultimately show ?thesis + by blast +qed + +nominal_datatype ('a, 'b, 'c) residual = + RIn "'a::fs_name" 'a "('a, 'b::fs_name, 'c::fs_name) psi" + | RBrIn "'a::fs_name" 'a "('a, 'b::fs_name, 'c::fs_name) psi" + | ROut 'a "('a, 'b, 'c) boundOutput" + | RBrOut 'a "('a, 'b, 'c) boundOutput" + | RTau "('a, 'b, 'c) psi" + +nominal_datatype 'a action = In "'a::fs_name" 'a ("_\_\" [90, 90] 90) + | BrIn "'a::fs_name" 'a ("\_\_\" [90, 90] 90) + | Out "'a::fs_name" "name list" 'a ("_\\*_\\_\" [90, 90, 90] 90) + | BrOut "'a::fs_name" "name list" 'a ("\_\\*_\\_\" [90, 90, 90] 90) + | Tau ("\" 90) + +nominal_primrec bn :: "('a::fs_name) action \ name list" + where + "bn (M\N\) = []" + | "bn (\M\N\) = []" + | "bn (M\\*xvec\\N\) = xvec" + | "bn (\M\\*xvec\\N\) = xvec" + | "bn (\) = []" + by(rule TrueI)+ + +lemma bnEqvt[eqvt]: + fixes p :: "name prm" + and \ :: "('a::fs_name) action" + +shows "(p \ bn \) = bn(p \ \)" + by(nominal_induct \ rule: action.strong_induct) auto + +nominal_primrec create_residual :: "('a::fs_name) action \ ('a, 'b::fs_name, 'c::fs_name) psi \ ('a, 'b, 'c) residual" ("_ \ _" [80, 80] 80) + where + "(M\N\) \ P = RIn M N P" + | "(\M\N\) \ P = RBrIn M N P" + | "M\\*xvec\\N\ \ P = ROut M (\\*xvec\(N \' P))" + | "(\M\\*xvec\\N\) \ P = RBrOut M (\\*xvec\(N \' P))" + | "\ \ P = (RTau P)" + by(rule TrueI)+ + +nominal_primrec subject :: "('a::fs_name) action \ 'a option" + where + "subject (M\N\) = Some M" + | "subject (\M\N\) = Some M" + | "subject (M\\*xvec\\N\) = Some M" + | "subject (\M\\*xvec\\N\) = Some M" + | "subject (\) = None" + by(rule TrueI)+ + +nominal_primrec object :: "('a::fs_name) action \ 'a option" + where + "object (M\N\) = Some N" + | "object (\M\N\) = Some N" + | "object (M\\*xvec\\N\) = Some N" + | "object (\M\\*xvec\\N\) = Some N" + | "object (\) = None" + by(rule TrueI)+ + +lemma optionFreshChain[simp]: + fixes xvec :: "name list" + and X :: "name set" + +shows "xvec \* (Some x) = xvec \* x" + and "X \* (Some x) = X \* x" + and "xvec \* None" + and "X \* None" + by(auto simp add: fresh_star_def fresh_some fresh_none) + +lemmas [simp] = fresh_some fresh_none + +lemma actionFresh[simp]: + fixes x :: name + and \ :: "('a::fs_name) action" + +shows "(x \ \) = (x \ (subject \) \ x \ (bn \) \ x \ (object \))" + by(nominal_induct \ rule: action.strong_induct) auto + +lemma actionFreshChain[simp]: + fixes X :: "name set" + and \ :: "('a::fs_name) action" + and xvec :: "name list" + +shows "(X \* \) = (X \* (subject \) \ X \* (bn \) \ X \* (object \))" + and "(xvec \* \) = (xvec \* (subject \) \ xvec \* (bn \) \ xvec \* (object \))" + by(auto simp add: fresh_star_def) + +lemma subjectEqvt[eqvt]: + fixes p :: "name prm" + and \ :: "('a::fs_name) action" + +shows "(p \ subject \) = subject(p \ \)" + by(nominal_induct \ rule: action.strong_induct) auto + +lemma okjectEqvt[eqvt]: + fixes p :: "name prm" + and \ :: "('a::fs_name) action" + +shows "(p \ object \) = object(p \ \)" + by(nominal_induct \ rule: action.strong_induct) auto + +lemma create_residualEqvt[eqvt]: + fixes p :: "name prm" + and \ :: "('a::fs_name) action" + and P :: "('a, 'b::fs_name, 'c::fs_name) psi" + +shows "(p \ (\ \ P)) = (p \ \) \ (p \ P)" + by(nominal_induct \ rule: action.strong_induct) + (auto simp add: eqvts) + +lemma residualFresh: + fixes x :: name + and \ :: "'a::fs_name action" + and P :: "('a, 'b::fs_name, 'c::fs_name) psi" + +shows "(x \ (\ \ P)) = (x \ (subject \) \ (x \ (set(bn(\))) \ (x \ object(\) \ x \ P)))" + by(nominal_induct \ rule: action.strong_induct) + (auto simp add: fresh_some fresh_none boundOutputFresh) + +lemma residualFresh2[simp]: + fixes x :: name + and \ :: "('a::fs_name) action" + and P :: "('a, 'b::fs_name, 'c::fs_name) psi" + +assumes "x \ \" + and "x \ P" + +shows "x \ \ \ P" + using assms + by(nominal_induct \ rule: action.strong_induct) auto + +lemma residualFreshChain2[simp]: + fixes xvec :: "name list" + and X :: "name set" + and \ :: "('a::fs_name) action" + and P :: "('a, 'b::fs_name, 'c::fs_name) psi" + +shows "\xvec \* \; xvec \* P\ \ xvec \* (\ \ P)" + and "\X \* \; X \* P\ \ X \* (\ \ P)" + by(auto simp add: fresh_star_def) + +lemma residualFreshSimp[simp]: + fixes x :: name + and M :: "'a::fs_name" + and N :: 'a + and P :: "('a, 'b::fs_name, 'c::fs_name) psi" + + +shows "x \ (M\N\ \ P) = (x \ M \ x \ N \ x \ P)" + and "x \ (\ M\N\ \ P) = (x \ M \ x \ N \ x \ P)" + and "x \ (M\\*xvec\\N\ \ P) = (x \ M \ x \ (\\*xvec\(N \' P)))" + and "x \ (\M\\*xvec\\N\ \ P) = (x \ M \ x \ (\\*xvec\(N \' P)))" + and "x \ (\ \ P) = (x \ P)" + by(auto simp add: residualFresh) + +lemma residualInject': + +shows "(\ \ P = RIn M N Q) = (P = Q \ \ = M\N\)" + and "(\ \ P = RBrIn M N Q) = (P = Q \ \ = \M\N\)" + and "(\ \ P = ROut M B) = (\xvec N. \ = M\\*xvec\\N\ \ B = \\*xvec\(N \' P))" + and "(\ \ P = RBrOut M B) = (\xvec N. \ = \M\\*xvec\\N\ \ B = \\*xvec\(N \' P))" + and "(\ \ P = RTau Q) = (\ = \ \ P = Q)" + and "(RIn M N Q = \ \ P) = (P = Q \ \ = M\N\)" + and "(RBrIn M N Q = \ \ P) = (P = Q \ \ = \M\N\)" + and "(ROut M B = \ \ P) = (\xvec N. \ = M\\*xvec\\N\ \ B = \\*xvec\(N \' P))" + and "(RBrOut M B = \ \ P) = (\xvec N. \ = \M\\*xvec\\N\ \ B = \\*xvec\(N \' P))" + and "(RTau Q = \ \ P) = (\ = \ \ P = Q)" +proof - + show "(\ \ P = RIn M N Q) = (P = Q \ \ = M\N\)" + by(nominal_induct \ rule: action.strong_induct) + (auto simp add: residual.inject action.inject) +next + show "(\ \ P = RBrIn M N Q) = (P = Q \ \ = \M\N\)" + by(nominal_induct \ rule: action.strong_induct) + (auto simp add: residual.inject action.inject) +next + show "(\ \ P = ROut M B) = (\xvec N. \ = M\\*xvec\\N\ \ B = \\*xvec\(N \' P))" + by(nominal_induct \ rule: action.strong_induct) + (auto simp add: residual.inject action.inject) +next + show "(\ \ P = RBrOut M B) = (\xvec N. \ = \M\\*xvec\\N\ \ B = \\*xvec\(N \' P))" + by(nominal_induct \ rule: action.strong_induct) + (auto simp add: residual.inject action.inject) +next + show "(\ \ P = RTau Q) = (\ = \ \ P = Q)" + by(nominal_induct \ rule: action.strong_induct) + (auto simp add: residual.inject action.inject) +next + show "(RIn M N Q = \ \ P) = (P = Q \ \ = M\N\)" + by(nominal_induct \ rule: action.strong_induct) + (auto simp add: residual.inject action.inject) +next + show "(RBrIn M N Q = \ \ P) = (P = Q \ \ = \M\N\)" + by(nominal_induct \ rule: action.strong_induct) + (auto simp add: residual.inject action.inject) +next + show "(ROut M B = \ \ P) = (\xvec N. \ = M\\*xvec\\N\ \ B = \\*xvec\(N \' P))" + by(nominal_induct \ rule: action.strong_induct) + (auto simp add: residual.inject action.inject) +next + show "(RBrOut M B = \ \ P) = (\xvec N. \ = \M\\*xvec\\N\ \ B = \\*xvec\(N \' P))" + by(nominal_induct \ rule: action.strong_induct) + (auto simp add: residual.inject action.inject) +next + show "(RTau Q = \ \ P) = (\ = \ \ P = Q)" + by(nominal_induct \ rule: action.strong_induct) + (auto simp add: residual.inject action.inject) +qed + +lemma residualFreshChainSimp[simp]: + fixes xvec :: "name list" + and X :: "name set" + and M :: "'a::fs_name" + and N :: 'a + and yvec :: "name list" + and P :: "('a, 'b::fs_name, 'c::fs_name) psi" + +shows "xvec \* (M\N\ \ P) = (xvec \* M \ xvec \* N \ xvec \* P)" + and "xvec \* (\M\N\ \ P) = (xvec \* M \ xvec \* N \ xvec \* P)" + and "xvec \* (M\\*yvec\\N\ \ P) = (xvec \* M \ xvec \* (\\*yvec\(N \' P)))" + and "xvec \* (\M\\*yvec\\N\ \ P) = (xvec \* M \ xvec \* (\\*yvec\(N \' P)))" + and "xvec \* (\ \ P) = (xvec \* P)" + and "X \* (M\N\ \ P) = (X \* M \ X \* N \ X \* P)" + and "X \* (\M\N\ \ P) = (X \* M \ X \* N \ X \* P)" + and "X \* (M\\*yvec\\N\ \ P) = (X \* M \ X \* (\\*yvec\(N \' P)))" + and "X \* (\M\\*yvec\\N\ \ P) = (X \* M \ X \* (\\*yvec\(N \' P)))" + and "X \* (\ \ P) = (X \* P)" + by(auto simp add: fresh_star_def) + +lemma residualFreshChainSimp2[simp]: + fixes xvec :: "name list" + and X :: "name set" + and M :: "'a::fs_name" + and N :: 'a + and yvec :: "name list" + and P :: "('a, 'b::fs_name, 'c::fs_name) psi" + +shows "xvec \* (RIn M N P) = (xvec \* M \ xvec \* N \ xvec \* P)" + and "xvec \* (RBrIn M N P) = (xvec \* M \ xvec \* N \ xvec \* P)" + and "xvec \* (ROut M B) = (xvec \* M \ xvec \* B)" + and "xvec \* (RBrOut M B) = (xvec \* M \ xvec \* B)" + and "xvec \* (RTau P) = (xvec \* P)" + and "X \* (RIn M N P) = (X \* M \ X \* N \ X \* P)" + and "X \* (RBrIn M N P) = (X \* M \ X \* N \ X \* P)" + and "X \* (ROut M B) = (X \* M \ X \* B)" + and "X \* (RBrOut M B) = (X \* M \ X \* B)" + and "X \* (RTau P) = (X \* P)" + by(auto simp add: fresh_star_def) + +lemma freshResidual3[dest]: + fixes x :: name + and \ :: "('a::fs_name) action" + and P :: "('a, 'b::fs_name, 'c::fs_name) psi" + +assumes "x \ bn \" + and "x \ \ \ P" + +shows "x \ \" and "x \ P" + using assms + by(nominal_induct rule: action.strong_induct) auto + +lemma freshResidualChain3[dest]: + fixes xvec :: "name list" + and \ :: "('a::fs_name) action" + and P :: "('a, 'b::fs_name, 'c::fs_name) psi" + +assumes "xvec \* (\ \ P)" + and "xvec \* bn \" + +shows "xvec \* \" and "xvec \* P" + using assms + by(nominal_induct rule: action.strong_induct) auto + +lemma freshResidual4[dest]: + fixes x :: name + and \ :: "('a::fs_name) action" + and P :: "('a, 'b::fs_name, 'c::fs_name) psi" + +assumes "x \ \ \ P" + +shows "x \ subject \" + using assms + by(nominal_induct rule: action.strong_induct) auto + +lemma freshResidualChain4[dest]: + fixes xvec :: "name list" + and \ :: "('a::fs_name) action" + and P :: "('a, 'b::fs_name, 'c::fs_name) psi" + +assumes "xvec \* (\ \ P)" + +shows "xvec \* subject \" + using assms + by(nominal_induct rule: action.strong_induct) auto + +lemma alphaOutputResidual: + fixes M :: "'a::fs_name" + and xvec :: "name list" + and N :: 'a + and P :: "('a, 'b::fs_name, 'c::fs_name) psi" + and p :: "name prm" + +assumes "(p \ xvec) \* N" + and "(p \ xvec) \* P" + and "set p \ set xvec \ set(p \ xvec)" + and "set xvec \ set yvec" + +shows "M\\*yvec\\N\ \ P = M\\*(p \ yvec)\\(p \ N)\ \ (p \ P)" + and "\M\\*yvec\\N\ \ P = \M\\*(p \ yvec)\\(p \ N)\ \ (p \ P)" + using assms + by(simp add: boundOutputChainAlpha'')+ + +lemmas[simp del] = create_residual.simps + +lemma residualInject'': + +assumes "bn \ = bn \" + +shows "(\ \ P = \ \ Q) = (\ = \ \ P = Q)" + using assms + by(nominal_induct \ rule: action.strong_induct) + (force simp add: residual.inject create_residual.simps residualInject' action.inject boundOutput.inject)+ + +lemmas residualInject = residual.inject create_residual.simps residualInject' residualInject'' + +lemma bnFreshResidual[simp]: + fixes \ :: "('a::fs_name) action" + +shows "(bn \) \* (\ \ P) = bn \ \* (subject \)" + by(nominal_induct \ rule: action.strong_induct) + (auto simp add: residualFresh fresh_some fresh_star_def) + +lemma actionCases[case_names cInput cBrInput cOutput cBrOutput cTau]: + fixes \ :: "('a::fs_name) action" + +assumes "\M N. \ = M\N\ \ Prop" + and "\M N. \ = \M\N\ \ Prop" + and "\M xvec N. \ = M\\*xvec\\N\ \ Prop" + and "\M xvec N. \ = \M\\*xvec\\N\ \ Prop" + and "\ = \ \ Prop" + +shows Prop + using assms + by(nominal_induct \ rule: action.strong_induct) auto + +lemma actionPar1Dest: + fixes \ :: "('a::fs_name) action" + and P :: "('a, 'b::fs_name, 'c::fs_name) psi" + and \ :: "('a::fs_name) action" + and Q :: "('a, 'b, 'c) psi" + and R :: "('a, 'b, 'c) psi" + +assumes "\ \ P = \ \ (Q \ R)" + and "bn \ \* bn \" + +obtains T p where "set p \ set(bn \) \ set(bn \)" and "P = T \ (p \ R)" and "\ \ T = \ \ Q" + using assms + by(cases rule: actionCases[where \=\]) + (force simp add: residualInject dest: boundOutputPar1Dest')+ + +lemma actionPar2Dest: + fixes \ :: "('a::fs_name) action" + and P :: "('a, 'b::fs_name, 'c::fs_name) psi" + and \ :: "('a::fs_name) action" + and Q :: "('a, 'b, 'c) psi" + and R :: "('a, 'b, 'c) psi" + +assumes "\ \ P = \ \ (Q \ R)" + and "bn \ \* bn \" + +obtains T p where "set p \ set(bn \) \ set(bn \)" and "P = (p \ Q) \ T" and "\ \ T = \ \ R" + using assms + by(cases rule: actionCases[where \=\]) + (force simp add: simp add: residualInject dest: boundOutputPar2Dest')+ + +lemma actionScopeDest: + fixes \ :: "('a::fs_name) action" + and P :: "('a, 'b::fs_name, 'c::fs_name) psi" + fixes \ :: "('a::fs_name) action" + and x :: name + and Q :: "('a, 'b, 'c) psi" + +assumes "\ \ P = \ \ \\x\Q" + and "x \ bn \" + and "x \ bn \" + +obtains R where "P = \\x\R" and "\ \ R = \ \ Q" + using assms boundOutputScopeDest + by(cases rule: actionCases[where \=\]) (force simp add: residualInject)+ + +lemma emptyFreshName: + fixes x :: name + and M :: "'a::fs_name" + +assumes "supp M = ({}::name set)" + +shows "x \ M" + using assms + by(auto simp add: fresh_def) + +lemma emptyFresh: + fixes xvec :: "name list" + and M :: "'a::fs_name" + +assumes "supp M = ({}::name set)" + +shows "xvec \* M" + using assms by (induct xvec, auto simp add: emptyFreshName) + +lemma permEmptyEq: + fixes p :: "name prm" + and M :: "'a::fs_name" + +assumes suppE: "supp M = ({}::name set)" + +shows "(p \ M) = M" +proof(induct p) + case Nil + then show ?case by simp +next + case(Cons a p) + have "p \ M = M" by(rule Cons) + then have "([a] \ p \ M) = [a] \ M" by simp + then have "((a#p) \ M) = [a] \ M" + by(simp add: pt2[OF pt_name_inst, symmetric]) + then show ?case using suppE perm_fresh_fresh + by(cases a) (simp add: fresh_def) +qed + +abbreviation + outputJudge ("_\_\" [110, 110] 110) where "M\N\ \ M\\*([])\\N\" + +abbreviation + brOutputJudge ("\_\_\" [110, 110] 110) where "\M\N\ \ \M\\*([])\\N\" + +declare [[unify_trace_bound=100]] + +locale env = substPsi substTerm substAssert substCond + + assertion SCompose' SImp' SBottom' SChanEq' SOutCon' SInCon' + for substTerm :: "('a::fs_name) \ name list \ 'a::fs_name list \ 'a" + and substAssert :: "('b::fs_name) \ name list \ 'a::fs_name list \ 'b" + and substCond :: "('c::fs_name) \ name list \ 'a::fs_name list \ 'c" + and SCompose' :: "'b \ 'b \ 'b" + and SImp' :: "'b \ 'c \ bool" + and SBottom' :: 'b + and SChanEq' :: "'a \ 'a \ 'c" + and SOutCon' :: "'a \ 'a \ 'c" + and SInCon' :: "'a \ 'a \ 'c" +begin +notation SCompose' (infixr "\" 90) +notation SImp' ("_ \ _" [85, 85] 85) +notation FrameImp ("_ \\<^sub>F _" [85, 85] 85) +abbreviation + FBottomJudge ("\\<^sub>F" 90) where "\\<^sub>F \ (FAssert SBottom')" +notation SChanEq' ("_ \ _" [90, 90] 90) +notation SOutCon' ("_ \ _" [90, 90] 90) +notation SInCon' ("_ \ _" [90, 90] 90) +notation substTerm ("_[_::=_]" [100, 100, 100] 100) +notation subs ("_[_::=_]" [100, 100, 100] 100) +notation AssertionStatEq ("_ \ _" [80, 80] 80) +notation FrameStatEq ("_ \\<^sub>F _" [80, 80] 80) +notation SBottom' ("\" 190) +abbreviation insertAssertion' ("insertAssertion") where "insertAssertion' \ assertionAux.insertAssertion (\)" + +inductive semantics :: "'b \ ('a, 'b, 'c) psi \ ('a, 'b, 'c) residual \ bool" + ("_ \ _ \ _" [50, 50, 50] 50) + where + cInput: "\\ \ M \ K; distinct xvec; set xvec \ supp N; xvec \* Tvec; + length xvec = length Tvec; + xvec \* \; xvec \* M; xvec \* K\ \ \ \ M\\*xvec N\.P \ K\(N[xvec::=Tvec])\ \ P[xvec::=Tvec]" + | cBrInput:"\\ \ K \ M; distinct xvec; set xvec \ supp N; xvec \* Tvec; + length xvec = length Tvec; + xvec \* \; xvec \* M; xvec \* K\ \ \ \ M\\*xvec N\.P \ \K\(N[xvec::=Tvec])\ \ P[xvec::=Tvec]" + | Output: "\\ \ M \ K\ \ \ \ M\N\.P \ K\N\ \ P" + | BrOutput: "\\ \ M \ K\ \ \ \ M\N\.P \ \K\N\ \ P" + | Case: "\\ \ P \ Rs; (\, P) \ set Cs; \ \ \; guarded P\ \ \ \ Cases Cs \ Rs" + | cPar1: "\(\ \ \\<^sub>Q) \ P \\ \ P'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + A\<^sub>Q \* P; A\<^sub>Q \* Q; A\<^sub>Q \* \; A\<^sub>Q \* \; A\<^sub>Q \* P'; distinct(bn \); + bn \ \* \; bn \ \* \\<^sub>Q; bn \ \* Q; bn \ \* P; bn \ \* (subject \)\ \ + \ \ P \ Q \\ \ (P' \ Q)" + | cPar2: "\(\ \ \\<^sub>P) \ Q \\ \ Q'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + A\<^sub>P \* P; A\<^sub>P \* Q; A\<^sub>P \* \; A\<^sub>P \* \; A\<^sub>P \* Q'; distinct(bn \); + bn \ \* \; bn \ \* \\<^sub>P; bn \ \* P; bn \ \* Q; bn \ \* (subject \)\ \ + \ \ P \ Q \\ \ (P \ Q')" + | cComm1: "\\ \ \\<^sub>Q \ P \ M\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \ \ \\<^sub>P \ Q \ K\\*xvec\\N\ \ Q'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + \ \ \\<^sub>P \ \\<^sub>Q \ M \ K; + A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* P'; + A\<^sub>P \* Q; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* xvec; + A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; A\<^sub>Q \* P; A\<^sub>Q \* N; A\<^sub>Q \* P'; + A\<^sub>Q \* Q; A\<^sub>Q \* K; A\<^sub>Q \* Q'; A\<^sub>Q \* xvec; distinct xvec; + xvec \* \; xvec \* \\<^sub>P; xvec \* \\<^sub>Q; xvec \* P; xvec \* M; + xvec \* Q; xvec \* K\ \ + \ \ P \ Q \ \ \ \\*xvec\(P' \ Q')" + | cComm2: "\\ \ \\<^sub>Q \ P \ M\\*xvec\\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \ \ \\<^sub>P \ Q \ K\N\ \ Q'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + \ \ \\<^sub>P \ \\<^sub>Q \ M \ K; + A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* P'; + A\<^sub>P \* Q; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* xvec; + A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; A\<^sub>Q \* P; A\<^sub>Q \* N; A\<^sub>Q \* P'; + A\<^sub>Q \* Q; A\<^sub>Q \* K; A\<^sub>Q \* Q'; A\<^sub>Q \* xvec; distinct xvec; + xvec \* \; xvec \* \\<^sub>P; xvec \* \\<^sub>Q; xvec \* P; xvec \* M; + xvec \* Q; xvec \* K\ \ + \ \ P \ Q \ \ \ \\*xvec\(P' \ Q')" + | cBrMerge: "\\ \ \\<^sub>Q \ P \ \M\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \ \ \\<^sub>P \ Q \ \M\N\ \ Q'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* N; A\<^sub>P \* P'; + A\<^sub>P \* Q; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; + A\<^sub>P \* M; A\<^sub>Q \* M; + A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; A\<^sub>Q \* P; A\<^sub>Q \* N; A\<^sub>Q \* P'; + A\<^sub>Q \* Q; A\<^sub>Q \* Q'\ \ + \ \ P \ Q \ \M\N\ \ (P' \ Q')" + (* Removed: A\<^sub>P \* M; A\<^sub>Q \* K; *) + | cBrComm1: "\\ \ \\<^sub>Q \ P \ \M\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \ \ \\<^sub>P \ Q \ \M\\*xvec\\N\ \ Q'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* N; A\<^sub>P \* P'; + A\<^sub>P \* Q; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* xvec; + A\<^sub>P \* M; A\<^sub>Q \* M; + A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; A\<^sub>Q \* P; A\<^sub>Q \* N; A\<^sub>Q \* P'; + A\<^sub>Q \* Q; A\<^sub>Q \* Q'; A\<^sub>Q \* xvec; distinct xvec; + xvec \* \; xvec \* \\<^sub>P; xvec \* \\<^sub>Q; xvec \* P; + xvec \* Q; xvec \* M\ \ + \ \ P \ Q \ \M\\*xvec\\N\ \ (P' \ Q')" + | cBrComm2: "\\ \ \\<^sub>Q \ P \ \M\\*xvec\\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \ \ \\<^sub>P \ Q \ \M\N\ \ Q'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* N; A\<^sub>P \* P'; + A\<^sub>P \* Q; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* xvec; + A\<^sub>P \* M; A\<^sub>Q \* M; + A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; A\<^sub>Q \* P; A\<^sub>Q \* N; A\<^sub>Q \* P'; + A\<^sub>Q \* Q; A\<^sub>Q \* Q'; A\<^sub>Q \* xvec; distinct xvec; + xvec \* \; xvec \* \\<^sub>P; xvec \* \\<^sub>Q; xvec \* P; + xvec \* Q; xvec \* M\ \ + \ \ P \ Q \ \M\\*xvec\\N\ \ (P' \ Q')" + | cBrClose: "\\ \ P \ \M\\*xvec\\N\ \ P'; + x \ supp M; + distinct xvec; xvec \* \; xvec \* P; + xvec \* M; + x \ \; x \ xvec\ \ + \ \ \\x\P \ \ \ \\x\(\\*xvec\P')" + | cOpen: "\\ \ P \ M\\*(xvec@yvec)\\N\ \ P'; x \ supp N; x \ xvec; x \ yvec; x \ M; x \ \; + distinct xvec; distinct yvec; + xvec \* \; xvec \* P; xvec \* M; xvec \* yvec; yvec \* \; yvec \* P; yvec \* M\ \ + \ \ \\x\P \ M\\*(xvec@x#yvec)\\N\ \ P'" + | cBrOpen: "\\ \ P \ \M\\*(xvec@yvec)\\N\ \ P'; x \ supp N; x \ xvec; x \ yvec; x \ M; x \ \; + distinct xvec; distinct yvec; + xvec \* \; xvec \* P; xvec \* M; xvec \* yvec; yvec \* \; yvec \* P; yvec \* M\ \ + \ \ \\x\P \ \M\\*(xvec@x#yvec)\\N\ \ P'" + | cScope: "\\ \ P \\ \ P'; x \ \; x \ \; bn \ \* \; bn \ \* P; bn \ \* (subject \); distinct(bn \)\ \ \ \ \\x\P \\ \ (\\x\P')" + | Bang: "\\ \ P \ !P \ Rs; guarded P\ \ \ \ !P \ Rs" + +abbreviation + semanticsBottomJudge ("_ \ _" [50, 50] 50) where "P \ Rs \ \ \ P \ Rs" + +equivariance env.semantics + +nominal_inductive2 env.semantics + avoids cInput: "set xvec" + | cBrInput: "set xvec" + | cPar1: "set A\<^sub>Q \ set(bn \)" + | cPar2: "set A\<^sub>P \ set(bn \)" + | cComm1: "set A\<^sub>P \ set A\<^sub>Q \ set xvec" + | cComm2: "set A\<^sub>P \ set A\<^sub>Q \ set xvec" + | cBrMerge: "set A\<^sub>P \ set A\<^sub>Q" + | cBrComm1: "set A\<^sub>P \ set A\<^sub>Q \ set xvec" + | cBrComm2: "set A\<^sub>P \ set A\<^sub>Q \ set xvec" + | cBrClose: "{x} \ set xvec" + | cOpen: "{x} \ set xvec \ set yvec" + | cBrOpen: "{x} \ set xvec \ set yvec" + | cScope: "{x} \ set(bn \)" + apply - + apply(force intro: substTerm.subst4Chain subst4Chain simp add: abs_fresh residualFresh)+ + apply(force intro: substTerm.subst4Chain subst4Chain simp add: abs_fresh residualFresh boundOutputFresh boundOutputFreshSet fresh_star_def resChainFresh)+ + done + +lemma nilTrans1: + fixes \ :: 'b + and M :: 'a + and xvec :: "name list" + and N :: 'a + and P :: "('a, 'b, 'c) psi" + +assumes "\ \ \ \ M\\*xvec\\N\ \ P" + +shows "False" + using assms + apply - + by (ind_cases "\ \ \ \ M\\*xvec\\N\ \ P") + +lemma nilTrans1': + fixes \ :: 'b + and M :: 'a + and xvec :: "name list" + and N :: 'a + and P :: "('a, 'b, 'c) psi" + +assumes "\ \ \ \ \M\\*xvec\\N\ \ P" + +shows "False" + using assms + apply - + by (ind_cases "\ \ \ \ \M\\*xvec\\N\ \ P") + +lemma nilTrans2: + fixes \ :: 'b + and Rs :: "('a, 'b, 'c) residual" + +assumes "\ \ \ \ Rs" + +shows "False" + using assms + apply(cases rule: semantics.cases) + by(auto simp add: residualInject)+ + +lemma nilTrans3: + fixes \ :: 'b + and M :: 'a + and M' :: 'a + and xvec :: "name list" + and yvec :: "name list" + and N :: 'a + and N' :: 'a + and P :: "('a, 'b, 'c) psi" + and P' :: "('a, 'b, 'c) psi" + +assumes "\ \ M\\*xvec N\.P \ M'\\*yvec\\N'\ \ P'" + +shows "False" + using assms + apply - + by(ind_cases "\ \ M\\*xvec N\.P \ M'\\*yvec\\N'\ \ P'") (auto simp add: residualInject) + +lemma nilTrans3': + fixes \ :: 'b + and M :: 'a + and M' :: 'a + and xvec :: "name list" + and yvec :: "name list" + and N :: 'a + and N' :: 'a + and P :: "('a, 'b, 'c) psi" + and P' :: "('a, 'b, 'c) psi" + +assumes "\ \ M\\*xvec N\.P \ \M'\\*yvec\\N'\ \ P'" + +shows "False" + using assms + apply - + by(ind_cases "\ \ M\\*xvec N\.P \ \M'\\*yvec\\N'\ \ P'") (auto simp add: residualInject) + +lemma nilTrans4: + fixes \ :: 'b + and Rs :: "('a, 'b, 'c) residual" + +assumes "\ \ M\\*xvec N\.P \\ \ P'" + +shows "False" + using assms + apply(cases rule: semantics.cases) + by(auto simp add: residualInject)+ + +lemma nilTrans5: + fixes \ :: 'b + fixes \' :: 'b + and M :: 'a + and xvec :: "name list" + and N :: 'a + and P :: "('a, 'b, 'c) psi" + +assumes "\ \ \\'\ \ M\\*xvec\\N\ \ P" + +shows "False" + using assms + apply - + by(ind_cases "\ \ \\'\ \ M\\*xvec\\N\ \ P") + +lemma nilTrans5': + fixes \ :: 'b + fixes \' :: 'b + and M :: 'a + and xvec :: "name list" + and N :: 'a + and P :: "('a, 'b, 'c) psi" + +assumes "\ \ \\'\ \ \M\\*xvec\\N\ \ P" + +shows "False" + using assms + apply - + by(ind_cases "\ \ \\'\ \ \M\\*xvec\\N\ \ P") + +lemma nilTrans6: + fixes \ :: 'b + and Rs :: "('a, 'b, 'c) residual" + +assumes "\ \ \\'\ \ Rs" + +shows "False" + using assms + apply(cases rule: semantics.cases) + by(auto simp add: residualInject)+ + +lemma nilTrans[dest]: + fixes \ :: 'b + and Rs :: "('a, 'b, 'c) residual" + and M :: 'a + and xvec :: "name list" + and N :: 'a + and P :: "('a, 'b, 'c) psi" + and K :: 'a + and yvec :: "name list" + and N' :: 'a + and P' :: "('a, 'b, 'c) psi" + and CsP :: "('c \ ('a, 'b, 'c) psi) list" + and \' :: 'b + +shows "\ \ \ \ Rs \ False" + and "\ \ M\\*xvec N\.P \K\\*yvec\\N'\ \ P' \ False" + and "\ \ M\\*xvec N\.P \\K\\*yvec\\N'\ \ P' \ False" + and "\ \ M\\*xvec N\.P \\ \ P' \ False" + and "\ \ M\N\.P \K\N'\ \ P' \ False" + and "\ \ M\N\.P \\K\N'\ \ P' \ False" + and "\ \ M\N\.P \\ \ P' \ False" + and "\ \ \\'\ \ Rs \ False" + apply - + apply(rule nilTrans2) + apply assumption + apply(cases rule: semantics.cases) apply(force simp add: residualInject)+ + apply(cases rule: semantics.cases) apply(force simp add: residualInject)+ + apply(rule nilTrans4) + apply assumption + apply(cases rule: semantics.cases) apply(force simp add: residualInject)+ + apply(cases rule: semantics.cases) apply(force simp add: residualInject)+ + apply(cases rule: semantics.cases) apply(force simp add: residualInject)+ + apply(rule nilTrans6) + by assumption + +lemma residualEq: + fixes \ :: "'a action" + and P :: "('a, 'b, 'c) psi" + and \ :: "'a action" + and Q :: "('a, 'b, 'c) psi" + +assumes "\ \ P = \ \ Q" + and "bn \ \* (bn \)" + and "distinct(bn \)" + and "distinct(bn \)" + and "bn \ \* (\ \ P)" + and "bn \ \* (\ \ Q)" + +obtains p where "set p \ set(bn \) \ set(bn(p \ \))" and "distinctPerm p" and "\ = p \ \" and "Q = p \ P" and "bn \ \* \" and "bn \ \* Q" and "bn(p \ \) \* \" and "bn(p \ \) \* P" + using assms +proof(nominal_induct \ rule: action.strong_induct) + case(In M N) + then show ?case by(simp add: residualInject) +next + case(BrIn M N) + then show ?case by(simp add: residualInject) +next + case(Out M xvec N) + then show ?case + using boundOutputChainEq'' by(force simp add: residualInject) +next + case(BrOut M xvec N) + then show ?case + using boundOutputChainEq'' by(force simp add: residualInject) +next + case Tau + then show ?case by(simp add: residualInject) +qed + +lemma semanticsInduct[consumes 3, case_names cAlpha cInput cBrInput cOutput cBrOutput cCase cPar1 cPar2 cComm1 cComm2 cBrMerge cBrComm1 cBrComm2 cBrClose cOpen cBrOpen cScope cBang]: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and \ :: "'a action" + and P' :: "('a, 'b, 'c) psi" + and Prop :: "'f::fs_name \ 'b \ ('a, 'b, 'c) psi \ + 'a action \ ('a, 'b, 'c) psi \ bool" + and C :: "'f::fs_name" + +assumes "\ \ P \\ \ P'" + and "bn \ \* (subject \)" + and "distinct(bn \)" + and rAlpha: "\\ P \ P' p C. \bn \ \* \; bn \ \* P; bn \ \* (subject \); + bn \ \* C; bn \ \* (bn(p \ \)); + set p \ set(bn \) \ set(bn(p \ \)); distinctPerm p; + (bn(p \ \)) \* \; (bn(p \ \)) \* P'; Prop C \ P \ P'\ \ + Prop C \ P (p \ \) (p \ P')" + and rInput: "\\ M K xvec N Tvec P C. + \\ \ M \ K; distinct xvec; set xvec \ supp N; + length xvec = length Tvec; xvec \* \; + xvec \* M; xvec \* K; xvec \* C\ \ + Prop C \ (M\\*xvec N\.P) + (K\(N[xvec::=Tvec])\) (P[xvec::=Tvec])" + and rBrInput:"\\ K M xvec N Tvec P C. + \\ \ K \ M; distinct xvec; set xvec \ supp N; + length xvec = length Tvec; xvec \* \; + xvec \* M; xvec \* K; xvec \* C\ \ + Prop C \ (M\\*xvec N\.P) + (\K\(N[xvec::=Tvec])\) (P[xvec::=Tvec])" + and rOutput: "\\ M K N P C. \\ \ M \ K\ \ Prop C \ (M\N\.P) (K\N\) P" + and rBrOutput:"\\ M K N P C. \\ \ M \ K\ \ Prop C \ (M\N\.P) (\K\N\) P" + and rCase: "\\ P \ P' \ Cs C. \\ \ P \\ \ P'; \C. Prop C \ P \ P'; (\, P) \ set Cs; \ \ \; guarded P\ \ + Prop C \ (Cases Cs) \ P'" + and rPar1: "\\ \\<^sub>Q P \ P' A\<^sub>Q Q C. + \\ \ \\<^sub>Q \ P \\ \ P'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + \C. Prop C (\ \ \\<^sub>Q) P \ P'; + A\<^sub>Q \* P; A\<^sub>Q \* Q; A\<^sub>Q \* \; A\<^sub>Q \* \; A\<^sub>Q \* P'; A\<^sub>Q \* C; distinct(bn \); bn \ \* Q; + bn \ \* \; bn \ \* \\<^sub>Q; bn \ \* P; bn \ \* subject \; bn \ \* C\ \ + Prop C \ (P \ Q) \ (P' \ Q)" + and rPar2: "\\ \\<^sub>P Q \ Q' A\<^sub>P P C. + \\ \ \\<^sub>P \ Q \\ \ Q'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \C. Prop C (\ \ \\<^sub>P) Q \ Q'; + A\<^sub>P \* P; A\<^sub>P \* Q; A\<^sub>P \* \; A\<^sub>P \* \; A\<^sub>P \* Q'; A\<^sub>P \* C; distinct(bn \); bn \ \* Q; + bn \ \* \; bn \ \* \\<^sub>P; bn \ \* P; bn \ \* subject \; bn \ \* C\ \ + Prop C \ (P \ Q) \ (P \ Q')" + and rComm1: "\\ \\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q K xvec Q' A\<^sub>Q C. + \\ \ \\<^sub>Q \ P \M\N\ \ P'; \C. Prop C (\ \ \\<^sub>Q) P (M\N\) P'; + extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \ \ \\<^sub>P \ Q \K\\*xvec\\N\ \ Q'; \C. Prop C (\ \ \\<^sub>P) Q (K\\*xvec\\N\) Q'; + extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + \ \ \\<^sub>P \ \\<^sub>Q \ M \ K; + A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* P'; + A\<^sub>P \* Q; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* xvec; A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; + A\<^sub>Q \* P; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* Q; A\<^sub>Q \* K; A\<^sub>Q \* Q'; distinct xvec; + A\<^sub>Q \* xvec; xvec \* \; xvec \* \\<^sub>P; xvec \* \\<^sub>Q; xvec \* P; xvec \* M; + xvec \* Q; xvec \* K; A\<^sub>P \* C; A\<^sub>Q \* C; xvec \* C\ \ + Prop C \ (P \ Q) (\) (\\*xvec\(P' \ Q'))" + and rComm2: "\\ \\<^sub>Q P M xvec N P' A\<^sub>P \\<^sub>P Q K Q' A\<^sub>Q C. + \\ \ \\<^sub>Q \ P \M\\*xvec\\N\ \ P'; \C. Prop C (\ \ \\<^sub>Q) P (M\\*xvec\\N\) P'; + extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \ \ \\<^sub>P \ Q \K\N\ \ Q'; \C. Prop C (\ \ \\<^sub>P) Q (K\N\) Q'; + extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + \ \ \\<^sub>P \ \\<^sub>Q \ M \ K; + A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* P'; + A\<^sub>P \* Q; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* xvec; A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; + A\<^sub>Q \* P; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* Q; A\<^sub>Q \* K; A\<^sub>Q \* Q'; distinct xvec; + A\<^sub>Q \* xvec; xvec \* \; xvec \* \\<^sub>P; xvec \* \\<^sub>Q; xvec \* P; xvec \* M; + xvec \* Q; xvec \* K; A\<^sub>P \* C; A\<^sub>Q \* C; xvec \* C\ \ + Prop C \ (P \ Q) (\) (\\*xvec\(P' \ Q'))" + and rBrMerge: "\\ \\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q Q' A\<^sub>Q C. + \\ \ \\<^sub>Q \ P \ \M\N\ \ P'; \C. Prop C (\ \ \\<^sub>Q) P (\M\N\) P'; + extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \ \ \\<^sub>P \ Q \ \M\N\ \ Q'; \C. Prop C (\ \ \\<^sub>P) Q (\M\N\) Q'; + extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* N; A\<^sub>P \* P'; + A\<^sub>P \* Q; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* M; A\<^sub>Q \* M; + A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; A\<^sub>Q \* P; A\<^sub>Q \* N; A\<^sub>Q \* P'; + A\<^sub>Q \* Q; A\<^sub>Q \* Q'; A\<^sub>P \* C; A\<^sub>Q \* C\ \ + Prop C \ (P \ Q) (\M\N\) (P' \ Q')" + and rBrComm1: "\\ \\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q xvec Q' A\<^sub>Q C. + \\ \ \\<^sub>Q \ P \\M\N\ \ P'; \C. Prop C (\ \ \\<^sub>Q) P (\M\N\) P'; + extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \ \ \\<^sub>P \ Q \\M\\*xvec\\N\ \ Q'; \C. Prop C (\ \ \\<^sub>P) Q (\M\\*xvec\\N\) Q'; + extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* N; A\<^sub>P \* P'; + A\<^sub>P \* Q; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* xvec; A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; + A\<^sub>Q \* P; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* Q; A\<^sub>Q \* Q'; distinct xvec; + A\<^sub>P \* M; A\<^sub>Q \* M; xvec \* M; + A\<^sub>Q \* xvec; xvec \* \; xvec \* \\<^sub>P; xvec \* \\<^sub>Q; xvec \* P; + xvec \* Q; A\<^sub>P \* C; A\<^sub>Q \* C; xvec \* C\ \ + Prop C \ (P \ Q) (\M\\*xvec\\N\) (P' \ Q')" + and rBrComm2: "\\ \\<^sub>Q P M xvec N P' A\<^sub>P \\<^sub>P Q Q' A\<^sub>Q C. + \\ \ \\<^sub>Q \ P \\M\\*xvec\\N\ \ P'; \C. Prop C (\ \ \\<^sub>Q) P (\M\\*xvec\\N\) P'; + extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \ \ \\<^sub>P \ Q \\M\N\ \ Q'; \C. Prop C (\ \ \\<^sub>P) Q (\M\N\) Q'; + extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* N; A\<^sub>P \* P'; + A\<^sub>P \* Q; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* xvec; A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; + A\<^sub>Q \* P; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* Q; A\<^sub>Q \* Q'; distinct xvec; + A\<^sub>P \* M; A\<^sub>Q \* M; xvec \* M; + A\<^sub>Q \* xvec; xvec \* \; xvec \* \\<^sub>P; xvec \* \\<^sub>Q; xvec \* P; + xvec \* Q; A\<^sub>P \* C; A\<^sub>Q \* C; xvec \* C\ \ + Prop C \ (P \ Q) (\M\\*xvec\\N\) (P' \ Q')" + and rBrClose: "\\ P M xvec N P' x C. + \\ \ P \ \M\\*xvec\\N\ \ P'; + \C. Prop C \ P (\M\\*xvec\\N\) P'; + x \ supp M; + distinct xvec; xvec \* \; xvec \* P; + xvec \* M; + x \ \; x \ xvec\ \ + Prop C \ (\\x\P) (\) (\\x\(\\*xvec\P'))" + and rOpen: "\\ P M xvec yvec N P' x C. + \\ \ P \M\\*(xvec@yvec)\\N\ \ P'; x \ supp N; \C. Prop C \ P (M\\*(xvec@yvec)\\N\) P'; + x \ \; x \ M; x \ xvec; x \ yvec; xvec \* \; xvec \* P; xvec \* M; distinct xvec; distinct yvec; + yvec \* \; yvec \* P; yvec \* M; yvec \* C; x \ C; xvec \* C\ \ + Prop C \ (\\x\P) (M\\*(xvec@x#yvec)\\N\) P'" + and rBrOpen: "\\ P M xvec yvec N P' x C. + \\ \ P \\M\\*(xvec@yvec)\\N\ \ P'; x \ supp N; \C. Prop C \ P (\M\\*(xvec@yvec)\\N\) P'; + x \ \; x \ M; x \ xvec; x \ yvec; xvec \* \; xvec \* P; xvec \* M; distinct xvec; distinct yvec; + yvec \* \; yvec \* P; yvec \* M; yvec \* C; x \ C; xvec \* C\ \ + Prop C \ (\\x\P) (\M\\*(xvec@x#yvec)\\N\) P'" + and rScope: "\\ P \ P' x C. + \\ \ P \\ \ P'; \C. Prop C \ P \ P'; + x \ \; x \ \; bn \ \* \; + bn \ \* P; bn \ \* (subject \); x \ C; bn \ \* C; distinct(bn \)\ \ + Prop C \ (\\x\P) \ (\\x\P')" + and rBang: "\\ P \ P' C. + \\ \ P \ !P \\ \ P'; guarded P; \C. Prop C \ (P \ !P) \ P'\ \ + Prop C \ (!P) \ P'" + +shows "Prop C \ P \ P'" + using \\ \ P \\ \ P'\ \bn \ \* (subject \)\ \distinct(bn \)\ +proof(nominal_induct x3=="\ \ P'" avoiding: \ C arbitrary: P' rule: semantics.strong_induct) + case(cInput \ M K xvec N Tvec P \ C P') + then show ?case by(force intro: rInput simp add: residualInject) +next + case(cBrInput \ M K xvec N Tvec P \ C P') + then show ?case + by(force simp add: rBrInput residualInject) +next + case(Output \ M K N P \ C P') + then show ?case by(force intro: rOutput simp add: residualInject) +next + case(BrOutput \ M K N P \ C P') + then show ?case by(force intro: rBrOutput simp add: residualInject) +next + case(Case \ P \ Cs \ C P') + then show ?case by(auto intro: rCase) +next + case(cPar1 \ \\<^sub>Q P \ P' Q A\<^sub>Q \' C P'') + note \\ \ (P' \ Q) = \' \ P''\ + moreover from \bn \ \* \'\ have "bn \ \* (bn \')" by auto + moreover note \distinct (bn \)\ \distinct(bn \')\ + moreover from \bn \ \* subject \\ \bn \' \* subject \'\ + have "bn \ \* (\ \ P' \ Q)" and "bn \' \* (\' \ P'')" by simp+ + ultimately obtain p where S: "(set p) \ (set(bn \)) \ (set(bn(p \ \)))" and "distinctPerm p" + and \Eq: "\' = p \ \" and P'eq: "P'' = p \ (P' \ Q)" and "(bn(p \ \)) \* \" + and "(bn(p \ \)) \* (P' \ Q)" + by(rule residualEq) + + note \\ \ \\<^sub>Q \ P \\ \ P'\ \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ \distinct A\<^sub>Q\ + moreover from \bn \ \* subject \\ \distinct(bn \)\ + have "\C. Prop C (\ \ \\<^sub>Q) P \ P'" by(metis cPar1) + moreover note \A\<^sub>Q \* P\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* \\ \A\<^sub>Q \* \\ \A\<^sub>Q \* P'\ \A\<^sub>Q \* C\ + \bn \ \* Q\ \distinct(bn \)\ \bn \ \* \\ \bn \ \* \\<^sub>Q\ \bn \ \* P\ \bn \ \* subject \\ \bn \ \* C\ + ultimately have "Prop C \ (P \ Q) \ (P' \ Q)" + by(metis rPar1) + + with \bn \ \* \\ \bn \ \* P\ \bn \ \* Q\ \bn \ \* subject \\ \bn \ \* C\ \bn \ \* bn \'\ S \distinctPerm p\ \bn(p \ \) \* \\ \bn(p \ \) \* (P' \ Q)\ \A\<^sub>Q \* C\ + have "Prop C \ (P \ Q) (p \ \) (p \ (P' \ Q))" + by - (rule rAlpha, auto) + with \Eq P'eq \distinctPerm p\ show ?case by simp +next + case(cPar2 \ \\<^sub>P Q \ Q' P A\<^sub>P \' C Q'') + note \\ \ (P \ Q') = \' \ Q''\ + moreover from \bn \ \* \'\ have "bn \ \* (bn \')" by auto + moreover note \distinct (bn \)\ \distinct(bn \')\ + moreover from \bn \ \* subject \\ \bn \' \* subject \'\ + have "bn \ \* (\ \ P \ Q')" and "bn \' \* (\' \ Q'')" by simp+ + ultimately obtain p where S: "(set p) \ (set(bn \)) \ (set(bn(p \ \)))" and "distinctPerm p" + and \Eq: "\' = p \ \" and Q'eq: "Q'' = p \ (P \ Q')" and "(bn(p \ \)) \* \" + and "(bn(p \ \)) \* (P \ Q')" + by(rule residualEq) + + note \\ \ \\<^sub>P \ Q \\ \ Q'\ \extractFrame P = \A\<^sub>P, \\<^sub>P\\ \distinct A\<^sub>P\ + moreover from \bn \ \* subject \\ \distinct(bn \)\ + have "\C. Prop C (\ \ \\<^sub>P) Q \ Q'" by(auto intro: cPar2) + + moreover note \A\<^sub>P \* P\ \A\<^sub>P \* Q\ \A\<^sub>P \* \\ \A\<^sub>P \* \\ \A\<^sub>P \* Q'\ \A\<^sub>P \* C\ + \bn \ \* Q\ \distinct(bn \)\ \bn \ \* \\ \bn \ \* \\<^sub>P\ \bn \ \* P\ \bn \ \* subject \\ \bn \ \* C\ + ultimately have "Prop C \ (P \ Q) \ (P \ Q')" + by(metis rPar2) + with \bn \ \* \\ \bn \ \* P\ \bn \ \* Q\ \bn \ \* subject \\ \bn \ \* C\ \bn \ \* (bn \')\ S \distinctPerm p\ \bn(p \ \) \* \\ \bn(p \ \) \* (P \ Q')\ + have "Prop C \ (P \ Q) (p \ \) (p \ (P \ Q'))" + by - (rule rAlpha, auto) + with \Eq Q'eq \distinctPerm p\ show ?case by simp +next + case(cComm1 \ \\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q K xvec Q' A\<^sub>Q \ C P'') + then have "Prop C \ (P \ Q) (\) (\\*xvec\(P' \ Q'))" + by(fastforce intro: rComm1) + then show ?case using \\ \ \\*xvec\(P' \ Q') = \ \ P''\ + by(simp add: residualInject) +next + case(cComm2 \ \\<^sub>Q P M xvec N P' A\<^sub>P \\<^sub>P Q K Q' A\<^sub>Q \ C P'') + then have "Prop C \ (P \ Q) (\) (\\*xvec\(P' \ Q'))" + by(fastforce intro: rComm2) + then show ?case using \\ \ \\*xvec\(P' \ Q') = \ \ P''\ + by(simp add: residualInject) +next + case (cBrMerge \ \\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q Q' A\<^sub>Q \ C P'') + then show ?case by(simp add: rBrMerge residualInject) +next + case(cBrComm1 \ \\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q xvec Q' A\<^sub>Q \ C P'') + from cBrComm1 \A\<^sub>P \* M\ \A\<^sub>Q \* M\ \xvec \* M\ have "Prop C \ (P \ Q) (\M\\*xvec\\N\) (P' \ Q')" + by(fastforce intro: rBrComm1) + + note \\M\\*xvec\\N\ \ P' \ Q' = \ \ P''\ + moreover from \xvec \* \\ have "bn (\M\\*xvec\\N\) \* (bn \)" by simp + moreover from \distinct xvec\ have "distinct (bn (\M\\*xvec\\N\))" by simp + moreover note \distinct (bn \)\ + moreover from \xvec \* M\ have "bn (\M\\*xvec\\N\) \* (\M\\*xvec\\N\ \ P' \ Q')" by simp + moreover from \bn \ \* subject \\ have "bn \ \* (\ \ P'')" by simp + ultimately obtain p where S: "(set p) \ (set(bn (\M\\*xvec\\N\))) \ (set(bn(p \ (\M\\*xvec\\N\))))" and "distinctPerm p" + and \Eq: "\ = p \ (\M\\*xvec\\N\)" and P'eq: "P'' = p \ (P' \ Q')" and "(bn(p \ (\M\\*xvec\\N\))) \* (\M\\*xvec\\N\)" + and "(bn(p \ (\M\\*xvec\\N\))) \* (P' \ Q')" + by(rule residualEq) + + from \xvec \* \\ have "bn (\M\\*xvec\\N\) \* \" by simp + moreover from \xvec \* P\ \xvec \* Q\ have "bn (\M\\*xvec\\N\) \* (P \ Q)" by simp + moreover from \xvec \* M\ have "bn (\M\\*xvec\\N\) \* subject (\M\\*xvec\\N\)" by simp + moreover from \xvec \* C\ have "bn (\M\\*xvec\\N\) \* C" by simp + moreover from \(bn(p \ (\M\\*xvec\\N\))) \* (\M\\*xvec\\N\)\ have "bn (\M\\*xvec\\N\) \* bn (p \ (\M\\*xvec\\N\))" by simp + moreover note \(set p) \ (set(bn (\M\\*xvec\\N\))) \ (set(bn(p \ (\M\\*xvec\\N\))))\ + moreover note \distinctPerm p\ + moreover note \(bn(p \ (\M\\*xvec\\N\))) \* (\M\\*xvec\\N\)\ + moreover note \(bn(p \ (\M\\*xvec\\N\))) \* (P' \ Q')\ + moreover note \Prop C \ (P \ Q) (\M\\*xvec\\N\) (P' \ Q')\ + + ultimately have propEqvt: "Prop C \ (P \ Q) (p \ (\M\\*xvec\\N\)) (p \ (P' \ Q'))" by(rule rAlpha) + + then show ?case by (simp add: \Eq P'eq propEqvt) +next + case(cBrComm2 \ \\<^sub>Q P M xvec N P' A\<^sub>P \\<^sub>P Q Q' A\<^sub>Q \ C P'') + from cBrComm2 \A\<^sub>P \* M\ \A\<^sub>Q \* M\ \xvec \* M\ have "Prop C \ (P \ Q) (\M\\*xvec\\N\) (P' \ Q')" + by(fastforce intro: rBrComm2) + + note \\M\\*xvec\\N\ \ P' \ Q' = \ \ P''\ + moreover from \xvec \* \\ have "bn (\M\\*xvec\\N\) \* (bn \)" by simp + moreover from \distinct xvec\ have "distinct (bn (\M\\*xvec\\N\))" by simp + moreover note \distinct (bn \)\ + moreover from \xvec \* M\ have "bn (\M\\*xvec\\N\) \* (\M\\*xvec\\N\ \ P' \ Q')" by simp + moreover from \bn \ \* subject \\ have "bn \ \* (\ \ P'')" by simp + ultimately obtain p where S: "(set p) \ (set(bn (\M\\*xvec\\N\))) \ (set(bn(p \ (\M\\*xvec\\N\))))" and "distinctPerm p" + and \Eq: "\ = p \ (\M\\*xvec\\N\)" and P'eq: "P'' = p \ (P' \ Q')" and "(bn(p \ (\M\\*xvec\\N\))) \* (\M\\*xvec\\N\)" + and "(bn(p \ (\M\\*xvec\\N\))) \* (P' \ Q')" + by(rule residualEq) + + from \xvec \* \\ have "bn (\M\\*xvec\\N\) \* \" by simp + moreover from \xvec \* P\ \xvec \* Q\ have "bn (\M\\*xvec\\N\) \* (P \ Q)" by simp + moreover from \xvec \* M\ have "bn (\M\\*xvec\\N\) \* subject (\M\\*xvec\\N\)" by simp + moreover from \xvec \* C\ have "bn (\M\\*xvec\\N\) \* C" by simp + moreover from \(bn(p \ (\M\\*xvec\\N\))) \* (\M\\*xvec\\N\)\ have "bn (\M\\*xvec\\N\) \* bn (p \ (\M\\*xvec\\N\))" by simp + moreover note \(set p) \ (set(bn (\M\\*xvec\\N\))) \ (set(bn(p \ (\M\\*xvec\\N\))))\ + moreover note \distinctPerm p\ + moreover note \(bn(p \ (\M\\*xvec\\N\))) \* (\M\\*xvec\\N\)\ + moreover note \(bn(p \ (\M\\*xvec\\N\))) \* (P' \ Q')\ + moreover note \Prop C \ (P \ Q) (\M\\*xvec\\N\) (P' \ Q')\ + + ultimately have propEqvt: "Prop C \ (P \ Q) (p \ (\M\\*xvec\\N\)) (p \ (P' \ Q'))" by(rule rAlpha) + + then show ?case by (simp add: \Eq P'eq propEqvt) +next + case (cBrClose \ P M xvec N P' x \ C P'') + then have "Prop C \ (\\x\P) (\) (\\x\(\\*xvec\P'))" + by(fastforce intro: rBrClose) + then show ?case using \\ \ \\x\(\\*xvec\P') = \ \ P''\ + by(simp add: residualInject) +next + case(cOpen \ P M xvec yvec N P' x \ C P'') + note \M\\*(xvec@x#yvec)\\N\ \ P' = \ \ P''\ + moreover from \xvec \* \\ \x \ \\ \yvec \* \\ have "(xvec@x#yvec) \* (bn \)" + by auto + moreover from \xvec \* yvec\ \x \ xvec\ \x \ yvec\ \distinct xvec\ \distinct yvec\ + have "distinct(xvec@x#yvec)" + by(auto simp add: fresh_star_def) (simp add: fresh_def name_list_supp) + moreover note \distinct(bn \)\ + moreover from \xvec \* M\ \x \ M\ \yvec \* M\ have "(xvec@x#yvec) \* M" by auto + then have "(xvec@x#yvec) \* (M\\*(xvec@x#yvec)\\N\ \ P')" by auto + moreover from \bn \ \* subject \\ have "bn \ \* (\ \ P'')" by simp + ultimately obtain p where S: "(set p) \ (set(xvec@x#yvec)) \ (set(p \ (xvec@x#yvec)))" and "distinctPerm p" + and \eq: "\ = (p \ M)\\*(p \ (xvec@x#yvec))\\(p \ N)\" and P'eq: "P'' = (p \ P')" + and A: "(xvec@x#yvec) \* ((p \ M)\\*(p \ (xvec@x#yvec))\\(p \ N)\)" + and B: "(p \ (xvec@x#yvec)) \* (M\\*(xvec@x#yvec)\\N\)" + and C: "(p \ (xvec@x#yvec)) \* P'" + by - (rule residualEq, (assumption | simp)+) + note \\ \ P \M\\*(xvec@yvec)\\N\ \ P'\ \x \ (supp N)\ + + moreover { + fix C + from \xvec \* M\ \yvec \* M\ have "(xvec@yvec) \* M" by simp + moreover from \distinct xvec\ \distinct yvec\ \xvec \* yvec\ have "distinct(xvec@yvec)" + by auto (simp add: fresh_star_def name_list_supp fresh_def) + ultimately have "Prop C \ P (M\\*(xvec@yvec)\\N\) P'" by(fastforce intro: cOpen) + } + + moreover note \x \ \\ \x \ M\ \x \ xvec\ \x \ yvec\ \xvec \* \\ \xvec \* P\ \xvec \* M\ + \yvec \* \\ \yvec \* P\ \yvec \* M\ \yvec \* C\ \x \ C\ \xvec \* C\ \distinct xvec\ \distinct yvec\ + ultimately have "Prop C \ (\\x\P) (M\\*(xvec@x#yvec)\\N\) P'" + by(metis rOpen) + + with \xvec \* \\ \yvec \* \\ \xvec \* P\ \yvec \* P\ \xvec \* M\ \yvec \* M\ + \yvec \* C\ S \distinctPerm p\ \x \ C\ \xvec \* C\ + \x \ \\ \x \ M\ \x \ xvec\ \x \ yvec\ A B C + have "Prop C \ (\\x\P) (p \ (M\\*(xvec@x#yvec)\\N\)) (p \ P')" + apply - + apply(rule rAlpha[where \="M\\*(xvec@x#yvec)\\N\"];clarsimp) + by (meson abs_fresh(1) abs_fresh_list_star' freshChainAppend freshSets(5) psiFreshVec(5)) + with \eq P'eq show ?case by simp +next + case(cBrOpen \ P M xvec yvec N P' x \ C P'') + note \\M\\*(xvec@x#yvec)\\N\ \ P' = \ \ P''\ + moreover from \xvec \* \\ \x \ \\ \yvec \* \\ have "(xvec@x#yvec) \* (bn \)" + by auto + moreover from \xvec \* yvec\ \x \ xvec\ \x \ yvec\ \distinct xvec\ \distinct yvec\ + have "distinct(xvec@x#yvec)" + by(clarsimp simp add: fresh_star_def; safe; simp add: fresh_def name_list_supp) + moreover note \distinct(bn \)\ + moreover from \xvec \* M\ \x \ M\ \yvec \* M\ have "(xvec@x#yvec) \* M" by auto + then have "(xvec@x#yvec) \* (\M\\*(xvec@x#yvec)\\N\ \ P')" by auto + moreover from \bn \ \* subject \\ have "bn \ \* (\ \ P'')" by simp + ultimately obtain p where S: "(set p) \ (set(xvec@x#yvec)) \ (set(p \ (xvec@x#yvec)))" and "distinctPerm p" + and \eq: "\ = \(p \ M)\\*(p \ (xvec@x#yvec))\\(p \ N)\" and P'eq: "P'' = (p \ P')" + and A: "(xvec@x#yvec) \* (\(p \ M)\\*(p \ (xvec@x#yvec))\\(p \ N)\)" + and B: "(p \ (xvec@x#yvec)) \* (\M\\*(xvec@x#yvec)\\N\)" + and C: "(p \ (xvec@x#yvec)) \* P'" + apply - + by(rule residualEq) (assumption|simp)+ + note \\ \ P \\M\\*(xvec@yvec)\\N\ \ P'\ \x \ (supp N)\ + + moreover { + fix C + from \xvec \* M\ \yvec \* M\ have "(xvec@yvec) \* M" by simp + moreover from \distinct xvec\ \distinct yvec\ \xvec \* yvec\ have "distinct(xvec@yvec)" + by auto (simp add: fresh_star_def name_list_supp fresh_def) + ultimately have "Prop C \ P (\M\\*(xvec@yvec)\\N\) P'" by(fastforce intro: cBrOpen) + } + + moreover note \x \ \\ \x \ M\ \x \ xvec\ \x \ yvec\ \xvec \* \\ \xvec \* P\ \xvec \* M\ + \yvec \* \\ \yvec \* P\ \yvec \* M\ \yvec \* C\ \x \ C\ \xvec \* C\ \distinct xvec\ \distinct yvec\ + ultimately have "Prop C \ (\\x\P) (\M\\*(xvec@x#yvec)\\N\) P'" + by(metis rBrOpen) + + with \xvec \* \\ \yvec \* \\ \xvec \* P\ \yvec \* P\ \xvec \* M\ \yvec \* M\ + \yvec \* C\ S \distinctPerm p\ \x \ C\ \xvec \* C\ + \x \ \\ \x \ M\ \x \ xvec\ \x \ yvec\ A B C + have "Prop C \ (\\x\P) (p \ (\M\\*(xvec@x#yvec)\\N\)) (p \ P')" + apply - + apply(rule rAlpha[where \="\M\\*(xvec@x#yvec)\\N\"]; clarsimp) + by (meson abs_fresh(1) abs_fresh_list_star' freshChainAppend freshSets(5) psiFreshVec(5)) + with \eq P'eq show ?case by simp +next + case(cScope \ P \ P' x \' C P'') + note \\ \ (\\x\P') = \' \ P''\ + moreover from \bn \ \* \'\ have "bn \ \* (bn \')" by auto + moreover note \distinct (bn \)\ \distinct(bn \')\ + moreover from \bn \ \* subject \\ \bn \' \* subject \'\ + have "bn \ \* (\ \ \\x\P')" and "bn \' \* (\' \ P'')" by simp+ + ultimately obtain p where S: "(set p) \ (set(bn \)) \ (set(bn(p \ \)))" and "distinctPerm p" + and \Eq: "\' = p \ \" and P'eq: "P'' = p \ (\\x\P')" and "(bn(p \ \)) \* \" + and "(bn(p \ \)) \* (\\x\P')" + by(rule residualEq) + + note \\ \ P \\ \ P'\ + moreover from \bn \ \* subject \\ \distinct(bn \)\ + have "\C. Prop C \ P \ P'" by(fastforce intro: cScope) + + moreover note \x \ \\ \x \ \\ \bn \ \* \\ \bn \ \* P\ \bn \ \* subject \\ + \x \ C\ \bn \ \* C\ \distinct(bn \)\ + ultimately have "Prop C \ (\\x\P) \ (\\x\P')" + by(rule rScope) + with \bn \ \* \\ \bn \ \* P\ \x \ \\ \bn \ \* subject \\ \bn \ \* C\ \bn \ \* (bn \')\ S \distinctPerm p\ \bn(p \ \) \* \\ \bn(p \ \) \* (\\x\P')\ + have "Prop C \ (\\x\P) (p \ \) (p \ (\\x\P'))" + by(fastforce intro: rAlpha) + with \Eq P'eq \distinctPerm p\ show ?case by simp +next + case(Bang \ P \ C P') + then show ?case by(fastforce intro: rBang) +qed + +lemma outputInduct[consumes 1, case_names cOutput cCase cPar1 cPar2 cOpen cScope cBang]: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and M :: 'a + and B :: "('a, 'b, 'c) boundOutput" + and Prop :: "'f::fs_name \ 'b \ ('a, 'b, 'c) psi \ + 'a \ ('a, 'b, 'c) boundOutput \ bool" + and C :: "'f::fs_name" + +assumes "\ \ P \ROut M B" + and rOutput: "\\ M K N P C. \\ \ M \ K\ \ Prop C \ (M\N\.P) K (N \' P)" + and rCase: "\\ P M B \ Cs C. + \\ \ P \(ROut M B); \C. Prop C \ P M B; (\, P) \ set Cs; \ \ \; guarded P\ \ + Prop C \ (Cases Cs) M B" + and rPar1: "\\ \\<^sub>Q P M xvec N P' A\<^sub>Q Q C. + \\ \ \\<^sub>Q \ P \M\\*xvec\\N\ \ P'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + \C. Prop C (\ \ \\<^sub>Q) P M (\\*xvec\N \' P'); + A\<^sub>Q \* P; A\<^sub>Q \* Q; A\<^sub>Q \* \; A\<^sub>Q \* M; + A\<^sub>Q \* xvec; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* C; xvec \* Q; + xvec \* \; xvec \* \\<^sub>Q; xvec \* P; xvec \* M; xvec \* C\ \ + Prop C \ (P \ Q) M (\\*xvec\N \' (P' \ Q))" + and rPar2: "\\ \\<^sub>P Q M xvec N Q' A\<^sub>P P C. + \\ \ \\<^sub>P \ Q \M\\*xvec\\N\ \ Q'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \C. Prop C (\ \ \\<^sub>P) Q M (\\*xvec\N \' Q'); + A\<^sub>P \* P; A\<^sub>P \* Q; A\<^sub>P \* \; A\<^sub>P \* M; + A\<^sub>P \* xvec; A\<^sub>P \* N; A\<^sub>P \* Q'; A\<^sub>P \* C; xvec \* P; + xvec \* \; xvec \* \\<^sub>P; xvec \* Q; xvec \* M; xvec \* C\ \ + Prop C \ (P \ Q) M (\\*xvec\N \' (P \ Q'))" + and rOpen: "\\ P M xvec yvec N P' x C. + \\ \ P \M\\*(xvec@yvec)\\N\ \ P'; x \ supp N; \C. Prop C \ P M (\\*(xvec@yvec)\N \' P'); + x \ \; x \ M; x \ xvec; x \ yvec; xvec \* \; xvec \* P; xvec \* M; + xvec \* yvec; yvec \* \; yvec \* P; yvec \* M; yvec \* C; x \ C; xvec \* C\ \ + Prop C \ (\\x\P) M (\\*(xvec@x#yvec)\N \' P')" + and rScope: "\\ P M xvec N P' x C. + \\ \ P \M\\*xvec\\N\ \ P'; \C. Prop C \ P M (\\*xvec\N \' P'); + x \ \; x \ M; x \ xvec; x \ N; xvec \* \; xvec \* P; xvec \* M; + x \ C; xvec \* C\ \ + Prop C \ (\\x\P) M (\\*xvec\N \' \\x\P')" + and rBang: "\\ P M B C. + \\ \ P \ !P \(ROut M B); guarded P; \C. Prop C \ (P \ !P) M B\ \ + Prop C \ (!P) M B" +shows "Prop C \ P M B" + using \\ \ P \(ROut M B)\ +proof(nominal_induct \ P Rs=="(ROut M B)" avoiding: C arbitrary: B rule: semantics.strong_induct) + case(cInput \ M K xvec N Tvec P C) + then show ?case by(simp add: residualInject) +next + case cBrInput + then show ?case by(simp add: residualInject) +next + case(Output \ M K N P C) + then show ?case by(force simp add: residualInject intro: rOutput) +next + case(BrOutput \ M N P C) + then show ?case by(simp add: residualInject) +next + case(Case \ P \ Cs C B) + then show ?case by(force intro: rCase) +next + case(cPar1 \ \\<^sub>Q P \ P' Q A\<^sub>Q C) + then show ?case by(force intro: rPar1 simp add: residualInject) +next + case(cPar2 \ \\<^sub>P Q \ Q' P A\<^sub>P C) + then show ?case by(force intro: rPar2 simp add: residualInject) +next + case cComm1 + then show ?case by(simp add: residualInject) +next + case cComm2 + then show ?case by(simp add: residualInject) +next + case cBrMerge + then show ?case by(simp add: residualInject) +next + case (cBrComm1 \ \\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q xvec Q' A\<^sub>Q C B) + then show ?case by(simp add: residualInject) +next + case (cBrComm2 \ \\<^sub>Q P M xvec N P' A\<^sub>P \\<^sub>P Q Q' A\<^sub>Q C B) + then show ?case by(simp add: residualInject) +next + case(cBrClose \ P M xvec N P' C B) + then show ?case by(simp add: residualInject) +next + case(cOpen \ P M xvec yvec N P' x C B) + then show ?case by(force intro: rOpen simp add: residualInject) +next + case cBrOpen + then show ?case by(simp add: residualInject) +next + case(cScope \ P M \ P' x C) + then show ?case by(force intro: rScope simp add: residualInject) +next + case(Bang \ P C B) + then show ?case by(force intro: rBang) +qed + +lemma brOutputInduct[consumes 1, case_names cBrOutput cCase cPar1 cPar2 cBrComm1 cBrComm2 cBrOpen cScope cBang]: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and M :: 'a + and B :: "('a, 'b, 'c) boundOutput" + and Prop :: "'f::fs_name \ 'b \ ('a, 'b, 'c) psi \ + 'a \ ('a, 'b, 'c) boundOutput \ bool" + and C :: "'f::fs_name" + +assumes "\ \ P \RBrOut M B" + and rBrOutput: "\\ M K N P C. \\ \ M \ K\ \ Prop C \ (M\N\.P) K (N \' P)" + and rCase: "\\ P M B \ Cs C. + \\ \ P \(RBrOut M B); \C. Prop C \ P M B; (\, P) \ set Cs; \ \ \; guarded P\ \ + Prop C \ (Cases Cs) M B" + and rPar1: "\\ \\<^sub>Q P M xvec N P' A\<^sub>Q Q C. + \\ \ \\<^sub>Q \ P \\M\\*xvec\\N\ \ P'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + \C. Prop C (\ \ \\<^sub>Q) P M (\\*xvec\N \' P'); + A\<^sub>Q \* P; A\<^sub>Q \* Q; A\<^sub>Q \* \; A\<^sub>Q \* M; + A\<^sub>Q \* xvec; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* C; xvec \* Q; + xvec \* \; xvec \* \\<^sub>Q; xvec \* P; xvec \* M; xvec \* C\ \ + Prop C \ (P \ Q) M (\\*xvec\N \' (P' \ Q))" + and rPar2: "\\ \\<^sub>P Q M xvec N Q' A\<^sub>P P C. + \\ \ \\<^sub>P \ Q \\M\\*xvec\\N\ \ Q'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \C. Prop C (\ \ \\<^sub>P) Q M (\\*xvec\N \' Q'); + A\<^sub>P \* P; A\<^sub>P \* Q; A\<^sub>P \* \; A\<^sub>P \* M; + A\<^sub>P \* xvec; A\<^sub>P \* N; A\<^sub>P \* Q'; A\<^sub>P \* C; xvec \* P; + xvec \* \; xvec \* \\<^sub>P; xvec \* Q; xvec \* M; xvec \* C\ \ + Prop C \ (P \ Q) M (\\*xvec\N \' (P \ Q'))" + and rBrComm1: "\\ \\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q xvec Q' A\<^sub>Q C. + \\ \ \\<^sub>Q \ P \\M\N\ \ P'; + extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \ \ \\<^sub>P \ Q \\M\\*xvec\\N\ \ Q'; \C. Prop C (\ \ \\<^sub>P) Q M (\\*xvec\N \' Q'); + extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* N; A\<^sub>P \* P'; + A\<^sub>P \* Q; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* xvec; A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; + A\<^sub>Q \* P; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* Q; A\<^sub>Q \* Q'; distinct xvec; + A\<^sub>P \* M; A\<^sub>Q \* M; xvec \* M; + A\<^sub>Q \* xvec; xvec \* \; xvec \* \\<^sub>P; xvec \* \\<^sub>Q; xvec \* P; + xvec \* Q; A\<^sub>P \* C; A\<^sub>Q \* C; xvec \* C\ \ + Prop C \ (P \ Q) M (\\*xvec\N \' (P' \ Q'))" (* Removed: "\C. Prop C (\ \ \\<^sub>Q) P (M\N\ \ P');" *) + and rBrComm2: "\\ \\<^sub>Q P M xvec N P' A\<^sub>P \\<^sub>P Q Q' A\<^sub>Q C. + \\ \ \\<^sub>Q \ P \\M\\*xvec\\N\ \ P'; \C. Prop C (\ \ \\<^sub>Q) P M (\\*xvec\N \' P'); + extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \ \ \\<^sub>P \ Q \\M\N\ \ Q'; + extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* N; A\<^sub>P \* P'; + A\<^sub>P \* Q; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* xvec; A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; + A\<^sub>Q \* P; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* Q; A\<^sub>Q \* Q'; distinct xvec; + A\<^sub>P \* M; A\<^sub>Q \* M; xvec \* M; + A\<^sub>Q \* xvec; xvec \* \; xvec \* \\<^sub>P; xvec \* \\<^sub>Q; xvec \* P; + xvec \* Q; A\<^sub>P \* C; A\<^sub>Q \* C; xvec \* C\ \ + Prop C \ (P \ Q) M (\\*xvec\N \' (P' \ Q'))" (* Removed: "\C. Prop C (\ \ \\<^sub>P) Q (K\N\) Q';" *) + and rBrOpen: "\\ P M xvec yvec N P' x C. + \\ \ P \\M\\*(xvec@yvec)\\N\ \ P'; x \ supp N; \C. Prop C \ P M (\\*(xvec@yvec)\N \' P'); + x \ \; x \ M; x \ xvec; x \ yvec; xvec \* \; xvec \* P; xvec \* M; + xvec \* yvec; yvec \* \; yvec \* P; yvec \* M; yvec \* C; x \ C; xvec \* C\ \ + Prop C \ (\\x\P) M (\\*(xvec@x#yvec)\N \' P')" + and rScope: "\\ P M xvec N P' x C. + \\ \ P \\M\\*xvec\\N\ \ P'; \C. Prop C \ P M (\\*xvec\N \' P'); + x \ \; x \ M; x \ xvec; x \ N; xvec \* \; xvec \* P; xvec \* M; + x \ C; xvec \* C\ \ + Prop C \ (\\x\P) M (\\*xvec\N \' \\x\P')" + and rBang: "\\ P M B C. + \\ \ P \ !P \(RBrOut M B); guarded P; \C. Prop C \ (P \ !P) M B\ \ + Prop C \ (!P) M B" +shows "Prop C \ P M B" + using \\ \ P \(RBrOut M B)\ +proof(nominal_induct \ P Rs=="(RBrOut M B)" avoiding: C arbitrary: B rule: semantics.strong_induct) + case(cInput \ M K xvec N Tvec P C) + then show ?case by(simp add: residualInject) +next + case cBrInput + then show ?case by(simp add: residualInject) +next + case(Output \ M K N P C) + then show ?case by(simp add: residualInject) +next + case(BrOutput \ M N P C) + then show ?case by(auto simp add: residualInject intro: rBrOutput) +next + case(Case \ P \ Cs C B) + then show ?case by(force intro: rCase) +next + case(cPar1 \ \\<^sub>Q P \ P' Q A\<^sub>Q C) + then show ?case by(force intro: rPar1 simp add: residualInject) +next + case(cPar2 \ \\<^sub>P Q \ Q' P A\<^sub>P C) + then show ?case by(force intro: rPar2 simp add: residualInject) +next + case cComm1 + then show ?case by(simp add: residualInject) +next + case cComm2 + then show ?case by(simp add: residualInject) +next + case cBrMerge + then show ?case by(simp add: residualInject) +next + case (cBrComm1 \ \\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q xvec Q' A\<^sub>Q C B) + then show ?case by(force intro: rBrComm1 simp add: residualInject) +next + case (cBrComm2 \ \\<^sub>Q P M xvec N P' A\<^sub>P \\<^sub>P Q Q' A\<^sub>Q C B) + then show ?case by(force intro: rBrComm2 simp add: residualInject) +next + case(cBrClose \ P M xvec N P' C B) + then show ?case by(simp add: residualInject) +next + case(cOpen \ P M xvec yvec N P' x C B) + then show ?case by(simp add: residualInject) +next + case(cBrOpen \ P M xvec yvec N P' x C B) + then show ?case by(force intro: rBrOpen simp add: residualInject) +next + case(cScope \ P M \ P' x C) + then show ?case by(force intro: rScope simp add: residualInject) +next + case(Bang \ P C B) + then show ?case by(force intro: rBang) +qed + +lemma boundOutputBindObject: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and M :: 'a + and yvec :: "name list" + and N :: 'a + and P' :: "('a, 'b, 'c) psi" + and y :: name + +assumes "\ \ P \\ \ P'" + and "bn \ \* subject \" + and "distinct(bn \)" + and "y \ set(bn \)" + +shows "y \ supp(object \)" + using assms +proof(nominal_induct avoiding: P' arbitrary: y rule: semanticsInduct) + case(cAlpha \ P \ P' p P'' y) + from \y \ set(bn(p \ \))\ have "(p \ y) \ (p \ set(bn(p \ \)))" + by(rule pt_set_bij2[OF pt_name_inst, OF at_name_inst]) + then have "(p \ y) \ set(bn \)" using \distinctPerm p\ + by(simp add: eqvts) + then have "(p \ y) \ supp(object \)" by(rule cAlpha) + then have "(p \ p \ y) \ (p \ supp(object \))" + by(rule pt_set_bij2[OF pt_name_inst, OF at_name_inst]) + then show ?case using \distinctPerm p\ + by(simp add: eqvts) +next + case cInput + then show ?case by simp +next + case cBrInput + then show ?case by simp +next + case cOutput + then show ?case by simp +next + case cBrOutput + then show ?case by simp +next + case cCase + then show ?case by simp +next + case cPar1 + then show ?case by simp +next + case cPar2 + then show ?case by simp +next + case cComm1 + then show ?case by simp +next + case cComm2 + then show ?case by simp +next + case cBrMerge + then show ?case by simp +next + case cBrComm1 + then show ?case by simp +next + case cBrComm2 + then show ?case by simp +next + case cBrClose + then show ?case by simp +next + case cOpen + then show ?case by(auto simp add: supp_list_cons supp_list_append supp_atm supp_some) +next + case cBrOpen + then show ?case by(auto simp add: supp_list_cons supp_list_append supp_atm supp_some) +next + case cScope + then show ?case by simp +next + case cBang + then show ?case by simp +qed + +lemma alphaBoundOutputChain': + fixes yvec :: "name list" + and xvec :: "name list" + and B :: "('a, 'b, 'c) boundOutput" + +assumes "length xvec = length yvec" + and "yvec \* B" + and "yvec \* xvec" + and "distinct yvec" + +shows "\\*xvec\B = \\*yvec\([xvec yvec] \\<^sub>v B)" + using assms +proof(induct rule: composePermInduct) + case cBase + show ?case by simp +next + case(cStep x xvec y yvec) + then show ?case + by (auto simp add: alphaBoundOutput[of y] eqvts) +qed + +lemma alphaBoundOutputChain'': + fixes yvec :: "name list" + and xvec :: "name list" + and N :: 'a + and P :: "('a, 'b, 'c) psi" + +assumes "length xvec = length yvec" + and "yvec \* N" + and "yvec \* P" + and "yvec \* xvec" + and "distinct yvec" + +shows "\\*xvec\(N \' P) = \\*yvec\(([xvec yvec] \\<^sub>v N) \' ([xvec yvec] \\<^sub>v P))" +proof - + from assms have "\\*xvec\(N \' P) = \\*yvec\([xvec yvec] \\<^sub>v (N \' P))" + by(simp add: alphaBoundOutputChain') + then show ?thesis by simp +qed + +lemma alphaDistinct: + fixes xvec :: "name list" + and N :: 'a + and P :: "('a, 'b, 'c) psi" + and yvec :: "name list" + and M :: 'a + and Q :: "('a, 'b, 'c) psi" + +assumes "\ \ P = \ \ Q" + and "distinct(bn \)" + and "\x. x \ set(bn \) \ x \ supp(object \)" + and "bn \ \* bn \" + and "bn \ \* (object \)" + and "bn \ \* Q" + +shows "distinct(bn \)" + using assms +proof - + { + fix xvec M yvec N + assume Eq: "\\*xvec\N \' P = \\*yvec\M \' Q" + assume "distinct xvec" and "xvec \* M" and "xvec \* yvec" and "xvec \* Q" + assume Mem: "\x. x \ set xvec \ x \ (supp N)" + have "distinct yvec" + proof - + from Eq have "length xvec = length yvec" + by(rule boundOutputChainEqLength) + with Eq \distinct xvec\ \xvec \* yvec\ \xvec \* M\ \xvec \* Q\ Mem show ?thesis + proof(induct n=="length xvec" arbitrary: xvec yvec M Q rule: nat.induct) + case(zero xvec yvec M Q) + then show ?case by simp + next + case(Suc n xvec yvec M Q) + have L: "length xvec = length yvec" and "Suc n = length xvec" by fact+ + then obtain x xvec' y yvec' where xEq: "xvec = x#xvec'" and yEq: "yvec = y#yvec'" + and L': "length xvec' = length yvec'" + by(cases xvec, auto, cases yvec, auto) + have xvecFreshyvec: "xvec \* yvec" and xvecDist: "distinct xvec" by fact+ + with xEq yEq have xineqy: "x \ y" and xvec'Freshyvec': "xvec' \* yvec'" + and xvec'Dist: "distinct xvec'" and xFreshxvec': "x \ xvec'" + and xFreshyvec': "x \ yvec'" and yFreshxvec': "y \ xvec'" + by auto + have Eq: "\\*xvec\N \' P = \\*yvec\M \' Q" by fact + with xEq yEq xineqy have Eq': "\\*xvec'\N \' P = \\*([(x, y)] \ yvec')\([(x, y)] \ M) \' ([(x, y)] \ Q)" + by(simp add: boundOutput.inject alpha eqvts) + moreover have Mem:"\x. x \ set xvec \ x \ supp N" by fact + with xEq have "\x. x \ set xvec' \ x \ supp N" by simp + moreover have "xvec \* M" by fact + with xEq xFreshxvec' yFreshxvec' have "xvec' \* ([(x, y)] \ M)" by simp + moreover have xvecFreshQ: "xvec \* Q" by fact + with xEq xFreshxvec' yFreshxvec' have "xvec' \* ([(x, y)] \ Q)" by simp + moreover have "Suc n = length xvec" by fact + with xEq have "n = length xvec'" by simp + moreover from xvec'Freshyvec' xFreshxvec' yFreshxvec' have "xvec' \* ([(x, y)] \ yvec')" + by simp + moreover from L' have "length xvec' = length([(x, y)] \ yvec')" by simp + ultimately have "distinct([(x, y)] \ yvec')" using xvec'Dist + apply - + apply(rule Suc) + by(assumption | simp)+ + then have "distinct yvec'" by simp + from Mem xEq have xSuppN: "x \ supp N" by simp + from L \distinct xvec\ \xvec \* yvec\ \xvec \* M\ \xvec \* Q\ + have "\\*yvec\M \' Q = \\*xvec\([yvec xvec] \\<^sub>v M) \' ([yvec xvec] \\<^sub>v Q)" + by(simp add: alphaBoundOutputChain'') + with Eq have "N = [yvec xvec] \\<^sub>v M" by simp + with xEq yEq have "N = [(y, x)] \ [yvec' xvec'] \\<^sub>v M" + by simp + with xSuppN have ySuppM: "y \ supp([yvec' xvec'] \\<^sub>v M)" + by(force simp add: calc_atm eqvts name_swap + dest: pt_set_bij2[where pi="[(x, y)]",OF pt_name_inst, OF at_name_inst]) + have "y \ yvec'" + proof - + { + assume "y \ supp yvec'" + then have "y \ set yvec'" + by(induct yvec') (auto simp add: supp_list_nil supp_list_cons supp_atm) + moreover from \xvec \* M\ xEq xFreshxvec' have "xvec' \* M" by simp + ultimately have "y \ [yvec' xvec'] \\<^sub>v M" using L' xvec'Freshyvec' xvec'Dist + by(force intro: freshChainPerm) + with ySuppM have "False" by(simp add: fresh_def) + } + then show ?thesis + by(simp add: fresh_def, rule notI) + qed + with \distinct yvec'\ yEq show ?case by simp + qed + qed + } note res = this + show ?thesis + apply(rule actionCases[where \=\]) + using assms res + by(auto simp add: residualInject supp_some) +qed + +lemma boundOutputDistinct: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and \ :: "'a action" + and P' :: "('a, 'b, 'c) psi" + +assumes "\ \ P \\ \ P'" + +shows "distinct(bn \)" + using assms +proof(nominal_induct \ P x3=="\ \ P'" avoiding: \ P' rule: semantics.strong_induct) + case cPar1 + then show ?case + by(force intro: alphaDistinct boundOutputBindObject) +next + case cPar2 + then show ?case + by(force intro: alphaDistinct boundOutputBindObject) +next + case (cBrComm1 \ \\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q xvec Q' A\<^sub>Q \ P'') + note \\ \ \\<^sub>P \ Q \ \M\\*xvec\\N\ \ Q'\ + moreover from \xvec \* M\ have "bn (\M\\*xvec\\N\) \* subject (\M\\*xvec\\N\)" by simp + moreover from \distinct xvec\ have "distinct (bn (\M\\*xvec\\N\))" by simp + ultimately have someX: "\ x. x \ set (bn (\M\\*xvec\\N\)) \ x \ supp (object (\M\\*xvec\\N\))" + by (drule boundOutputBindObject) (assumption) + + note \\M\\*xvec\\N\ \ P' \ Q' = \ \ P''\ + moreover from \distinct xvec\ have "distinct (bn (\M\\*xvec\\N\))" by simp + moreover note \\x. x \ set (bn (\M\\*xvec\\N\)) \ x \ supp (object (\M\\*xvec\\N\))\ + moreover from \xvec \* \\ have "bn (\M\\*xvec\\N\) \* bn \" by simp + moreover from \xvec \* \\ have "bn (\M\\*xvec\\N\) \* object \" by simp + moreover from \xvec \* P''\ have "bn (\M\\*xvec\\N\) \* P''" by simp + ultimately show ?case + by(rule alphaDistinct) +next + case (cBrComm2 \ \\<^sub>Q P M xvec N P' A\<^sub>P \\<^sub>P Q Q' A\<^sub>Q \ P'') + note \\ \ \\<^sub>Q \ P \ \M\\*xvec\\N\ \ P'\ + moreover from \xvec \* M\ have "bn (\M\\*xvec\\N\) \* subject (\M\\*xvec\\N\)" by simp + moreover from \distinct xvec\ have "distinct (bn (\M\\*xvec\\N\))" by simp + ultimately have someX: "\ x. x \ set (bn (\M\\*xvec\\N\)) \ x \ supp (object (\M\\*xvec\\N\))" + by (drule boundOutputBindObject) (assumption) + + note \\M\\*xvec\\N\ \ P' \ Q' = \ \ P''\ + moreover from \distinct xvec\ have "distinct (bn (\M\\*xvec\\N\))" by simp + moreover note \\x. x \ set (bn (\M\\*xvec\\N\)) \ x \ supp (object (\M\\*xvec\\N\))\ + moreover from \xvec \* \\ have "bn (\M\\*xvec\\N\) \* bn \" by simp + moreover from \xvec \* \\ have "bn (\M\\*xvec\\N\) \* object \" by simp + moreover from \xvec \* P''\ have "bn (\M\\*xvec\\N\) \* P''" by simp + ultimately show ?case + by(rule alphaDistinct) +next + case(cBrClose \ P M xvec N P' \ P'') + then show ?case by(simp add: residualInject) +next + case(cOpen \ P M xvec yvec N P' x \ P'') + note \M\\*(xvec@x#yvec)\\N\ \ P' = \ \ P''\ + moreover from \xvec \* yvec\ \x \ xvec\ \x \ yvec\ \distinct xvec\ \distinct yvec\ + have "distinct(bn(M\\*(xvec@x#yvec)\\N\))" + by auto (simp add: fresh_star_def fresh_def name_list_supp) + moreover { + fix y + from \\ \ P \M\\*(xvec@yvec)\\N\ \ P'\ \x \ supp N\ \x \ xvec\ \x \ yvec\ \x \ M\ \x \ \\ \distinct xvec\ \distinct yvec\ \xvec \* \\ \xvec \* P\ \xvec \* M\ \xvec \* yvec\ \yvec \* \\ \yvec \* P\ \yvec \* M\ + have "\ \ \\x\P \M\\*(xvec@x#yvec)\\N\ \ P'" by(rule semantics.cOpen) + moreover moreover from \xvec \* M\ \x \ M\ \yvec \* M\ + have "bn(M\\*(xvec@x#yvec)\\N\) \* (subject(M\\*(xvec@x#yvec)\\N\))" + by simp + moreover note \distinct(bn(M\\*(xvec@x#yvec)\\N\))\ + moreover assume "y \ set(bn(M\\*(xvec@x#yvec)\\N\))" + + ultimately have "y \ supp(object(M\\*(xvec@x#yvec)\\N\))" + by(metis boundOutputBindObject) + } + moreover from \xvec \* \\ \x \ \\ \yvec \* \\ + have "bn(M\\*(xvec@x#yvec)\\N\) \* bn \" and "bn(M\\*(xvec@x#yvec)\\N\) \* object \" by simp+ + moreover from \xvec \* P''\ \x \ P''\ \yvec \* P''\ + have "bn(M\\*(xvec@x#yvec)\\N\) \* P''" by simp + ultimately show ?case by(rule alphaDistinct) +next + case(cBrOpen \ P M xvec yvec N P' x \ P'') + note \\M\\*(xvec@x#yvec)\\N\ \ P' = \ \ P''\ + moreover from \xvec \* yvec\ \x \ xvec\ \x \ yvec\ \distinct xvec\ \distinct yvec\ + have "distinct(bn(\M\\*(xvec@x#yvec)\\N\))" + by auto (simp add: fresh_star_def fresh_def name_list_supp) + moreover { + fix y + from \\ \ P \\M\\*(xvec@yvec)\\N\ \ P'\ \x \ supp N\ \x \ xvec\ \x \ yvec\ \x \ M\ \x \ \\ \distinct xvec\ \distinct yvec\ \xvec \* \\ \xvec \* P\ \xvec \* M\ \xvec \* yvec\ \yvec \* \\ \yvec \* P\ \yvec \* M\ + have "\ \ \\x\P \\M\\*(xvec@x#yvec)\\N\ \ P'" by(rule semantics.cBrOpen) + moreover moreover from \xvec \* M\ \x \ M\ \yvec \* M\ + have "bn(\M\\*(xvec@x#yvec)\\N\) \* (subject(\M\\*(xvec@x#yvec)\\N\))" + by simp + moreover note \distinct(bn(\M\\*(xvec@x#yvec)\\N\))\ + moreover assume "y \ set(bn(\M\\*(xvec@x#yvec)\\N\))" + + ultimately have "y \ supp(object(\M\\*(xvec@x#yvec)\\N\))" + by(metis boundOutputBindObject) + } + moreover from \xvec \* \\ \x \ \\ \yvec \* \\ + have "bn(\M\\*(xvec@x#yvec)\\N\) \* bn \" and "bn(\M\\*(xvec@x#yvec)\\N\) \* object \" by simp+ + moreover from \xvec \* P''\ \x \ P''\ \yvec \* P''\ + have "bn(\M\\*(xvec@x#yvec)\\N\) \* P''" by simp + ultimately show ?case by(rule alphaDistinct) +next + case cScope + then show ?case + by - (rule alphaDistinct, auto intro: boundOutputBindObject) +qed (simp_all add: residualInject) + +lemma inputDistinct: + fixes \ :: 'b + and M :: 'a + and xvec :: "name list" + and N :: 'a + and P :: "('a, 'b, 'c) psi" + and Rs :: "('a, 'b, 'c) residual" + +assumes "\ \ M\\*xvec N\.P \ Rs" + +shows "distinct xvec" + using assms + by(nominal_induct \ P=="M\\*xvec N\.P" Rs avoiding: xvec N P rule: semantics.strong_induct) + (auto simp add: psi.inject intro: alphaInputDistinct) + +lemma outputInduct'[consumes 2, case_names cAlpha cOutput cCase cPar1 cPar2 cOpen cScope cBang]: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and M :: 'a + and yvec :: "name list" + and N :: 'a + and P' :: "('a, 'b, 'c) psi" + and Prop :: "'f::fs_name \ 'b \ ('a, 'b, 'c) psi \ + 'a \ name list \ 'a \ ('a, 'b, 'c) psi \ bool" + and C :: "'f::fs_name" + +assumes "\ \ P \M\\*xvec\\N\ \ P'" + and "xvec \* M" + and rAlpha: "\\ P M xvec N P' p C. \xvec \* \; xvec \* P; xvec \* M; xvec \* C; xvec \* (p \ xvec); + set p \ set xvec \ set(p \ xvec); distinctPerm p; + (p \ xvec) \* N; (p \ xvec) \* P'; Prop C \ P M xvec N P'\ \ + Prop C \ P M (p \ xvec) (p \ N) (p \ P')" + and rOutput: "\\ M K N P C. \\ \ M \ K\ \ Prop C \ (M\N\.P) K ([]) N P" + and rCase: "\\ P M xvec N P' \ Cs C. \\ \ P \M\\*xvec\\N\ \ P'; \C. Prop C \ P M xvec N P'; (\, P) \ set Cs; \ \ \; guarded P\ \ + Prop C \ (Cases Cs) M xvec N P'" + and rPar1: "\\ \\<^sub>Q P M xvec N P' A\<^sub>Q Q C. + \\ \ \\<^sub>Q \ P \M\\*xvec\\N\ \ P'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + \C. Prop C (\ \ \\<^sub>Q) P M xvec N P'; + A\<^sub>Q \* P; A\<^sub>Q \* Q; A\<^sub>Q \* \; A\<^sub>Q \* M; + A\<^sub>Q \* xvec; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* C; xvec \* Q; + xvec \* \; xvec \* \\<^sub>Q; xvec \* P; xvec \* M; xvec \* C\ \ + Prop C \ (P \ Q) M xvec N (P' \ Q)" + and rPar2: "\\ \\<^sub>P Q M xvec N Q' A\<^sub>P P C. + \\ \ \\<^sub>P \ Q \M\\*xvec\\N\ \ Q'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \C. Prop C (\ \ \\<^sub>P) Q M xvec N Q'; + A\<^sub>P \* P; A\<^sub>P \* Q; A\<^sub>P \* \; A\<^sub>P \* M; + A\<^sub>P \* xvec; A\<^sub>P \* N; A\<^sub>P \* Q'; A\<^sub>P \* C; xvec \* Q; + xvec \* \; xvec \* \\<^sub>P; xvec \* P; xvec \* M; xvec \* C\ \ + Prop C \ (P \ Q) M xvec N (P \ Q')" + and rOpen: "\\ P M xvec yvec N P' x C. + \\ \ P \M\\*(xvec@yvec)\\N\ \ P'; x \ supp N; \C. Prop C \ P M (xvec@yvec) N P'; + x \ \; x \ M; x \ xvec; x \ yvec; xvec \* \; xvec \* P; xvec \* M; + yvec \* \; yvec \* P; yvec \* M; yvec \* C; x \ C; xvec \* C\ \ + Prop C \ (\\x\P) M (xvec@x#yvec) N P'" + and rScope: "\\ P M xvec N P' x C. + \\ \ P \M\\*xvec\\N\ \ P'; \C. Prop C \ P M xvec N P'; + x \ \; x \ M; x \ xvec; x \ N; xvec \* \; + xvec \* P; xvec \* M; x \ C; xvec \* C\ \ + Prop C \ (\\x\P) M xvec N (\\x\P')" + and rBang: "\\ P M xvec N P' C. + \\ \ P \ !P \M\\*xvec\\N\ \ P'; guarded P; \C. Prop C \ (P \ !P) M xvec N P'\ \ + Prop C \ (!P) M xvec N P'" +shows "Prop C \ P M xvec N P'" +proof - + note \\ \ P \M\\*xvec\\N\ \ P'\ + moreover from \xvec \* M\ have "bn(M\\*xvec\\N\) \* subject(M\\*xvec\\N\)" by simp + moreover from \\ \ P \M\\*xvec\\N\ \ P'\ have "distinct(bn(M\\*xvec\\N\))" + by(rule boundOutputDistinct) + ultimately show ?thesis + proof(nominal_induct \ P \=="M\\*xvec\\N\" P' avoiding: C arbitrary: M xvec N rule: semanticsInduct) + case(cAlpha \ P \ P' p C M xvec N) + from \(p \ \) = M\\*xvec\\N\\ have "(p \ p \ \) = p \ (M\\*xvec\\N\)" + by(simp add: fresh_bij) + with \distinctPerm p\ have A: "\ = (p \ M)\\*(p \ xvec)\\(p \ N)\" + by(simp add: eqvts) + with \bn \ \* \\ \bn \ \* P\ \bn \ \* subject \ \ \bn \ \* C\ \bn \ \* bn(p \ \)\ \distinctPerm p\ + have "(p \ xvec) \* \" and "(p \ xvec) \* P" and "(p \ xvec) \* (p \ M)" and "(p \ xvec) \* C" and "(p \ xvec) \* (p \ p \ xvec)" + by auto + moreover from A \set p \ set(bn \) \ set(bn(p \ \))\ \distinctPerm p\ + have S: "set p \ set(p \ xvec) \ set(p \ p \ xvec)" by simp + moreover note \distinctPerm p\ + moreover from A \bn(p \ \) \* \\ \bn(p \ \) \* P'\ + have "(p \ p \ xvec) \* (p \ N)" and "(p \ p \ xvec) \* P'" by simp+ + moreover from A have "Prop C \ P (p \ M) (p \ xvec) (p \ N) P'" + by(rule cAlpha) + ultimately have "Prop C \ P (p \ M) (p \ p \ xvec) (p \ p \ N) (p \ P')" + by(rule rAlpha) + moreover from A \bn \ \* subject \\ have "(p \ xvec) \* (p \ M)" by simp + then have "xvec \* M" by(simp add: fresh_star_bij) + from A \bn(p \ \) \* \\ \distinctPerm p\ have "xvec \* (p \ M)" by simp + then have "(p \ xvec) \* (p \ p \ M)" by(simp add: fresh_star_bij) + with \distinctPerm p\ have "(p \ xvec) \* M" by simp + with \xvec \* M\ S \distinctPerm p\ have "(p \ M) = M" by simp + ultimately show ?case using S \distinctPerm p\ by simp + next + case cInput + then show ?case by(simp add: residualInject) + next + case cBrInput + then show ?case by simp + next + case cOutput + then show ?case by(force dest: rOutput simp add: action.inject) + next + case cBrOutput + then show ?case by simp + next + case cCase + then show ?case by(force intro: rCase) + next + case cPar1 + then show ?case by(force intro: rPar1) + next + case cPar2 + then show ?case by(force intro: rPar2) + next + case cComm1 + then show ?case by(simp add: action.inject) + next + case cComm2 + then show ?case by(simp add: action.inject) + next + case cBrMerge + then show ?case by(simp add: action.inject) + next + case cBrComm1 + then show ?case by simp + next + case cBrComm2 + then show ?case by simp + next + case cBrClose + then show ?case by simp + next + case cOpen + then show ?case by(auto intro: rOpen simp add: action.inject) + next + case cBrOpen + then show ?case by simp + next + case cScope + then show ?case by(auto intro: rScope) + next + case cBang + then show ?case by(auto intro: rBang) + qed +qed + +lemma brOutputInduct'[consumes 2, case_names cAlpha cBrOutput cCase cPar1 cPar2 cBrComm1 cBrComm2 cBrOpen cScope cBang]: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and M :: 'a + and yvec :: "name list" + and N :: 'a + and P' :: "('a, 'b, 'c) psi" + and Prop :: "'f::fs_name \ 'b \ ('a, 'b, 'c) psi \ + 'a \ name list \ 'a \ ('a, 'b, 'c) psi \ bool" + and C :: "'f::fs_name" + +assumes "\ \ P \\M\\*xvec\\N\ \ P'" + and "xvec \* M" + and rAlpha: "\\ P M xvec N P' p C. \xvec \* \; xvec \* P; xvec \* M; xvec \* C; xvec \* (p \ xvec); + set p \ set xvec \ set(p \ xvec); distinctPerm p; + (p \ xvec) \* N; (p \ xvec) \* P'; Prop C \ P M xvec N P'\ \ + Prop C \ P M (p \ xvec) (p \ N) (p \ P')" + and rBrOutput: "\\ M K N P C. \\ \ M \ K\ \ Prop C \ (M\N\.P) K ([]) N P" + and rCase: "\\ P M xvec N P' \ Cs C. \\ \ P \\M\\*xvec\\N\ \ P'; \C. Prop C \ P M xvec N P'; (\, P) \ set Cs; \ \ \; guarded P\ \ + Prop C \ (Cases Cs) M xvec N P'" + and rPar1: "\\ \\<^sub>Q P M xvec N P' A\<^sub>Q Q C. + \\ \ \\<^sub>Q \ P \\M\\*xvec\\N\ \ P'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + \C. Prop C (\ \ \\<^sub>Q) P M xvec N P'; + A\<^sub>Q \* P; A\<^sub>Q \* Q; A\<^sub>Q \* \; A\<^sub>Q \* M; + A\<^sub>Q \* xvec; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* C; xvec \* Q; + xvec \* \; xvec \* \\<^sub>Q; xvec \* P; xvec \* M; xvec \* C\ \ + Prop C \ (P \ Q) M xvec N (P' \ Q)" + and rPar2: "\\ \\<^sub>P Q M xvec N Q' A\<^sub>P P C. + \\ \ \\<^sub>P \ Q \\M\\*xvec\\N\ \ Q'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \C. Prop C (\ \ \\<^sub>P) Q M xvec N Q'; + A\<^sub>P \* P; A\<^sub>P \* Q; A\<^sub>P \* \; A\<^sub>P \* M; + A\<^sub>P \* xvec; A\<^sub>P \* N; A\<^sub>P \* Q'; A\<^sub>P \* C; xvec \* Q; + xvec \* \; xvec \* \\<^sub>P; xvec \* P; xvec \* M; xvec \* C\ \ + Prop C \ (P \ Q) M xvec N (P \ Q')" + and rBrComm1: "\\ \\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q xvec Q' A\<^sub>Q C. + \\ \ \\<^sub>Q \ P \\M\N\ \ P'; + extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \ \ \\<^sub>P \ Q \\M\\*xvec\\N\ \ Q'; \C. Prop C (\ \ \\<^sub>P) Q M xvec N Q'; + extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* N; A\<^sub>P \* P'; + A\<^sub>P \* Q; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* xvec; A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; + A\<^sub>Q \* P; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* Q; A\<^sub>Q \* Q'; distinct xvec; + A\<^sub>P \* M; A\<^sub>Q \* M; xvec \* M; + A\<^sub>Q \* xvec; xvec \* \; xvec \* \\<^sub>P; xvec \* \\<^sub>Q; xvec \* P; + xvec \* Q; A\<^sub>P \* C; A\<^sub>Q \* C; xvec \* C\ \ + Prop C \ (P \ Q) M xvec N (P' \ Q')" (* Removed: "\C. Prop C (\ \ \\<^sub>Q) P (M\N\ \ P');" *) + and rBrComm2: "\\ \\<^sub>Q P M xvec N P' A\<^sub>P \\<^sub>P Q Q' A\<^sub>Q C. + \\ \ \\<^sub>Q \ P \\M\\*xvec\\N\ \ P'; \C. Prop C (\ \ \\<^sub>Q) P M xvec N P'; + extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \ \ \\<^sub>P \ Q \\M\N\ \ Q'; + extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* N; A\<^sub>P \* P'; + A\<^sub>P \* Q; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* xvec; A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; + A\<^sub>Q \* P; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* Q; A\<^sub>Q \* Q'; distinct xvec; + A\<^sub>P \* M; A\<^sub>Q \* M; xvec \* M; + A\<^sub>Q \* xvec; xvec \* \; xvec \* \\<^sub>P; xvec \* \\<^sub>Q; xvec \* P; + xvec \* Q; A\<^sub>P \* C; A\<^sub>Q \* C; xvec \* C\ \ + Prop C \ (P \ Q) M xvec N (P' \ Q')" (* Removed: "\C. Prop C (\ \ \\<^sub>P) Q (K\N\) Q';" *) + and rOpen: "\\ P M xvec yvec N P' x C. + \\ \ P \\M\\*(xvec@yvec)\\N\ \ P'; x \ supp N; \C. Prop C \ P M (xvec@yvec) N P'; + x \ \; x \ M; x \ xvec; x \ yvec; xvec \* \; xvec \* P; xvec \* M; + yvec \* \; yvec \* P; yvec \* M; yvec \* C; x \ C; xvec \* C\ \ + Prop C \ (\\x\P) M (xvec@x#yvec) N P'" + and rScope: "\\ P M xvec N P' x C. + \\ \ P \\M\\*xvec\\N\ \ P'; \C. Prop C \ P M xvec N P'; + x \ \; x \ M; x \ xvec; x \ N; xvec \* \; + xvec \* P; xvec \* M; x \ C; xvec \* C\ \ + Prop C \ (\\x\P) M xvec N (\\x\P')" + and rBang: "\\ P M xvec N P' C. + \\ \ P \ !P \\M\\*xvec\\N\ \ P'; guarded P; \C. Prop C \ (P \ !P) M xvec N P'\ \ + Prop C \ (!P) M xvec N P'" +shows "Prop C \ P M xvec N P'" +proof - + note \\ \ P \\M\\*xvec\\N\ \ P'\ + moreover from \xvec \* M\ have "bn(\M\\*xvec\\N\) \* subject(\M\\*xvec\\N\)" by simp + moreover from \\ \ P \\M\\*xvec\\N\ \ P'\ have "distinct(bn(\M\\*xvec\\N\))" + by(rule boundOutputDistinct) + ultimately show ?thesis + proof(nominal_induct \ P \=="\M\\*xvec\\N\" P' avoiding: C arbitrary: M xvec N rule: semanticsInduct) + case(cAlpha \ P \ P' p C M xvec N) + from \(p \ \) = \M\\*xvec\\N\\ have "(p \ p \ \) = p \ (\M\\*xvec\\N\)" + by(simp add: fresh_bij) + with \distinctPerm p\ have A: "\ = \(p \ M)\\*(p \ xvec)\\(p \ N)\" + by(simp add: eqvts) + with \bn \ \* \\ \bn \ \* P\ \bn \ \* subject \ \ \bn \ \* C\ \bn \ \* bn(p \ \)\ \distinctPerm p\ + have "(p \ xvec) \* \" and "(p \ xvec) \* P" and "(p \ xvec) \* (p \ M)" and "(p \ xvec) \* C" and "(p \ xvec) \* (p \ p \ xvec)" + by auto + moreover from A \set p \ set(bn \) \ set(bn(p \ \))\ \distinctPerm p\ + have S: "set p \ set(p \ xvec) \ set(p \ p \ xvec)" by simp + moreover note \distinctPerm p\ + moreover from A \bn(p \ \) \* \\ \bn(p \ \) \* P'\ + have "(p \ p \ xvec) \* (p \ N)" and "(p \ p \ xvec) \* P'" by simp+ + moreover from A have "Prop C \ P (p \ M) (p \ xvec) (p \ N) P'" + by(rule cAlpha) + ultimately have "Prop C \ P (p \ M) (p \ p \ xvec) (p \ p \ N) (p \ P')" + by(rule rAlpha) + moreover from A \bn \ \* subject \\ have "(p \ xvec) \* (p \ M)" by simp + then have "xvec \* M" by(simp add: fresh_star_bij) + from A \bn(p \ \) \* \\ \distinctPerm p\ have "xvec \* (p \ M)" by simp + then have "(p \ xvec) \* (p \ p \ M)" by(simp add: fresh_star_bij) + with \distinctPerm p\ have "(p \ xvec) \* M" by simp + with \xvec \* M\ S \distinctPerm p\ have "(p \ M) = M" by simp + ultimately show ?case using S \distinctPerm p\ by simp + next + case cInput + then show ?case by(simp add: residualInject) + next + case cBrInput + then show ?case by simp + next + case cOutput + then show ?case by simp + next + case cBrOutput + then show ?case by(simp add: rBrOutput action.inject) + next + case cCase + then show ?case by(force intro: rCase) + next + case cPar1 + then show ?case by(force intro: rPar1) + next + case cPar2 + then show ?case by(force intro: rPar2) + next + case cComm1 + then show ?case by(simp add: action.inject) + next + case cComm2 + then show ?case by(simp add: action.inject) + next + case cBrMerge + then show ?case by(simp add: action.inject) + next + case cBrComm1 + then show ?case + by(auto intro: rBrComm1 simp add: action.inject) + next + case cBrComm2 + then show ?case + by(auto intro: rBrComm2 simp add: action.inject) + next + case cBrClose + then show ?case by simp + next + case cOpen + then show ?case by simp + next + case cBrOpen + then show ?case by(auto intro: rOpen simp add: action.inject) + next + case cScope + then show ?case by(auto intro: rScope) + next + case cBang + then show ?case by(auto intro: rBang) + qed +qed + +lemma inputInduct[consumes 1, case_names cInput cCase cPar1 cPar2 cScope cBang]: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and M :: 'a + and N :: 'a + and P' :: "('a, 'b, 'c) psi" + and Prop :: "'f::fs_name \ 'b \ ('a, 'b, 'c) psi \ + 'a \ 'a \ ('a, 'b, 'c) psi \ bool" + and C :: "'f::fs_name" + +assumes Trans: "\ \ P \M\N\ \ P'" + and rInput: "\\ M K xvec N Tvec P C. + \\ \ M \ K; distinct xvec; set xvec \ supp N; + length xvec = length Tvec; xvec \* \; + xvec \* M; xvec \* K; xvec \* C\ \ + Prop C \ (M\\*xvec N\.P) + K (N[xvec::=Tvec]) (P[xvec::=Tvec])" + and rCase: "\\ P M N P' \ Cs C. \\ \ P \M\N\ \ P'; \C. Prop C \ P M N P'; (\, P) \ set Cs; \ \ \; guarded P\ \ + Prop C \ (Cases Cs) M N P'" + and rPar1: "\\ \\<^sub>Q P M N P' A\<^sub>Q Q C. + \\ \ \\<^sub>Q \ P \M\N\ \ P'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + \C. Prop C (\ \ \\<^sub>Q) P M N P'; distinct A\<^sub>Q; + A\<^sub>Q \* P; A\<^sub>Q \* Q; A\<^sub>Q \* \; A\<^sub>Q \* M; A\<^sub>Q \* N; + A\<^sub>Q \* P'; A\<^sub>Q \* C\ \ + Prop C \ (P \ Q) M N (P' \ Q)" + and rPar2: "\\ \\<^sub>P Q M N Q' A\<^sub>P P C. + \\ \ \\<^sub>P \ Q \M\N\ \ Q'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \C. Prop C (\ \ \\<^sub>P) Q M N Q'; distinct A\<^sub>P; + A\<^sub>P \* P; A\<^sub>P \* Q; A\<^sub>P \* \; A\<^sub>P \* M; A\<^sub>P \* N; + A\<^sub>P \* Q'; A\<^sub>P \* C\ \ + Prop C \ (P \ Q) M N (P \ Q')" + and rScope: "\\ P M N P' x C. + \\ \ P \M\N\ \ P'; \C. Prop C \ P M N P'; x \ \; x \ M; x \ N; x \ C\ \ + Prop C \ (\\x\P) M N (\\x\P')" + and rBang: "\\ P M N P' C. + \\ \ P \ !P \M\N\ \ P'; guarded P; \C. Prop C \ (P \ !P) M N P'\ \ Prop C \ (!P) M N P'" +shows "Prop C \ P M N P'" + using Trans +proof(nominal_induct \ P Rs=="M\N\ \ P'" avoiding: C arbitrary: P' rule: semantics.strong_induct) + case(cInput \ M K xvec N Tvec P C) + then show ?case + by(force intro: rInput simp add: residualInject action.inject) +next + case(cBrInput \ M K xvec N Tvec P C) + then show ?case + by (simp add: residualInject) +next + case(Output \ M K N P C) + then show ?case by(simp add: residualInject) +next + case BrOutput + then show ?case by(simp add: residualInject) +next + case(Case \ P \ CS C P') + then show ?case by(force intro: rCase) +next + case(cPar1 \ \\<^sub>Q P \ P' Q A\<^sub>Q C P'') + then show ?case by(force intro: rPar1 simp add: residualInject) +next + case(cPar2 \ \P Q \ Q' xvec P C Q'') + then show ?case by(force intro: rPar2 simp add: residualInject) +next + case(cComm1 \ \Q P M N P' xvec \P Q K zvec Q' yvec C PQ) + then show ?case by(simp add: residualInject) +next + case(cComm2 \ \Q P M zvec N P' xvec \P Q K yvec Q' C PQ) + then show ?case by(simp add: residualInject) +next + case cBrMerge + then show ?case by(simp add: residualInject) +next + case cBrComm1 + then show ?case by(simp add: residualInject) +next + case cBrComm2 + then show ?case by(simp add: residualInject) +next + case(cBrClose \ P M xvec N P' C P'') + then show ?case by(simp add: residualInject) +next + case(cOpen \ P M xvec N P' x yvec C P'') + then show ?case by(simp add: residualInject) +next + case(cBrOpen \ P M xvec N P' x yvec C P'') + then show ?case by(simp add: residualInject) +next + case(cScope \ P \ P' x C P'') + then show ?case by(force intro: rScope simp add: residualInject) +next + case(Bang \ P C P') + then show ?case by(force intro: rBang) +qed + +lemma brInputInduct[consumes 1, case_names cBrInput cCase cPar1 cPar2 cBrMerge cScope cBang]: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and M :: 'a + and N :: 'a + and P' :: "('a, 'b, 'c) psi" + and Prop :: "'f::fs_name \ 'b \ ('a, 'b, 'c) psi \ + 'a \ 'a \ ('a, 'b, 'c) psi \ bool" + and C :: "'f::fs_name" + +assumes Trans: "\ \ P \\M\N\ \ P'" + and rBrInput: "\\ K M xvec N Tvec P C. + \\ \ K \ M; distinct xvec; set xvec \ supp N; + length xvec = length Tvec; xvec \* \; + xvec \* M; xvec \* K; xvec \* C\ \ + Prop C \ (M\\*xvec N\.P) + K (N[xvec::=Tvec]) (P[xvec::=Tvec])" + and rCase: "\\ P M N P' \ Cs C. \\ \ P \\M\N\ \ P'; \C. Prop C \ P M N P'; (\, P) \ set Cs; \ \ \; guarded P\ \ + Prop C \ (Cases Cs) M N P'" + and rPar1: "\\ \\<^sub>Q P M N P' A\<^sub>Q Q C. + \\ \ \\<^sub>Q \ P \\M\N\ \ P'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + \C. Prop C (\ \ \\<^sub>Q) P M N P'; distinct A\<^sub>Q; + A\<^sub>Q \* P; A\<^sub>Q \* Q; A\<^sub>Q \* \; A\<^sub>Q \* M; A\<^sub>Q \* N; + A\<^sub>Q \* P'; A\<^sub>Q \* C\ \ + Prop C \ (P \ Q) M N (P' \ Q)" + and rPar2: "\\ \\<^sub>P Q M N Q' A\<^sub>P P C. + \\ \ \\<^sub>P \ Q \\M\N\ \ Q'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \C. Prop C (\ \ \\<^sub>P) Q M N Q'; distinct A\<^sub>P; + A\<^sub>P \* P; A\<^sub>P \* Q; A\<^sub>P \* \; A\<^sub>P \* M; A\<^sub>P \* N; + A\<^sub>P \* Q'; A\<^sub>P \* C\ \ + Prop C \ (P \ Q) M N (P \ Q')" + and rBrMerge: "\\ \\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q Q' A\<^sub>Q C. + \\ \ \\<^sub>Q \ P \ \M\N\ \ P'; \C. Prop C (\ \ \\<^sub>Q) P M N P'; + extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \ \ \\<^sub>P \ Q \ \M\N\ \ Q'; \C. Prop C (\ \ \\<^sub>P) Q M N Q'; + extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* N; A\<^sub>P \* P'; + A\<^sub>P \* Q; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* M; A\<^sub>Q \* M; + A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; A\<^sub>Q \* P; A\<^sub>Q \* N; A\<^sub>Q \* P'; + A\<^sub>Q \* Q; A\<^sub>Q \* Q'; A\<^sub>P \* C; A\<^sub>Q \* C; + A\<^sub>P \* M; A\<^sub>Q \* M\ \ + Prop C \ (P \ Q) M N (P' \ Q')" + and rScope: "\\ P M N P' x C. + \\ \ P \\M\N\ \ P'; \C. Prop C \ P M N P'; x \ \; x \ M; x \ N; x \ C\ \ + Prop C \ (\\x\P) M N (\\x\P')" + and rBang: "\\ P M N P' C. + \\ \ P \ !P \\M\N\ \ P'; guarded P; \C. Prop C \ (P \ !P) M N P'\ \ Prop C \ (!P) M N P'" +shows "Prop C \ P M N P'" + using Trans +proof(nominal_induct \ P Rs=="\M\N\ \ P'" avoiding: C arbitrary: P' rule: semantics.strong_induct) + case(cInput \ M K xvec N Tvec P C) + then show ?case by (simp add: residualInject) +next + case(cBrInput \ K M xvec N Tvec P C) + then show ?case + by(auto intro: rBrInput simp add: residualInject action.inject) +next + case(Output \ M K N P C) + then show ?case by(simp add: residualInject) +next + case BrOutput + then show ?case by(simp add: residualInject) +next + case(Case \ P \ CS C P') + then show ?case by(force intro: rCase) +next + case(cPar1 \ \\<^sub>Q P \ P' Q A\<^sub>Q C P'') + then show ?case by(force intro: rPar1 simp add: residualInject) +next + case(cPar2 \ \P Q \ Q' xvec P C Q'') + then show ?case by(force intro: rPar2 simp add: residualInject) +next + case(cComm1 \ \Q P M N P' xvec \P Q K zvec Q' yvec C PQ) + then show ?case by(simp add: residualInject) +next + case(cComm2 \ \Q P M zvec N P' xvec \P Q K yvec Q' C PQ) + then show ?case by(simp add: residualInject) +next + case cBrMerge + then show ?case by(auto intro: rBrMerge simp add: residualInject action.inject) +next + case cBrComm1 + then show ?case by(simp add: residualInject) +next + case cBrComm2 + then show ?case by(simp add: residualInject) +next + case(cBrClose \ P M xvec N P' C P'') + then show ?case by(simp add: residualInject) +next + case(cOpen \ P M xvec N P' x yvec C P'') + then show ?case by(simp add: residualInject) +next + case(cBrOpen \ P M xvec N P' x yvec C P'') + then show ?case by(simp add: residualInject) +next + case(cScope \ P \ P' x C P'') + then show ?case by(force intro: rScope simp add: residualInject) +next + case(Bang \ P C P') + then show ?case by(force intro: rBang) +qed + +lemma tauInduct[consumes 1, case_names cCase cPar1 cPar2 cComm1 cComm2 cBrClose cScope cBang]: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and Rs :: "('a, 'b, 'c) residual" + and Prop :: "'f::fs_name \ 'b \ ('a, 'b, 'c) psi \ + ('a, 'b, 'c) psi \ bool" + and C :: "'f::fs_name" + +assumes Trans: "\ \ P \\ \ P'" + and rCase: "\\ P P' \ Cs C. \\ \ P \\ \ P'; \C. Prop C \ P P'; (\, P) \ set Cs; \ \ \; guarded P\ \ + Prop C \ (Cases Cs) P'" + and rPar1: "\\ \\<^sub>Q P P' A\<^sub>Q Q C. + \\ \ \\<^sub>Q \ P \\ \ P'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + \C. Prop C (\ \ \\<^sub>Q) P P'; + A\<^sub>Q \* P; A\<^sub>Q \* Q; A\<^sub>Q \* \; + A\<^sub>Q \* P'; A\<^sub>Q \* C\ \ + Prop C \ (P \ Q) (P' \ Q)" + and rPar2: "\\ \\<^sub>P Q Q' A\<^sub>P P C. + \\ \ \\<^sub>P \ Q \\ \ Q'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \C. Prop C (\ \ \\<^sub>P) Q Q'; + A\<^sub>P \* P; A\<^sub>P \* Q; A\<^sub>P \* \; + A\<^sub>P \* Q'; A\<^sub>P \* C\ \ + Prop C \ (P \ Q) (P \ Q')" + and rComm1: "\\ \\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q K xvec Q' A\<^sub>Q C. + \\ \ \\<^sub>Q \ P \M\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \ \ \\<^sub>P \ Q \K\\*xvec\\N\ \ Q'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + \ \ \\<^sub>P \ \\<^sub>Q \ M \ K; + A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* P'; + A\<^sub>P \* Q; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* xvec; A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; + A\<^sub>Q \* P; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* Q; A\<^sub>Q \* K; A\<^sub>Q \* Q'; + A\<^sub>Q \* xvec; xvec \* \; xvec \* \\<^sub>P; xvec \* \\<^sub>Q; xvec \* P; xvec \* M; + xvec \* Q; xvec \* K; A\<^sub>P \* C; A\<^sub>Q \* C; xvec \* C\ \ + Prop C \ (P \ Q) (\\*xvec\(P' \ Q'))" + and rComm2: "\\ \\<^sub>Q P M xvec N P' A\<^sub>P \\<^sub>P Q K Q' A\<^sub>Q C. + \\ \ \\<^sub>Q \ P \M\\*xvec\\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \ \ \\<^sub>P \ Q \K\N\ \ Q'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + \ \ \\<^sub>P \ \\<^sub>Q \ M \ K; + A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* P'; + A\<^sub>P \* Q; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* xvec; A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; + A\<^sub>Q \* P; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* Q; A\<^sub>Q \* K; A\<^sub>Q \* Q'; + A\<^sub>Q \* xvec; xvec \* \; xvec \* \\<^sub>P; xvec \* \\<^sub>Q; xvec \* P; xvec \* M; + xvec \* Q; xvec \* K; A\<^sub>P \* C; A\<^sub>Q \* C; xvec \* C\ \ + Prop C \ (P \ Q) (\\*xvec\(P' \ Q'))" + and rBrClose: "\\ P M xvec N P' A\<^sub>P \\<^sub>P x C. + \\ \ P \ \M\\*xvec\\N\ \ P'; + x \ supp M; + distinct xvec; xvec \* \; xvec \* P; + xvec \* M; + x \ \; x \ xvec\ \ + Prop C \ (\\x\P) (\\x\(\\*xvec\P'))" + and rScope: "\\ P P' x C. + \\ \ P \\ \ P'; \C. Prop C \ P P'; x \ \; x \ C\ \ + Prop C \ (\\x\P) (\\x\P')" + and rBang: "\\ P P' C. + \\ \ P \ !P \\ \ P'; guarded P; \C. Prop C \ (P \ !P) P'\ \ Prop C \ (!P) P'" +shows "Prop C \ P P'" + using Trans +proof(nominal_induct \ P Rs=="\ \ P'" avoiding: C arbitrary: P' rule: semantics.strong_induct) + case(cInput M K xvec N Tvec P C) + then show ?case by(simp add: residualInject) +next + case cBrInput + then show ?case by(simp add: residualInject) +next + case(Output \ M K N P C) + then show ?case by(simp add: residualInject) +next + case BrOutput + then show ?case by(simp add: residualInject) +next + case(Case \ P \ Cs C P') + then show ?case by(force intro: rCase simp add: residualInject) +next + case(cPar1 \ \\<^sub>Q P \ P' A\<^sub>Q Q C P'') + then show ?case by(force intro: rPar1 simp add: residualInject) +next + case(cPar2 \ \\<^sub>P Q \ Q' A\<^sub>P P C Q'') + then show ?case by(force intro: rPar2 simp add: residualInject) +next + case(cComm1 \ \\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q K xvec Q' A\<^sub>Q C PQ) + then show ?case by(force intro: rComm1 simp add: residualInject) +next + case(cComm2 \ \\<^sub>Q P M xvec N P' A\<^sub>P \P Q' A\<^sub>Q C PQ) + then show ?case by(force intro: rComm2 simp add: residualInject) +next + case cBrMerge + then show ?case by(simp add: residualInject) +next + case cBrComm1 + then show ?case by(simp add: residualInject) +next + case cBrComm2 + then show ?case by(simp add: residualInject) +next + case cBrClose + then show ?case by(force intro: rBrClose simp add: residualInject) +next + case(cOpen \ P M xvec N P' x yvec C P'') + then show ?case by(simp add: residualInject) +next + case(cBrOpen \ P M xvec N P' x yvec C P'') + then show ?case by(simp add: residualInject) +next + case(cScope \ P \ P' x C P'') + then show ?case by(force intro: rScope simp add: residualInject) +next + case(Bang \ P C P') + then show ?case by(force intro: rBang simp add: residualInject) +qed + +lemma semanticsFrameInduct[consumes 3, case_names cAlpha cInput cBrInput cOutput cBrOutput cCase cPar1 cPar2 cComm1 cComm2 cBrMerge cBrComm1 cBrComm2 cBrClose cOpen cBrOpen cScope cBang]: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and Rs :: "('a, 'b, 'c) residual" + and A\<^sub>P :: "name list" + and \\<^sub>P :: 'b + and Prop :: "'f::fs_name \ 'b \ ('a, 'b, 'c) psi \ + ('a, 'b, 'c) residual \ name list \ 'b \ bool" + and C :: "'f::fs_name" + +assumes Trans: "\ \ P \ Rs" + and FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" + and "distinct A\<^sub>P" + and rAlpha: "\\ P A\<^sub>P \\<^sub>P p Rs C. \A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* (p \ A\<^sub>P); A\<^sub>P \* Rs; A\<^sub>P \* C; + set p \ set A\<^sub>P \ set(p \ A\<^sub>P); distinctPerm p; + Prop C \ P Rs A\<^sub>P \\<^sub>P\ \ Prop C \ P Rs (p \ A\<^sub>P) (p \ \\<^sub>P)" + and rInput: "\\ M K xvec N Tvec P C. + \\ \ M \ K; distinct xvec; set xvec \ supp N; + length xvec = length Tvec; xvec \* \; + xvec \* M; xvec \* K; xvec \* C\ \ + Prop C \ (M\\*xvec N\.P) + (K\(N[xvec::=Tvec])\ \ (P[xvec::=Tvec])) ([]) (\)" + and rBrInput: "\\ M K xvec N Tvec P C. + \\ \ K \ M; distinct xvec; set xvec \ supp N; + length xvec = length Tvec; xvec \* \; + xvec \* M; xvec \* K; xvec \* C\ \ + Prop C \ (M\\*xvec N\.P) + (\K\(N[xvec::=Tvec])\ \ (P[xvec::=Tvec])) ([]) (\)" + and rOutput: "\\ M K N P C. \ \ M \ K \ Prop C \ (M\N\.P) (K\N\ \ P) ([]) (\)" + and rBrOutput: "\\ M K N P C. \ \ M \ K \ Prop C \ (M\N\.P) (\K\N\ \ P) ([]) (\)" + and rCase: "\\ P Rs \ Cs A\<^sub>P \\<^sub>P C. \\ \ P \ Rs; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; \C. Prop C \ P Rs A\<^sub>P \\<^sub>P; + (\, P) \ set Cs; \ \ \; guarded P; \\<^sub>P \ \; (supp \\<^sub>P) = ({}::name set); + A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* Rs; A\<^sub>P \* C\ \ Prop C \ (Cases Cs) Rs ([]) (\)" + and rPar1: "\\ \\<^sub>Q P \ P' A\<^sub>Q Q A\<^sub>P \\<^sub>P C. + \\ \ \\<^sub>Q \ P \\ \ P'; + extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + \C. Prop C (\ \ \\<^sub>Q) P (\ \ P') A\<^sub>P \\<^sub>P; distinct(bn \); + A\<^sub>P \* P; A\<^sub>P \* Q; A\<^sub>P \* \; A\<^sub>P \* \; A\<^sub>P \* P'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* \\<^sub>Q; + A\<^sub>Q \* P; A\<^sub>Q \* Q; A\<^sub>Q \* \; A\<^sub>Q \* \; A\<^sub>Q \* P'; A\<^sub>Q \* \\<^sub>P; + bn \ \* \; bn \ \* P; bn \ \* Q; bn \ \* subject \; bn \ \* \\<^sub>P; bn \ \* \\<^sub>Q; + A\<^sub>P \* C; A\<^sub>Q \* C; bn \ \* C\ \ + Prop C \ (P \ Q) (\ \ (P' \ Q)) (A\<^sub>P@A\<^sub>Q) (\\<^sub>P \ \\<^sub>Q)" + and rPar2: "\\ \\<^sub>P Q \ Q' A\<^sub>P P A\<^sub>Q \\<^sub>Q C. + \\ \ \\<^sub>P \ Q \\ \ Q'; + extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + \C. Prop C (\ \ \\<^sub>P) Q (\ \ Q') A\<^sub>Q \\<^sub>Q; distinct(bn \); + A\<^sub>P \* P; A\<^sub>P \* Q; A\<^sub>P \* \; A\<^sub>P \* \; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* \\<^sub>Q; + A\<^sub>Q \* P; A\<^sub>Q \* Q; A\<^sub>Q \* \; A\<^sub>Q \* \; A\<^sub>Q \* Q'; A\<^sub>Q \* \\<^sub>P; + bn \ \* \; bn \ \* P; bn \ \* Q; bn \ \* subject \; bn \ \* \\<^sub>P; bn \ \* \\<^sub>Q; + A\<^sub>P \* C; A\<^sub>Q \* C; bn \ \* C\ \ + Prop C \ (P \ Q) (\ \ (P \ Q')) (A\<^sub>P@A\<^sub>Q) (\\<^sub>P \ \\<^sub>Q)" + and rComm1: "\\ \\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q K xvec Q' A\<^sub>Q C. + \\ \ \\<^sub>Q \ P \M\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \C. Prop C (\ \ \\<^sub>Q) P ((M\N\) \ P') A\<^sub>P \\<^sub>P; + \ \ \\<^sub>P \ Q \K\\*xvec\\N\ \ Q'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + \ \ \\<^sub>P \ \\<^sub>Q \ M \ K; + \C. Prop C (\ \ \\<^sub>P) Q (K\\*xvec\\N\ \ Q') A\<^sub>Q \\<^sub>Q; distinct xvec; + A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* P'; + A\<^sub>P \* Q; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* xvec; A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; + A\<^sub>Q \* P; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* Q; A\<^sub>Q \* K; A\<^sub>Q \* Q'; + A\<^sub>Q \* xvec; xvec \* \; xvec \* \\<^sub>P; xvec \* \\<^sub>Q; xvec \* P; xvec \* M; + xvec \* Q; xvec \* K; A\<^sub>P \* C; A\<^sub>Q \* C; xvec \* C\ \ + Prop C \ (P \ Q) (\ \ \\*xvec\(P' \ Q')) (A\<^sub>P@A\<^sub>Q) (\\<^sub>P \ \\<^sub>Q)" + and rComm2: "\\ \\<^sub>Q P M xvec N P' A\<^sub>P \\<^sub>P Q K Q' A\<^sub>Q C. + \\ \ \\<^sub>Q \ P \M\\*xvec\\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \C. Prop C (\ \ \\<^sub>Q) P (M\\*xvec\\N\ \ P') A\<^sub>P \\<^sub>P; + \ \ \\<^sub>P \ Q \K\N\ \ Q'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + \C. Prop C (\ \ \\<^sub>P) Q (K\N\ \ Q') A\<^sub>Q \\<^sub>Q; + \ \ \\<^sub>P \ \\<^sub>Q \ M \ K; distinct xvec; + A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* P'; + A\<^sub>P \* Q; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* xvec; A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; + A\<^sub>Q \* P; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* Q; A\<^sub>Q \* K; A\<^sub>Q \* Q'; + A\<^sub>Q \* xvec; xvec \* \; xvec \* \\<^sub>P; xvec \* \\<^sub>Q; xvec \* P; xvec \* M; + xvec \* Q; xvec \* K; A\<^sub>P \* C; A\<^sub>Q \* C; xvec \* C\ \ + Prop C \ (P \ Q) (\ \ \\*xvec\(P' \ Q')) (A\<^sub>P@A\<^sub>Q) (\\<^sub>P \ \\<^sub>Q)" + and rBrMerge: "\\ \\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q Q' A\<^sub>Q C. + \\ \ \\<^sub>Q \ P \ \M\N\ \ P'; \C. Prop C (\ \ \\<^sub>Q) P (\M\N\ \ P') A\<^sub>P \\<^sub>P; + extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \ \ \\<^sub>P \ Q \ \M\N\ \ Q'; \C. Prop C (\ \ \\<^sub>P) Q (\M\N\ \ Q') A\<^sub>Q \\<^sub>Q; + extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* N; A\<^sub>P \* P'; + A\<^sub>P \* Q; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* M; A\<^sub>Q \* M; + A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; A\<^sub>Q \* P; A\<^sub>Q \* N; A\<^sub>Q \* P'; + A\<^sub>Q \* Q; A\<^sub>Q \* Q'; A\<^sub>P \* C; A\<^sub>Q \* C; + A\<^sub>P \* M; A\<^sub>Q \* M\ \ + Prop C \ (P \ Q) (\M\N\ \ (P' \ Q')) (A\<^sub>P@A\<^sub>Q) (\\<^sub>P \ \\<^sub>Q)" + and rBrComm1:"\\ \\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q xvec Q' A\<^sub>Q C. + \\ \ \\<^sub>Q \ P \\M\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \C. Prop C (\ \ \\<^sub>Q) P ((\M\N\) \ P') A\<^sub>P \\<^sub>P; + \ \ \\<^sub>P \ Q \\M\\*xvec\\N\ \ Q'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + \C. Prop C (\ \ \\<^sub>P) Q (\M\\*xvec\\N\ \ Q') A\<^sub>Q \\<^sub>Q; distinct xvec; + A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* N; A\<^sub>P \* P'; + A\<^sub>P \* Q; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* xvec; A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; + A\<^sub>Q \* P; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* Q; A\<^sub>Q \* Q'; + A\<^sub>Q \* xvec; xvec \* \; xvec \* \\<^sub>P; xvec \* \\<^sub>Q; xvec \* P; + xvec \* Q; A\<^sub>P \* C; A\<^sub>Q \* C; xvec \* C; + A\<^sub>P \* M; A\<^sub>Q \* M; xvec \* M\ \ + Prop C \ (P \ Q) (\M\\*xvec\\N\ \ (P' \ Q')) (A\<^sub>P@A\<^sub>Q) (\\<^sub>P \ \\<^sub>Q)" + and rBrComm2:"\\ \\<^sub>Q P M xvec N P' A\<^sub>P \\<^sub>P Q Q' A\<^sub>Q C. + \\ \ \\<^sub>Q \ P \\M\\*xvec\\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \C. Prop C (\ \ \\<^sub>Q) P (\M\\*xvec\\N\ \ P') A\<^sub>P \\<^sub>P; + \ \ \\<^sub>P \ Q \\M\N\ \ Q'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + \C. Prop C (\ \ \\<^sub>P) Q (\M\N\ \ Q') A\<^sub>Q \\<^sub>Q; distinct xvec; + A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* N; A\<^sub>P \* P'; + A\<^sub>P \* Q; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* xvec; A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; + A\<^sub>Q \* P; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* Q; A\<^sub>Q \* Q'; + A\<^sub>Q \* xvec; xvec \* \; xvec \* \\<^sub>P; xvec \* \\<^sub>Q; xvec \* P; + xvec \* Q; A\<^sub>P \* C; A\<^sub>Q \* C; xvec \* C; + A\<^sub>P \* M; A\<^sub>Q \* M; xvec \* M\ \ + Prop C \ (P \ Q) (\M\\*xvec\\N\ \ (P' \ Q')) (A\<^sub>P@A\<^sub>Q) (\\<^sub>P \ \\<^sub>Q)" + and rBrClose: "\\ P M xvec N P' A\<^sub>P \\<^sub>P x C. + \\ \ P \ \M\\*xvec\\N\ \ P'; + x \ supp M; + \C. Prop C \ P (\M\\*xvec\\N\ \ P') A\<^sub>P \\<^sub>P; + extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* P'; A\<^sub>P \* xvec; + distinct xvec; xvec \* \; xvec \* \\<^sub>P; xvec \* P; + xvec \* M; + x \ \; x \ xvec; x \ A\<^sub>P; + A\<^sub>P \* C; xvec \* C; x \ C\ \ + Prop C \ (\\x\P) (\ \ (\\x\(\\*xvec\P'))) (x#A\<^sub>P) \\<^sub>P" + and rOpen: "\\ P M xvec yvec N P' x A\<^sub>P \\<^sub>P C. + \\ \ P \M\\*(xvec@yvec)\\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \C. Prop C \ P (M\\*(xvec@yvec)\\N\ \ P') A\<^sub>P \\<^sub>P; x \ supp N; x \ \; x \ M; + x \ A\<^sub>P; x \ xvec; x \ yvec; A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* P'; + A\<^sub>P \* xvec; A\<^sub>P \* yvec; xvec \* yvec; distinct xvec; distinct yvec; + xvec \* \; xvec \* P; xvec \* M; xvec \* \\<^sub>P; yvec \* \\<^sub>P; + yvec \* \; yvec \* P; yvec \* M; A\<^sub>P \* C; x \ C; xvec \* C; yvec \* C\ \ + Prop C \ (\\x\P) (M\\*(xvec@x#yvec)\\N\ \ P') (x#A\<^sub>P) \\<^sub>P" + and rBrOpen: "\\ P M xvec yvec N P' x A\<^sub>P \\<^sub>P C. + \\ \ P \\M\\*(xvec@yvec)\\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \C. Prop C \ P (\M\\*(xvec@yvec)\\N\ \ P') A\<^sub>P \\<^sub>P; x \ supp N; x \ \; x \ M; + x \ A\<^sub>P; x \ xvec; x \ yvec; A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* P'; + A\<^sub>P \* xvec; A\<^sub>P \* yvec; xvec \* yvec; distinct xvec; distinct yvec; + xvec \* \; xvec \* P; xvec \* M; xvec \* \\<^sub>P; yvec \* \\<^sub>P; + yvec \* \; yvec \* P; yvec \* M; A\<^sub>P \* C; x \ C; xvec \* C; yvec \* C\ \ + Prop C \ (\\x\P) (\M\\*(xvec@x#yvec)\\N\ \ P') (x#A\<^sub>P) \\<^sub>P" + and rScope: "\\ P \ P' x A\<^sub>P \\<^sub>P C. + \\ \ P \\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \C. Prop C \ P (\ \ P') A\<^sub>P \\<^sub>P; + x \ \; x \ \; x \ A\<^sub>P; A\<^sub>P \* \; A\<^sub>P \* P; + A\<^sub>P \* \; A\<^sub>P \* P'; distinct(bn \); + bn \ \* \; bn \ \* P; bn \ \* subject \; bn \ \* \\<^sub>P; + A\<^sub>P \* C; x \ C; bn \ \* C\ \ + Prop C \ (\\x\P) (\ \ (\\x\P')) (x#A\<^sub>P) \\<^sub>P" + and rBang: "\\ P Rs A\<^sub>P \\<^sub>P C. + \\ \ P \ !P \ Rs; guarded P; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \C. Prop C \ (P \ !P) Rs A\<^sub>P (\\<^sub>P \ \); \\<^sub>P \ \; supp \\<^sub>P = ({}::name set); + A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* Rs; A\<^sub>P \* C\ \ Prop C \ (!P) Rs ([]) (\)" +shows "Prop C \ P Rs A\<^sub>P \\<^sub>P" + using Trans FrP \distinct A\<^sub>P\ +proof(nominal_induct avoiding: A\<^sub>P \\<^sub>P C rule: semantics.strong_induct) + case(cInput \ M K xvec N Tvec P A\<^sub>P \\<^sub>P C) + from \extractFrame (M\\*xvec N\.P) = \A\<^sub>P, \\<^sub>P\\ + have "A\<^sub>P = []" and "\\<^sub>P = \" + by auto + with \\ \ M \ K\ \distinct xvec\ \set xvec \ supp N\ \length xvec = length Tvec\ + \xvec \* \\ \xvec \* M\ \xvec \* K\ \xvec \* C\ + show ?case by(blast intro: rInput) +next + case(cBrInput \ K M xvec N Tvec P A\<^sub>P \\<^sub>P C) + from \extractFrame (M\\*xvec N\.P) = \A\<^sub>P, \\<^sub>P\\ + have "A\<^sub>P = []" and "\\<^sub>P = \" + by auto + with \\ \ K \ M\ \distinct xvec\ \set xvec \ supp N\ \length xvec = length Tvec\ + \xvec \* \\ \xvec \* M\ \xvec \* K\ \xvec \* C\ + show ?case by(blast intro: rBrInput) +next + case(Output \ M K N P A\<^sub>P \\<^sub>P) + from \extractFrame (M\N\.P) = \A\<^sub>P, \\<^sub>P\\ + have "A\<^sub>P = []" and "\\<^sub>P = \" + by auto + with \\ \ M \ K\ show ?case + by(blast intro: rOutput) +next + case(BrOutput \ M K N P A\<^sub>P \\<^sub>P) + from \extractFrame (M\N\.P) = \A\<^sub>P, \\<^sub>P\\ + have "A\<^sub>P = []" and "\\<^sub>P = \" + by auto + with \\ \ M \ K\ show ?case + by(blast intro: rBrOutput) +next + case(Case \ P Rs \ Cs A\<^sub>c\<^sub>P \\<^sub>c\<^sub>P C) + obtain A\<^sub>P \\<^sub>P where FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" and "distinct A\<^sub>P" + and "A\<^sub>P \* (\, P, Rs, C)" + by(rule freshFrame) + then have "A\<^sub>P \* \" and "A\<^sub>P \* P" and "A\<^sub>P \* Rs" and "A\<^sub>P \* C" + by simp+ + note \\ \ P \ Rs\ FrP \distinct A\<^sub>P\ + moreover from FrP \distinct A\<^sub>P\ \\A\<^sub>P \\<^sub>P C. \extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P\ \ Prop C \ P Rs A\<^sub>P \\<^sub>P\ + have "\C. Prop C \ P Rs A\<^sub>P \\<^sub>P" by simp + moreover note \(\, P) \ set Cs\ \\ \ \\ \guarded P\ + moreover from \guarded P\ FrP have "\\<^sub>P \ \" and "supp \\<^sub>P = ({}::name set)" by(metis guardedStatEq)+ + moreover note \A\<^sub>P \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* Rs\ \A\<^sub>P \* C\ + ultimately have "Prop C \ (Cases Cs) Rs ([]) (\)" + by(rule rCase) + then show ?case using \extractFrame(Cases Cs) = \A\<^sub>c\<^sub>P, \\<^sub>c\<^sub>P\\ by simp +next + case(cPar1 \ \\<^sub>Q P \ P' Q A\<^sub>Q A\<^sub>P\<^sub>Q \\<^sub>P\<^sub>Q C) + obtain A\<^sub>P \\<^sub>P where FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" and "distinct A\<^sub>P" + "A\<^sub>P \* (P, Q, \, \, P', A\<^sub>Q, A\<^sub>P\<^sub>Q, C, \\<^sub>Q)" + by(rule freshFrame) + then have "A\<^sub>P \* P" and "A\<^sub>P \* Q" and "A\<^sub>P \* \" and "A\<^sub>P \* \" and "A\<^sub>P \* P'" + and "A\<^sub>P \* A\<^sub>Q" and "A\<^sub>P \* A\<^sub>P\<^sub>Q" and "A\<^sub>P \* C" and "A\<^sub>P \* \\<^sub>Q" + by simp+ + + have FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" by fact + + from \A\<^sub>Q \* P\ \A\<^sub>P \* A\<^sub>Q\ FrP have "A\<^sub>Q \* \\<^sub>P" + by(force dest: extractFrameFreshChain) + + from \bn \ \* P\ \A\<^sub>P \* \\ FrP have "bn \ \* \\<^sub>P" + by(force dest: extractFrameFreshChain) + + from \extractFrame(P \ Q) = \A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q\\ FrP FrQ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>Q \* \\<^sub>P\ + have "\(A\<^sub>P@A\<^sub>Q), \\<^sub>P \ \\<^sub>Q\ = \A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q\" + by simp + moreover from \distinct A\<^sub>P\ \distinct A\<^sub>Q\ \A\<^sub>P \* A\<^sub>Q\ have "distinct(A\<^sub>P@A\<^sub>Q)" + by(auto simp add: fresh_star_def fresh_def name_list_supp) + ultimately obtain p where S: "set p \ set(A\<^sub>P@A\<^sub>Q) \ set((p \ A\<^sub>P)@(p \ A\<^sub>Q))" and "distinctPerm p" + and \eq: "\\<^sub>P\<^sub>Q = p \ (\\<^sub>P \ \\<^sub>Q)" and Aeq: "A\<^sub>P\<^sub>Q = (p \ A\<^sub>P)@(p \ A\<^sub>Q)" + using \A\<^sub>P \* A\<^sub>P\<^sub>Q\ \A\<^sub>Q \* A\<^sub>P\<^sub>Q\ \distinct A\<^sub>P\<^sub>Q\ + apply - + apply(rule frameChainEq') + by (assumption | simp add: eqvts)+ + note \\ \ \\<^sub>Q \ P \\ \ P'\ FrP \distinct A\<^sub>P\ FrQ \distinct A\<^sub>Q\ + + moreover from FrP \distinct A\<^sub>P\ \\A\<^sub>P \\<^sub>P C. \extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P\ \ Prop C (\ \ \\<^sub>Q) P (\ \ P') A\<^sub>P \\<^sub>P\ + have "\C. Prop C (\ \ \\<^sub>Q) P (\ \ P') A\<^sub>P \\<^sub>P" by simp + + moreover note \A\<^sub>P \* P\ \A\<^sub>P \* Q\ \A\<^sub>P \* \\ \A\<^sub>P \* \\ \A\<^sub>P \* P'\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* \\<^sub>Q\ + \A\<^sub>Q \* P\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* \\ \A\<^sub>Q \* \\ \A\<^sub>Q \* P'\ \A\<^sub>Q \* \\<^sub>P\ \distinct(bn \)\ + \bn \ \* \\ \bn \ \* P\ \bn \ \* Q\ \bn \ \* subject \\ \bn \ \* \\<^sub>P\ \bn \ \* \\<^sub>Q\ + \A\<^sub>P \* C\ \A\<^sub>Q \* C\ \bn \ \* C\ + ultimately have "Prop C \ (P \ Q) (\ \ (P' \ Q)) (A\<^sub>P@A\<^sub>Q) (\\<^sub>P \ \\<^sub>Q)" + by(metis rPar1) + with \A\<^sub>P \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* Q\ \A\<^sub>P \* \\ \A\<^sub>P \* P'\ \A\<^sub>P \* A\<^sub>P\<^sub>Q\ \A\<^sub>P \* C\ + \A\<^sub>Q \* \\ \A\<^sub>Q \* P\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* \\ \A\<^sub>Q \* P'\ \A\<^sub>Q \* A\<^sub>P\<^sub>Q\ \A\<^sub>Q \* C\ + S \distinctPerm p\ Aeq + have "Prop C \ (P \ Q) (\ \ (P' \ Q)) (p \ (A\<^sub>P@A\<^sub>Q)) (p \ (\\<^sub>P \ \\<^sub>Q))" + apply - + apply(rule rAlpha) + by(assumption | simp add: eqvts)+ + with \eq Aeq show ?case by(simp add: eqvts) +next + case(cPar2 \ \\<^sub>P Q \ Q' P A\<^sub>P A\<^sub>P\<^sub>Q \\<^sub>P\<^sub>Q C) + obtain A\<^sub>Q \\<^sub>Q where FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" and "distinct A\<^sub>Q" + "A\<^sub>Q \* (P, Q, \, \, Q', A\<^sub>P, A\<^sub>P\<^sub>Q, C, \\<^sub>P)" + by(rule freshFrame) + then have "A\<^sub>Q \* P" and "A\<^sub>Q \* Q" and "A\<^sub>Q \* \" and "A\<^sub>Q \* \" and "A\<^sub>Q \* Q'" + and "A\<^sub>Q \* A\<^sub>P" and "A\<^sub>Q \* A\<^sub>P\<^sub>Q" and "A\<^sub>Q \* C" and "A\<^sub>Q \* \\<^sub>P" + by simp+ + + from \A\<^sub>Q \* A\<^sub>P\ have "A\<^sub>P \* A\<^sub>Q" by simp + have FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" by fact + + from \A\<^sub>P \* Q\ \A\<^sub>Q \* A\<^sub>P\ FrQ have "A\<^sub>P \* \\<^sub>Q" + by(force dest: extractFrameFreshChain) + from \bn \ \* Q\ \A\<^sub>Q \* \\ FrQ have "bn \ \* \\<^sub>Q" + by(force dest: extractFrameFreshChain) + + from \extractFrame(P \ Q) = \A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q\\ FrP FrQ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>Q \* \\<^sub>P\ + have "\(A\<^sub>P@A\<^sub>Q), \\<^sub>P \ \\<^sub>Q\ = \A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q\" + by simp + moreover from \distinct A\<^sub>P\ \distinct A\<^sub>Q\ \A\<^sub>P \* A\<^sub>Q\ have "distinct(A\<^sub>P@A\<^sub>Q)" + by(auto simp add: fresh_star_def fresh_def name_list_supp) + ultimately obtain p where S: "(set p \ (set(A\<^sub>P@A\<^sub>Q)) \ (set A\<^sub>P\<^sub>Q))" and "distinctPerm p" + and \eq: "\\<^sub>P\<^sub>Q = p \ (\\<^sub>P \ \\<^sub>Q)" and Aeq: "A\<^sub>P\<^sub>Q = ((p \ A\<^sub>P)@(p \ A\<^sub>Q))" + using \A\<^sub>P \* A\<^sub>P\<^sub>Q\ \A\<^sub>Q \* A\<^sub>P\<^sub>Q\ \distinct A\<^sub>P\<^sub>Q\ + apply - + apply(rule frameChainEq') + by(assumption | simp add: eqvts)+ + + note \\ \ \\<^sub>P \ Q \\ \ Q'\ FrP \distinct A\<^sub>P\ FrQ \distinct A\<^sub>Q\ + + moreover from FrQ \distinct A\<^sub>Q\ \\A\<^sub>Q \\<^sub>Q C. \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q\ \ Prop C (\ \ \\<^sub>P) Q (\ \ Q') A\<^sub>Q \\<^sub>Q\ + have "\C. Prop C (\ \ \\<^sub>P) Q (\ \ Q') A\<^sub>Q \\<^sub>Q" by simp + + moreover note \A\<^sub>P \* P\ \A\<^sub>P \* Q\ \A\<^sub>P \* \\ \A\<^sub>P \* \\ \A\<^sub>P \* Q'\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* \\<^sub>Q\ + \A\<^sub>Q \* P\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* \\ \A\<^sub>Q \* \\ \A\<^sub>Q \* Q'\ \A\<^sub>Q \* \\<^sub>P\ \distinct(bn \)\ + \bn \ \* \\ \bn \ \* P\ \bn \ \* Q\ \bn \ \* subject \\ \bn \ \* \\<^sub>P\ \bn \ \* \\<^sub>Q\ + \A\<^sub>P \* C\ \A\<^sub>Q \* C\ \bn \ \* C\ + ultimately have "Prop C \ (P \ Q) (\ \ (P \ Q')) (A\<^sub>P@A\<^sub>Q) (\\<^sub>P \ \\<^sub>Q)" + by(metis rPar2) + + with \A\<^sub>P \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* Q\ \A\<^sub>P \* \\ \A\<^sub>P \* Q'\ \A\<^sub>P \* A\<^sub>P\<^sub>Q\ \A\<^sub>P \* C\ + \A\<^sub>Q \* \\ \A\<^sub>Q \* P\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* \\ \A\<^sub>Q \* Q'\ \A\<^sub>Q \* A\<^sub>P\<^sub>Q\ \A\<^sub>Q \* C\ + S \distinctPerm p\ Aeq + have "Prop C \ (P \ Q) (\ \ (P \ Q')) (p \ (A\<^sub>P@A\<^sub>Q)) (p \ (\\<^sub>P \ \\<^sub>Q))" + apply - + apply(rule rAlpha) + by(assumption | simp add: eqvts)+ + with \eq Aeq show ?case by(simp add: eqvts) +next + case(cComm1 \ \\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q K xvec Q' A\<^sub>Q A\<^sub>P\<^sub>Q \\<^sub>P\<^sub>Q C) + from \distinct A\<^sub>P\ \distinct A\<^sub>Q\ \A\<^sub>P \* A\<^sub>Q\ have "distinct(A\<^sub>P@A\<^sub>Q)" + by(auto simp add: fresh_star_def fresh_def name_list_supp) + from cComm1 have "Prop C \ (P \ Q) (\ \ \\*xvec\(P' \ Q')) (A\<^sub>P@A\<^sub>Q) (\\<^sub>P \ \\<^sub>Q)" + by(metis rComm1) + moreover from \extractFrame(P \ Q) = \A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q\\ \extractFrame P = \A\<^sub>P, \\<^sub>P\\ \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ + \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>Q \* \\<^sub>P\ + have "\(A\<^sub>P@A\<^sub>Q), (\\<^sub>P \ \\<^sub>Q)\ = \A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q\" + by simp + with \A\<^sub>P \* A\<^sub>P\<^sub>Q\ \A\<^sub>Q \* A\<^sub>P\<^sub>Q\ \distinct(A\<^sub>P@A\<^sub>Q)\ \distinct A\<^sub>P\<^sub>Q\ + obtain p where S: "(set p \ (set(A\<^sub>P@A\<^sub>Q)) \ (set A\<^sub>P\<^sub>Q))" and "distinctPerm p" + and \eq: "\\<^sub>P\<^sub>Q = p \ (\\<^sub>P \ \\<^sub>Q)" and Aeq: "A\<^sub>P\<^sub>Q = p \ (A\<^sub>P@A\<^sub>Q)" + apply - + apply(rule frameChainEq') + by(assumption | simp)+ + moreover note \A\<^sub>P \* \\ \A\<^sub>Q \* \\ \A\<^sub>P \* P\ \A\<^sub>Q \* P\ \A\<^sub>P \* Q\ \A\<^sub>Q \* Q\ \A\<^sub>P \* xvec\ + \A\<^sub>Q \* xvec\ \A\<^sub>P \* P'\ \A\<^sub>Q \* P'\ \A\<^sub>P \* Q'\ \A\<^sub>Q \* Q'\ \A\<^sub>P \* A\<^sub>P\<^sub>Q\ \A\<^sub>Q \* A\<^sub>P\<^sub>Q\ + \A\<^sub>P \* C\ \A\<^sub>Q \* C\ + ultimately have "Prop C \ (P \ Q) (\ \ \\*xvec\(P' \ Q')) (p \ (A\<^sub>P@A\<^sub>Q)) (p \ (\\<^sub>P \ \\<^sub>Q))" + by(fastforce simp add: rAlpha) + with \eq Aeq show ?case by simp +next + case(cComm2 \ \\<^sub>Q P M xvec N P' A\<^sub>P \\<^sub>P Q K Q' A\<^sub>Q A\<^sub>P\<^sub>Q \\<^sub>P\<^sub>Q C) + from \distinct A\<^sub>P\ \distinct A\<^sub>Q\ \A\<^sub>P \* A\<^sub>Q\ have "distinct(A\<^sub>P@A\<^sub>Q)" + by(auto simp add: fresh_star_def fresh_def name_list_supp) + from cComm2 have "Prop C \ (P \ Q) (\ \ \\*xvec\(P' \ Q')) (A\<^sub>P@A\<^sub>Q) (\\<^sub>P \ \\<^sub>Q)" + by(metis rComm2) + moreover from \extractFrame(P \ Q) = \A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q\\ \extractFrame P = \A\<^sub>P, \\<^sub>P\\ \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ + \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>Q \* \\<^sub>P\ + have "\(A\<^sub>P@A\<^sub>Q), (\\<^sub>P \ \\<^sub>Q)\ = \A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q\" + by simp + with \A\<^sub>P \* A\<^sub>P\<^sub>Q\ \A\<^sub>Q \* A\<^sub>P\<^sub>Q\ \distinct(A\<^sub>P@A\<^sub>Q)\ \distinct A\<^sub>P\<^sub>Q\ + obtain p where S: "(set p \ (set(A\<^sub>P@A\<^sub>Q)) \ (set A\<^sub>P\<^sub>Q))" and "distinctPerm p" + and \eq: "\\<^sub>P\<^sub>Q = p \ (\\<^sub>P \ \\<^sub>Q)" and Aeq: "A\<^sub>P\<^sub>Q = p \ (A\<^sub>P@A\<^sub>Q)" + apply - + apply(rule frameChainEq') + by(assumption | simp)+ + moreover note \A\<^sub>P \* \\ \A\<^sub>Q \* \\ \A\<^sub>P \* P\ \A\<^sub>Q \* P\ \A\<^sub>P \* Q\ \A\<^sub>Q \* Q\ \A\<^sub>P \* xvec\ + \A\<^sub>Q \* xvec\ \A\<^sub>P \* P'\ \A\<^sub>Q \* P'\ \A\<^sub>P \* Q'\ \A\<^sub>Q \* Q'\ \A\<^sub>P \* A\<^sub>P\<^sub>Q\ \A\<^sub>Q \* A\<^sub>P\<^sub>Q\ + \A\<^sub>P \* C\ \A\<^sub>Q \* C\ + ultimately have "Prop C \ (P \ Q) (\ \ \\*xvec\(P' \ Q')) (p \ (A\<^sub>P@A\<^sub>Q)) (p \ (\\<^sub>P \ \\<^sub>Q))" + by(fastforce intro: rAlpha) + with \eq Aeq show ?case by simp +next + case(cBrMerge \ \\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q Q' A\<^sub>Q A\<^sub>P\<^sub>Q \\<^sub>P\<^sub>Q C) + from \distinct A\<^sub>P\ \distinct A\<^sub>Q\ \A\<^sub>P \* A\<^sub>Q\ have "distinct(A\<^sub>P@A\<^sub>Q)" + by(auto simp add: fresh_star_def fresh_def name_list_supp) + from cBrMerge have "Prop C \ (P \ Q) (\M\N\ \ (P' \ Q')) (A\<^sub>P@A\<^sub>Q) (\\<^sub>P \ \\<^sub>Q)" + by(fastforce intro!: rBrMerge) + moreover from \extractFrame(P \ Q) = \A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q\\ \extractFrame P = \A\<^sub>P, \\<^sub>P\\ \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ + \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>Q \* \\<^sub>P\ + have "\(A\<^sub>P@A\<^sub>Q), (\\<^sub>P \ \\<^sub>Q)\ = \A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q\" + by simp + with \A\<^sub>P \* A\<^sub>P\<^sub>Q\ \A\<^sub>Q \* A\<^sub>P\<^sub>Q\ \distinct(A\<^sub>P@A\<^sub>Q)\ \distinct A\<^sub>P\<^sub>Q\ + obtain p where S: "(set p \ (set(A\<^sub>P@A\<^sub>Q)) \ (set A\<^sub>P\<^sub>Q))" and "distinctPerm p" + and \eq: "\\<^sub>P\<^sub>Q = p \ (\\<^sub>P \ \\<^sub>Q)" and Aeq: "A\<^sub>P\<^sub>Q = p \ (A\<^sub>P@A\<^sub>Q)" + apply - + apply(rule frameChainEq') + by(assumption | simp)+ + moreover note \A\<^sub>P \* \\ \A\<^sub>Q \* \\ \A\<^sub>P \* P\ \A\<^sub>Q \* P\ \A\<^sub>P \* Q\ \A\<^sub>Q \* Q\ + \A\<^sub>P \* P'\ \A\<^sub>Q \* P'\ \A\<^sub>P \* Q'\ \A\<^sub>Q \* Q'\ \A\<^sub>P \* A\<^sub>P\<^sub>Q\ \A\<^sub>Q \* A\<^sub>P\<^sub>Q\ + \A\<^sub>P \* C\ \A\<^sub>Q \* C\ \A\<^sub>P \* M\ \A\<^sub>Q \* M\ \A\<^sub>P \* N\ \A\<^sub>Q \* N\ + ultimately have "Prop C \ (P \ Q) (\M\N\ \ (P' \ Q')) (p \ (A\<^sub>P@A\<^sub>Q)) (p \ (\\<^sub>P \ \\<^sub>Q))" + by(fastforce intro: rAlpha) + with \eq Aeq show ?case by simp +next + case(cBrComm1 \ \\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q xvec Q' A\<^sub>Q A\<^sub>P\<^sub>Q \\<^sub>P\<^sub>Q C) + from \distinct A\<^sub>P\ \distinct A\<^sub>Q\ \A\<^sub>P \* A\<^sub>Q\ have "distinct(A\<^sub>P@A\<^sub>Q)" + by(auto simp add: fresh_star_def fresh_def name_list_supp) + from cBrComm1 have "Prop C \ (P \ Q) (\M\\*xvec\\N\ \ (P' \ Q')) (A\<^sub>P@A\<^sub>Q) (\\<^sub>P \ \\<^sub>Q)" + by(metis rBrComm1) + moreover from \extractFrame(P \ Q) = \A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q\\ \extractFrame P = \A\<^sub>P, \\<^sub>P\\ \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ + \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>Q \* \\<^sub>P\ + have "\(A\<^sub>P@A\<^sub>Q), (\\<^sub>P \ \\<^sub>Q)\ = \A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q\" + by simp + with \A\<^sub>P \* A\<^sub>P\<^sub>Q\ \A\<^sub>Q \* A\<^sub>P\<^sub>Q\ \distinct(A\<^sub>P@A\<^sub>Q)\ \distinct A\<^sub>P\<^sub>Q\ + obtain p where S: "(set p \ (set(A\<^sub>P@A\<^sub>Q)) \ (set A\<^sub>P\<^sub>Q))" and "distinctPerm p" + and \eq: "\\<^sub>P\<^sub>Q = p \ (\\<^sub>P \ \\<^sub>Q)" and Aeq: "A\<^sub>P\<^sub>Q = p \ (A\<^sub>P@A\<^sub>Q)" + apply - + apply(rule frameChainEq') + by(assumption | simp)+ + moreover note \A\<^sub>P \* \\ \A\<^sub>Q \* \\ \A\<^sub>P \* P\ \A\<^sub>Q \* P\ \A\<^sub>P \* Q\ \A\<^sub>Q \* Q\ \A\<^sub>P \* xvec\ + \A\<^sub>Q \* xvec\ \A\<^sub>P \* P'\ \A\<^sub>Q \* P'\ \A\<^sub>P \* Q'\ \A\<^sub>Q \* Q'\ \A\<^sub>P \* A\<^sub>P\<^sub>Q\ \A\<^sub>Q \* A\<^sub>P\<^sub>Q\ + \A\<^sub>P \* C\ \A\<^sub>Q \* C\ \A\<^sub>P \* M\ \A\<^sub>Q \* M\ \A\<^sub>P \* N\ \A\<^sub>Q \* N\ + ultimately have "Prop C \ (P \ Q) (\M\\*xvec\\N\ \ (P' \ Q')) (p \ (A\<^sub>P@A\<^sub>Q)) (p \ (\\<^sub>P \ \\<^sub>Q))" + by(fastforce intro: rAlpha) + with \eq Aeq show ?case by simp +next + case(cBrComm2 \ \\<^sub>Q P M xvec N P' A\<^sub>P \\<^sub>P Q Q' A\<^sub>Q A\<^sub>P\<^sub>Q \\<^sub>P\<^sub>Q C) + from \distinct A\<^sub>P\ \distinct A\<^sub>Q\ \A\<^sub>P \* A\<^sub>Q\ have "distinct(A\<^sub>P@A\<^sub>Q)" + by(auto simp add: fresh_star_def fresh_def name_list_supp) + from cBrComm2 have "Prop C \ (P \ Q) (\M\\*xvec\\N\ \ (P' \ Q')) (A\<^sub>P@A\<^sub>Q) (\\<^sub>P \ \\<^sub>Q)" + by(metis rBrComm2) + moreover from \extractFrame(P \ Q) = \A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q\\ \extractFrame P = \A\<^sub>P, \\<^sub>P\\ \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ + \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>Q \* \\<^sub>P\ + have "\(A\<^sub>P@A\<^sub>Q), (\\<^sub>P \ \\<^sub>Q)\ = \A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q\" + by simp + with \A\<^sub>P \* A\<^sub>P\<^sub>Q\ \A\<^sub>Q \* A\<^sub>P\<^sub>Q\ \distinct(A\<^sub>P@A\<^sub>Q)\ \distinct A\<^sub>P\<^sub>Q\ + obtain p where S: "(set p \ (set(A\<^sub>P@A\<^sub>Q)) \ (set A\<^sub>P\<^sub>Q))" and "distinctPerm p" + and \eq: "\\<^sub>P\<^sub>Q = p \ (\\<^sub>P \ \\<^sub>Q)" and Aeq: "A\<^sub>P\<^sub>Q = p \ (A\<^sub>P@A\<^sub>Q)" + apply - + apply(rule frameChainEq') + by(assumption | simp)+ + moreover note \A\<^sub>P \* \\ \A\<^sub>Q \* \\ \A\<^sub>P \* P\ \A\<^sub>Q \* P\ \A\<^sub>P \* Q\ \A\<^sub>Q \* Q\ \A\<^sub>P \* xvec\ + \A\<^sub>Q \* xvec\ \A\<^sub>P \* P'\ \A\<^sub>Q \* P'\ \A\<^sub>P \* Q'\ \A\<^sub>Q \* Q'\ \A\<^sub>P \* A\<^sub>P\<^sub>Q\ \A\<^sub>Q \* A\<^sub>P\<^sub>Q\ + \A\<^sub>P \* C\ \A\<^sub>Q \* C\ \A\<^sub>P \* M\ \A\<^sub>Q \* M\ \A\<^sub>P \* N\ \A\<^sub>Q \* N\ + ultimately have "Prop C \ (P \ Q) (\M\\*xvec\\N\ \ (P' \ Q')) (p \ (A\<^sub>P@A\<^sub>Q)) (p \ (\\<^sub>P \ \\<^sub>Q))" + by(fastforce intro: rAlpha) + with \eq Aeq show ?case by simp +next + case(cBrClose \ P M xvec N P' x A\<^sub>P' \\<^sub>P' C) + obtain A\<^sub>P \\<^sub>P where FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" and "distinct A\<^sub>P" + and "A\<^sub>P \* (\, P, M, xvec, N, P', A\<^sub>P', \\<^sub>P', C, x)" + by(rule freshFrame) + then have "A\<^sub>P \* \" and "A\<^sub>P \* P" and "A\<^sub>P \* M" and "A\<^sub>P \* xvec" and "A\<^sub>P \* N" and "A\<^sub>P \* P'" + and "A\<^sub>P \* A\<^sub>P'" and "A\<^sub>P \* \\<^sub>P'" and "A\<^sub>P \* C" and "x \ A\<^sub>P" + by simp+ + from FrP \A\<^sub>P \* xvec\ \xvec \* P\ have "xvec \* \\<^sub>P" + by(force dest: extractFrameFreshChain) + from \A\<^sub>P \* xvec\ \A\<^sub>P \* P'\ \x \ A\<^sub>P\ + have "A\<^sub>P \* (\ \ \\x\(\\*xvec\P'))" by simp + from \extractFrame P = \A\<^sub>P, \\<^sub>P\\ \extractFrame (\\x\P) = \A\<^sub>P', \\<^sub>P'\\ + have "\(x#A\<^sub>P), \\<^sub>P\ = \A\<^sub>P', \\<^sub>P'\" by simp + with \A\<^sub>P \* A\<^sub>P'\ \x \ A\<^sub>P'\ \x \ A\<^sub>P\ \distinct A\<^sub>P\ \distinct A\<^sub>P'\ + obtain p where S: "(set p \ (set (x#A\<^sub>P)) \ (set A\<^sub>P'))" and "distinctPerm p" + and \eq: "\\<^sub>P' = p \ \\<^sub>P" and Aeq: "A\<^sub>P' = p \ (x#A\<^sub>P)" + apply - + apply(rule frameChainEq') + by(assumption | simp)+ + + from \x \ A\<^sub>P'\ Aeq have "x \ (p \ (x#A\<^sub>P))" by simp + moreover from S Aeq \distinct A\<^sub>P'\ \x \ A\<^sub>P'\ \A\<^sub>P \* A\<^sub>P'\ have "(p \ x) \ A\<^sub>P" + by simp + from cBrClose FrP \distinct A\<^sub>P\ \A\<^sub>P \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* M\ \A\<^sub>P \* xvec\ \A\<^sub>P \* N\ \A\<^sub>P \* P'\ + \A\<^sub>P \* A\<^sub>P'\ \A\<^sub>P \* \\<^sub>P'\ \A\<^sub>P \* C\ \x \ A\<^sub>P\ \xvec \* \\<^sub>P\ \A\<^sub>P \* C\ \xvec \* C\ \x \ C\ + have "Prop C \ (\\x\P) (\ \ \\x\(\\*xvec\P')) (x#A\<^sub>P) \\<^sub>P" + by(force intro: rBrClose) + + moreover from Aeq \A\<^sub>P \* A\<^sub>P'\ have "A\<^sub>P \* (p \ A\<^sub>P)" by simp + moreover from Aeq \(set p \ (set (x#A\<^sub>P)) \ (set A\<^sub>P'))\ + have "(set p \ (set (x#A\<^sub>P)) \ (set (p \ (x#A\<^sub>P))))" by simp + moreover from \A\<^sub>P \* P\ \x \ A\<^sub>P\ have "A\<^sub>P \* (\\x\P)" by simp + moreover from S \x \ A\<^sub>P'\ have "p \ x \ x" + using Aeq by fastforce + moreover from \x \ A\<^sub>P'\ Aeq have "x \ p \ A\<^sub>P" by simp + moreover note \A\<^sub>P \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* (\ \ \\x\(\\*xvec\P'))\ + \A\<^sub>P \* C\ \distinctPerm p\ + \x \ \\ \x \ C\ \(p \ x) \ A\<^sub>P\ + ultimately + have "Prop C \ (\\x\P) (\ \ \\x\(\\*xvec\P')) (p \ (x#A\<^sub>P)) (p \ \\<^sub>P)" + by(fastforce intro!: rAlpha simp add: abs_fresh) + with \eq Aeq show ?case by simp +next + case(cOpen \ P M xvec yvec N P' x A\<^sub>x\<^sub>P \\<^sub>x\<^sub>P C) + obtain A\<^sub>P \\<^sub>P where FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" and "distinct A\<^sub>P" + and "A\<^sub>P \* (\, P, M, xvec, yvec, N, P', A\<^sub>x\<^sub>P, \\<^sub>x\<^sub>P, C, x)" + by(rule freshFrame) + then have "A\<^sub>P \* \" and "A\<^sub>P \* P" and "A\<^sub>P \* M" and "A\<^sub>P \* xvec"and "A\<^sub>P \* yvec" and "A\<^sub>P \* N" and "A\<^sub>P \* P'" + and "A\<^sub>P \* A\<^sub>x\<^sub>P" and "A\<^sub>P \* \\<^sub>x\<^sub>P" and "A\<^sub>P \* C" and "x \ A\<^sub>P" + by simp+ + + from \xvec \* P\ \A\<^sub>P \* xvec\ FrP have "xvec \* \\<^sub>P" + by(force dest: extractFrameFreshChain) + from \yvec \* P\ \A\<^sub>P \* yvec\ FrP have "yvec \* \\<^sub>P" + by(force dest: extractFrameFreshChain) + + from \extractFrame(\\x\P) = \A\<^sub>x\<^sub>P, \\<^sub>x\<^sub>P\\ FrP + have "\(x#A\<^sub>P), \\<^sub>P\ = \A\<^sub>x\<^sub>P, \\<^sub>x\<^sub>P\" + by simp + moreover from \x \ A\<^sub>P\ \distinct A\<^sub>P\ have "distinct(x#A\<^sub>P)" by simp + ultimately obtain p where S: "set p \ set (x#A\<^sub>P) \ set (p \ (x#A\<^sub>P))" and "distinctPerm p" + and \eq: "\\<^sub>x\<^sub>P = p \ \\<^sub>P" and Aeq: "A\<^sub>x\<^sub>P = (p \ x)#(p \ A\<^sub>P)" + using \A\<^sub>P \* A\<^sub>x\<^sub>P\\x \ A\<^sub>x\<^sub>P\ \distinct A\<^sub>x\<^sub>P\ + apply - + apply(rule frameChainEq') + by(assumption | simp)+ + + note \\ \ P \M\\*(xvec@yvec)\\N\ \ P'\ FrP \distinct A\<^sub>P\ + moreover from FrP \distinct A\<^sub>P\ \\A\<^sub>P \\<^sub>P C. \extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P\ \ Prop C \ P (M\\*(xvec@yvec)\\N\ \ P') A\<^sub>P \\<^sub>P\ + have "\C. Prop C \ P (M\\*(xvec@yvec)\\N\ \ P') A\<^sub>P \\<^sub>P" by simp + moreover note \x \ \\ \x \ M\ \x \ xvec\ \x \ yvec\ \x \ supp N\ \x \ A\<^sub>P\ \A\<^sub>P \* \\ \A\<^sub>P \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* M\ \A\<^sub>P \* xvec\ \A\<^sub>P \* yvec\ \A\<^sub>P \* N\ \A\<^sub>P \* P'\ + \xvec \* \\ \xvec \* P\ \xvec \* M\ \xvec \* \\<^sub>P\ \yvec \* \\ \yvec \* P\ \yvec \* M\ \yvec \* \\<^sub>P\ + \A\<^sub>P \* C\ \x \ C\ \xvec \* C\ \yvec \* C\ \xvec \* yvec\ \distinct xvec\ \distinct yvec\ + ultimately have "Prop C \ (\\x\P) (M\\*(xvec@x#yvec)\\N\ \ P') (x#A\<^sub>P) \\<^sub>P" + by(metis rOpen) + + with \A\<^sub>P \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* M\ \A\<^sub>P \* xvec\ \A\<^sub>P \* yvec\ \A\<^sub>P \* N\ \A\<^sub>P \* P'\ \A\<^sub>P \* A\<^sub>x\<^sub>P\ \A\<^sub>P \* C\ \x \ A\<^sub>x\<^sub>P\ \A\<^sub>P \* A\<^sub>x\<^sub>P\ \x \ A\<^sub>P\ + \x \ \\ \x \ M\ \x \ C\ \x \ xvec\ \x \ yvec\ Aeq + S \distinctPerm p\ + have "Prop C \ (\\x\P) (M\\*(xvec@x#yvec)\\N\ \ P') (p \ (x#A\<^sub>P)) (p \ \\<^sub>P)" + apply - + apply(rule rAlpha[where A\<^sub>P="x#A\<^sub>P"]) + by(assumption | simp add: abs_fresh fresh_star_def boundOutputFresh)+ + with \eq Aeq show ?case by(simp add: eqvts) +next + case(cBrOpen \ P M xvec yvec N P' x A\<^sub>x\<^sub>P \\<^sub>x\<^sub>P C) + obtain A\<^sub>P \\<^sub>P where FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" and "distinct A\<^sub>P" + and "A\<^sub>P \* (\, P, M, xvec, yvec, N, P', A\<^sub>x\<^sub>P, \\<^sub>x\<^sub>P, C, x)" + by(rule freshFrame) + then have "A\<^sub>P \* \" and "A\<^sub>P \* P" and "A\<^sub>P \* M" and "A\<^sub>P \* xvec"and "A\<^sub>P \* yvec" and "A\<^sub>P \* N" and "A\<^sub>P \* P'" + and "A\<^sub>P \* A\<^sub>x\<^sub>P" and "A\<^sub>P \* \\<^sub>x\<^sub>P" and "A\<^sub>P \* C" and "x \ A\<^sub>P" + by simp+ + + from \xvec \* P\ \A\<^sub>P \* xvec\ FrP have "xvec \* \\<^sub>P" + by(force dest: extractFrameFreshChain) + from \yvec \* P\ \A\<^sub>P \* yvec\ FrP have "yvec \* \\<^sub>P" + by(force dest: extractFrameFreshChain) + + from \extractFrame(\\x\P) = \A\<^sub>x\<^sub>P, \\<^sub>x\<^sub>P\\ FrP + have "\(x#A\<^sub>P), \\<^sub>P\ = \A\<^sub>x\<^sub>P, \\<^sub>x\<^sub>P\" + by simp + moreover from \x \ A\<^sub>P\ \distinct A\<^sub>P\ have "distinct(x#A\<^sub>P)" by simp + ultimately obtain p where S: "set p \ set (x#A\<^sub>P) \ set (p \ (x#A\<^sub>P))" and "distinctPerm p" + and \eq: "\\<^sub>x\<^sub>P = p \ \\<^sub>P" and Aeq: "A\<^sub>x\<^sub>P = (p \ x)#(p \ A\<^sub>P)" + using \A\<^sub>P \* A\<^sub>x\<^sub>P\\x \ A\<^sub>x\<^sub>P\ \distinct A\<^sub>x\<^sub>P\ + apply - + apply(rule frameChainEq') + by(assumption | simp)+ + + note \\ \ P \\M\\*(xvec@yvec)\\N\ \ P'\ FrP \distinct A\<^sub>P\ + moreover from FrP \distinct A\<^sub>P\ \\A\<^sub>P \\<^sub>P C. \extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P\ \ Prop C \ P (\M\\*(xvec@yvec)\\N\ \ P') A\<^sub>P \\<^sub>P\ + have "\C. Prop C \ P (\M\\*(xvec@yvec)\\N\ \ P') A\<^sub>P \\<^sub>P" by simp + moreover note \x \ \\ \x \ M\ \x \ xvec\ \x \ yvec\ \x \ supp N\ \x \ A\<^sub>P\ \A\<^sub>P \* \\ \A\<^sub>P \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* M\ \A\<^sub>P \* xvec\ \A\<^sub>P \* yvec\ \A\<^sub>P \* N\ \A\<^sub>P \* P'\ + \xvec \* \\ \xvec \* P\ \xvec \* M\ \xvec \* \\<^sub>P\ \yvec \* \\ \yvec \* P\ \yvec \* M\ \yvec \* \\<^sub>P\ + \A\<^sub>P \* C\ \x \ C\ \xvec \* C\ \yvec \* C\ \xvec \* yvec\ \distinct xvec\ \distinct yvec\ + ultimately have "Prop C \ (\\x\P) (\M\\*(xvec@x#yvec)\\N\ \ P') (x#A\<^sub>P) \\<^sub>P" + by(metis rBrOpen) + + with \A\<^sub>P \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* M\ \A\<^sub>P \* xvec\ \A\<^sub>P \* yvec\ \A\<^sub>P \* N\ \A\<^sub>P \* P'\ \A\<^sub>P \* A\<^sub>x\<^sub>P\ \A\<^sub>P \* C\ \x \ A\<^sub>x\<^sub>P\ \A\<^sub>P \* A\<^sub>x\<^sub>P\ \x \ A\<^sub>P\ + \x \ \\ \x \ M\ \x \ C\ \x \ xvec\ \x \ yvec\ Aeq + S \distinctPerm p\ + have "Prop C \ (\\x\P) (\M\\*(xvec@x#yvec)\\N\ \ P') (p \ (x#A\<^sub>P)) (p \ \\<^sub>P)" + apply - + apply(rule rAlpha[where A\<^sub>P="x#A\<^sub>P"]) + by(assumption | simp add: abs_fresh fresh_star_def boundOutputFresh)+ + with \eq Aeq show ?case by(simp add: eqvts) +next + case(cScope \ P \ P' x A\<^sub>x\<^sub>P \\<^sub>x\<^sub>P C) + obtain A\<^sub>P \\<^sub>P where FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" and "distinct A\<^sub>P" + and "A\<^sub>P \* (\, P, \, P', A\<^sub>x\<^sub>P, \\<^sub>x\<^sub>P, C, x)" + by(rule freshFrame) + then have "A\<^sub>P \* \" and "A\<^sub>P \* P" and "A\<^sub>P \* \" and "A\<^sub>P \* P'" + and "A\<^sub>P \* A\<^sub>x\<^sub>P" and "A\<^sub>P \* \\<^sub>x\<^sub>P" and "A\<^sub>P \* C" and "x \ A\<^sub>P" + by simp+ + + from \bn \ \* P\ \A\<^sub>P \* \\ FrP have "bn \ \* \\<^sub>P" + by(force dest: extractFrameFreshChain) + + from \extractFrame(\\x\P) = \A\<^sub>x\<^sub>P, \\<^sub>x\<^sub>P\\ FrP + have "\(x#A\<^sub>P), \\<^sub>P\ = \A\<^sub>x\<^sub>P, \\<^sub>x\<^sub>P\" + by simp + moreover from \x \ A\<^sub>P\ \distinct A\<^sub>P\ have "distinct(x#A\<^sub>P)" by simp + ultimately obtain p where S: "set p \ set (x#A\<^sub>P) \ set (p \ (x#A\<^sub>P))" and "distinctPerm p" + and \eq: "\\<^sub>x\<^sub>P = p \ \\<^sub>P" and Aeq: "A\<^sub>x\<^sub>P = (p \ x)#(p \ A\<^sub>P)" + using \A\<^sub>P \* A\<^sub>x\<^sub>P\\x \ A\<^sub>x\<^sub>P\ \distinct A\<^sub>x\<^sub>P\ + apply - + apply(rule frameChainEq') + by(assumption | simp)+ + + note \\ \ P \\ \ P'\ FrP \distinct A\<^sub>P\ + moreover from FrP \distinct A\<^sub>P\ \\A\<^sub>P \\<^sub>P C. \extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P\ \ Prop C \ P (\ \ P') A\<^sub>P \\<^sub>P\ + have "\C. Prop C \ P (\ \ P') A\<^sub>P \\<^sub>P" by simp + moreover note \x \ \\ \x \ \\ \x \ A\<^sub>P\ \A\<^sub>P \* \\ \A\<^sub>P \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* \\ \A\<^sub>P \* P'\ \distinct(bn \)\ + \bn \ \* \\ \bn \ \* P\ \bn \ \* subject \\ \bn \ \* \\<^sub>P\ \A\<^sub>P \* C\ \x \ C\ \bn \ \* C\ + ultimately have "Prop C \ (\\x\P) (\ \ (\\x\P')) (x#A\<^sub>P) \\<^sub>P" + by(metis rScope) + + with \A\<^sub>P \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* \\ \A\<^sub>P \* P'\ \A\<^sub>P \* A\<^sub>x\<^sub>P\ \A\<^sub>P \* C\ \x \ A\<^sub>x\<^sub>P\ \A\<^sub>P \* A\<^sub>x\<^sub>P\ \x \ A\<^sub>P\ + \x \ \\ \x \ \\ \x \ C\ Aeq + S \distinctPerm p\ + have "Prop C \ (\\x\P) (\ \ (\\x\P')) (p \ (x#A\<^sub>P)) (p \ \\<^sub>P)" + apply - + apply(rule rAlpha[where A\<^sub>P="x#A\<^sub>P"]) + by(assumption | simp add: abs_fresh fresh_star_def)+ + with \eq Aeq show ?case by(simp add: eqvts) +next + case(Bang \ P Rs A\<^sub>b\<^sub>P \\<^sub>b\<^sub>P C) + + obtain A\<^sub>P \\<^sub>P where FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" and "distinct A\<^sub>P" + and "A\<^sub>P \* (\, P, Rs, C)" + by(rule freshFrame) + then have "A\<^sub>P \* \" and "A\<^sub>P \* P" and "A\<^sub>P \* Rs" and "A\<^sub>P \* C" + by simp+ + + note \\ \ P \ !P \ Rs\ \guarded P\ FrP \distinct A\<^sub>P\ + moreover from FrP have "extractFrame (P \ !P) = \A\<^sub>P, \\<^sub>P \ \\" + by simp + with \distinct A\<^sub>P\ \\A\<^sub>P \\<^sub>P C. \extractFrame (P \ !P) = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P\ \ Prop C \ (P \ !P) Rs A\<^sub>P \\<^sub>P\ + have "\C. Prop C \ (P \ !P) Rs A\<^sub>P (\\<^sub>P \ \)" by simp + moreover from \guarded P\ FrP have "\\<^sub>P \ \" and "supp \\<^sub>P = ({}::name set)" by(metis guardedStatEq)+ + moreover note \A\<^sub>P \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* Rs\ \A\<^sub>P \* C\ + ultimately have "Prop C \ (!P) Rs ([]) (\)" + by(rule rBang) + then show ?case using \extractFrame(!P) = \A\<^sub>b\<^sub>P, \\<^sub>b\<^sub>P\\ by simp +qed + +lemma semanticsFrameInduct'[consumes 5, case_names cAlpha cFrameAlpha cInput cBrInput cOutput cBrOutput cCase cPar1 cPar2 cComm1 cComm2 cBrMerge cBrComm1 cBrComm2 cBrClose cOpen cBrOpen cScope cBang]: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and Rs :: "('a, 'b, 'c) residual" + and A\<^sub>P :: "name list" + and \\<^sub>P :: 'b + and Prop :: "'f::fs_name \ 'b \ ('a, 'b, 'c) psi \ 'a action \ + ('a, 'b, 'c) psi \ name list \ 'b \ bool" + and C :: "'f::fs_name" + +assumes Trans: "\ \ P \\ \ P'" + and FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" + and "distinct A\<^sub>P" + and "bn \ \* subject \" + and "distinct(bn \)" + and rAlpha: "\\ P \ P' p A\<^sub>P \\<^sub>P C. \bn \ \* \; bn \ \* P; bn \ \* subject \; bn \ \* \\<^sub>P; + bn \ \* C; bn \ \* (p \ \); A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* \; A\<^sub>P \* P'; A\<^sub>P \* C; + set p \ set(bn \) \ set(bn(p \ \)); distinctPerm p; + bn(p \ \) \* \; (bn(p \ \)) \* P'; Prop C \ P \ P' A\<^sub>P \\<^sub>P\ \ + Prop C \ P (p \ \) (p \ P') A\<^sub>P \\<^sub>P" + and rFrameAlpha: "\\ P A\<^sub>P \\<^sub>P p \ P' C. \A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* (p \ A\<^sub>P); A\<^sub>P \* \; A\<^sub>P \* P'; A\<^sub>P \* C; + set p \ set A\<^sub>P \ set(p \ A\<^sub>P); distinctPerm p; A\<^sub>P \* subject \; + Prop C \ P \ P' A\<^sub>P \\<^sub>P\ \ Prop C \ P \ P' (p \ A\<^sub>P) (p \ \\<^sub>P)" + and rInput: "\\ M K xvec N Tvec P C. + \\ \ M \ K; distinct xvec; set xvec \ supp N; + length xvec = length Tvec; xvec \* \; + xvec \* M; xvec \* K; xvec \* C\ \ + Prop C \ (M\\*xvec N\.P) + (K\(N[xvec::=Tvec])\) (P[xvec::=Tvec]) ([]) (\)" + and rBrInput: "\\ M K xvec N Tvec P C. + \\ \ K \ M; distinct xvec; set xvec \ supp N; + length xvec = length Tvec; xvec \* \; + xvec \* M; xvec \* K; xvec \* C\ \ + Prop C \ (M\\*xvec N\.P) + (\K\(N[xvec::=Tvec])\) (P[xvec::=Tvec]) ([]) (\)" + and rOutput: "\\ M K N P C. \ \ M \ K \ Prop C \ (M\N\.P) (K\N\) P ([]) (\)" + and rBrOutput: "\\ M K N P C. \ \ M \ K \ Prop C \ (M\N\.P) (\K\N\) P ([]) (\)" + and rCase: "\\ P \ P' \ Cs A\<^sub>P \\<^sub>P C. \\ \ P \\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; \C. Prop C \ P \ P' A\<^sub>P \\<^sub>P; + (\, P) \ set Cs; \ \ \; guarded P; \\<^sub>P \ \; (supp \\<^sub>P) = ({}::name set); + A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* \; A\<^sub>P \* P'; A\<^sub>P \* C\ \ Prop C \ (Cases Cs) \ P' ([]) (\)" + and rPar1: "\\ \\<^sub>Q P \ P' A\<^sub>Q Q A\<^sub>P \\<^sub>P C. + \\ \ \\<^sub>Q \ P \\ \ P'; + extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + \C. Prop C (\ \ \\<^sub>Q) P \ P' A\<^sub>P \\<^sub>P; + A\<^sub>P \* P; A\<^sub>P \* Q; A\<^sub>P \* \; A\<^sub>P \* \; A\<^sub>P \* P'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* \\<^sub>Q; + A\<^sub>Q \* P; A\<^sub>Q \* Q; A\<^sub>Q \* \; A\<^sub>Q \* \; A\<^sub>Q \* P'; A\<^sub>Q \* \\<^sub>P; + bn \ \* \; bn \ \* P; bn \ \* Q; bn \ \* subject \; bn \ \* \\<^sub>P; bn \ \* \\<^sub>Q; + A\<^sub>P \* C; A\<^sub>Q \* C; bn \ \* C\ \ + Prop C \ (P \ Q) \ (P' \ Q) (A\<^sub>P@A\<^sub>Q) (\\<^sub>P \ \\<^sub>Q)" + and rPar2: "\\ \\<^sub>P Q \ Q' A\<^sub>P P A\<^sub>Q \\<^sub>Q C. + \\ \ \\<^sub>P \ Q \\ \ Q'; + extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + \C. Prop C (\ \ \\<^sub>P) Q \ Q' A\<^sub>Q \\<^sub>Q; + A\<^sub>P \* P; A\<^sub>P \* Q; A\<^sub>P \* \; A\<^sub>P \* \; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* \\<^sub>Q; + A\<^sub>Q \* P; A\<^sub>Q \* Q; A\<^sub>Q \* \; A\<^sub>Q \* \; A\<^sub>Q \* Q'; A\<^sub>Q \* \\<^sub>P; + bn \ \* \; bn \ \* P; bn \ \* Q; bn \ \* subject \; bn \ \* \\<^sub>P; bn \ \* \\<^sub>Q; + A\<^sub>P \* C; A\<^sub>Q \* C; bn \ \* C\ \ + Prop C \ (P \ Q) \ (P \ Q') (A\<^sub>P@A\<^sub>Q) (\\<^sub>P \ \\<^sub>Q)" + and rComm1: "\\ \\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q K xvec Q' A\<^sub>Q C. + \\ \ \\<^sub>Q \ P \M\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \C. Prop C (\ \ \\<^sub>Q) P (M\N\) P' A\<^sub>P \\<^sub>P; + \ \ \\<^sub>P \ Q \K\\*xvec\\N\ \ Q'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + \ \ \\<^sub>P \ \\<^sub>Q \ M \ K; distinct xvec; + \C. Prop C (\ \ \\<^sub>P) Q (K\\*xvec\\N\) Q' A\<^sub>Q \\<^sub>Q; + A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* P'; + A\<^sub>P \* Q; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* xvec; A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; + A\<^sub>Q \* P; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* Q; A\<^sub>Q \* K; A\<^sub>Q \* Q'; + A\<^sub>Q \* xvec; xvec \* \; xvec \* \\<^sub>P; xvec \* \\<^sub>Q; xvec \* P; xvec \* M; + xvec \* Q; xvec \* K; A\<^sub>P \* C; A\<^sub>Q \* C; xvec \* C\ \ + Prop C \ (P \ Q) (\) (\\*xvec\(P' \ Q')) (A\<^sub>P@A\<^sub>Q) (\\<^sub>P \ \\<^sub>Q)" + and rComm2: "\\ \\<^sub>Q P M xvec N P' A\<^sub>P \\<^sub>P Q K Q' A\<^sub>Q C. + \\ \ \\<^sub>Q \ P \M\\*xvec\\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \C. Prop C (\ \ \\<^sub>Q) P (M\\*xvec\\N\) P' A\<^sub>P \\<^sub>P; + \ \ \\<^sub>P \ Q \K\N\ \ Q'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + \C. Prop C (\ \ \\<^sub>P) Q (K\N\) Q' A\<^sub>Q \\<^sub>Q; + \ \ \\<^sub>P \ \\<^sub>Q \ M \ K; distinct xvec; + A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* P'; + A\<^sub>P \* Q; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* xvec; A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; + A\<^sub>Q \* P; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* Q; A\<^sub>Q \* K; A\<^sub>Q \* Q'; + A\<^sub>Q \* xvec; xvec \* \; xvec \* \\<^sub>P; xvec \* \\<^sub>Q; xvec \* P; xvec \* M; + xvec \* Q; xvec \* K; A\<^sub>P \* C; A\<^sub>Q \* C; xvec \* C\ \ + Prop C \ (P \ Q) (\) (\\*xvec\(P' \ Q')) (A\<^sub>P@A\<^sub>Q) (\\<^sub>P \ \\<^sub>Q)" + and rBrMerge: "\\ \\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q Q' A\<^sub>Q C. + \\ \ \\<^sub>Q \ P \ \M\N\ \ P'; \C. Prop C (\ \ \\<^sub>Q) P (\M\N\) P' A\<^sub>P \\<^sub>P; + extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \ \ \\<^sub>P \ Q \ \M\N\ \ Q'; \C. Prop C (\ \ \\<^sub>P) Q (\M\N\) Q' A\<^sub>Q \\<^sub>Q; + extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* N; A\<^sub>P \* P'; + A\<^sub>P \* Q; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* M; A\<^sub>Q \* M; + A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; A\<^sub>Q \* P; A\<^sub>Q \* N; A\<^sub>Q \* P'; + A\<^sub>Q \* Q; A\<^sub>Q \* Q'; A\<^sub>P \* C; A\<^sub>Q \* C\ \ + Prop C \ (P \ Q) (\M\N\) (P' \ Q') (A\<^sub>P@A\<^sub>Q) (\\<^sub>P \ \\<^sub>Q)" + and rBrComm1:"\\ \\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q xvec Q' A\<^sub>Q C. + \\ \ \\<^sub>Q \ P \\M\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \C. Prop C (\ \ \\<^sub>Q) P (\M\N\) P' A\<^sub>P \\<^sub>P; + \ \ \\<^sub>P \ Q \\M\\*xvec\\N\ \ Q'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + distinct xvec; + \C. Prop C (\ \ \\<^sub>P) Q (\M\\*xvec\\N\) Q' A\<^sub>Q \\<^sub>Q; + A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* N; A\<^sub>P \* P'; + A\<^sub>P \* Q; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* xvec; A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; + A\<^sub>Q \* P; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* Q; A\<^sub>Q \* Q'; + A\<^sub>Q \* xvec; xvec \* \; xvec \* \\<^sub>P; xvec \* \\<^sub>Q; xvec \* P; + xvec \* Q; A\<^sub>P \* C; A\<^sub>Q \* C; xvec \* C; + A\<^sub>P \* M; A\<^sub>Q \* M; xvec \* M\ \ + Prop C \ (P \ Q) (\M\\*xvec\\N\) (P' \ Q') (A\<^sub>P@A\<^sub>Q) (\\<^sub>P \ \\<^sub>Q)" + and rBrComm2:"\\ \\<^sub>Q P M xvec N P' A\<^sub>P \\<^sub>P Q Q' A\<^sub>Q C. + \\ \ \\<^sub>Q \ P \\M\\*xvec\\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \C. Prop C (\ \ \\<^sub>Q) P (\M\\*xvec\\N\) P' A\<^sub>P \\<^sub>P; + \ \ \\<^sub>P \ Q \\M\N\ \ Q'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + \C. Prop C (\ \ \\<^sub>P) Q (\M\N\) Q' A\<^sub>Q \\<^sub>Q; + distinct xvec; + A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* N; A\<^sub>P \* P'; + A\<^sub>P \* Q; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* xvec; A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; + A\<^sub>Q \* P; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* Q; A\<^sub>Q \* Q'; + A\<^sub>Q \* xvec; xvec \* \; xvec \* \\<^sub>P; xvec \* \\<^sub>Q; xvec \* P; + xvec \* Q; A\<^sub>P \* C; A\<^sub>Q \* C; xvec \* C; + A\<^sub>P \* M; A\<^sub>Q \* M; xvec \* M\ \ + Prop C \ (P \ Q) (\M\\*xvec\\N\) (P' \ Q') (A\<^sub>P@A\<^sub>Q) (\\<^sub>P \ \\<^sub>Q)" + and rBrClose: "\\ P M xvec N P' A\<^sub>P \\<^sub>P x C. + \\ \ P \ \M\\*xvec\\N\ \ P'; + x \ supp M; + \C. Prop C \ P (\M\\*xvec\\N\) P' A\<^sub>P \\<^sub>P; + extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* P'; A\<^sub>P \* xvec; + distinct xvec; xvec \* \; xvec \* \\<^sub>P; xvec \* P; + xvec \* M; + x \ \; x \ xvec; x \ A\<^sub>P; + A\<^sub>P \* C; xvec \* C; x \ C\ \ + Prop C \ (\\x\P) (\) (\\x\(\\*xvec\P')) (x#A\<^sub>P) \\<^sub>P" + and rOpen: "\\ P M xvec yvec N P' x A\<^sub>P \\<^sub>P y C. + \\ \ P \M\\*(xvec@yvec)\\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \C. Prop C \ P (M\\*(xvec@yvec)\\N\) P' A\<^sub>P \\<^sub>P; x \ supp N; x \ \; x \ M; + x \ A\<^sub>P; x \ xvec; x \ yvec; A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* P'; + A\<^sub>P \* xvec; A\<^sub>P \* yvec; xvec \* yvec; distinct xvec; distinct yvec; + xvec \* \; xvec \* P; xvec \* M; xvec \* \\<^sub>P; + yvec \* \; yvec \* P; yvec \* M; A\<^sub>P \* C; x \ C; xvec \* C; yvec \* C; + y \ x; y \ \; y \ P; y \ M; y \ xvec; y \ yvec; y \ N; y \ P'; y \ A\<^sub>P; y \ \\<^sub>P; y \ C\ \ + Prop C \ (\\x\P) (M\\*(xvec@y#yvec)\\([(x, y)] \ N)\) ([(x, y)] \ P') (x#A\<^sub>P) \\<^sub>P" + and rBrOpen: "\\ P M xvec yvec N P' x A\<^sub>P \\<^sub>P y C. + \\ \ P \\M\\*(xvec@yvec)\\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \C. Prop C \ P (\M\\*(xvec@yvec)\\N\) P' A\<^sub>P \\<^sub>P; x \ supp N; x \ \; x \ M; + x \ A\<^sub>P; x \ xvec; x \ yvec; A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* P'; + A\<^sub>P \* xvec; A\<^sub>P \* yvec; xvec \* yvec; distinct xvec; distinct yvec; + xvec \* \; xvec \* P; xvec \* M; xvec \* \\<^sub>P; + yvec \* \; yvec \* P; yvec \* M; A\<^sub>P \* C; x \ C; xvec \* C; yvec \* C; + y \ x; y \ \; y \ P; y \ M; y \ xvec; y \ yvec; y \ N; y \ P'; y \ A\<^sub>P; y \ \\<^sub>P; y \ C\ \ + Prop C \ (\\x\P) (\M\\*(xvec@y#yvec)\\([(x, y)] \ N)\) ([(x, y)] \ P') (x#A\<^sub>P) \\<^sub>P" + and rScope: "\\ P \ P' x A\<^sub>P \\<^sub>P C. + \\ \ P \\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \C. Prop C \ P \ P' A\<^sub>P \\<^sub>P; + x \ \; x \ \; x \ A\<^sub>P; A\<^sub>P \* \; A\<^sub>P \* P; + A\<^sub>P \* \; A\<^sub>P \* P'; + bn \ \* \; bn \ \* P; bn \ \* subject \; bn \ \* \\<^sub>P; + A\<^sub>P \* C; x \ C; bn \ \* C\ \ + Prop C \ (\\x\P) \ (\\x\P') (x#A\<^sub>P) \\<^sub>P" + and rBang: "\\ P \ P' A\<^sub>P \\<^sub>P C. + \\ \ P \ !P \\ \ P'; guarded P; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \C. Prop C \ (P \ !P) \ P' A\<^sub>P (\\<^sub>P \ \); \\<^sub>P \ \; supp \\<^sub>P = ({}::name set); + A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* \; A\<^sub>P \* P'; A\<^sub>P \* C\ \ Prop C \ (!P) \ P' ([]) (\)" +shows "Prop C \ P \ P' A\<^sub>P \\<^sub>P" + using Trans FrP \distinct A\<^sub>P\ \bn \ \* subject \\ \distinct(bn \)\ +proof(nominal_induct \ P Rs=="\ \ P'" A\<^sub>P \\<^sub>P avoiding: C \ P' rule: semanticsFrameInduct) + case cAlpha + then show ?case using rFrameAlpha + by auto +next + case cInput + then show ?case using rInput + by(auto simp add: residualInject) +next + case cBrInput + then show ?case using rBrInput + by(auto simp add: residualInject) +next + case cOutput + then show ?case using rOutput + by(auto simp add: residualInject) +next + case cBrOutput + then show ?case using rBrOutput + by(auto simp add: residualInject) +next + case cCase + then show ?case using rCase + by(auto simp add: residualInject) +next + case(cPar1 \ \\<^sub>Q P \ P' A\<^sub>Q Q A\<^sub>P \\<^sub>P C \' P'') + note \\ \ (P' \ Q) = \' \ P''\ + moreover from \bn \ \* \'\ have "bn \ \* (bn \')" by auto + moreover note \distinct (bn \)\ \distinct(bn \')\ + moreover from \bn \ \* subject \\ \bn \' \* subject \'\ + have "bn \ \* (\ \ P' \ Q)" and "bn \' \* (\' \ P'')" by simp+ + ultimately obtain p where S: "(set p) \ (set(bn \)) \ (set(bn(p \ \)))" and "distinctPerm p" + and \Eq: "\' = p \ \" and P'eq: "P'' = p \ (P' \ Q)" and "(bn(p \ \)) \* \" + and "(bn(p \ \)) \* (P' \ Q)" + by(rule residualEq) + + note \\ \ \\<^sub>Q \ P \\ \ P'\ \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ \distinct A\<^sub>Q\ + moreover from \bn \ \* subject \\ \distinct(bn \)\ \A\<^sub>P \* \\ + have "\C. Prop C (\ \ \\<^sub>Q) P \ P' A\<^sub>P \\<^sub>P" by(fastforce intro: cPar1) + + moreover note \A\<^sub>Q \* P\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* \\ \A\<^sub>Q \* \\ \A\<^sub>Q \* P'\ \A\<^sub>Q \* C\ \A\<^sub>P \* P\ \A\<^sub>P \* Q\ \A\<^sub>P \* \\ \A\<^sub>P \* \\ \A\<^sub>P \* P'\ \A\<^sub>P \* C\ + \bn \ \* Q\ \distinct(bn \)\ \bn \ \* \\ \bn \ \* \\<^sub>Q\ \bn \ \* P\ \bn \ \* subject \\ \bn \ \* C\ + \extractFrame P = \A\<^sub>P, \\<^sub>P\\ \distinct A\<^sub>P\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>Q \* \\<^sub>P\ \bn \ \* \\<^sub>P\ + ultimately have "Prop C \ (P \ Q) \ (P' \ Q) (A\<^sub>P@A\<^sub>Q) (\\<^sub>P \ \\<^sub>Q)" + by(metis rPar1) + with \bn \ \* \\ \bn \ \* P\ \bn \ \* Q\ \bn \ \* subject \\ \bn \ \* C\ \bn \ \* (bn \')\ S \distinctPerm p\ \bn(p \ \) \* \\ \bn(p \ \) \* (P' \ Q)\ \bn \ \* \\<^sub>P\ \bn \ \* \\<^sub>Q\ \A\<^sub>P \* \\ \A\<^sub>Q \* \\ \A\<^sub>P \* \'\ \A\<^sub>Q \* \'\ \Eq \bn \ \* \\<^sub>P\ \bn \ \* \'\ \A\<^sub>P \* \\ \A\<^sub>Q \* \\ \A\<^sub>P \* P\ \A\<^sub>Q \* P\ \A\<^sub>P \* Q\ \A\<^sub>Q \* Q\ \A\<^sub>P \* P'\ \A\<^sub>Q \* P'\ \A\<^sub>P \* C\ \A\<^sub>Q \* C\ + have "Prop C \ (P \ Q) (p \ \) (p \ (P' \ Q)) (A\<^sub>P@A\<^sub>Q) (\\<^sub>P \ \\<^sub>Q)" + by(fastforce intro!: rAlpha) + with \Eq P'eq \distinctPerm p\ show ?case by simp +next + case(cPar2 \ \\<^sub>P Q \ Q' A\<^sub>P P A\<^sub>Q \\<^sub>Q C \' Q'') + note \\ \ (P \ Q') = \' \ Q''\ + moreover from \bn \ \* \'\ have "bn \ \* (bn \')" by auto + moreover note \distinct (bn \)\ \distinct(bn \')\ + moreover from \bn \ \* subject \\ \bn \' \* subject \'\ + have "bn \ \* (\ \ P \ Q')" and "bn \' \* (\' \ Q'')" by simp+ + ultimately obtain p where S: "(set p) \ (set(bn \)) \ (set(bn(p \ \)))" and "distinctPerm p" + and \Eq: "\' = p \ \" and Q'eq: "Q'' = p \ (P \ Q')" and "(bn(p \ \)) \* \" + and "(bn(p \ \)) \* (P \ Q')" + by(rule residualEq) + + note \\ \ \\<^sub>P \ Q \\ \ Q'\ \extractFrame P = \A\<^sub>P, \\<^sub>P\\ \distinct A\<^sub>P\ + moreover from \bn \ \* subject \\ \distinct(bn \)\ \A\<^sub>Q \* \\ + have "\C. Prop C (\ \ \\<^sub>P) Q \ Q' A\<^sub>Q \\<^sub>Q" by(fastforce intro!: cPar2) + + moreover note \A\<^sub>Q \* P\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* \\ \A\<^sub>Q \* \\ \A\<^sub>Q \* Q'\ \A\<^sub>Q \* C\ \A\<^sub>P \* P\ \A\<^sub>P \* Q\ \A\<^sub>P \* \\ \A\<^sub>P \* \\ \A\<^sub>P \* Q'\ \A\<^sub>P \* C\ + \bn \ \* Q\ \distinct(bn \)\ \bn \ \* \\ \bn \ \* \\<^sub>Q\ \bn \ \* P\ \bn \ \* subject \\ \bn \ \* C\ + \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ \distinct A\<^sub>Q\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>Q \* \\<^sub>P\ \bn \ \* \\<^sub>P\ + ultimately have "Prop C \ (P \ Q) \ (P \ Q') (A\<^sub>P@A\<^sub>Q) (\\<^sub>P \ \\<^sub>Q)" + by(fastforce intro!: rPar2) + with \bn \ \* \\ \bn \ \* P\ \bn \ \* Q\ \bn \ \* subject \\ \bn \ \* C\ \bn \ \* (bn \')\ S \distinctPerm p\ \bn(p \ \) \* \\ \bn(p \ \) \* (P \ Q')\ \bn \ \* \\<^sub>P\ \bn \ \* \\<^sub>Q\ \A\<^sub>P \* \\ \A\<^sub>Q \* \\ \A\<^sub>P \* \'\ \A\<^sub>Q \* \'\ \Eq \bn \ \* \'\ \A\<^sub>P \* \\ \A\<^sub>Q \* \\ \A\<^sub>P \* P\ \A\<^sub>Q \* P\ \A\<^sub>P \* Q\ \A\<^sub>Q \* Q\ \A\<^sub>P \* Q'\ \A\<^sub>Q \* Q'\ \A\<^sub>P \* C\ \A\<^sub>Q \* C\ + have "Prop C \ (P \ Q) (p \ \) (p \ (P \ Q')) (A\<^sub>P@A\<^sub>Q) (\\<^sub>P \ \\<^sub>Q)" + by(fastforce intro!: rAlpha) + with \Eq Q'eq \distinctPerm p\ show ?case by simp +next + case(cComm1 \ \\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q K xvec Q' A\<^sub>Q C \ P'') + then show ?case using rComm1 + apply - + apply(drule meta_spec[where x="M\N\"]) + apply(drule meta_spec[where x="K\\*xvec\\N\"]) + by(auto simp add: residualInject) +next + case(cComm2 \ \\<^sub>Q P M xvec N P' A\<^sub>P \\<^sub>P Q K Q' A\<^sub>Q C \ Q'') + then show ?case using rComm2 + apply - + apply(drule meta_spec[where x="M\\*xvec\\N\"]) + apply(drule meta_spec[where x="K\N\"]) + by(auto simp add: residualInject) +next + case(cBrMerge \ \\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q Q' A\<^sub>Q C \ P'') + then show ?case using rBrMerge + apply - + apply(drule meta_spec[where x="\M\N\"]) + apply(drule meta_spec[where x="\M\N\"]) + by(auto simp add: residualInject) +next + case(cBrComm1 \ \\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q xvec Q' A\<^sub>Q C \ P'') + have "bn (\M\N\) \* subject (\M\N\)" by simp + moreover have "distinct (bn (\M\N\))" by simp + moreover have "\M\N\ \ P' = \M\N\ \ P'" by simp + moreover note cBrComm1 + ultimately have inProp: "\C. Prop C (\ \ \\<^sub>Q) P (\M\N\) P' A\<^sub>P \\<^sub>P" by simp + + note \xvec \* M\ \distinct xvec\ cBrComm1 + then have outProp: "\C. Prop C (\ \ \\<^sub>P) Q (\M\\*xvec\\N\) Q' A\<^sub>Q \\<^sub>Q" by simp + + note inProp outProp cBrComm1 + then have bigProp: "Prop C \ (P \ Q) (\M\\*xvec\\N\) (P' \ Q') (A\<^sub>P @ A\<^sub>Q) (\\<^sub>P \ \\<^sub>Q)" by (simp add: rBrComm1) + + note \\M\\*xvec\\N\ \ (P' \ Q') = \ \ P''\ + moreover from \xvec \* \\ have "bn (\M\\*xvec\\N\) \* (bn \)" by simp + moreover from \distinct xvec\ have "distinct (bn (\M\\*xvec\\N\))" by simp + moreover note \distinct(bn \)\ + moreover from \xvec \* M\ \bn \ \* subject \\ + have "bn (\M\\*xvec\\N\) \* (\M\\*xvec\\N\ \ P' \ Q')" and "bn \ \* (\ \ P'')" by simp+ + ultimately obtain p where S: "(set p) \ (set(bn (\M\\*xvec\\N\))) \ (set(bn(p \ (\M\\*xvec\\N\))))" and "distinctPerm p" + and \Eq: "\ = p \ (\M\\*xvec\\N\)" and P'eq: "P'' = p \ (P' \ Q')" + and "(bn(p \ (\M\\*xvec\\N\))) \* (\M\\*xvec\\N\)" and "(bn(p \ (\M\\*xvec\\N\))) \* (P' \ Q')" + by(rule residualEq) simp + + from \xvec \* \\ have "bn (\M\\*xvec\\N\) \* \" by simp + moreover from \xvec \* P\ \xvec \* Q\ have "bn (\M\\*xvec\\N\) \* (P \ Q)" by simp + moreover note \xvec \* M\ + moreover from \xvec \* \\<^sub>P\ \xvec \* \\<^sub>Q\ have "bn (\M\\*xvec\\N\) \* (\\<^sub>P \ \\<^sub>Q)" by auto + moreover from \xvec \* C\ have "bn (\M\\*xvec\\N\) \* C" by simp + moreover from \xvec \* \\ \Eq have "bn (\M\\*xvec\\N\) \* (p \ (\M\\*xvec\\N\))" by simp + moreover from \A\<^sub>P \* \\ \A\<^sub>Q \* \\ have "(A\<^sub>P @ A\<^sub>Q) \* \" by simp + moreover from \A\<^sub>P \* P\ \A\<^sub>Q \* Q\ \A\<^sub>P \* Q\ \A\<^sub>Q \* P\ have "(A\<^sub>P @ A\<^sub>Q) \* (P \ Q)" by simp + moreover from \A\<^sub>P \* \\ \A\<^sub>Q \* \\ have "(A\<^sub>P @ A\<^sub>Q) \* \" by simp + moreover from \A\<^sub>P \* P'\ \A\<^sub>Q \* Q'\ \A\<^sub>P \* Q'\ \A\<^sub>Q \* P'\ have "(A\<^sub>P @ A\<^sub>Q) \* (P' \ Q')" by simp + moreover from \A\<^sub>P \* C\ \A\<^sub>Q \* C\ have "(A\<^sub>P @ A\<^sub>Q) \* C" by simp + moreover note S \distinctPerm p\ \(bn(p \ (\M\\*xvec\\N\))) \* (\M\\*xvec\\N\)\ + \(bn(p \ (\M\\*xvec\\N\))) \* (P' \ Q')\ bigProp + \A\<^sub>P \* M\ \A\<^sub>Q \* M\ \A\<^sub>P \* xvec\ \A\<^sub>Q \* xvec\ \A\<^sub>P \* N\ \A\<^sub>Q \* N\ + + ultimately have "Prop C \ (P \ Q) (p \ (\M\\*xvec\\N\)) (p \ (P' \ Q')) (A\<^sub>P @ A\<^sub>Q) (\\<^sub>P \ \\<^sub>Q)" + by(fastforce intro!: rAlpha) + then show ?case using \Eq P'eq by simp +next + case(cBrComm2 \ \\<^sub>Q P M xvec N P' A\<^sub>P \\<^sub>P Q Q' A\<^sub>Q C \ Q'') + have "bn (\M\N\) \* subject (\M\N\)" by simp + moreover have "distinct (bn (\M\N\))" by simp + moreover have "\M\N\ \ Q' = \M\N\ \ Q'" by simp + moreover note cBrComm2 + ultimately have inProp: "\C. Prop C (\ \ \\<^sub>P) Q (\M\N\) Q' A\<^sub>Q \\<^sub>Q" by simp + + from \xvec \* M\ have "bn (\M\\*xvec\\N\) \* subject (\M\\*xvec\\N\)" by simp + moreover from \distinct xvec\ have "distinct (bn (\M\\*xvec\\N\))" by simp + moreover have "\M\\*xvec\\N\ \ P' = \M\\*xvec\\N\ \ P'" by simp + moreover note cBrComm2 + ultimately have outProp: "\C. Prop C (\ \ \\<^sub>Q) P (\M\\*xvec\\N\) P' A\<^sub>P \\<^sub>P" by simp + + note inProp outProp cBrComm2 + then have bigProp: "Prop C \ (P \ Q) (\M\\*xvec\\N\) (P' \ Q') (A\<^sub>P @ A\<^sub>Q) (\\<^sub>P \ \\<^sub>Q)" by (simp add: rBrComm2) + + note \\M\\*xvec\\N\ \ (P' \ Q') = \ \ Q''\ + moreover from \xvec \* \\ have "bn (\M\\*xvec\\N\) \* (bn \)" by simp + moreover note \distinct (bn (\M\\*xvec\\N\))\ \distinct(bn \)\ + moreover from \bn (\M\\*xvec\\N\) \* subject (\M\\*xvec\\N\)\ \bn \ \* subject \\ + have "bn (\M\\*xvec\\N\) \* (\M\\*xvec\\N\ \ P' \ Q')" and "bn \ \* (\ \ Q'')" by simp+ + ultimately obtain p where S: "(set p) \ (set(bn (\M\\*xvec\\N\))) \ (set(bn(p \ (\M\\*xvec\\N\))))" and "distinctPerm p" + and \Eq: "\ = p \ (\M\\*xvec\\N\)" and P'eq: "Q'' = p \ (P' \ Q')" and "(bn(p \ (\M\\*xvec\\N\))) \* (\M\\*xvec\\N\)" + and "(bn(p \ (\M\\*xvec\\N\))) \* (P' \ Q')" + by(rule residualEq) + + from \xvec \* \\ have "bn (\M\\*xvec\\N\) \* \" by simp + moreover from \xvec \* P\ \xvec \* Q\ have "bn (\M\\*xvec\\N\) \* (P \ Q)" by simp + moreover note \bn (\M\\*xvec\\N\) \* subject (\M\\*xvec\\N\)\ + moreover from \xvec \* \\<^sub>P\ \xvec \* \\<^sub>Q\ have "bn (\M\\*xvec\\N\) \* (\\<^sub>P \ \\<^sub>Q)" by auto + moreover from \xvec \* C\ have "bn (\M\\*xvec\\N\) \* C" by simp + moreover from \xvec \* \\ \Eq have "bn (\M\\*xvec\\N\) \* (p \ (\M\\*xvec\\N\))" by simp + moreover from \A\<^sub>P \* \\ \A\<^sub>Q \* \\ have "(A\<^sub>P @ A\<^sub>Q) \* \" by simp + moreover from \A\<^sub>P \* P\ \A\<^sub>Q \* Q\ \A\<^sub>P \* Q\ \A\<^sub>Q \* P\ have "(A\<^sub>P @ A\<^sub>Q) \* (P \ Q)" by simp + moreover from \A\<^sub>P \* \\ \A\<^sub>Q \* \\ have "(A\<^sub>P @ A\<^sub>Q) \* \" by simp + moreover from \A\<^sub>P \* P'\ \A\<^sub>Q \* Q'\ \A\<^sub>P \* Q'\ \A\<^sub>Q \* P'\ have "(A\<^sub>P @ A\<^sub>Q) \* (P' \ Q')" by simp + moreover from \A\<^sub>P \* C\ \A\<^sub>Q \* C\ have "(A\<^sub>P @ A\<^sub>Q) \* C" by simp + moreover note S \distinctPerm p\ \(bn(p \ (\M\\*xvec\\N\))) \* (\M\\*xvec\\N\)\ + \(bn(p \ (\M\\*xvec\\N\))) \* (P' \ Q')\ bigProp + \A\<^sub>P \* M\ \A\<^sub>Q \* M\ \A\<^sub>P \* xvec\ \A\<^sub>Q \* xvec\ \A\<^sub>P \* N\ \A\<^sub>Q \* N\ + + ultimately have "Prop C \ (P \ Q) (p \ (\M\\*xvec\\N\)) (p \ (P' \ Q')) (A\<^sub>P @ A\<^sub>Q) (\\<^sub>P \ \\<^sub>Q)" + by(fastforce intro!: rAlpha) + then show ?case using \Eq P'eq by simp +next + case(cBrClose \ P M xvec N P' A\<^sub>P \\<^sub>P x C \ P'') + note \\ \ \\x\(\\*xvec\P') = \ \ P''\ + moreover have "bn (\) \* (bn \)" by simp + moreover have "distinct (bn (\))" by simp + moreover note \distinct (bn \)\ + moreover have "(bn (\) \* (\ \ \\x\(\\*xvec\P')))" by simp + moreover from \bn \ \* subject \\ have "bn \ \* (\ \ P'')" by simp + ultimately obtain p where S: "(set p) \ (set(bn (\))) \ (set(bn(p \ (\))))" + and \Eq: "\ = p \ (\)" and P'eq: "P'' = p \ (\\x\(\\*xvec\P'))" + and "bn (\) \* \" and "bn (\) \* P''" + and "(bn(p \ (\))) \* (\)" and "(bn(p \ (\))) \* (\\x\(\\*xvec\P'))" + by(rule residualEq) simp + moreover from cBrClose have "\C. Prop C \ P (\M\\*xvec\\N\) P' A\<^sub>P \\<^sub>P" by simp + moreover with cBrClose have "Prop C \ (\\x\P) (\) (\\x\(\\*xvec\P')) (x#A\<^sub>P) \\<^sub>P" + by(simp add: rBrClose) + with S have "Prop C \ (\\x\P) (p \ \) (p \ \\x\(\\*xvec\P')) (x#A\<^sub>P) \\<^sub>P" by simp + then show ?case using \Eq P'eq + by simp +next + case(cOpen \ P M xvec yvec N P' x A\<^sub>P \\<^sub>P C \ P'') + note \M\\*(xvec@x#yvec)\\N\ \ P' = \ \ P''\ + moreover from \xvec \* \\ \x \ \\ \yvec \* \\ have "(xvec@x#yvec) \* (bn \)" + by auto + moreover from \xvec \* yvec\ \x \ xvec\ \x \ yvec\ \distinct xvec\ \distinct yvec\ + have "distinct(xvec@x#yvec)" + by(auto simp add: fresh_star_def) (simp add: fresh_def name_list_supp) + moreover note \distinct(bn \)\ + moreover from \xvec \* M\ \x \ M\ \yvec \* M\ have "(xvec@x#yvec) \* M" by auto + then have "(xvec@x#yvec) \* (M\\*(xvec@x#yvec)\\N\ \ P')" by auto + moreover from \bn \ \* subject \\ have "bn \ \* (\ \ P'')" by simp + ultimately obtain p where S: "(set p) \ (set(xvec@x#yvec)) \ (set(p \ (xvec@x#yvec)))" and "distinctPerm p" + and \eq: "\ = (p \ M)\\*(p \ (xvec@x#yvec))\\(p \ N)\" and P'eq: "P'' = (p \ P')" + and A: "(xvec@x#yvec) \* ((p \ M)\\*(p \ (xvec@x#yvec))\\(p \ N)\)" + and B: "(p \ (xvec@x#yvec)) \* (M\\*(xvec@x#yvec)\\N\)" + and C: "(p \ (xvec@x#yvec)) \* P'" + apply - + apply(rule residualEq) + by(assumption | simp)+ + + note \\ \ P \M\\*(xvec@yvec)\\N\ \ P'\ \x \ (supp N)\ + + moreover { + fix C + from \xvec \* M\ \yvec \* M\ have "(xvec@yvec) \* M" by simp + moreover from \distinct xvec\ \distinct yvec\ \xvec \* yvec\ have "distinct(xvec@yvec)" + by (auto simp add: fresh_star_def name_list_supp fresh_def) + ultimately have "Prop C \ P (M\\*(xvec@yvec)\\N\) P' A\<^sub>P \\<^sub>P" using \A\<^sub>P \* xvec\ \A\<^sub>P \* yvec\ \A\<^sub>P \* M\ \A\<^sub>P \* N\ + by(fastforce intro!: cOpen) + } + moreover obtain y::name where "y \ \" and "y \ x" and "y \ P" and "y \ xvec" and "y \ yvec" and "y \ \" and "y \ P'" and "y \ A\<^sub>P" and "y \ \\<^sub>P" and "y \ M" and "y \ N" and "y \ C" and "y \ p" + by(generate_fresh "name") auto + moreover note \x \ \\ \x \ M\ \x \ xvec\ \x \ yvec\ \xvec \* \\ \xvec \* P\ \xvec \* M\ + \yvec \* \\ \yvec \* P\ \yvec \* M\ \yvec \* C\ \x \ C\ \xvec \* C\ \distinct xvec\ \distinct yvec\ + \extractFrame P = \A\<^sub>P, \\<^sub>P\\ \distinct A\<^sub>P\ \x \ A\<^sub>P\ \xvec \* yvec\ \xvec \* \\<^sub>P\ + \A\<^sub>P \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* M\ \A\<^sub>P \* xvec\ \A\<^sub>P \* yvec\ \A\<^sub>P \* N\ \A\<^sub>P \* P'\ \A\<^sub>P \* C\ + ultimately have "Prop C \ (\\x\P) (M\\*(xvec@y#yvec)\\([(x, y)] \ N)\) ([(x, y)] \ P') (x#A\<^sub>P) \\<^sub>P" + by(metis rOpen) + moreover have "(([(x, y)] \ p) \ [(x, y)] \ M) = [(x, y)] \ p \ M" + by(subst perm_compose[symmetric]) simp + with \y \ M\ \x \ \\ \eq \y \ p\ \x \ M\ have D: "(([(x, y)] \ p) \ M) = p \ M" + by(auto simp add: eqvts freshChainSimps) + moreover have "(([(x, y)] \ p) \ [(x, y)] \ xvec) = [(x, y)] \ p \ xvec" + by(subst perm_compose[symmetric]) simp + with \y \ xvec\ \x \ \\ \eq \y \ p\ \x \ xvec\ have E: "(([(x, y)] \ p) \ xvec) = p \ xvec" + by(auto simp add: eqvts freshChainSimps) + moreover have "(([(x, y)] \ p) \ [(x, y)] \ yvec) = [(x, y)] \ p \ yvec" + by(subst perm_compose[symmetric]) simp + with \y \ yvec\ \x \ \\ \eq \y \ p\ \x \ yvec\ have F: "(([(x, y)] \ p) \ yvec) = p \ yvec" + by(auto simp add: eqvts freshChainSimps) + moreover have "(([(x, y)] \ p) \ [(x, y)] \ x) = [(x, y)] \ p \ x" + by(subst perm_compose[symmetric]) simp + with \y \ x\ \y \ p\ have G: "(([(x, y)] \ p) \ y) = p \ x" + apply(simp add: freshChainSimps calc_atm) + apply(subgoal_tac "y \ p \ x") + apply(clarsimp) + using A \eq + apply(simp add: eqvts) + apply(subst fresh_atm[symmetric]) + apply(simp only: freshChainSimps) + by simp + moreover have "(([(x, y)] \ p) \ [(x, y)] \ N) = [(x, y)] \ p \ N" + by(subst perm_compose[symmetric]) simp + with \y \ N\ \x \ \\ \y \ p\ \eq have H: "(([(x, y)] \ p) \ [(x, y)] \ N) = p \ N" + by(auto simp add: eqvts freshChainSimps) + moreover have "(([(x, y)] \ p) \ [(x, y)] \ P') = [(x, y)] \ p \ P'" + by(subst perm_compose[symmetric]) simp + with \y \ P'\ \x \ P''\ \y \ p\ P'eq have I: "(([(x, y)] \ p) \ [(x, y)] \ P') = p \ P'" + by(auto simp add: eqvts freshChainSimps) + from \y \ p\ \y \ x\ have "y \ p \ x" + apply(subst fresh_atm[symmetric]) + apply(simp only: freshChainSimps) + by simp + moreover from S have "([(x, y)] \ set p) \ [(x, y)] \ (set(xvec@x#yvec) \ set(p \ (xvec@x#yvec)))" + by(simp) + with \y \ p \ x\ \(([(x, y)] \ p) \ y) = p \ x\ \x \ xvec\ \y \ xvec\ \x \ yvec\ \y \ yvec\ \y \ p\ \x \ \\ \eq have + "set([(x, y)] \ p) \ set(xvec@y#yvec) \ set(([(x, y)] \ p) \ (xvec@y#yvec))" + by(simp add: eqvts calc_atm perm_compose) + moreover note \xvec \* \\ \yvec \* \\ \xvec \* P\ \yvec \* P\ \xvec \* M\ \yvec \* M\ + \yvec \* C\ S \distinctPerm p\ \x \ C\ \xvec \* C\ \xvec \* \\<^sub>P\ \yvec \* \\<^sub>P\ \x \ \\ + \A\<^sub>P \* xvec\ \x \ A\<^sub>P\ \A\<^sub>P \* yvec\ \A\<^sub>P \* M\ \x \ xvec\ \x \ yvec\ \x \ M\ \x \ A\<^sub>P\ \A\<^sub>P \* N\ + A B C \eq \A\<^sub>P \* \\ \y \ \\ \y \ x\ \y \ P\ \y \ M\ \y \ \\<^sub>P\ \y \ C\ \xvec \* \\ \x \ \\ \yvec \* \\ \y \ \\ \A\<^sub>P \* P\ \A\<^sub>P \* \\ \y \ A\<^sub>P\ \y \ N\ \A\<^sub>P \* P'\ \y \ P'\ \A\<^sub>P \* C\ P'eq + ultimately have "Prop C \ (\\x\P) (([(x, y)] \ p) \ (M\\*(xvec@y#yvec)\\([(x, y)] \ N)\)) (([(x, y)] \ p) \ [(x, y)] \ P') (x#A\<^sub>P) \\<^sub>P" + apply - + apply(rule rAlpha[where \="M\\*(xvec@y#yvec)\\([(x, y)] \ N)\"]) + apply(assumption | simp)+ + apply(simp add: eqvts) + apply(assumption | simp add: abs_fresh)+ + apply(simp add: fresh_left calc_atm) + apply(assumption | simp)+ + apply(simp add: fresh_left calc_atm) + apply(assumption | simp)+ + by(simp add: eqvts fresh_left)+ + with \eq P'eq D E F G H I show ?case + by(simp add: eqvts) +next + case(cBrOpen \ P M xvec yvec N P' x A\<^sub>P \\<^sub>P C \ P'') + note \\M\\*(xvec@x#yvec)\\N\ \ P' = \ \ P''\ + moreover from \xvec \* \\ \x \ \\ \yvec \* \\ have "(xvec@x#yvec) \* (bn \)" + by auto + moreover from \xvec \* yvec\ \x \ xvec\ \x \ yvec\ \distinct xvec\ \distinct yvec\ + have "distinct(xvec@x#yvec)" + by(auto simp add: fresh_star_def) (simp add: fresh_def name_list_supp) + moreover note \distinct(bn \)\ + moreover from \xvec \* M\ \x \ M\ \yvec \* M\ have "(xvec@x#yvec) \* M" by auto + then have "(xvec@x#yvec) \* (\M\\*(xvec@x#yvec)\\N\ \ P')" by auto + moreover from \bn \ \* subject \\ have "bn \ \* (\ \ P'')" by simp + ultimately obtain p where S: "(set p) \ (set(xvec@x#yvec)) \ (set(p \ (xvec@x#yvec)))" and "distinctPerm p" + and \eq: "\ = \(p \ M)\\*(p \ (xvec@x#yvec))\\(p \ N)\" and P'eq: "P'' = (p \ P')" + and A: "(xvec@x#yvec) \* (\(p \ M)\\*(p \ (xvec@x#yvec))\\(p \ N)\)" + and B: "(p \ (xvec@x#yvec)) \* (\M\\*(xvec@x#yvec)\\N\)" + and C: "(p \ (xvec@x#yvec)) \* P'" + apply - + apply(rule residualEq) + by(assumption | simp)+ + + note \\ \ P \\M\\*(xvec@yvec)\\N\ \ P'\ \x \ (supp N)\ + + moreover { + fix C + from \xvec \* M\ \yvec \* M\ have "(xvec@yvec) \* M" by simp + moreover from \distinct xvec\ \distinct yvec\ \xvec \* yvec\ have "distinct(xvec@yvec)" + by auto (simp add: fresh_star_def name_list_supp fresh_def) + ultimately have "Prop C \ P (\M\\*(xvec@yvec)\\N\) P' A\<^sub>P \\<^sub>P" using \A\<^sub>P \* xvec\ \A\<^sub>P \* yvec\ \A\<^sub>P \* M\ \A\<^sub>P \* N\ + by(fastforce intro!: cBrOpen) + } + moreover obtain y::name where "y \ \" and "y \ x" and "y \ P" and "y \ xvec" and "y \ yvec" and "y \ \" and "y \ P'" and "y \ A\<^sub>P" and "y \ \\<^sub>P" and "y \ M" and "y \ N" and "y \ C" and "y \ p" + by(generate_fresh "name") auto + moreover note \x \ \\ \x \ M\ \x \ xvec\ \x \ yvec\ \xvec \* \\ \xvec \* P\ \xvec \* M\ + \yvec \* \\ \yvec \* P\ \yvec \* M\ \yvec \* C\ \x \ C\ \xvec \* C\ \distinct xvec\ \distinct yvec\ + \extractFrame P = \A\<^sub>P, \\<^sub>P\\ \distinct A\<^sub>P\ \x \ A\<^sub>P\ \xvec \* yvec\ \xvec \* \\<^sub>P\ + \A\<^sub>P \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* M\ \A\<^sub>P \* xvec\ \A\<^sub>P \* yvec\ \A\<^sub>P \* N\ \A\<^sub>P \* P'\ \A\<^sub>P \* C\ + ultimately have "Prop C \ (\\x\P) (\M\\*(xvec@y#yvec)\\([(x, y)] \ N)\) ([(x, y)] \ P') (x#A\<^sub>P) \\<^sub>P" + by(metis rBrOpen) + moreover have "(([(x, y)] \ p) \ [(x, y)] \ M) = [(x, y)] \ p \ M" + by(subst perm_compose[symmetric]) simp + with \y \ M\ \x \ \\ \eq \y \ p\ \x \ M\ have D: "(([(x, y)] \ p) \ M) = p \ M" + by(auto simp add: eqvts freshChainSimps) + moreover have "(([(x, y)] \ p) \ [(x, y)] \ xvec) = [(x, y)] \ p \ xvec" + by(subst perm_compose[symmetric]) simp + with \y \ xvec\ \x \ \\ \eq \y \ p\ \x \ xvec\ have E: "(([(x, y)] \ p) \ xvec) = p \ xvec" + by(auto simp add: eqvts freshChainSimps) + moreover have "(([(x, y)] \ p) \ [(x, y)] \ yvec) = [(x, y)] \ p \ yvec" + by(subst perm_compose[symmetric]) simp + with \y \ yvec\ \x \ \\ \eq \y \ p\ \x \ yvec\ have F: "(([(x, y)] \ p) \ yvec) = p \ yvec" + by(auto simp add: eqvts freshChainSimps) + moreover have "(([(x, y)] \ p) \ [(x, y)] \ x) = [(x, y)] \ p \ x" + by(subst perm_compose[symmetric]) simp + with \y \ x\ \y \ p\ have G: "(([(x, y)] \ p) \ y) = p \ x" + apply(simp add: freshChainSimps calc_atm) + apply(subgoal_tac "y \ p \ x") + apply(clarsimp) + using A \eq + apply(simp add: eqvts) + apply(subst fresh_atm[symmetric]) + apply(simp only: freshChainSimps) + by simp + moreover have "(([(x, y)] \ p) \ [(x, y)] \ N) = [(x, y)] \ p \ N" + by(subst perm_compose[symmetric]) simp + with \y \ N\ \x \ \\ \y \ p\ \eq have H: "(([(x, y)] \ p) \ [(x, y)] \ N) = p \ N" + by(auto simp add: eqvts freshChainSimps) + moreover have "(([(x, y)] \ p) \ [(x, y)] \ P') = [(x, y)] \ p \ P'" + by(subst perm_compose[symmetric]) simp + with \y \ P'\ \x \ P''\ \y \ p\ P'eq have I: "(([(x, y)] \ p) \ [(x, y)] \ P') = p \ P'" + by(auto simp add: eqvts freshChainSimps) + from \y \ p\ \y \ x\ have "y \ p \ x" + apply(subst fresh_atm[symmetric]) + apply(simp only: freshChainSimps) + by simp + moreover from S have "([(x, y)] \ set p) \ [(x, y)] \ (set(xvec@x#yvec) \ set(p \ (xvec@x#yvec)))" + by(simp) + with \y \ p \ x\ \(([(x, y)] \ p) \ y) = p \ x\ \x \ xvec\ \y \ xvec\ \x \ yvec\ \y \ yvec\ \y \ p\ \x \ \\ \eq have + "set([(x, y)] \ p) \ set(xvec@y#yvec) \ set(([(x, y)] \ p) \ (xvec@y#yvec))" + by(simp add: eqvts calc_atm perm_compose) + moreover note \xvec \* \\ \yvec \* \\ \xvec \* P\ \yvec \* P\ \xvec \* M\ \yvec \* M\ + \yvec \* C\ S \distinctPerm p\ \x \ C\ \xvec \* C\ \xvec \* \\<^sub>P\ \yvec \* \\<^sub>P\ \x \ \\ + \A\<^sub>P \* xvec\ \x \ A\<^sub>P\ \A\<^sub>P \* yvec\ \A\<^sub>P \* M\ \x \ xvec\ \x \ yvec\ \x \ M\ \x \ A\<^sub>P\ \A\<^sub>P \* N\ + A B C \eq \A\<^sub>P \* \\ \y \ \\ \y \ x\ \y \ P\ \y \ M\ \y \ \\<^sub>P\ \y \ C\ \xvec \* \\ \x \ \\ \yvec \* \\ \y \ \\ \A\<^sub>P \* P\ \A\<^sub>P \* \\ \y \ A\<^sub>P\ \y \ N\ \A\<^sub>P \* P'\ \y \ P'\ \A\<^sub>P \* C\ P'eq + ultimately have "Prop C \ (\\x\P) (([(x, y)] \ p) \ (\M\\*(xvec@y#yvec)\\([(x, y)] \ N)\)) (([(x, y)] \ p) \ [(x, y)] \ P') (x#A\<^sub>P) \\<^sub>P" + apply - + apply(rule rAlpha[where \="\M\\*(xvec@y#yvec)\\([(x, y)] \ N)\"]) + apply(assumption | simp)+ (* slow proof step *) + apply(simp add: eqvts) + apply(assumption | simp add: abs_fresh)+ + apply(simp add: fresh_left calc_atm) + apply(assumption | simp)+ + apply(simp add: fresh_left calc_atm) + apply(assumption | simp)+ + by(simp add: eqvts fresh_left)+ + with \eq P'eq D E F G H I show ?case + by(simp add: eqvts) +next + case(cScope \ P \ P' x A\<^sub>P \\<^sub>P C \' P'') + note \\ \ (\\x\P') = \' \ P''\ + moreover from \bn \ \* \'\ have "bn \ \* (bn \')" by auto + moreover note \distinct (bn \)\ \distinct(bn \')\ + moreover from \bn \ \* subject \\ \bn \' \* subject \'\ + have "bn \ \* (\ \ \\x\P')" and "bn \' \* (\' \ P'')" by simp+ + ultimately obtain p where S: "(set p) \ (set(bn \)) \ (set(bn(p \ \)))" and "distinctPerm p" + and \Eq: "\' = p \ \" and P'eq: "P'' = p \ (\\x\P')" and "(bn(p \ \)) \* \" + and "(bn(p \ \)) \* (\\x\P')" + by(rule residualEq) + + note \\ \ P \\ \ P'\ + moreover from \bn \ \* subject \\ \distinct(bn \)\ + have "\C. Prop C \ P \ P' A\<^sub>P \\<^sub>P" by(fastforce intro!: cScope) + + moreover note \x \ \\ \x \ \\ \bn \ \* \\ \bn \ \* P\ \bn \ \* subject \\ \bn \ \* \\<^sub>P\ + \x \ C\ \bn \ \* C\ \distinct(bn \)\ \extractFrame P = \A\<^sub>P, \\<^sub>P\\ + \distinct A\<^sub>P\ \x \ A\<^sub>P\ \A\<^sub>P \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* \\ \A\<^sub>P \* P'\ \A\<^sub>P \* C\ + ultimately have "Prop C \ (\\x\P) \ (\\x\P') (x#A\<^sub>P) \\<^sub>P" + by(metis rScope) + with \bn \ \* \\ \bn \ \* P\ \x \ \\ \bn \ \* subject \\ \bn \ \* C\ \bn \ \* (bn \')\ S \distinctPerm p\ \bn(p \ \) \* \\ \bn(p \ \) \* (\\x\P')\ \A\<^sub>P \* \\ \A\<^sub>P \* \'\ \Eq \x \ \'\ \bn \ \* \\<^sub>P\ \bn \ \* \'\ \x \ \\ \A\<^sub>P \* \\ \x \ A\<^sub>P\ \A\<^sub>P \* P\ \A\<^sub>P \* P'\ \x \ C\ \A\<^sub>P \* C\ + have "Prop C \ (\\x\P) (p \ \) (p \ (\\x\P')) (x#A\<^sub>P) \\<^sub>P" + by(fastforce intro!: rAlpha simp add: abs_fresh) + with \Eq P'eq \distinctPerm p\ show ?case by simp +next + case(cBang \ P A\<^sub>P \\<^sub>P C \ P') + then show ?case by(fastforce intro!: rBang) +qed + +lemma inputFrameInduct[consumes 3, case_names cAlpha cInput cCase cPar1 cPar2 cScope cBang]: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and M :: 'a + and N :: 'a + and P' :: "('a, 'b, 'c) psi" + and Prop :: "'f::fs_name \ 'b \ ('a, 'b, 'c) psi \ + 'a \ 'a \ ('a, 'b, 'c) psi \ name list \ 'b \ bool" + and C :: "'f::fs_name" + +assumes Trans: "\ \ P \M\N\ \ P'" + and FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" + and "distinct A\<^sub>P" + and rAlpha: "\\ P M N P' A\<^sub>P \\<^sub>P p C. \A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* P'; A\<^sub>P \* (p \ A\<^sub>P); A\<^sub>P \* C; + set p \ set A\<^sub>P \ set(p \ A\<^sub>P); distinctPerm p; + Prop C \ P M N P' A\<^sub>P \\<^sub>P\ \ Prop C \ P M N P' (p \ A\<^sub>P) (p \ \\<^sub>P)" + and rInput: "\\ M K xvec N Tvec P C. + \\ \ M \ K; distinct xvec; set xvec \ supp N; + length xvec = length Tvec; xvec \* \; + xvec \* M; xvec \* K; xvec \* C\ \ + Prop C \ (M\\*xvec N\.P) + K (N[xvec::=Tvec]) (P[xvec::=Tvec]) ([]) (\)" + and rCase: "\\ P M N P' \ Cs A\<^sub>P \\<^sub>P C. \\ \ P \M\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; \C. Prop C \ P M N P' A\<^sub>P \\<^sub>P; + (\, P) \ set Cs; \ \ \; guarded P; \\<^sub>P \ \; (supp \\<^sub>P) = ({}::name set); + A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* P'; A\<^sub>P \* C\ \ Prop C \ (Cases Cs) M N P' ([]) (\)" + and rPar1: "\\ \\<^sub>Q P M N P' A\<^sub>Q Q A\<^sub>P \\<^sub>P C. + \\ \ \\<^sub>Q \ P \M\N\ \ P'; + extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + \C. Prop C (\ \ \\<^sub>Q) P M N P' A\<^sub>P \\<^sub>P; + A\<^sub>P \* P; A\<^sub>P \* Q; A\<^sub>P \* \; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* P'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* \\<^sub>Q; + A\<^sub>Q \* P; A\<^sub>Q \* Q; A\<^sub>Q \* \; A\<^sub>Q \* M; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* \\<^sub>P; + A\<^sub>P \* C; A\<^sub>Q \* C\ \ + Prop C \ (P \ Q) M N (P' \ Q) (A\<^sub>P@A\<^sub>Q) (\\<^sub>P \ \\<^sub>Q)" + and rPar2: "\\ \\<^sub>P Q M N Q' A\<^sub>P P A\<^sub>Q \\<^sub>Q C. + \\ \ \\<^sub>P \ Q \M\N\ \ Q'; + extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + \C. Prop C (\ \ \\<^sub>P) Q M N Q' A\<^sub>Q \\<^sub>Q; + A\<^sub>P \* P; A\<^sub>P \* Q; A\<^sub>P \* \; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* \\<^sub>Q; + A\<^sub>Q \* P; A\<^sub>Q \* Q; A\<^sub>Q \* \; A\<^sub>Q \* M; A\<^sub>Q \* N; A\<^sub>Q \* Q'; A\<^sub>Q \* \\<^sub>P; + A\<^sub>P \* C; A\<^sub>Q \* C\ \ + Prop C \ (P \ Q) M N (P \ Q') (A\<^sub>P@A\<^sub>Q) (\\<^sub>P \ \\<^sub>Q)" + and rScope: "\\ P M N P' x A\<^sub>P \\<^sub>P C. + \\ \ P \M\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \C. Prop C \ P M N P' A\<^sub>P \\<^sub>P; x \ \; x \ M; x \ N; + x \ A\<^sub>P; A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* P'; + A\<^sub>P \* C; x \ C\ \ + Prop C \ (\\x\P) M N (\\x\P') (x#A\<^sub>P) \\<^sub>P" + and rBang: "\\ P M N P' A\<^sub>P \\<^sub>P C. + \\ \ P \ !P \M\N\ \ P'; guarded P; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \C. Prop C \ (P \ !P) M N P' A\<^sub>P (\\<^sub>P \ \); \\<^sub>P \ \; (supp \\<^sub>P) = ({}::name set); + A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* P'; A\<^sub>P \* C\ \ Prop C \ (!P) M N P' ([]) (\)" +shows "Prop C \ P M N P' A\<^sub>P \\<^sub>P" + using Trans FrP \distinct A\<^sub>P\ +proof(nominal_induct \ P Rs=="M\N\ \ P'" A\<^sub>P \\<^sub>P avoiding: C arbitrary: P' rule: semanticsFrameInduct) + case cAlpha + then show ?case by (simp add: rAlpha) +next + case cInput + then show ?case by(auto simp add: rInput residualInject) +next + case cBrInput + then show ?case by(simp add: residualInject) +next + case cOutput + then show ?case by(simp add: residualInject) +next + case cBrOutput + then show ?case by(simp add: residualInject) +next + case cCase + then show ?case by(simp add: rCase residualInject) +next + case cPar1 + then show ?case by(auto simp add: rPar1 residualInject) +next + case cPar2 + then show ?case by(auto simp add: rPar2 residualInject) +next + case cComm1 + then show ?case by(simp add: residualInject) +next + case cComm2 + then show ?case by(simp add: residualInject) +next + case cBrMerge + then show ?case by(simp add: residualInject) +next + case cBrComm1 + then show ?case by(simp add: residualInject) +next + case cBrComm2 + then show ?case by(simp add: residualInject) +next + case cBrClose + then show ?case by(simp add: residualInject) +next + case cOpen + then show ?case by(simp add: residualInject) +next + case cBrOpen + then show ?case by(simp add: residualInject) +next + case cScope + then show ?case by(auto simp add: rScope residualInject) +next + case cBang + then show ?case by(simp add: rBang residualInject) +qed + +lemma brinputFrameInduct[consumes 3, case_names cAlpha cBrInput cCase cPar1 cPar2 cBrMerge cScope cBang]: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and M :: 'a + and N :: 'a + and P' :: "('a, 'b, 'c) psi" + and Prop :: "'f::fs_name \ 'b \ ('a, 'b, 'c) psi \ + 'a \ 'a \ ('a, 'b, 'c) psi \ name list \ 'b \ bool" + and C :: "'f::fs_name" + +assumes Trans: "\ \ P \\M\N\ \ P'" + and FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" + and "distinct A\<^sub>P" + and rAlpha: "\\ P M N P' A\<^sub>P \\<^sub>P p C. \A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* P'; A\<^sub>P \* (p \ A\<^sub>P); A\<^sub>P \* C; + set p \ set A\<^sub>P \ set(p \ A\<^sub>P); distinctPerm p; + Prop C \ P M N P' A\<^sub>P \\<^sub>P\ \ Prop C \ P M N P' (p \ A\<^sub>P) (p \ \\<^sub>P)" + and rBrInput: "\\ M K xvec N Tvec P C. + \\ \ K \ M; distinct xvec; set xvec \ supp N; + length xvec = length Tvec; xvec \* \; + xvec \* M; xvec \* K; xvec \* C\ \ + Prop C \ (M\\*xvec N\.P) + K (N[xvec::=Tvec]) (P[xvec::=Tvec]) ([]) (\)" + and rCase: "\\ P M N P' \ Cs A\<^sub>P \\<^sub>P C. \\ \ P \\M\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; \C. Prop C \ P M N P' A\<^sub>P \\<^sub>P; + (\, P) \ set Cs; \ \ \; guarded P; \\<^sub>P \ \; (supp \\<^sub>P) = ({}::name set); + A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* P'; A\<^sub>P \* C\ \ Prop C \ (Cases Cs) M N P' ([]) (\)" + and rPar1: "\\ \\<^sub>Q P M N P' A\<^sub>Q Q A\<^sub>P \\<^sub>P C. + \\ \ \\<^sub>Q \ P \\M\N\ \ P'; + extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + \C. Prop C (\ \ \\<^sub>Q) P M N P' A\<^sub>P \\<^sub>P; + A\<^sub>P \* P; A\<^sub>P \* Q; A\<^sub>P \* \; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* P'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* \\<^sub>Q; + A\<^sub>Q \* P; A\<^sub>Q \* Q; A\<^sub>Q \* \; A\<^sub>Q \* M; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* \\<^sub>P; + A\<^sub>P \* C; A\<^sub>Q \* C\ \ + Prop C \ (P \ Q) M N (P' \ Q) (A\<^sub>P@A\<^sub>Q) (\\<^sub>P \ \\<^sub>Q)" + and rPar2: "\\ \\<^sub>P Q M N Q' A\<^sub>P P A\<^sub>Q \\<^sub>Q C. + \\ \ \\<^sub>P \ Q \\M\N\ \ Q'; + extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + \C. Prop C (\ \ \\<^sub>P) Q M N Q' A\<^sub>Q \\<^sub>Q; + A\<^sub>P \* P; A\<^sub>P \* Q; A\<^sub>P \* \; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* \\<^sub>Q; + A\<^sub>Q \* P; A\<^sub>Q \* Q; A\<^sub>Q \* \; A\<^sub>Q \* M; A\<^sub>Q \* N; A\<^sub>Q \* Q'; A\<^sub>Q \* \\<^sub>P; + A\<^sub>P \* C; A\<^sub>Q \* C\ \ + Prop C \ (P \ Q) M N (P \ Q') (A\<^sub>P@A\<^sub>Q) (\\<^sub>P \ \\<^sub>Q)" + and rBrMerge: "\\ \\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q Q' A\<^sub>Q C. + \\ \ \\<^sub>Q \ P \ \M\N\ \ P'; \C. Prop C (\ \ \\<^sub>Q) P M N P' A\<^sub>P \\<^sub>P; + extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \ \ \\<^sub>P \ Q \ \M\N\ \ Q'; \C. Prop C (\ \ \\<^sub>P) Q M N Q' A\<^sub>Q \\<^sub>Q; + extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* N; A\<^sub>P \* P'; + A\<^sub>P \* Q; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* M; A\<^sub>Q \* M; + A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; A\<^sub>Q \* P; A\<^sub>Q \* N; A\<^sub>Q \* P'; + A\<^sub>Q \* Q; A\<^sub>Q \* Q'; A\<^sub>P \* C; A\<^sub>Q \* C\ \ + Prop C \ (P \ Q) M N (P' \ Q') (A\<^sub>P@A\<^sub>Q) (\\<^sub>P \ \\<^sub>Q)" + and rScope: "\\ P M N P' x A\<^sub>P \\<^sub>P C. + \\ \ P \\M\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \C. Prop C \ P M N P' A\<^sub>P \\<^sub>P; x \ \; x \ M; x \ N; + x \ A\<^sub>P; A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* P'; + A\<^sub>P \* C; x \ C\ \ + Prop C \ (\\x\P) M N (\\x\P') (x#A\<^sub>P) \\<^sub>P" + and rBang: "\\ P M N P' A\<^sub>P \\<^sub>P C. + \\ \ P \ !P \\M\N\ \ P'; guarded P; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \C. Prop C \ (P \ !P) M N P' A\<^sub>P (\\<^sub>P \ \); \\<^sub>P \ \; (supp \\<^sub>P) = ({}::name set); + A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* P'; A\<^sub>P \* C\ \ Prop C \ (!P) M N P' ([]) (\)" +shows "Prop C \ P M N P' A\<^sub>P \\<^sub>P" + using Trans FrP \distinct A\<^sub>P\ +proof(nominal_induct \ P Rs=="\M\N\ \ P'" A\<^sub>P \\<^sub>P avoiding: C arbitrary: P' rule: semanticsFrameInduct) + case cAlpha + then show ?case by (simp add: rAlpha) +next + case cInput + then show ?case by(simp add: residualInject) +next + case cBrInput + then show ?case by(auto simp add: rBrInput residualInject) +next + case cOutput + then show ?case by(simp add: residualInject) +next + case cBrOutput + then show ?case by(simp add: residualInject) +next + case cCase + then show ?case by(simp add: rCase residualInject) +next + case cPar1 + then show ?case by(auto simp add: rPar1 residualInject) +next + case cPar2 + then show ?case by(auto simp add: rPar2 residualInject) +next + case cComm1 + then show ?case by(simp add: residualInject) +next + case cComm2 + then show ?case by(simp add: residualInject) +next + case cBrMerge + then show ?case by(auto simp add: rBrMerge residualInject) +next + case cBrComm1 + then show ?case by(simp add: residualInject) +next + case cBrComm2 + then show ?case by(simp add: residualInject) +next + case cBrClose + then show ?case by(simp add: residualInject) +next + case cOpen + then show ?case by(simp add: residualInject) +next + case cBrOpen + then show ?case by(simp add: residualInject) +next + case cScope + then show ?case by(auto simp add: rScope residualInject) +next + case cBang + then show ?case by(simp add: rBang residualInject) +qed + +lemma outputFrameInduct[consumes 3, case_names cAlpha cOutput cCase cPar1 cPar2 cOpen cScope cBang]: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and M :: 'a + and B :: "('a, 'b, 'c) boundOutput" + and A\<^sub>P :: "name list" + and \\<^sub>P :: 'b + and Prop :: "'f::fs_name \ 'b \ ('a, 'b, 'c) psi \ + 'a \ ('a, 'b, 'c) boundOutput \ name list \ 'b \ bool" + and C :: "'f::fs_name" + +assumes Trans: "\ \ P \ROut M B" + and FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" + and "distinct A\<^sub>P" + and rAlpha: "\\ P M A\<^sub>P \\<^sub>P p B C. \A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* (p \ A\<^sub>P); A\<^sub>P \* B; A\<^sub>P \* C; + set p \ set A\<^sub>P \ set(p \ A\<^sub>P); distinctPerm p; + Prop C \ P M B A\<^sub>P \\<^sub>P\ \ Prop C \ P M B (p \ A\<^sub>P) (p \ \\<^sub>P)" + and rOutput: "\\ M K N P C. \ \ M \ K \ Prop C \ (M\N\.P) K (N \' P) ([]) (\)" + and rCase: "\\ P M B \ Cs A\<^sub>P \\<^sub>P C. \\ \ P \(ROut M B); extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; \C. Prop C \ P M B A\<^sub>P \\<^sub>P; + (\, P) \ set Cs; \ \ \; guarded P; \\<^sub>P \ \; (supp \\<^sub>P) = ({}::name set); + A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* B; A\<^sub>P \* C\ \ Prop C \ (Cases Cs) M B ([]) (\)" + and rPar1: "\\ \\<^sub>Q P M xvec N P' A\<^sub>Q Q A\<^sub>P \\<^sub>P C. + \\ \ \\<^sub>Q \ P \M\\*xvec\\N\ \ P'; + extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + \C. Prop C (\ \ \\<^sub>Q) P M (\\*xvec\N \' P') A\<^sub>P \\<^sub>P; + A\<^sub>P \* P; A\<^sub>P \* Q; A\<^sub>P \* \; A\<^sub>P \* M; A\<^sub>P \* xvec; A\<^sub>P \* N; A\<^sub>P \* P'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* \\<^sub>Q; + A\<^sub>Q \* P; A\<^sub>Q \* Q; A\<^sub>Q \* \; A\<^sub>Q \* M; A\<^sub>Q \* xvec; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* \\<^sub>P; + xvec \* \; xvec \* P; xvec \* Q; xvec \* M; xvec \* \\<^sub>P; xvec \* \\<^sub>Q; + A\<^sub>P \* C; A\<^sub>Q \* C; xvec \* C\ \ + Prop C \ (P \ Q) M (\\*xvec\N \' (P' \ Q)) (A\<^sub>P@A\<^sub>Q) (\\<^sub>P \ \\<^sub>Q)" + and rPar2: "\\ \\<^sub>P Q M xvec N Q' A\<^sub>P P A\<^sub>Q \\<^sub>Q C. + \\ \ \\<^sub>P \ Q \M\\*xvec\\N\ \ Q'; + extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + \C. Prop C (\ \ \\<^sub>P) Q M (\\*xvec\N \' Q') A\<^sub>Q \\<^sub>Q; + A\<^sub>P \* P; A\<^sub>P \* Q; A\<^sub>P \* \; A\<^sub>P \* M; A\<^sub>P \* xvec; A\<^sub>P \* N; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* \\<^sub>Q; + A\<^sub>Q \* P; A\<^sub>Q \* Q; A\<^sub>Q \* \; A\<^sub>Q \* M; A\<^sub>Q \* xvec; A\<^sub>Q \* N; A\<^sub>Q \* Q'; A\<^sub>Q \* \\<^sub>P; + xvec \* \; xvec \* P; xvec \* Q; xvec \* M; xvec \* \\<^sub>P; xvec \* \\<^sub>Q; + A\<^sub>P \* C; A\<^sub>Q \* C; xvec \* C\ \ + Prop C \ (P \ Q) M (\\*xvec\N \' (P \ Q')) (A\<^sub>P@A\<^sub>Q) (\\<^sub>P \ \\<^sub>Q)" + and rOpen: "\\ P M xvec yvec N P' x A\<^sub>P \\<^sub>P C. + \\ \ P \M\\*(xvec@yvec)\\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \C. Prop C \ P M (\\*(xvec@yvec)\N \' P') A\<^sub>P \\<^sub>P; x \ supp N; x \ \; x \ M; + x \ A\<^sub>P; x \ xvec; x \ yvec; A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* P'; + A\<^sub>P \* xvec; A\<^sub>P \* yvec; + xvec \* \; xvec \* P; xvec \* M; xvec \* \\<^sub>P; + yvec \* \; yvec \* P; yvec \* M; A\<^sub>P \* C; x \ C; xvec \* C; yvec \* C\ \ + Prop C \ (\\x\P) M (\\*(xvec@x#yvec)\N \' P') (x#A\<^sub>P) \\<^sub>P" + and rScope: "\\ P M xvec N P' x A\<^sub>P \\<^sub>P C. + \\ \ P \M\\*xvec\\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \C. Prop C \ P M (\\*xvec\N \' P') A\<^sub>P \\<^sub>P; + x \ \; x \ M; x \ xvec; x \ N; x \ A\<^sub>P; A\<^sub>P \* \; A\<^sub>P \* P; + A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* P'; A\<^sub>P \* xvec; + xvec \* \; xvec \* P; xvec \* M; xvec \* \\<^sub>P; + A\<^sub>P \* C; x \ C; xvec \* C\ \ + Prop C \ (\\x\P) M (\\*xvec\N \' (\\x\P')) (x#A\<^sub>P) \\<^sub>P" + and rBang: "\\ P M B A\<^sub>P \\<^sub>P C. + \\ \ P \ !P \ROut M B; guarded P; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \C. Prop C \ (P \ !P) M B A\<^sub>P (\\<^sub>P \ \); \\<^sub>P \ \; supp \\<^sub>P = ({}::name set); + A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* C\ \ Prop C \ (!P) M B ([]) (\)" +shows "Prop C \ P M B A\<^sub>P \\<^sub>P" +proof - + { + fix B + assume "\ \ P \ROut M B" + then have "Prop C \ P M B A\<^sub>P \\<^sub>P" using FrP \distinct A\<^sub>P\ + proof(nominal_induct \ P Rs=="ROut M B" A\<^sub>P \\<^sub>P avoiding: C arbitrary: B rule: semanticsFrameInduct) + case cAlpha + then show ?case by(auto intro: rAlpha) + next + case cInput + then show ?case by(simp add: residualInject) + next + case cBrInput + then show ?case by(simp add: residualInject) + next + case cOutput + then show ?case by(force intro: rOutput simp add: residualInject) + next + case cBrOutput + then show ?case by(simp add: residualInject) + next + case cCase + then show ?case by(force intro: rCase simp add: residualInject) + next + case cPar1 + then show ?case + by(auto intro!: rPar1 simp add: residualInject) + next + case cPar2 + then show ?case + by(auto intro!: rPar2 simp add: residualInject) + next + case cComm1 + then show ?case by(simp add: residualInject) + next + case cComm2 + then show ?case by(simp add: residualInject) + next + case cBrMerge + then show ?case by(simp add: residualInject) + next + case cBrComm1 + then show ?case by(simp add: residualInject) + next + case cBrComm2 + then show ?case by(simp add: residualInject) + next + case cBrClose + then show ?case by(simp add: residualInject) + next + case cOpen + then show ?case by(auto intro: rOpen simp add: residualInject) + next + case cBrOpen + then show ?case by(simp add: residualInject) + next + case cScope + then show ?case by(force intro: rScope simp add: residualInject) + next + case cBang + then show ?case by(force intro: rBang simp add: residualInject) + qed + } + with Trans show ?thesis by(simp add: residualInject) +qed + +lemma broutputFrameInduct[consumes 3, case_names cAlpha cBrOutput cCase cPar1 cPar2 cBrComm1 cBrComm2 cBrOpen cScope cBang]: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and M :: 'a + and B :: "('a, 'b, 'c) boundOutput" + and A\<^sub>P :: "name list" + and \\<^sub>P :: 'b + and Prop :: "'f::fs_name \ 'b \ ('a, 'b, 'c) psi \ + 'a \ ('a, 'b, 'c) boundOutput \ name list \ 'b \ bool" + and C :: "'f::fs_name" + +assumes Trans: "\ \ P \RBrOut M B" + and FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" + and "distinct A\<^sub>P" + and rAlpha: "\\ P M A\<^sub>P \\<^sub>P p B C. \A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* (p \ A\<^sub>P); A\<^sub>P \* B; A\<^sub>P \* C; + set p \ set A\<^sub>P \ set(p \ A\<^sub>P); distinctPerm p; + Prop C \ P M B A\<^sub>P \\<^sub>P\ \ Prop C \ P M B (p \ A\<^sub>P) (p \ \\<^sub>P)" + and rBrOutput: "\\ M K N P C. \ \ M \ K \ Prop C \ (M\N\.P) K (N \' P) ([]) (\)" + and rCase: "\\ P M B \ Cs A\<^sub>P \\<^sub>P C. \\ \ P \(RBrOut M B); extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; \C. Prop C \ P M B A\<^sub>P \\<^sub>P; + (\, P) \ set Cs; \ \ \; guarded P; \\<^sub>P \ \; (supp \\<^sub>P) = ({}::name set); + A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* B; A\<^sub>P \* C\ \ Prop C \ (Cases Cs) M B ([]) (\)" + and rPar1: "\\ \\<^sub>Q P M xvec N P' A\<^sub>Q Q A\<^sub>P \\<^sub>P C. + \\ \ \\<^sub>Q \ P \\M\\*xvec\\N\ \ P'; + extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + \C. Prop C (\ \ \\<^sub>Q) P M (\\*xvec\N \' P') A\<^sub>P \\<^sub>P; + A\<^sub>P \* P; A\<^sub>P \* Q; A\<^sub>P \* \; A\<^sub>P \* M; A\<^sub>P \* xvec; A\<^sub>P \* N; A\<^sub>P \* P'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* \\<^sub>Q; + A\<^sub>Q \* P; A\<^sub>Q \* Q; A\<^sub>Q \* \; A\<^sub>Q \* M; A\<^sub>Q \* xvec; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* \\<^sub>P; + xvec \* \; xvec \* P; xvec \* Q; xvec \* M; xvec \* \\<^sub>P; xvec \* \\<^sub>Q; + A\<^sub>P \* C; A\<^sub>Q \* C; xvec \* C\ \ + Prop C \ (P \ Q) M (\\*xvec\N \' (P' \ Q)) (A\<^sub>P@A\<^sub>Q) (\\<^sub>P \ \\<^sub>Q)" + and rPar2: "\\ \\<^sub>P Q M xvec N Q' A\<^sub>P P A\<^sub>Q \\<^sub>Q C. + \\ \ \\<^sub>P \ Q \\M\\*xvec\\N\ \ Q'; + extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + \C. Prop C (\ \ \\<^sub>P) Q M (\\*xvec\N \' Q') A\<^sub>Q \\<^sub>Q; + A\<^sub>P \* P; A\<^sub>P \* Q; A\<^sub>P \* \; A\<^sub>P \* M; A\<^sub>P \* xvec; A\<^sub>P \* N; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* \\<^sub>Q; + A\<^sub>Q \* P; A\<^sub>Q \* Q; A\<^sub>Q \* \; A\<^sub>Q \* M; A\<^sub>Q \* xvec; A\<^sub>Q \* N; A\<^sub>Q \* Q'; A\<^sub>Q \* \\<^sub>P; + xvec \* \; xvec \* P; xvec \* Q; xvec \* M; xvec \* \\<^sub>P; xvec \* \\<^sub>Q; + A\<^sub>P \* C; A\<^sub>Q \* C; xvec \* C\ \ + Prop C \ (P \ Q) M (\\*xvec\N \' (P \ Q')) (A\<^sub>P@A\<^sub>Q) (\\<^sub>P \ \\<^sub>Q)" + and rBrComm1:"\\ \\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q xvec Q' A\<^sub>Q C. + \\ \ \\<^sub>Q \ P \\M\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \ \ \\<^sub>P \ Q \\M\\*xvec\\N\ \ Q'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + distinct xvec; + \C. Prop C (\ \ \\<^sub>P) Q M (\\*xvec\N \' Q') A\<^sub>Q \\<^sub>Q; + A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* N; A\<^sub>P \* P'; + A\<^sub>P \* Q; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* xvec; A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; + A\<^sub>Q \* P; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* Q; A\<^sub>Q \* Q'; + A\<^sub>Q \* xvec; xvec \* \; xvec \* \\<^sub>P; xvec \* \\<^sub>Q; xvec \* P; + xvec \* Q; A\<^sub>P \* C; A\<^sub>Q \* C; xvec \* C; + A\<^sub>P \* M; A\<^sub>Q \* M; xvec \* M\ \ + Prop C \ (P \ Q) M (\\*xvec\N \' (P' \ Q')) (A\<^sub>P@A\<^sub>Q) (\\<^sub>P \ \\<^sub>Q)" (* Not applicable: \C. Prop C (\ \ \\<^sub>Q) P (M\N\) P' A\<^sub>P \\<^sub>P; *) + and rBrComm2:"\\ \\<^sub>Q P M xvec N P' A\<^sub>P \\<^sub>P Q Q' A\<^sub>Q C. + \\ \ \\<^sub>Q \ P \\M\\*xvec\\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \C. Prop C (\ \ \\<^sub>Q) P M (\\*xvec\N \' P') A\<^sub>P \\<^sub>P; + \ \ \\<^sub>P \ Q \\M\N\ \ Q'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + distinct xvec; + A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* N; A\<^sub>P \* P'; + A\<^sub>P \* Q; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* xvec; A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; + A\<^sub>Q \* P; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* Q; A\<^sub>Q \* Q'; + A\<^sub>Q \* xvec; xvec \* \; xvec \* \\<^sub>P; xvec \* \\<^sub>Q; xvec \* P; + xvec \* Q; A\<^sub>P \* C; A\<^sub>Q \* C; xvec \* C; + A\<^sub>P \* M; A\<^sub>Q \* M; xvec \* M\ \ + Prop C \ (P \ Q) M (\\*xvec\N \' (P' \ Q')) (A\<^sub>P@A\<^sub>Q) (\\<^sub>P \ \\<^sub>Q)" (* Not applicable: \C. Prop C (\ \ \\<^sub>P) Q (K\N\) Q' A\<^sub>Q \\<^sub>Q; *) + and rBrOpen: "\\ P M xvec yvec N P' x A\<^sub>P \\<^sub>P C. + \\ \ P \\M\\*(xvec@yvec)\\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \C. Prop C \ P M (\\*(xvec@yvec)\N \' P') A\<^sub>P \\<^sub>P; x \ supp N; x \ \; x \ M; + x \ A\<^sub>P; x \ xvec; x \ yvec; A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* P'; + A\<^sub>P \* xvec; A\<^sub>P \* yvec; + xvec \* \; xvec \* P; xvec \* M; xvec \* \\<^sub>P; + yvec \* \; yvec \* P; yvec \* M; A\<^sub>P \* C; x \ C; xvec \* C; yvec \* C\ \ + Prop C \ (\\x\P) M (\\*(xvec@x#yvec)\N \' P') (x#A\<^sub>P) \\<^sub>P" + and rScope: "\\ P M xvec N P' x A\<^sub>P \\<^sub>P C. + \\ \ P \\M\\*xvec\\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \C. Prop C \ P M (\\*xvec\N \' P') A\<^sub>P \\<^sub>P; + x \ \; x \ M; x \ xvec; x \ N; x \ A\<^sub>P; A\<^sub>P \* \; A\<^sub>P \* P; + A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* P'; A\<^sub>P \* xvec; + xvec \* \; xvec \* P; xvec \* M; xvec \* \\<^sub>P; + A\<^sub>P \* C; x \ C; xvec \* C\ \ + Prop C \ (\\x\P) M (\\*xvec\N \' (\\x\P')) (x#A\<^sub>P) \\<^sub>P" + and rBang: "\\ P M B A\<^sub>P \\<^sub>P C. + \\ \ P \ !P \RBrOut M B; guarded P; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \C. Prop C \ (P \ !P) M B A\<^sub>P (\\<^sub>P \ \); \\<^sub>P \ \; supp \\<^sub>P = ({}::name set); + A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* C\ \ Prop C \ (!P) M B ([]) (\)" +shows "Prop C \ P M B A\<^sub>P \\<^sub>P" +proof - + { + fix B + assume "\ \ P \RBrOut M B" + then have "Prop C \ P M B A\<^sub>P \\<^sub>P" using FrP \distinct A\<^sub>P\ + proof(nominal_induct \ P Rs=="RBrOut M B" A\<^sub>P \\<^sub>P avoiding: C arbitrary: B rule: semanticsFrameInduct) + case cAlpha + then show ?case by(auto intro: rAlpha) + next + case cInput + then show ?case by(simp add: residualInject) + next + case cBrInput + then show ?case by(simp add: residualInject) + next + case cOutput + then show ?case by(simp add: residualInject) + next + case cBrOutput + then show ?case by(force intro: rBrOutput simp add: residualInject) + next + case cCase + then show ?case by(force intro: rCase simp add: residualInject) + next + case cPar1 + then show ?case by(auto intro!: rPar1 simp add: residualInject) + next + case cPar2 + then show ?case by(auto intro!: rPar2 simp add: residualInject) + next + case cComm1 + then show ?case by(simp add: residualInject) + next + case cComm2 + then show ?case by(simp add: residualInject) + next + case cBrMerge + then show ?case by(simp add: residualInject) + next + case cBrComm1 + then show ?case by(auto intro: rBrComm1 simp add: residualInject) + next + case cBrComm2 + then show ?case by(auto intro: rBrComm2 simp add: residualInject) + next + case cBrClose + then show ?case by(simp add: residualInject) + next + case cOpen + then show ?case by(simp add: residualInject) + next + case cBrOpen + then show ?case by(auto intro: rBrOpen simp add: residualInject) + next + case cScope + then show ?case by(force intro: rScope simp add: residualInject) + next + case cBang + then show ?case by(force intro: rBang simp add: residualInject) + qed + } + with Trans show ?thesis by(simp add: residualInject) +qed + +lemma tauFrameInduct[consumes 3, case_names cAlpha cCase cPar1 cPar2 cComm1 cComm2 cBrClose cScope cBang]: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and P' :: "('a, 'b, 'c) psi" + and Prop :: "'f::fs_name \ 'b \ ('a, 'b, 'c) psi \ + ('a, 'b, 'c) psi \ name list \ 'b \ bool" + and C :: "'f::fs_name" + +assumes Trans: "\ \ P \\ \ P'" + and FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" + and "distinct A\<^sub>P" + and rAlpha: "\\ P P' A\<^sub>P \\<^sub>P p C. \A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* P'; A\<^sub>P \* (p \ A\<^sub>P); A\<^sub>P \* C; + set p \ set A\<^sub>P \ set (p \ A\<^sub>P); distinctPerm p; + Prop C \ P P' A\<^sub>P \\<^sub>P\ \ Prop C \ P P' (p \ A\<^sub>P) (p \ \\<^sub>P)" + and rCase: "\\ P P' \ Cs A\<^sub>P \\<^sub>P C. \\ \ P \\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; \C. Prop C \ P P' A\<^sub>P \\<^sub>P; + (\, P) \ set Cs; \ \ \; guarded P; \\<^sub>P \ \; (supp \\<^sub>P) = ({}::name set); + A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* P'; A\<^sub>P \* C\ \ Prop C \ (Cases Cs) P' ([]) (\)" + and rPar1: "\\ \\<^sub>Q P P' A\<^sub>Q Q A\<^sub>P \\<^sub>P C. + \\ \ \\<^sub>Q \ P \\ \ P'; + extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + \C. Prop C (\ \ \\<^sub>Q) P P' A\<^sub>P \\<^sub>P; + A\<^sub>P \* P; A\<^sub>P \* Q; A\<^sub>P \* \; A\<^sub>P \* P'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* \\<^sub>Q; + A\<^sub>Q \* P; A\<^sub>Q \* Q; A\<^sub>Q \* \; A\<^sub>Q \* P'; A\<^sub>Q \* \\<^sub>P; + A\<^sub>P \* C; A\<^sub>Q \* C\ \ + Prop C \ (P \ Q) (P' \ Q) (A\<^sub>P@A\<^sub>Q) (\\<^sub>P \ \\<^sub>Q)" + and rPar2: "\\ \\<^sub>P Q Q' A\<^sub>P P A\<^sub>Q \\<^sub>Q C. + \\ \ \\<^sub>P \ Q \\ \ Q'; + extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + \C. Prop C (\ \ \\<^sub>P) Q Q' A\<^sub>Q \\<^sub>Q; + A\<^sub>P \* P; A\<^sub>P \* Q; A\<^sub>P \* \; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* \\<^sub>Q; + A\<^sub>Q \* P; A\<^sub>Q \* Q; A\<^sub>Q \* \; A\<^sub>Q \* Q'; A\<^sub>Q \* \\<^sub>P; + A\<^sub>P \* C; A\<^sub>Q \* C\ \ + Prop C \ (P \ Q) (P \ Q') (A\<^sub>P@A\<^sub>Q) (\\<^sub>P \ \\<^sub>Q)" + and rComm1: "\\ \\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q K xvec Q' A\<^sub>Q C. + \\ \ \\<^sub>Q \ P \M\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \ \ \\<^sub>P \ Q \K\\*xvec\\N\ \ Q'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + \ \ \\<^sub>P \ \\<^sub>Q \ M \ K; distinct xvec; + A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* P'; + A\<^sub>P \* Q; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* xvec; A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; + A\<^sub>Q \* P; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* Q; A\<^sub>Q \* K; A\<^sub>Q \* Q'; + A\<^sub>Q \* xvec; xvec \* \; xvec \* \\<^sub>P; xvec \* \\<^sub>Q; xvec \* P; xvec \* M; + xvec \* Q; xvec \* K; A\<^sub>P \* C; A\<^sub>Q \* C; xvec \* C\ \ + Prop C \ (P \ Q) (\\*xvec\(P' \ Q')) (A\<^sub>P@A\<^sub>Q) (\\<^sub>P \ \\<^sub>Q)" + and rComm2: "\\ \\<^sub>Q P M xvec N P' A\<^sub>P \\<^sub>P Q K Q' A\<^sub>Q C. + \\ \ \\<^sub>Q \ P \M\\*xvec\\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \ \ \\<^sub>P \ Q \K\N\ \ Q'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + \ \ \\<^sub>P \ \\<^sub>Q \ M \ K; distinct xvec; + A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* P'; + A\<^sub>P \* Q; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* xvec; A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; + A\<^sub>Q \* P; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* Q; A\<^sub>Q \* K; A\<^sub>Q \* Q'; + A\<^sub>Q \* xvec; xvec \* \; xvec \* \\<^sub>P; xvec \* \\<^sub>Q; xvec \* P; xvec \* M; + xvec \* Q; xvec \* K; A\<^sub>P \* C; A\<^sub>Q \* C; xvec \* C\ \ + Prop C \ (P \ Q) (\\*xvec\(P' \ Q')) (A\<^sub>P@A\<^sub>Q) (\\<^sub>P \ \\<^sub>Q)" + and rBrClose: "\\ P M xvec N P' A\<^sub>P \\<^sub>P x C. + \\ \ P \ \M\\*xvec\\N\ \ P'; + x \ supp M; + extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* P'; A\<^sub>P \* xvec; + distinct xvec; xvec \* \; xvec \* \\<^sub>P; xvec \* P; + xvec \* M; + x \ \; x \ xvec; x \ A\<^sub>P; + A\<^sub>P \* C; xvec \* C; x \ C\ \ + Prop C \ (\\x\P) (\\x\(\\*xvec\P')) (x#A\<^sub>P) \\<^sub>P" + and rScope: "\\ P P' x A\<^sub>P \\<^sub>P C. + \\ \ P \\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \C. Prop C \ P P' A\<^sub>P \\<^sub>P; x \ \; + x \ A\<^sub>P; A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* P'; + A\<^sub>P \* C; x \ C\ \ + Prop C \ (\\x\P) (\\x\P') (x#A\<^sub>P) \\<^sub>P" + and rBang: "\\ P P' A\<^sub>P \\<^sub>P C. + \\ \ P \ !P \\ \ P'; guarded P; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \C. Prop C \ (P \ !P) P' A\<^sub>P (\\<^sub>P \ \); \\<^sub>P \ \; supp \\<^sub>P = ({}::name set); + A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* P'; A\<^sub>P \* C\ \ Prop C \ (!P) P' ([]) (\)" +shows "Prop C \ P P' A\<^sub>P \\<^sub>P" + using Trans FrP \distinct A\<^sub>P\ +proof(nominal_induct \ P Rs=="\ \ P'" A\<^sub>P \\<^sub>P avoiding: C arbitrary: P' rule: semanticsFrameInduct) + case cAlpha + then show ?case by(force intro: rAlpha simp add: residualInject) +next + case cInput + then show ?case by(simp add: residualInject) +next + case cBrInput + then show ?case by(simp add: residualInject) +next + case cOutput + then show ?case by(simp add: residualInject) +next + case cBrOutput + then show ?case by(simp add: residualInject) +next + case cCase + then show ?case by(force intro: rCase simp add: residualInject) +next + case cPar1 + then show ?case by(force intro: rPar1 simp add: residualInject) +next + case cPar2 + then show ?case by(force intro: rPar2 simp add: residualInject) +next + case cComm1 + then show ?case by(force intro: rComm1 simp add: residualInject) +next + case cComm2 + then show ?case by(force intro: rComm2 simp add: residualInject) +next + case cBrMerge + then show ?case by(simp add: residualInject) +next + case cBrComm1 + then show ?case by(simp add: residualInject) +next + case cBrComm2 + then show ?case by(simp add: residualInject) +next + case cBrClose + then show ?case by(force intro: rBrClose simp add: residualInject) +next + case cOpen + then show ?case by(simp add: residualInject) +next + case cBrOpen + then show ?case by(simp add: residualInject) +next + case cScope + then show ?case by(force intro: rScope simp add: residualInject) +next + case cBang + then show ?case by(force intro: rBang simp add: residualInject) +qed + +lemma inputFreshDerivative: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and M :: 'a + and N :: 'a + and P' :: "('a, 'b, 'c) psi" + and x :: name + +assumes "\ \ P \M\N\ \ P'" + and "x \ P" + and "x \ N" + +shows "x \ P'" +proof - + have "bn(M\N\) \* subject(M\N\)" and "distinct(bn(M\N\))" by simp+ + with \\ \ P \M\N\ \ P'\ show ?thesis using \x \ P\ \x \ N\ + proof(nominal_induct \ P \=="M\N\" P' avoiding: x rule: semanticsInduct) + case(cAlpha \ P \ P' p x) + then show ?case by simp + next + case(cInput \ M' K xvec N' Tvec P x) + from \K\(N'[xvec::=Tvec])\ = M\N\\ have "M = K" and NeqN': "N = N'[xvec::=Tvec]" by(simp add: action.inject)+ + note \length xvec = length Tvec\ \distinct xvec\ then + moreover have "x \ Tvec" using \set xvec \ supp N'\ \x \ N\ NeqN' + by(blast intro: substTerm.subst3) + moreover from \xvec \* x\ \x \ M'\\*xvec N'\.P\ + have "x \ P" by(simp add: inputChainFresh) (simp add: name_list_supp fresh_def) + ultimately show ?case using \xvec \* x\ by auto + next + case cBrInput + then show ?case by simp + next + case(cOutput \ M K N P x) + then show ?case by simp + next + case cBrOutput + then show ?case by simp + next + case(cCase \ P P' \ Cs x) + then show ?case by(induct Cs, auto) + next + case(cPar1 \ \\<^sub>Q P P' xvec Q x) + then show ?case by simp + next + case(cPar2 \ \\<^sub>P Q Q' xvec P x) + then show ?case by simp + next + case(cComm1 \ \\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q K xvec Q' A\<^sub>Q x) + then show ?case by simp + next + case(cComm2 \ \\<^sub>Q P M xwec N P' A\<^sub>P \\<^sub>P Q K Q' A\<^sub>Q x) + then show ?case by simp + next + case cBrMerge + then show ?case by simp + next + case cBrComm1 + then show ?case by simp + next + case cBrComm2 + then show ?case by simp + next + case cBrClose + then show ?case by simp + next + case(cOpen \ P M xvec yvec N P' x y) + then show ?case by simp + next + case(cBrOpen \ P M xvec yvec N P' x y) + then show ?case by simp + next + case(cScope \ P P' x y) + then show ?case by(simp add: abs_fresh) + next + case(cBang \ P P' x) + then show ?case by simp + qed +qed + +lemma brinputFreshDerivative: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and M :: 'a + and N :: 'a + and P' :: "('a, 'b, 'c) psi" + and x :: name + +assumes "\ \ P \\M\N\ \ P'" + and "x \ P" + and "x \ N" + +shows "x \ P'" +proof - + have "bn(\M\N\) \* subject(\M\N\)" and "distinct(bn(\M\N\))" by simp+ + with \\ \ P \\M\N\ \ P'\ show ?thesis using \x \ P\ \x \ N\ + proof(nominal_induct \ P \=="\M\N\" P' avoiding: x rule: semanticsInduct) + case(cAlpha \ P \ P' p x) + then show ?case by simp + next + case(cInput \ M' K xvec N' Tvec P x) + then show ?case by simp + next + case(cBrInput \ M' K xvec N' Tvec P x) + from \\M'\(N'[xvec::=Tvec])\ = \M\N\\ have "M' = M" and NeqN': "N = N'[xvec::=Tvec]" by(simp add: action.inject)+ + note \length xvec = length Tvec\ \distinct xvec\ then + moreover have "x \ Tvec" using \set xvec \ supp N'\ \x \ N\ NeqN' + by(blast intro: substTerm.subst3) + moreover from \xvec \* x\ \x \ K\\*xvec N'\.P\ + have "x \ P" by(simp add: inputChainFresh) (simp add: name_list_supp fresh_def) + ultimately show ?case using \xvec \* x\ by auto + next + case(cOutput \ M K N P x) + then show ?case by simp + next + case cBrOutput + then show ?case by simp + next + case(cCase \ P P' \ Cs x) + then show ?case by(induct Cs, auto) + next + case(cPar1 \ \\<^sub>Q P P' xvec Q x) + then show ?case by simp + next + case(cPar2 \ \\<^sub>P Q Q' xvec P x) + then show ?case by simp + next + case(cComm1 \ \\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q K xvec Q' A\<^sub>Q x) + then show ?case by simp + next + case(cComm2 \ \\<^sub>Q P M xwec N P' A\<^sub>P \\<^sub>P Q K Q' A\<^sub>Q x) + then show ?case by simp + next + case cBrMerge + then show ?case by simp + next + case cBrComm1 + then show ?case by simp + next + case cBrComm2 + then show ?case by simp + next + case cBrClose + then show ?case by simp + next + case(cOpen \ P M xvec yvec N P' x y) + then show ?case by simp + next + case(cBrOpen \ P M xvec yvec N P' x y) + then show ?case by simp + next + case(cScope \ P P' x y) + then show ?case by(simp add: abs_fresh) + next + case(cBang \ P P' x) + then show ?case by simp + qed +qed + +lemma inputFreshChainDerivative: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and M :: 'a + and N :: 'a + and P' :: "('a, 'b, 'c) psi" + and xvec :: "name list" + +assumes "\ \ P \M\N\ \ P'" + and "xvec \* P" + and "xvec \* N" + +shows "xvec \* P'" + using assms + by(induct xvec) + (auto intro: inputFreshDerivative) + +lemma brinputFreshChainDerivative: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and M :: 'a + and N :: 'a + and P' :: "('a, 'b, 'c) psi" + and xvec :: "name list" + +assumes "\ \ P \\M\N\ \ P'" + and "xvec \* P" + and "xvec \* N" + +shows "xvec \* P'" + using assms + by(induct xvec) + (auto intro: brinputFreshDerivative) + +lemma outputFreshDerivativeN: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and M :: 'a + and xvec :: "name list" + and N :: 'a + and P' :: "('a, 'b, 'c) psi" + and x :: name + +assumes "\ \ P \M\\*xvec\\N\ \ P'" + and "xvec \* M" + and "distinct xvec" + and "x \ P" + and "x \ xvec" + +shows "x \ N" +proof - + note \\ \ P \M\\*xvec\\N\ \ P'\ + moreover from \xvec \* M\ have "bn(M\\*xvec\\N\) \* subject(M\\*xvec\\N\)" by simp + moreover from \distinct xvec\ have "distinct(bn(M\\*xvec\\N\))" by simp + ultimately show freshN: "x \ N" using \x \ P\ \x \ xvec\ + proof(nominal_induct \ P \=="M\\*xvec\\N\" P' avoiding: x arbitrary: M xvec N rule: semanticsInduct) + case(cAlpha \ P \ P' p x M xvec N) + have S: "set p \ set(bn \) \ set(bn(p \ \))" by fact + from \(p \ \) = M\\*xvec\\N\\ have "(p \ p \ \) = p \ (M\\*xvec\\N\)" by(simp add: fresh_star_bij) + with \distinctPerm p\ have "\ = (p \ M)\\*(p \ xvec)\\(p \ N)\" by simp + moreover from \(p \ \) = M\\*xvec\\N\\ \x \ xvec\ have "x \ (bn(p \ \))" by simp + with \(bn \) \* x\ \x \ xvec\ S have "x \ (p \ xvec)" + by(fastforce dest: pt_fresh_bij1[OF pt_name_inst, OF at_name_inst, where pi=p and x=xvec]) + ultimately have "x \ (p \ N)" using \x \ P\ by(metis cAlpha) + then have "(p \ x) \ (p \ p \ N)" by(simp add: pt_fresh_bij1[OF pt_name_inst, OF at_name_inst]) + with \distinctPerm p\ \bn(\) \* x\ \x \ (bn(p \ \))\S show ?case by simp + next + case cInput + then show ?case by simp + next + case cBrInput + then show ?case by simp + next + case cOutput + then show ?case by(simp add: action.inject) + next + case cBrOutput + then show ?case by(simp add: action.inject) + next + case (cCase \ P P' \ Cs x M xvec N) + then show ?case by(auto simp add: action.inject dest: memFresh) + next + case cPar1 + then show ?case by simp + next + case cPar2 + then show ?case by simp + next + case cComm1 + then show ?case by simp + next + case cComm2 + then show ?case by simp + next + case cBrMerge + then show ?case by simp + next + case cBrComm1 + then show ?case by simp + next + case cBrComm2 + then show ?case by simp + next + case cBrClose + then show ?case by simp + next + case(cOpen \ P M xvec yvec N P' x y M' zvec N') + from \M\\*(xvec@x#yvec)\\N\ = M'\\*zvec\\N'\\ have "zvec = xvec@x#yvec" and "N = N'" + by(simp add: action.inject)+ + from \y \ \\x\P\ \x \ y\ have "y \ P" by(simp add: abs_fresh) + moreover from \y \ zvec\ \zvec = xvec@x#yvec\have "y \ (xvec@yvec)" + by simp + ultimately have "y \ N" by(fastforce intro!: cOpen) + with \N = N'\ show ?case by simp + next + case cBrOpen + then show ?case by simp + next + case cScope + then show ?case by(auto simp add: abs_fresh) + next + case cBang + then show ?case by simp + qed +qed + +lemma broutputFreshDerivativeN: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and M :: 'a + and xvec :: "name list" + and N :: 'a + and P' :: "('a, 'b, 'c) psi" + and x :: name + +assumes "\ \ P \\M\\*xvec\\N\ \ P'" + and "xvec \* M" + and "distinct xvec" + and "x \ P" + and "x \ xvec" + +shows "x \ N" +proof - + note \\ \ P \\M\\*xvec\\N\ \ P'\ + moreover from \xvec \* M\ have "bn(\M\\*xvec\\N\) \* subject(\M\\*xvec\\N\)" by simp + moreover from \distinct xvec\ have "distinct(bn(\M\\*xvec\\N\))" by simp + ultimately show freshN: "x \ N" using \x \ P\ \x \ xvec\ + proof(nominal_induct \ P \=="\M\\*xvec\\N\" P' avoiding: x arbitrary: M xvec N rule: semanticsInduct) + case(cAlpha \ P \ P' p x M xvec N) + have S: "set p \ set(bn \) \ set(bn(p \ \))" by fact + from \(p \ \) = \M\\*xvec\\N\\ have "(p \ p \ \) = p \ (\M\\*xvec\\N\)" by(simp add: fresh_star_bij) + with \distinctPerm p\ have "\ = \(p \ M)\\*(p \ xvec)\\(p \ N)\" by simp + moreover from \(p \ \) = \M\\*xvec\\N\\ \x \ xvec\ have "x \ (bn(p \ \))" by simp + with \(bn \) \* x\ \x \ xvec\ S have "x \ (p \ xvec)" + by(fastforce dest: pt_fresh_bij1[OF pt_name_inst, OF at_name_inst, where pi=p and x=xvec]) + ultimately have "x \ (p \ N)" using \x \ P\ by(metis cAlpha) + then have "(p \ x) \ (p \ p \ N)" by(simp add: pt_fresh_bij1[OF pt_name_inst, OF at_name_inst]) + with \distinctPerm p\ \bn(\) \* x\ \x \ (bn(p \ \))\S show ?case by simp + next + case cInput + then show ?case by simp + next + case cBrInput + then show ?case by simp + next + case cOutput + then show ?case by(simp add: action.inject) + next + case cBrOutput + then show ?case by(simp add: action.inject) + next + case cCase + then show ?case by(auto simp add: action.inject dest: memFresh) + next + case cPar1 + then show ?case by simp + next + case cPar2 + then show ?case by simp + next + case cComm1 + then show ?case by simp + next + case cComm2 + then show ?case by simp + next + case cBrMerge + then show ?case by simp + next + case cBrComm1 + then show ?case by simp + next + case cBrComm2 + then show ?case by simp + next + case cBrClose + then show ?case by simp + next + case(cOpen \ P M xvec yvec N P' x y M' zvec N') + then show ?case by simp + next + case(cBrOpen \ P M xvec yvec N P' x y M' zvec N') + from \\M\\*(xvec@x#yvec)\\N\ = \M'\\*zvec\\N'\\ have "zvec = xvec@x#yvec" and "N = N'" + by(simp add: action.inject)+ + from \y \ \\x\P\ \x \ y\ have "y \ P" by(simp add: abs_fresh) + moreover from \y \ zvec\ \zvec = xvec@x#yvec\have "y \ (xvec@yvec)" + by simp + ultimately have "y \ N" by(fastforce intro!: cBrOpen) + with \N = N'\ show ?case by simp + next + case cScope + then show ?case by(auto simp add: abs_fresh) + next + case cBang + then show ?case by simp + qed +qed + +lemma outputFreshDerivativeP: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and M :: 'a + and xvec :: "name list" + and N :: 'a + and P' :: "('a, 'b, 'c) psi" + and x :: name + +assumes "\ \ P \M\\*xvec\\N\ \ P'" + and "xvec \* M" + and "distinct xvec" + and "x \ P" + and "x \ xvec" + +shows "x \ P'" +proof - + note \\ \ P \M\\*xvec\\N\ \ P'\ + moreover from \xvec \* M\ have "bn(M\\*xvec\\N\) \* subject(M\\*xvec\\N\)" by simp + moreover from \distinct xvec\ have "distinct(bn(M\\*xvec\\N\))" by simp + ultimately show "x \ P'" using \x \ P\ \x \ xvec\ + proof(nominal_induct \ P \=="M\\*xvec\\N\" P' avoiding: x arbitrary: M xvec N rule: semanticsInduct) + case(cAlpha \ P \ P' p x M xvec N) + have S: "set p \ set(bn \) \ set(bn(p \ \))" by fact + from \(p \ \) = M\\*xvec\\N\\ have "(p \ p \ \) = p \ (M\\*xvec\\N\)" by(simp add: fresh_star_bij) + with \distinctPerm p\ have "\ = (p \ M)\\*(p \ xvec)\\(p \ N)\" by simp + moreover from \(p \ \) = M\\*xvec\\N\\ \x \ xvec\ have "x \ (bn(p \ \))" by simp + with \(bn \) \* x\ \x \ xvec\ S have "x \ (p \ xvec)" + by(fastforce dest: pt_fresh_bij1[OF pt_name_inst, OF at_name_inst, where pi=p and x=xvec]) + ultimately have "x \ P'" using \x \ P\ by(metis cAlpha) + then have "(p \ x) \ (p \ P')" by(simp add: pt_fresh_bij1[OF pt_name_inst, OF at_name_inst]) + with \distinctPerm p\ \bn(\) \* x\ \x \ (bn(p \ \))\S show ?case by simp + next + case cInput + then show ?case by simp + next + case cBrInput + then show ?case by simp + next + case cOutput + then show ?case by(simp add: action.inject) + next + case cBrOutput + then show ?case by(simp add: action.inject) + next + case cCase + then show ?case by(auto simp add: action.inject dest: memFresh) + next + case cPar1 + then show ?case by simp + next + case cPar2 + then show ?case by simp + next + case cComm1 + then show ?case by simp + next + case cComm2 + then show ?case by simp + next + case cBrMerge + then show ?case by simp + next + case cBrComm1 + then show ?case by simp + next + case cBrComm2 + then show ?case by simp + next + case cBrClose + then show ?case by simp + next + case(cOpen \ P M xvec yvec N P' x y M' zvec N') + from \M\\*(xvec@x#yvec)\\N\ = M'\\*zvec\\N'\\ have "zvec = xvec@x#yvec" + by(simp add: action.inject) + from \y \ \\x\P\ \x \ y\ have "y \ P" by(simp add: abs_fresh) + moreover from \y \ zvec\ \zvec = xvec@x#yvec\have "y \ (xvec@yvec)" + by simp + ultimately show "y \ P'" + by(fastforce intro!: cOpen) + next + case cBrOpen + then show ?case by simp + next + case cScope + then show ?case by(auto simp add: abs_fresh) + next + case cBang + then show ?case by simp + qed +qed + +lemma broutputFreshDerivativeP: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and M :: 'a + and xvec :: "name list" + and N :: 'a + and P' :: "('a, 'b, 'c) psi" + and x :: name + +assumes "\ \ P \\M\\*xvec\\N\ \ P'" + and "xvec \* M" + and "distinct xvec" + and "x \ P" + and "x \ xvec" + +shows "x \ P'" +proof - + note \\ \ P \\M\\*xvec\\N\ \ P'\ + moreover from \xvec \* M\ have "bn(\M\\*xvec\\N\) \* subject(\M\\*xvec\\N\)" by simp + moreover from \distinct xvec\ have "distinct(bn(\M\\*xvec\\N\))" by simp + ultimately show "x \ P'" using \x \ P\ \x \ xvec\ + proof(nominal_induct \ P \=="\M\\*xvec\\N\" P' avoiding: x arbitrary: M xvec N rule: semanticsInduct) + case(cAlpha \ P \ P' p x M xvec N) + have S: "set p \ set(bn \) \ set(bn(p \ \))" by fact + from \(p \ \) = \M\\*xvec\\N\\ have "(p \ p \ \) = p \ (\M\\*xvec\\N\)" by(simp add: fresh_star_bij) + with \distinctPerm p\ have "\ = \(p \ M)\\*(p \ xvec)\\(p \ N)\" by simp + moreover from \(p \ \) = \M\\*xvec\\N\\ \x \ xvec\ have "x \ (bn(p \ \))" by simp + with \(bn \) \* x\ \x \ xvec\ S have "x \ (p \ xvec)" + by(fastforce dest: pt_fresh_bij1[OF pt_name_inst, OF at_name_inst, where pi=p and x=xvec]) + ultimately have "x \ P'" using \x \ P\ by(metis cAlpha) + then have "(p \ x) \ (p \ P')" by(simp add: pt_fresh_bij1[OF pt_name_inst, OF at_name_inst]) + with \distinctPerm p\ \bn(\) \* x\ \x \ (bn(p \ \))\S show ?case by simp + next + case cInput + then show ?case by simp + next + case cBrInput + then show ?case by simp + next + case cOutput + then show ?case by(simp add: action.inject) + next + case cBrOutput + then show ?case by(simp add: action.inject) + next + case cCase + then show ?case by(auto simp add: action.inject dest: memFresh) + next + case cPar1 + then show ?case by simp + next + case cPar2 + then show ?case by simp + next + case cComm1 + then show ?case by simp + next + case cComm2 + then show ?case by simp + next + case cBrMerge + then show ?case by simp + next + case (cBrComm1 \ \\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q xvec Q' A\<^sub>Q x M' zvec N') + from \x \ (P \ Q)\ have "x \ P" and "x \ Q" by simp+ + + from \\ \ \\<^sub>P \ Q \ \M\\*xvec\\N\ \ Q'\ \xvec \* M\ \distinct xvec\ \x \ Q\ \xvec \* x\ + have "x \ N" by(simp add: broutputFreshDerivativeN) + + with \\ \ \\<^sub>Q \ P \ \M\N\ \ P'\ \x \ P\ have "x \ P'" by(simp add: brinputFreshDerivative) + + then show ?case using cBrComm1 by simp + next + case (cBrComm2 \ \\<^sub>Q P M xvec N P' A\<^sub>P \\<^sub>P Q Q' A\<^sub>Q x M' zvec N') + from \x \ (P \ Q)\ have "x \ P" and "x \ Q" by simp+ + + from \\ \ \\<^sub>Q \ P \ \M\\*xvec\\N\ \ P'\ \xvec \* M\ \distinct xvec\ \x \ P\ \xvec \* x\ + have "x \ N" by(simp add: broutputFreshDerivativeN) + + with \\ \ \\<^sub>P \ Q \ \M\N\ \ Q'\ \x \ Q\ have "x \ Q'" by(simp add: brinputFreshDerivative) + + then show ?case using cBrComm2 by simp + next + case cBrClose + then show ?case by simp + next + case cOpen + then show ?case by simp + next + case(cBrOpen \ P M xvec yvec N P' x y M' zvec N') + from \\M\\*(xvec@x#yvec)\\N\ = \M'\\*zvec\\N'\\ have "zvec = xvec@x#yvec" + by(simp add: action.inject) + from \y \ \\x\P\ \x \ y\ have "y \ P" by(simp add: abs_fresh) + moreover from \y \ zvec\ \zvec = xvec@x#yvec\have "y \ (xvec@yvec)" + by simp + ultimately show "y \ P'" + by(fastforce intro: cBrOpen) + next + case cScope + then show ?case by(auto simp add: abs_fresh) + next + case cBang + then show ?case by simp + qed +qed + +lemma outputFreshDerivative: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and M :: 'a + and xvec :: "name list" + and N :: 'a + and P' :: "('a, 'b, 'c) psi" + and x :: name + +assumes "\ \ P \M\\*xvec\\N\ \ P'" + and "xvec \* M" + and "distinct xvec" + and "x \ P" + and "x \ xvec" + +shows "x \ N" + and "x \ P'" + using assms + by(auto simp add: outputFreshDerivativeN outputFreshDerivativeP) + +lemma broutputFreshDerivative: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and M :: 'a + and xvec :: "name list" + and N :: 'a + and P' :: "('a, 'b, 'c) psi" + and x :: name + +assumes "\ \ P \\M\\*xvec\\N\ \ P'" + and "xvec \* M" + and "distinct xvec" + and "x \ P" + and "x \ xvec" + +shows "x \ N" + and "x \ P'" + using assms + by(auto simp add: broutputFreshDerivativeN broutputFreshDerivativeP) + +lemma outputFreshChainDerivative: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and M :: 'a + and xvec :: "name list" + and N :: 'a + and P' :: "('a, 'b, 'c) psi" + and yvec :: "name list" + +assumes "\ \ P \M\\*xvec\\N\ \ P'" + and "xvec \* M" + and "distinct xvec" + and "yvec \* P" + and "yvec \* xvec" + +shows "yvec \* N" + and "yvec \* P'" + using assms + by(induct yvec) (auto intro: outputFreshDerivative) + +lemma broutputFreshChainDerivative: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and M :: 'a + and xvec :: "name list" + and N :: 'a + and P' :: "('a, 'b, 'c) psi" + and yvec :: "name list" + +assumes "\ \ P \\M\\*xvec\\N\ \ P'" + and "xvec \* M" + and "distinct xvec" + and "yvec \* P" + and "yvec \* xvec" + +shows "yvec \* N" + and "yvec \* P'" + using assms + by(induct yvec) (auto intro: broutputFreshDerivative) + +lemma tauFreshDerivative: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and P' :: "('a, 'b, 'c) psi" + and x :: name + +assumes "\ \ P \\ \ P'" + and "x \ P" + +shows "x \ P'" +proof - + have "bn(\) \* subject(\)" and "distinct(bn(\))" by simp+ + with \\ \ P \\ \ P'\ show ?thesis using \x \ P\ + proof(nominal_induct \ P \=="(\::('a action))" P' avoiding: x rule: semanticsInduct) + case cAlpha + then show ?case by simp + next + case cInput + then show ?case by simp + next + case cBrInput + then show ?case by simp + next + case cOutput + then show ?case by simp + next + case cBrOutput + then show ?case by simp + next + case cCase + then show ?case by(auto dest: memFresh) + next + case cPar1 + then show ?case by simp + next + case cPar2 + then show ?case by simp + next + case cComm1 + then show ?case + by(auto dest: inputFreshDerivative outputFreshDerivative simp add: resChainFresh) + next + case cComm2 + then show ?case + by(auto dest: inputFreshDerivative outputFreshDerivative simp add: resChainFresh) + next + case cBrMerge + then show ?case by simp + next + case cBrComm1 + then show ?case by simp + next + case cBrComm2 + then show ?case by simp + next + case cBrClose + then show ?case + by(auto dest: brinputFreshDerivative broutputFreshDerivative simp add: resChainFresh abs_fresh) + next + case cOpen + then show ?case by simp + next + case cBrOpen + then show ?case by simp + next + case cScope + then show ?case by(simp add: abs_fresh) + next + case cBang + then show ?case by simp + qed +qed + +lemma tauFreshChainDerivative: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and M :: 'a + and N :: 'a + and P' :: "('a, 'b, 'c) psi" + and xvec :: "name list" + +assumes "\ \ P \\ \ P'" + and "xvec \* P" + +shows "xvec \* P'" + using assms + by(induct xvec) (auto intro: tauFreshDerivative) + +lemma freeFreshDerivative: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and \ :: "'a action" + and P' :: "('a, 'b, 'c) psi" + and x :: name + +assumes "\ \ P \\ \ P'" + and "bn \ \* subject \" + and "distinct(bn \)" + and "x \ \" + and "x \ P" + +shows "x \ P'" + using assms + apply - + by(rule actionCases[where \=\]) + (auto intro: inputFreshDerivative brinputFreshDerivative + tauFreshDerivative + outputFreshDerivative broutputFreshDerivative) + +lemma freeFreshChainDerivative: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and \ :: "'a action" + and P' :: "('a, 'b, 'c) psi" + and xvec :: "name list" + +assumes "\ \ P \\ \ P'" + and "bn \ \* subject \" + and "distinct(bn \)" + and "xvec \* P" + and "xvec \* \" + +shows "xvec \* P'" + using assms + by(auto intro: freeFreshDerivative simp add: fresh_star_def) + +lemma Input: + fixes \ :: 'b + and M :: 'a + and K :: 'a + and xvec :: "name list" + and N :: 'a + and Tvec :: "'a list" + +assumes "\ \ M \ K" + and "distinct xvec" + and "set xvec \ supp N" + and "length xvec = length Tvec" + +shows "\ \ M\\*xvec N\.P \K\N[xvec::=Tvec]\ \ P[xvec::=Tvec]" +proof - + obtain p where xvecFreshPsi: "((p::name prm) \ (xvec::name list)) \* \" + and xvecFreshM: "(p \ xvec) \* M" + and xvecFreshN: "(p \ xvec) \* N" + and xvecFreshK: "(p \ xvec) \* K" + and xvecFreshTvec: "(p \ xvec) \* Tvec" + and xvecFreshP: "(p \ xvec) \* P" + and S: "(set p) \ (set xvec) \ (set(p \ xvec))" + and dp: "distinctPerm p" + apply - + by(rule name_list_avoiding[where xvec=xvec and c="(\, M, K, N, P, Tvec)"]) + (auto simp add: eqvts fresh_star_prod) + note \\ \ M \ K\ + moreover from \distinct xvec\ have "distinct(p \ xvec)" + by simp + moreover from \(set xvec) \ (supp N)\ have "(p \ (set xvec)) \ (p \ (supp N))" + by simp + then have "set(p \ xvec) \ supp(p \ N)" + by(simp add: eqvts) + moreover from \length xvec = length Tvec\ have "length(p \ xvec) = length Tvec" + by simp + ultimately have "\ \ M\\*(p \ xvec) (p \ N)\.(p \ P) \K\(p \ N)[(p \ xvec)::=Tvec]\ \ (p \ P)[(p \ xvec)::=Tvec]" + using xvecFreshPsi xvecFreshM xvecFreshK xvecFreshTvec + by(metis cInput) + then show ?thesis using xvecFreshN xvecFreshP S \length xvec = length Tvec\ dp + by(auto simp add: inputChainAlpha' substTerm.renaming renaming) +qed + +lemma BrInput: + fixes \ :: 'b + and M :: 'a + and K :: 'a + and xvec :: "name list" + and N :: 'a + and Tvec :: "'a list" + +assumes "\ \ K \ M" + and "distinct xvec" + and "set xvec \ supp N" + and "length xvec = length Tvec" + +shows "\ \ M\\*xvec N\.P \\K\N[xvec::=Tvec]\ \ P[xvec::=Tvec]" +proof - + obtain p where xvecFreshPsi: "((p::name prm) \ (xvec::name list)) \* \" + and xvecFreshM: "(p \ xvec) \* M" + and xvecFreshN: "(p \ xvec) \* N" + and xvecFreshK: "(p \ xvec) \* K" + and xvecFreshTvec: "(p \ xvec) \* Tvec" + and xvecFreshP: "(p \ xvec) \* P" + and S: "(set p) \ (set xvec) \ (set(p \ xvec))" + and dp: "distinctPerm p" + apply - + by(rule name_list_avoiding[where xvec=xvec and c="(\, M, K, N, P, Tvec)"]) + (auto simp add: eqvts fresh_star_prod) + note \\ \ K \ M\ + moreover from \distinct xvec\ have "distinct(p \ xvec)" + by simp + moreover from \(set xvec) \ (supp N)\ have "(p \ (set xvec)) \ (p \ (supp N))" + by simp + then have "set(p \ xvec) \ supp(p \ N)" + by(simp add: eqvts) + moreover from \length xvec = length Tvec\ have "length(p \ xvec) = length Tvec" + by simp + ultimately have "\ \ M\\*(p \ xvec) (p \ N)\.(p \ P) \\K\(p \ N)[(p \ xvec)::=Tvec]\ \ (p \ P)[(p \ xvec)::=Tvec]" + using xvecFreshPsi xvecFreshM xvecFreshK xvecFreshTvec + by(metis cBrInput) + then show ?thesis using xvecFreshN xvecFreshP S \length xvec = length Tvec\ dp + by(auto simp add: inputChainAlpha' substTerm.renaming renaming) +qed + +lemma residualAlpha: + fixes p :: "name prm" + and \ :: "'a action" + and P :: "('a, 'b, 'c) psi" + +assumes "bn(p \ \) \* object \" + and "bn(p \ \) \* P" + and "bn \ \* subject \" + and "bn(p \ \) \* subject \" + and "set p \ set(bn \) \ set(bn(p \ \))" + +shows "\ \ P = (p \ \) \ (p \ P)" + using assms + apply - + apply(rule actionCases[where \=\]) + apply(simp only: eqvts bn.simps) + apply simp + apply(simp only: eqvts bn.simps) + apply simp + apply simp + apply(simp add: boundOutputChainAlpha'' residualInject) + apply simp + apply(simp add: boundOutputChainAlpha'' residualInject) + by simp + +lemma Par1: + fixes \ :: 'b + and \\<^sub>Q :: 'b + and P :: "('a, 'b, 'c) psi" + and \ :: "'a action" + and P' :: "('a, 'b, 'c) psi" + and A\<^sub>Q :: "name list" + and Q :: "('a, 'b, 'c) psi" + +assumes Trans: "\ \ \\<^sub>Q \ P \\ \ P'" + and "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" + and "bn \ \* Q" + and "A\<^sub>Q \* \" + and "A\<^sub>Q \* P" + and "A\<^sub>Q \* \" + +shows "\ \ P \ Q \\ \ (P' \ Q)" +proof - + { + fix \ :: 'b + and \\<^sub>Q :: 'b + and P :: "('a, 'b, 'c) psi" + and \ :: "'a action" + and P' :: "('a, 'b, 'c) psi" + and A\<^sub>Q :: "name list" + and Q :: "('a, 'b, 'c) psi" + + assume "\ \ \\<^sub>Q \ P \\ \ P'" + and "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" + and "bn \ \* Q" + and "bn \ \* subject \" + and "A\<^sub>Q \* \" + and "A\<^sub>Q \* P" + and "A\<^sub>Q \* \" + and "distinct A\<^sub>Q" + + have "\ \ P \ Q \\ \ (P' \ Q)" + proof - + from \\ \ \\<^sub>Q \ P \\ \ P'\ have "distinct(bn \)" by(rule boundOutputDistinct) + obtain q::"name prm" where "bn(q \ \) \* \" and "bn(q \ \) \* P" and "bn(q \ \) \* Q" and "bn(q \ \) \* \" + and "bn(q \ \) \* A\<^sub>Q" and "bn(q \ \) \* P'" and "bn(q \ \) \* \\<^sub>Q" + and Sq: "(set q) \ (set (bn \)) \ (set(bn(q \ \)))" + apply - + by(rule name_list_avoiding[where xvec="bn \" and c="(\, P, Q, \, A\<^sub>Q, \\<^sub>Q, P')"]) (auto simp add: eqvts) + obtain p::"name prm" where "(p \ A\<^sub>Q) \* \" and "(p \ A\<^sub>Q) \* P" and "(p \ A\<^sub>Q) \* Q" and "(p \ A\<^sub>Q) \* \" + and "(p \ A\<^sub>Q) \* \" and "(p \ A\<^sub>Q) \* (q \ \)" and "(p \ A\<^sub>Q) \* P'" + and "(p \ A\<^sub>Q) \* (q \ P')" and "(p \ A\<^sub>Q) \* \\<^sub>Q" and Sp: "(set p) \ (set A\<^sub>Q) \ (set(p \ A\<^sub>Q))" + apply - + by(rule name_list_avoiding[where xvec=A\<^sub>Q and c="(\, P, Q, \, bn \, q \ \, P', (q \ P'), \\<^sub>Q)"]) auto + from \distinct(bn \)\ have "distinct(bn(q \ \))" + by - (rule actionCases[where \=\], auto simp add: eqvts) + from \A\<^sub>Q \* \\ \bn(q \ \) \* A\<^sub>Q\ Sq have "A\<^sub>Q \* (q \ \)" + apply - + apply(rule actionCases[where \=\]) + apply(simp only: bn.simps eqvts, simp) + apply(simp only: bn.simps eqvts, simp) + apply(simp add: freshChainSimps) + apply(simp add: freshChainSimps) + by simp + from \bn \ \* subject \\ have "(q \ (bn \)) \* (q \ (subject \))" + by(simp add: fresh_star_bij) + then have "bn(q \ \) \* subject(q \ \)" by(simp add: eqvts) + from \\ \ \\<^sub>Q \ P \\ \ P'\ \bn(q \ \) \* \\ \bn(q \ \) \* P'\ \bn \ \* (subject \)\ Sq + have Trans: "\ \ \\<^sub>Q \ P \(q \ \) \ (q \ P')" + by(force simp add: residualAlpha) + then have "A\<^sub>Q \* (q \ P')" using \bn(q \ \) \* subject(q \ \)\ \distinct(bn(q \ \))\ \A\<^sub>Q \* P\ \A\<^sub>Q \* (q \ \)\ + by(auto intro: freeFreshChainDerivative) + from Trans have "(p \ (\ \ \\<^sub>Q)) \ (p \ P) \p \ ((q \ \) \ (q \ P'))" + by(rule semantics.eqvt) + with \A\<^sub>Q \* \\ \A\<^sub>Q \* P\ \A\<^sub>Q \* (q \ \)\ \(p \ A\<^sub>Q) \* (q \ \)\ \A\<^sub>Q \* (q \ P')\ + \(p \ A\<^sub>Q) \* \\ \(p \ A\<^sub>Q) \* P\ \(p \ A\<^sub>Q) \* (q \ P')\ Sp + have "\ \ (p \ \\<^sub>Q) \ P \(q \ \) \ (q \ P')" by(simp add: eqvts) + moreover from \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ \(p \ A\<^sub>Q) \* \\<^sub>Q\ Sp have "extractFrame Q = \(p \ A\<^sub>Q), (p \ \\<^sub>Q)\" + by(simp add: frameChainAlpha' eqvts) + moreover from \(bn(q \ \)) \* \\<^sub>Q\ \(bn(q \ \)) \* A\<^sub>Q\ \(p \ A\<^sub>Q) \* (q \ \)\ Sp + have "(bn(q \ \)) \* (p \ \\<^sub>Q)" + by(simp add: freshAlphaPerm) + moreover from \distinct A\<^sub>Q\ have "distinct(p \ A\<^sub>Q)" by simp + ultimately have "\ \ P \ Q \(q \ \) \ ((q \ P') \ Q)" + using \(p \ A\<^sub>Q) \* P\ \(p \ A\<^sub>Q) \* Q\ \(p \ A\<^sub>Q) \* \\ \(p \ A\<^sub>Q) \* (q \ \)\ + \(p \ A\<^sub>Q) \* (q \ P')\ \(bn(q \ \)) \* \\ \(bn(q \ \)) \* Q\ \(bn(q \ \)) \* P\ + \(bn(q \ \)) \* (subject (q \ \))\ \distinct(bn(q \ \))\ + by(metis cPar1) + + then show ?thesis using \bn(q \ \) \* \\ \bn(q \ \) \* P'\ \bn \ \* subject \\ \bn(q \ \) \* Q\ \bn \ \* Q\ Sq + by(force simp add: residualAlpha) + qed + } + note Goal = this + from \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ \A\<^sub>Q \* \\ \A\<^sub>Q \* P\ \A\<^sub>Q \* \\ + obtain A\<^sub>Q' where FrQ: "extractFrame Q = \A\<^sub>Q', \\<^sub>Q\" and "distinct A\<^sub>Q'" and "A\<^sub>Q' \* \" and "A\<^sub>Q' \* P" and "A\<^sub>Q' \* \" + apply - + by(rule distinctFrame[where C="(\, P, \)"]) auto + show ?thesis + proof(induct rule: actionCases[where \=\]) + case(cInput M N) + from Trans FrQ \A\<^sub>Q' \* \\ \A\<^sub>Q' \* P\ \A\<^sub>Q' \* \\ \distinct A\<^sub>Q'\ \bn \ \* Q\ + show ?case using \\ = M\N\\ by(force intro: Goal) + next + case(cBrInput M N) + from Trans FrQ \A\<^sub>Q' \* \\ \A\<^sub>Q' \* P\ \A\<^sub>Q' \* \\ \distinct A\<^sub>Q'\ \bn \ \* Q\ + show ?case using \\ = \M\N\\ by(force intro: Goal) + next + case cTau + from Trans FrQ \A\<^sub>Q' \* \\ \A\<^sub>Q' \* P\ \A\<^sub>Q' \* \\ \distinct A\<^sub>Q'\ \bn \ \* Q\ + show ?case using \\ = \\ by(force intro: Goal) + next + case(cOutput M xvec N) + from \\ = M\\*xvec\\N\\ \A\<^sub>Q' \* \\ \bn \ \* Q\ have "xvec \* A\<^sub>Q'" and "xvec \* Q" + by simp+ + obtain p where "(p \ xvec) \* N" and "(p \ xvec) \* P'" and "(p \ xvec) \* Q" + and "(p \ xvec) \* M" and "(p \ xvec) \* A\<^sub>Q'" + and S: "set p \ set xvec \ set(p \ xvec)" + apply - + by(rule name_list_avoiding[where xvec=xvec and c="(N, P', Q, M, A\<^sub>Q')"]) auto + from Trans \\=M\\*xvec\\N\\ have "\ \ \\<^sub>Q \ P \M\\*xvec\\N\ \ P'" by simp + with \(p \ xvec) \* N\ \(p \ xvec) \* P'\ S + have "\ \ \\<^sub>Q \ P \M\\*(p \ xvec)\\(p \ N)\ \ (p \ P')" + by(simp add: boundOutputChainAlpha'' create_residual.simps) + moreover from \xvec \* A\<^sub>Q'\ \(p \ xvec) \* A\<^sub>Q'\ \A\<^sub>Q' \* \\ S + have "A\<^sub>Q' \* (p \ \)" by(simp add: freshChainSimps del: actionFreshChain) + ultimately have "\ \ P \ Q \M\\*(p \ xvec)\\(p \ N)\ \ (p \ P') \ Q" + using FrQ \A\<^sub>Q' \* \\ \A\<^sub>Q' \* P\ \distinct A\<^sub>Q'\ \(p \ xvec) \* Q\ \A\<^sub>Q' \* \\ + \(p \ xvec) \* M\ \\ = M\\*xvec\\N\\ + by(force intro: Goal) + with \(p \ xvec) \* N\ \(p \ xvec) \* P'\ \(p \ xvec) \* Q\ \xvec \* Q\ S \\ = M\\*xvec\\N\\ + show ?case + by(simp add: boundOutputChainAlpha'' eqvts create_residual.simps) + next + case(cBrOutput M xvec N) + from \\ = \M\\*xvec\\N\\ \A\<^sub>Q' \* \\ \bn \ \* Q\ have "xvec \* A\<^sub>Q'" and "xvec \* Q" + by simp+ + obtain p where "(p \ xvec) \* N" and "(p \ xvec) \* P'" and "(p \ xvec) \* Q" + and "(p \ xvec) \* M" and "(p \ xvec) \* A\<^sub>Q'" + and S: "set p \ set xvec \ set(p \ xvec)" + by(rule name_list_avoiding[where xvec=xvec and c="(N, P', Q, M, A\<^sub>Q')"]) auto + from Trans \\=\M\\*xvec\\N\\ have "\ \ \\<^sub>Q \ P \\M\\*xvec\\N\ \ P'" by simp + with \(p \ xvec) \* N\ \(p \ xvec) \* P'\ S + have "\ \ \\<^sub>Q \ P \\M\\*(p \ xvec)\\(p \ N)\ \ (p \ P')" + by(simp add: boundOutputChainAlpha'' create_residual.simps) + moreover from \xvec \* A\<^sub>Q'\ \(p \ xvec) \* A\<^sub>Q'\ \A\<^sub>Q' \* \\ S + have "A\<^sub>Q' \* (p \ \)" by(simp add: freshChainSimps del: actionFreshChain) + ultimately have "\ \ P \ Q \\M\\*(p \ xvec)\\(p \ N)\ \ (p \ P') \ Q" + using FrQ \A\<^sub>Q' \* \\ \A\<^sub>Q' \* P\ \distinct A\<^sub>Q'\ \(p \ xvec) \* Q\ \A\<^sub>Q' \* \\ + \(p \ xvec) \* M\ \\ = \M\\*xvec\\N\\ + by(force intro: Goal) + with \(p \ xvec) \* N\ \(p \ xvec) \* P'\ \(p \ xvec) \* Q\ \xvec \* Q\ S \\ = \M\\*xvec\\N\\ + show ?case + by(simp add: boundOutputChainAlpha'' eqvts create_residual.simps) + qed +qed + +lemma Par2: + fixes \ :: 'b + and \\<^sub>P :: 'b + and Q :: "('a, 'b, 'c) psi" + and \ :: "'a action" + and Q' :: "('a, 'b, 'c) psi" + and A\<^sub>P :: "name list" + and P :: "('a, 'b, 'c) psi" + +assumes Trans: "\ \ \\<^sub>P \ Q \\ \ Q'" + and "extractFrame P = \A\<^sub>P, \\<^sub>P\" + and "bn \ \* P" + and "A\<^sub>P \* \" + and "A\<^sub>P \* Q" + and "A\<^sub>P \* \" + +shows "\ \ P \ Q \\ \ (P \ Q')" +proof - + { + fix \ :: 'b + and \\<^sub>P :: 'b + and Q :: "('a, 'b, 'c) psi" + and \ :: "'a action" + and Q' :: "('a, 'b, 'c) psi" + and A\<^sub>P :: "name list" + and P :: "('a, 'b, 'c) psi" + + assume "\ \ \\<^sub>P \ Q \\ \ Q'" + and "extractFrame P = \A\<^sub>P, \\<^sub>P\" + and "bn \ \* P" + and "bn \ \* subject \" + and "A\<^sub>P \* \" + and "A\<^sub>P \* Q" + and "A\<^sub>P \* \" + and "distinct A\<^sub>P" + + have "\ \ P \ Q \\ \ (P \ Q')" + proof - + from \\ \ \\<^sub>P \ Q \\ \ Q'\ have "distinct(bn \)" by(rule boundOutputDistinct) + obtain q::"name prm" where "bn(q \ \) \* \" and "bn(q \ \) \* P" and "bn(q \ \) \* Q" and "bn(q \ \) \* \" + and "bn(q \ \) \* A\<^sub>P" and "bn(q \ \) \* Q'" and "bn(q \ \) \* \\<^sub>P" + and Sq: "(set q) \ (set (bn \)) \ (set(bn(q \ \)))" + by(rule name_list_avoiding[where xvec="bn \" and c="(\, P, Q, \, A\<^sub>P, \\<^sub>P, Q')"]) (auto simp add: eqvts) + obtain p::"name prm" where "(p \ A\<^sub>P) \* \" and "(p \ A\<^sub>P) \* P" and "(p \ A\<^sub>P) \* Q" and "(p \ A\<^sub>P) \* \" + and "(p \ A\<^sub>P) \* \" and "(p \ A\<^sub>P) \* (q \ \)" and "(p \ A\<^sub>P) \* Q'" + and "(p \ A\<^sub>P) \* (q \ Q')" and "(p \ A\<^sub>P) \* \\<^sub>P" + and Sp: "(set p) \ (set A\<^sub>P) \ (set(p \ A\<^sub>P))" + by(rule name_list_avoiding[where xvec=A\<^sub>P and c="(\, P, Q, \, q \ \, Q', (q \ Q'), \\<^sub>P)"]) auto + from \distinct(bn \)\ have "distinct(bn(q \ \))" + apply - + by(rule actionCases[where \=\]) (auto simp add: eqvts) + from \A\<^sub>P \* \\ \bn(q \ \) \* A\<^sub>P\ Sq have "A\<^sub>P \* (q \ \)" + apply - + apply(rule actionCases[where \=\]) + apply(simp only: bn.simps eqvts, simp) + apply(simp only: bn.simps eqvts, simp) + apply(simp add: freshChainSimps) + apply(simp add: freshChainSimps) + by simp + from \bn \ \* subject \\ have "(q \ (bn \)) \* (q \ (subject \))" + by(simp add: fresh_star_bij) + then have "bn(q \ \) \* subject(q \ \)" by(simp add: eqvts) + from \\ \ \\<^sub>P \ Q \\ \ Q'\ \bn(q \ \) \* \\ \bn(q \ \) \* Q'\ \bn \ \* (subject \)\ Sq + have Trans: "\ \ \\<^sub>P \ Q \(q \ \) \ (q \ Q')" + by(force simp add: residualAlpha) + then have "A\<^sub>P \* (q \ Q')" using \bn(q \ \) \* subject(q \ \)\ \distinct(bn(q \ \))\ \A\<^sub>P \* Q\ \A\<^sub>P \* (q \ \)\ + by(auto intro: freeFreshChainDerivative) + from Trans have "(p \ (\ \ \\<^sub>P)) \ (p \ Q) \p \ ((q \ \) \ (q \ Q'))" + by(rule semantics.eqvt) + with \A\<^sub>P \* \\ \A\<^sub>P \* Q\ \A\<^sub>P \* (q \ \)\ \(p \ A\<^sub>P) \* (q \ \)\ \A\<^sub>P \* (q \ Q')\ + \(p \ A\<^sub>P) \* \\ \(p \ A\<^sub>P) \* Q\ \(p \ A\<^sub>P) \* (q \ Q')\ Sp + have "\ \ (p \ \\<^sub>P) \ Q \(q \ \) \ (q \ Q')" by(simp add: eqvts) + moreover from \extractFrame P = \A\<^sub>P, \\<^sub>P\\ \(p \ A\<^sub>P) \* \\<^sub>P\ Sp have "extractFrame P = \(p \ A\<^sub>P), (p \ \\<^sub>P)\" + by(simp add: frameChainAlpha' eqvts) + moreover from \(bn(q \ \)) \* \\<^sub>P\ \(bn(q \ \)) \* A\<^sub>P\ \(p \ A\<^sub>P) \* (q \ \)\ Sp + have "(bn(q \ \)) \* (p \ \\<^sub>P)" + by(simp add: freshAlphaPerm) + moreover from \distinct A\<^sub>P\ have "distinct(p \ A\<^sub>P)" by simp + ultimately have "\ \ P \ Q \(q \ \) \ (P \ (q \ Q'))" + using \(p \ A\<^sub>P) \* P\ \(p \ A\<^sub>P) \* Q\ \(p \ A\<^sub>P) \* \\ \(p \ A\<^sub>P) \* (q \ \)\ + \(p \ A\<^sub>P) \* (q \ Q')\ \(bn(q \ \)) \* \\ \(bn(q \ \)) \* Q\ \(bn(q \ \)) \* P\ + \(bn(q \ \)) \* (subject (q \ \))\ \distinct(bn(q \ \))\ + by(metis cPar2) + + then show ?thesis using \bn(q \ \) \* \\ \bn(q \ \) \* Q'\ \bn \ \* subject \\ \bn(q \ \) \* P\ \bn \ \* P\ Sq + by(force simp add: residualAlpha) + qed + } + note Goal = this + from \extractFrame P = \A\<^sub>P, \\<^sub>P\\ \A\<^sub>P \* \\ \A\<^sub>P \* Q\ \A\<^sub>P \* \\ + obtain A\<^sub>P' where FrP: "extractFrame P = \A\<^sub>P', \\<^sub>P\" and "distinct A\<^sub>P'" and "A\<^sub>P' \* \" and "A\<^sub>P' \* Q" and "A\<^sub>P' \* \" + apply - + by(rule distinctFrame[where C="(\, Q, \)"]) auto + show ?thesis + proof(induct rule: actionCases[where \=\]) + case(cInput M N) + from Trans FrP \A\<^sub>P' \* \\ \A\<^sub>P' \* Q\ \A\<^sub>P' \* \\ \distinct A\<^sub>P'\ \bn \ \* P\ + show ?case using \\ = M\N\\ by(force intro: Goal) + next + case(cBrInput M N) + from Trans FrP \A\<^sub>P' \* \\ \A\<^sub>P' \* Q\ \A\<^sub>P' \* \\ \distinct A\<^sub>P'\ \bn \ \* P\ + show ?case using \\ = \M\N\\ by(force intro: Goal) + next + case cTau + from Trans FrP \A\<^sub>P' \* \\ \A\<^sub>P' \* Q\ \A\<^sub>P' \* \\ \distinct A\<^sub>P'\ \bn \ \* P\ + show ?case using \\ = \\ by(force intro: Goal) + next + case(cOutput M xvec N) + from \\ = M\\*xvec\\N\\ \A\<^sub>P' \* \\ \bn \ \* P\ have "xvec \* A\<^sub>P'" and "xvec \* P" + by simp+ + obtain p where "(p \ xvec) \* N" and "(p \ xvec) \* Q'" and "(p \ xvec) \* P" + and "(p \ xvec) \* M" and "(p \ xvec) \* A\<^sub>P'" + and S: "set p \ set xvec \ set(p \ xvec)" + by(rule name_list_avoiding[where xvec=xvec and c="(N, Q', P, M, A\<^sub>P')"]) auto + from Trans \\=M\\*xvec\\N\\ have "\ \ \\<^sub>P \ Q \M\\*xvec\\N\ \ Q'" by simp + with \(p \ xvec) \* N\ \(p \ xvec) \* Q'\ S + have "\ \ \\<^sub>P \ Q \M\\*(p \ xvec)\\(p \ N)\ \ (p \ Q')" + by(simp add: boundOutputChainAlpha'' create_residual.simps) + moreover from \xvec \* A\<^sub>P'\ \(p \ xvec) \* A\<^sub>P'\ \A\<^sub>P' \* \\ S + have "A\<^sub>P' \* (p \ \)" by(simp add: freshChainSimps del: actionFreshChain) + ultimately have "\ \ P \ Q \M\\*(p \ xvec)\\(p \ N)\ \ P \ (p \ Q')" + using FrP \A\<^sub>P' \* \\ \A\<^sub>P' \* Q\ \distinct A\<^sub>P'\ \(p \ xvec) \* P\ \A\<^sub>P' \* \\ + \(p \ xvec) \* M\ \\ = M\\*xvec\\N\\ + by(force intro: Goal) + with \(p \ xvec) \* N\ \(p \ xvec) \* Q'\ \(p \ xvec) \* P\ \xvec \* P\ S \\ = M\\*xvec\\N\\ + show ?case + by(simp add: boundOutputChainAlpha'' eqvts create_residual.simps) + next + case(cBrOutput M xvec N) + from \\ = \M\\*xvec\\N\\ \A\<^sub>P' \* \\ \bn \ \* P\ have "xvec \* A\<^sub>P'" and "xvec \* P" + by simp+ + obtain p where "(p \ xvec) \* N" and "(p \ xvec) \* Q'" and "(p \ xvec) \* P" + and "(p \ xvec) \* M" and "(p \ xvec) \* A\<^sub>P'" + and S: "set p \ set xvec \ set(p \ xvec)" + by(rule name_list_avoiding[where xvec=xvec and c="(N, Q', P, M, A\<^sub>P')"]) auto + from Trans \\=\M\\*xvec\\N\\ have "\ \ \\<^sub>P \ Q \\M\\*xvec\\N\ \ Q'" by simp + with \(p \ xvec) \* N\ \(p \ xvec) \* Q'\ S + have "\ \ \\<^sub>P \ Q \\M\\*(p \ xvec)\\(p \ N)\ \ (p \ Q')" + by(simp add: boundOutputChainAlpha'' create_residual.simps) + moreover from \xvec \* A\<^sub>P'\ \(p \ xvec) \* A\<^sub>P'\ \A\<^sub>P' \* \\ S + have "A\<^sub>P' \* (p \ \)" by(simp add: freshChainSimps del: actionFreshChain) + ultimately have "\ \ P \ Q \\M\\*(p \ xvec)\\(p \ N)\ \ P \ (p \ Q')" + using FrP \A\<^sub>P' \* \\ \A\<^sub>P' \* Q\ \distinct A\<^sub>P'\ \(p \ xvec) \* P\ \A\<^sub>P' \* \\ + \(p \ xvec) \* M\ \\ = \M\\*xvec\\N\\ + by(force intro: Goal) + with \(p \ xvec) \* N\ \(p \ xvec) \* Q'\ \(p \ xvec) \* P\ \xvec \* P\ S \\ = \M\\*xvec\\N\\ + show ?case + by(simp add: boundOutputChainAlpha'' eqvts create_residual.simps) + qed +qed + +lemma Open: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and M :: 'a + and xvec :: "name list" + and yvec :: "name list" + and N :: 'a + and P' :: "('a, 'b, 'c) psi" + and x :: name + +assumes Trans: "\ \ P \M\\*(xvec@yvec)\\N\ \ P'" + and "x \ supp N" + and "x \ \" + and "x \ M" + and "x \ xvec" + and "x \ yvec" + +shows "\ \ \\x\P \M\\*(xvec@x#yvec)\\N\ \ P'" +proof - + from Trans have "distinct(xvec@yvec)" by(force dest: boundOutputDistinct) + then have "xvec \* yvec" by(induct xvec) auto + + obtain p where "(p \ yvec) \* \" and "(p \ yvec) \* P" and "(p \ yvec) \* M" + and "(p \ yvec) \* yvec" and "(p \ yvec) \* N" and "(p \ yvec) \* P'" + and "x \ (p \ yvec)" and "(p \ yvec) \* xvec" + and Sp: "(set p) \ (set yvec) \ (set(p \ yvec))" + by(rule name_list_avoiding[where xvec=yvec and c="(\, P, M, xvec, yvec, N, P', x)"]) + (auto simp add: eqvts fresh_star_prod) + obtain q where "(q \ xvec) \* \" and "(q \ xvec) \* P" and "(q \ xvec) \* M" + and "(q \ xvec) \* xvec" and "(q \ xvec) \* N" and "(q \ xvec) \* P'" + and "x \ (q \ xvec)" and "(q \ xvec) \* yvec" + and "(q \ xvec) \* p" and "(q \ xvec) \* (p \ yvec)" + and Sq: "(set q) \ (set xvec) \ (set(q \ xvec))" + by(rule name_list_avoiding[where xvec=xvec and c="(\, P, M, xvec, yvec, p \ yvec, N, P', x, p)"]) + (auto simp add: eqvts fresh_star_prod) + + note \\ \ P \M\\*(xvec@yvec)\\N\ \ P'\ + moreover from \(p \ yvec) \* N\ \(q \ xvec) \* N\ \xvec \* yvec\ \(q \ xvec) \* yvec\ \(q \ xvec) \* (p \ yvec)\ \(p \ yvec) \* xvec\ Sp Sq + have "((p@q) \ (xvec @ yvec)) \* N" + apply(simp only: eqvts) + apply(simp only: pt2[OF pt_name_inst]) + by simp + moreover from \(p \ yvec) \* P'\ \(q \ xvec) \* P'\ \xvec \* yvec\ \(q \ xvec) \* yvec\ \(q \ xvec) \* (p \ yvec)\ \(p \ yvec) \* xvec\ Sp Sq + have "((p@q) \ (xvec @ yvec)) \* P'" by(simp del: freshAlphaPerm add: eqvts pt2[OF pt_name_inst]) + moreover from Sp Sq \xvec \* yvec\ \(q \ xvec) \* yvec\ \(q \ xvec) \* (p \ yvec)\ \(p \ yvec) \* xvec\ + have Spq: "set(p@q) \ set(xvec@yvec) \ set((p@q) \ (xvec@yvec))" + by(simp add: pt2[OF pt_name_inst] eqvts) blast + ultimately have "\ \ P \M\\*((p@q) \ (xvec@yvec))\\((p@q) \ N)\ \ ((p@q) \ P')" + apply(simp add: create_residual.simps) + by(erule rev_mp) (subst boundOutputChainAlpha, auto) + + with Sp Sq \xvec \* yvec\ \(q \ xvec) \* yvec\ \(q \ xvec) \* (p \ yvec)\ \(p \ yvec) \* xvec\ + have "\ \ P \M\\*((q \ xvec)@(p \ yvec))\\((p@q) \ N)\ \ ((p@q) \ P')" + by(simp add: eqvts pt2[OF pt_name_inst] del: freshAlphaPerm) + moreover from \x \ supp N\ have "((p@q) \ x) \ (p@q) \ (supp N)" + by(simp add: pt_set_bij[OF pt_name_inst, OF at_name_inst]) + with \x \ xvec\ \x \ yvec\ \x \ (q \ xvec)\ \x \ (p \ yvec)\ Sp Sq + have "x \ supp((p@q)\ N)" by(simp add: eqvts pt2[OF pt_name_inst]) + moreover from \distinct(xvec@yvec)\ have "distinct(q \ xvec)" and "distinct(p \ yvec)" + by auto + moreover note \x \ (q \ xvec)\ \x \ (p \ yvec)\ \x \ M\ \x \ \\ + \(q \ xvec) \* \\ \(q \ xvec) \* P\ \(q \ xvec) \* M\ \(q \ xvec) \* (p \ yvec)\ + \(p \ yvec) \* \\ \(p \ yvec) \* P\ \(p \ yvec) \* M\ \distinct(q \ xvec)\ + ultimately have "\ \ \\x\P \M\\*((q \ xvec)@x#(p \ yvec))\\((p@q) \ N)\ \ ((p@q) \ P')" + by(metis cOpen) + with \x \ xvec\ \x \ yvec\ \x \ (q \ xvec)\ \x \ (p \ yvec)\ + \xvec \* yvec\ \(q \ xvec) \* yvec\ \(q \ xvec) \* (p \ yvec)\ \(p \ yvec) \* xvec\ Sp Sq + have "\ \ \\x\P \M\\*((p@q) \ (xvec@x#yvec))\\((p@q) \ N)\ \ ((p@q) \ P')" + by(simp add: eqvts pt2[OF pt_name_inst] del: freshAlphaPerm) + then show ?thesis using \((p@q) \ (xvec @ yvec)) \* N\ \((p@q) \ (xvec @ yvec)) \* P'\ Spq + apply(simp add: create_residual.simps) + by(erule rev_mp) (subst boundOutputChainAlpha, auto) +qed + +lemma BrOpen: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and M :: 'a + and xvec :: "name list" + and yvec :: "name list" + and N :: 'a + and P' :: "('a, 'b, 'c) psi" + and x :: name + +assumes Trans: "\ \ P \\M\\*(xvec@yvec)\\N\ \ P'" + and "x \ supp N" + and "x \ \" + and "x \ M" + and "x \ xvec" + and "x \ yvec" + +shows "\ \ \\x\P \\M\\*(xvec@x#yvec)\\N\ \ P'" +proof - + from Trans have "distinct(xvec@yvec)" by(force dest: boundOutputDistinct) + then have "xvec \* yvec" by(induct xvec) auto + + obtain p where "(p \ yvec) \* \" and "(p \ yvec) \* P" and "(p \ yvec) \* M" + and "(p \ yvec) \* yvec" and "(p \ yvec) \* N" and "(p \ yvec) \* P'" + and "x \ (p \ yvec)" and "(p \ yvec) \* xvec" + and Sp: "(set p) \ (set yvec) \ (set(p \ yvec))" + by(rule name_list_avoiding[where xvec=yvec and c="(\, P, M, xvec, yvec, N, P', x)"]) + (auto simp add: eqvts fresh_star_prod) + obtain q where "(q \ xvec) \* \" and "(q \ xvec) \* P" and "(q \ xvec) \* M" + and "(q \ xvec) \* xvec" and "(q \ xvec) \* N" and "(q \ xvec) \* P'" + and "x \ (q \ xvec)" and "(q \ xvec) \* yvec" + and "(q \ xvec) \* p" and "(q \ xvec) \* (p \ yvec)" + and Sq: "(set q) \ (set xvec) \ (set(q \ xvec))" + by(rule name_list_avoiding[where xvec=xvec and c="(\, P, M, xvec, yvec, p \ yvec, N, P', x, p)"]) + (auto simp add: eqvts fresh_star_prod) + + note \\ \ P \\M\\*(xvec@yvec)\\N\ \ P'\ + moreover from \(p \ yvec) \* N\ \(q \ xvec) \* N\ \xvec \* yvec\ \(q \ xvec) \* yvec\ \(q \ xvec) \* (p \ yvec)\ \(p \ yvec) \* xvec\ Sp Sq + have "((p@q) \ (xvec @ yvec)) \* N" + apply(simp only: eqvts) + apply(simp only: pt2[OF pt_name_inst]) + by simp + moreover from \(p \ yvec) \* P'\ \(q \ xvec) \* P'\ \xvec \* yvec\ \(q \ xvec) \* yvec\ \(q \ xvec) \* (p \ yvec)\ \(p \ yvec) \* xvec\ Sp Sq + have "((p@q) \ (xvec @ yvec)) \* P'" by(simp del: freshAlphaPerm add: eqvts pt2[OF pt_name_inst]) + moreover from Sp Sq \xvec \* yvec\ \(q \ xvec) \* yvec\ \(q \ xvec) \* (p \ yvec)\ \(p \ yvec) \* xvec\ + have Spq: "set(p@q) \ set(xvec@yvec) \ set((p@q) \ (xvec@yvec))" + by(simp add: pt2[OF pt_name_inst] eqvts) blast + ultimately have "\ \ P \\M\\*((p@q) \ (xvec@yvec))\\((p@q) \ N)\ \ ((p@q) \ P')" + apply(simp add: create_residual.simps) + by(erule rev_mp) (subst boundOutputChainAlpha, auto) + + with Sp Sq \xvec \* yvec\ \(q \ xvec) \* yvec\ \(q \ xvec) \* (p \ yvec)\ \(p \ yvec) \* xvec\ + have "\ \ P \\M\\*((q \ xvec)@(p \ yvec))\\((p@q) \ N)\ \ ((p@q) \ P')" + by(simp add: eqvts pt2[OF pt_name_inst] del: freshAlphaPerm) + moreover from \x \ supp N\ have "((p@q) \ x) \ (p@q) \ (supp N)" + by(simp add: pt_set_bij[OF pt_name_inst, OF at_name_inst]) + with \x \ xvec\ \x \ yvec\ \x \ (q \ xvec)\ \x \ (p \ yvec)\ Sp Sq + have "x \ supp((p@q)\ N)" by(simp add: eqvts pt2[OF pt_name_inst]) + moreover from \distinct(xvec@yvec)\ have "distinct(q \ xvec)" and "distinct(p \ yvec)" + by auto + moreover note \x \ (q \ xvec)\ \x \ (p \ yvec)\ \x \ M\ \x \ \\ + \(q \ xvec) \* \\ \(q \ xvec) \* P\ \(q \ xvec) \* M\ \(q \ xvec) \* (p \ yvec)\ + \(p \ yvec) \* \\ \(p \ yvec) \* P\ \(p \ yvec) \* M\ \distinct(q \ xvec)\ + ultimately have "\ \ \\x\P \\M\\*((q \ xvec)@x#(p \ yvec))\\((p@q) \ N)\ \ ((p@q) \ P')" + by(metis cBrOpen) + with \x \ xvec\ \x \ yvec\ \x \ (q \ xvec)\ \x \ (p \ yvec)\ + \xvec \* yvec\ \(q \ xvec) \* yvec\ \(q \ xvec) \* (p \ yvec)\ \(p \ yvec) \* xvec\ Sp Sq + have "\ \ \\x\P \\M\\*((p@q) \ (xvec@x#yvec))\\((p@q) \ N)\ \ ((p@q) \ P')" + by(simp add: eqvts pt2[OF pt_name_inst] del: freshAlphaPerm) + then show ?thesis using \((p@q) \ (xvec @ yvec)) \* N\ \((p@q) \ (xvec @ yvec)) \* P'\ Spq + apply(simp add: create_residual.simps) + by(erule rev_mp) (subst boundOutputChainAlpha, auto) +qed + +lemma Scope: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and \ :: "'a action" + and P' :: "('a, 'b, 'c) psi" + and x :: name + +assumes "\ \ P \\ \ P'" + and "x \ \" + and "x \ \" + +shows "\ \ \\x\P \\ \ \\x\P'" +proof - + { + fix \ P M xvec N P' x + + assume "\ \ P \M\\*xvec\\N\ \ P'" + and "(x::name) \ \" + and "x \ M" + and "x \ xvec" + and "x \ N" + + obtain p::"name prm" where "(p \ xvec) \* \" and "(p \ xvec) \* P" and "(p \ xvec) \* M" and "(p \ xvec) \* xvec" + and "(p \ xvec) \* N" and "(p \ xvec) \* P'" and "x \ (p \ xvec)" + and S: "(set p) \ (set xvec) \ (set(p \ xvec))" + by(rule name_list_avoiding[where xvec=xvec and c="(\, P, M, xvec, N, P', x)"]) + (auto simp add: eqvts fresh_star_prod) + from \\ \ P \M\\*xvec\\N\ \ P'\ \(p \ xvec) \* N\ \(p \ xvec) \* P'\ S + have "\ \ P \M\\*(p \ xvec)\\(p \ N)\ \ (p \ P')" + by(simp add: boundOutputChainAlpha'' create_residual.simps) + moreover then have "distinct(p \ xvec)" by(force dest: boundOutputDistinct) + moreover note \x \ \\ \x \ M\ \x \ (p \ xvec)\ + moreover from \x \ xvec\ \x \ p \ xvec\ \x \ N\ S have "x \ (p \ N)" + by(simp add: fresh_left del: freshAlphaSwap) + ultimately have "\ \ \\x\P \M\\*(p \ xvec)\\(p \ N)\ \ \\x\(p \ P')" using \(p \ xvec) \* \\ \(p \ xvec) \* P\ \(p \ xvec) \* M\ + by(force intro: cScope) + moreover from \x \ xvec\ \x \ p \ xvec\ S have "p \ x = x" by simp + ultimately have "\ \ \\x\P \M\\*(p \ xvec)\\(p \ N)\ \ (p \ (\\x\P'))" by simp + moreover from \(p \ xvec) \* P'\ \x \ xvec\ \x \ (p \ xvec)\ have "(p \ xvec) \* \\x\P'" + by(simp add: abs_fresh_star) + ultimately have "\ \ \\x\P \M\\*xvec\\N\ \ \\x\P'" using \(p \ xvec) \* N\ S + by(simp add: boundOutputChainAlpha'' create_residual.simps) + } + then have + (* This is simply for naming purposes, as we have two raw proof blocks. *) + outputCase: "\\ P M xvec N P' x. + \\ \ P \M\\*xvec\\N\ \ P'; + (x::name) \ \; + x \ M; + x \ xvec; + x \ N\ \ + \ \ \\x\P \M\\*xvec\\N\ \ \\x\P'" by simp + + { + fix \ P M xvec N P' x + + assume "\ \ P \\M\\*xvec\\N\ \ P'" + and "(x::name) \ \" + and "x \ M" + and "x \ xvec" + and "x \ N" + + obtain p::"name prm" where "(p \ xvec) \* \" and "(p \ xvec) \* P" and "(p \ xvec) \* M" and "(p \ xvec) \* xvec" + and "(p \ xvec) \* N" and "(p \ xvec) \* P'" and "x \ (p \ xvec)" + and S: "(set p) \ (set xvec) \ (set(p \ xvec))" + by(rule name_list_avoiding[where xvec=xvec and c="(\, P, M, xvec, N, P', x)"]) + (auto simp add: eqvts fresh_star_prod) + from \\ \ P \\M\\*xvec\\N\ \ P'\ \(p \ xvec) \* N\ \(p \ xvec) \* P'\ S + have "\ \ P \\M\\*(p \ xvec)\\(p \ N)\ \ (p \ P')" + by(simp add: boundOutputChainAlpha'' create_residual.simps) + moreover then have "distinct(p \ xvec)" by(force dest: boundOutputDistinct) + moreover note \x \ \\ \x \ M\ \x \ (p \ xvec)\ + moreover from \x \ xvec\ \x \ p \ xvec\ \x \ N\ S have "x \ (p \ N)" + by(simp add: fresh_left del: freshAlphaSwap) + ultimately have "\ \ \\x\P \\M\\*(p \ xvec)\\(p \ N)\ \ \\x\(p \ P')" using \(p \ xvec) \* \\ \(p \ xvec) \* P\ \(p \ xvec) \* M\ + by(force simp add: cScope) + moreover from \x \ xvec\ \x \ p \ xvec\ S have "p \ x = x" by simp + ultimately have "\ \ \\x\P \\M\\*(p \ xvec)\\(p \ N)\ \ (p \ (\\x\P'))" by simp + moreover from \(p \ xvec) \* P'\ \x \ xvec\ \x \ (p \ xvec)\ have "(p \ xvec) \* \\x\P'" + by(simp add: abs_fresh_star) + ultimately have "\ \ \\x\P \\M\\*xvec\\N\ \ \\x\P'" using \(p \ xvec) \* N\ S + by(simp add: boundOutputChainAlpha'' create_residual.simps) + } + then have + (* This is simply for naming purposes, as we have two raw proof blocks. *) + broutputCase: "\\ P M xvec N P' x. + \\ \ P \\M\\*xvec\\N\ \ P'; + (x::name) \ \; + x \ M; + x \ xvec; + x \ N\ \ + \ \ \\x\P \\M\\*xvec\\N\ \ \\x\P'" by simp + + show ?thesis + proof(induct rule: actionCases[where \=\]) + case(cInput M N) + with assms show ?case by(force intro: cScope) + next + case(cBrInput M N) + with assms show ?case by(force intro: cScope) + next + case(cOutput M xvec N) + with assms show ?case by(force intro: outputCase) + next + case(cBrOutput M xvec N) + with assms show ?case by(force intro: broutputCase) + next + case cTau + with assms show ?case by(force intro: cScope) + qed +qed + +lemma inputSwapFrameSubject: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and M :: 'a + and N :: 'a + and P' :: "('a, 'b, 'c) psi" + and x :: name + and y :: name + +assumes "\ \ P \M\N\ \ P'" + and "x \ P" + and "y \ P" + +shows "([(x, y)] \ \) \ P \ ([(x, y)] \ M)\N\ \ P'" + using assms +proof(nominal_induct avoiding: x y rule: inputInduct) + case(cInput \ M K xvec N Tvec P x y) + from \x \ M\\*xvec N\.P\ have "x \ M" by simp + from \y \ M\\*xvec N\.P\ have "y \ M" by simp + from \\ \ M \ K\ have "([(x, y)] \ \) \ ([(x, y)] \ M) \ ([(x, y)] \ K)" + by(rule chanEqClosed) + with \x \ M\ \y \ M\ have "([(x, y)] \ \) \ M \ ([(x, y)] \ K)" + by(simp) + then show ?case using \distinct xvec\ \set xvec \ supp N\ \length xvec = length Tvec\ + by(rule Input) +next + case(cCase \ P M N P' \ Cs x y) + from \x \ Cases Cs\ \y \ Cases Cs\ \(\, P) \ set Cs\ have "x \ \" and "x \ P" and "y \ \" and "y \ P" + by(auto dest: memFresh) + from \x \ P\ \y \ P\ have "([(x ,y)] \ \) \ P \ ([(x, y)] \ M)\N\ \ P'" by(rule cCase) + moreover note \(\, P) \ set Cs\ + moreover from \\ \ \\ have "([(x, y)] \ \) \ ([(x, y)] \ \)" by(rule statClosed) + with \x \ \\ \y \ \\ have "([(x, y)] \ \) \ \" by simp + ultimately show ?case using \guarded P\ by(rule Case) +next + case(cPar1 \ \\<^sub>Q P M N P' A\<^sub>Q Q x y) + from \x \ P \ Q\ have "x \ P" and "x \ Q" by simp+ + from \y \ P \ Q\ have "y \ P" and "y \ Q" by simp+ + from \x \ P\ \y \ P\ \\x y. \x \ P; y \ P\ \ ([(x, y)] \ (\ \ \\<^sub>Q)) \ P \([(x, y)] \ M)\N\ \ P'\ + have "([(x, y)] \ \) \ ([(x, y)] \ \\<^sub>Q) \ P \([(x, y)] \ M)\N\ \ P'" + by(simp add: eqvts) + + moreover from \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ have "([(x, y)] \ (extractFrame Q)) = ([(x, y)] \ \A\<^sub>Q, \\<^sub>Q\)" + by simp + with \A\<^sub>Q \* x\ \x \ Q\ \A\<^sub>Q \* y\ \y \ Q\ have "\A\<^sub>Q, ([(x, y)] \ \\<^sub>Q)\ = extractFrame Q" + by(simp add: eqvts) + moreover from \A\<^sub>Q \* \\ have "([(x, y)] \ A\<^sub>Q) \* ([(x, y)] \ \)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>Q \* x\ \A\<^sub>Q \* y\ have "A\<^sub>Q \* ([(x, y)] \ \)" by simp + moreover from \A\<^sub>Q \* M\ have "([(x, y)] \ A\<^sub>Q) \* ([(x, y)] \ M)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>Q \* x\ \A\<^sub>Q \* y\ have "A\<^sub>Q \* ([(x, y)] \ M)" by simp + ultimately show ?case using \A\<^sub>Q \* P\ \A\<^sub>Q \* N\ + by(force intro!: Par1) +next + case(cPar2 \ \\<^sub>P Q M N Q' A\<^sub>P P x y) + from \x \ P \ Q\ have "x \ P" and "x \ Q" by simp+ + from \y \ P \ Q\ have "y \ P" and "y \ Q" by simp+ + from \x \ Q\ \y \ Q\ \\x y. \x \ Q; y \ Q\ \ ([(x, y)] \ (\ \ \\<^sub>P)) \ Q \([(x, y)] \ M)\N\ \ Q'\ + have "([(x, y)] \ \) \ ([(x, y)] \ \\<^sub>P) \ Q \([(x, y)] \ M)\N\ \ Q'" + by(simp add: eqvts) + + moreover from \extractFrame P = \A\<^sub>P, \\<^sub>P\\ have "([(x, y)] \ (extractFrame P)) = ([(x, y)] \ \A\<^sub>P, \\<^sub>P\)" + by simp + with \A\<^sub>P \* x\ \x \ P\ \A\<^sub>P \* y\ \y \ P\ have "\A\<^sub>P, ([(x, y)] \ \\<^sub>P)\ = extractFrame P" + by(simp add: eqvts) + moreover from \A\<^sub>P \* \\ have "([(x, y)] \ A\<^sub>P) \* ([(x, y)] \ \)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>P \* x\ \A\<^sub>P \* y\ have "A\<^sub>P \* ([(x, y)] \ \)" by simp + moreover from \A\<^sub>P \* M\ have "([(x, y)] \ A\<^sub>P) \* ([(x, y)] \ M)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>P \* x\ \A\<^sub>P \* y\ have "A\<^sub>P \* ([(x, y)] \ M)" by simp + ultimately show ?case using \A\<^sub>P \* Q\ \A\<^sub>P \* N\ + by(force intro: Par2) +next + case(cScope \ P M N P' z x y) + from \x \ \\z\P\ \z \ x\ have "x \ P" by(simp add: abs_fresh) + from \y \ \\z\P\ \z \ y\ have "y \ P" by(simp add: abs_fresh) + from \x \ P\ \y \ P\ \\x y. \x \ P; y \ P\ \ ([(x, y)] \ \) \ P \([(x, y)] \ M)\N\ \ P'\ + have "([(x, y)] \ \) \ P \([(x, y)] \ M)\N\ \ P'" by simp + moreover with \z \ \\ have "([(x, y)] \ z) \ [(x, y)] \ \" + by(simp add: pt_fresh_bij[OF pt_name_inst, OF at_name_inst]) + with \z \ x\ \z \ y\ have "z \ [(x, y)] \ \" by simp + moreover with \z \ M\ have "([(x, y)] \ z) \ [(x, y)] \ M" + by(simp add: pt_fresh_bij[OF pt_name_inst, OF at_name_inst]) + with \z \ x\ \z \ y\ have "z \ [(x, y)] \ M" by simp + ultimately show ?case using \z \ N\ + by(force intro!: Scope) +next + case(cBang \ P M N P' x y) + then show ?case by(force intro: Bang) +qed + +lemma brinputSwapFrameSubject: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and M :: 'a + and N :: 'a + and P' :: "('a, 'b, 'c) psi" + and x :: name + and y :: name + +assumes "\ \ P \\M\N\ \ P'" + and "x \ P" + and "y \ P" + +shows "([(x, y)] \ \) \ P \ \([(x, y)] \ M)\N\ \ P'" + using assms +proof(nominal_induct avoiding: x y rule: brInputInduct) + case(cBrInput \ K M xvec N Tvec P x y) + from \x \ M\\*xvec N\.P\ have "x \ M" by simp + from \y \ M\\*xvec N\.P\ have "y \ M" by simp + from \\ \ K \ M\ have "([(x, y)] \ \) \ ([(x, y)] \ K) \ ([(x, y)] \ M)" + by(rule chanInConClosed) + with \x \ M\ \y \ M\ have "([(x, y)] \ \) \ ([(x, y)] \ K) \ M" + by(simp) + then show ?case using \distinct xvec\ \set xvec \ supp N\ \length xvec = length Tvec\ + by(rule BrInput) +next + case(cCase \ P M N P' \ Cs x y) + from \x \ Cases Cs\ \y \ Cases Cs\ \(\, P) \ set Cs\ have "x \ \" and "x \ P" and "y \ \" and "y \ P" + by(auto dest: memFresh) + from \x \ P\ \y \ P\ have "([(x ,y)] \ \) \ P \ \([(x, y)] \ M)\N\ \ P'" by(rule cCase) + moreover note \(\, P) \ set Cs\ + moreover from \\ \ \\ have "([(x, y)] \ \) \ ([(x, y)] \ \)" by(rule statClosed) + with \x \ \\ \y \ \\ have "([(x, y)] \ \) \ \" by simp + ultimately show ?case using \guarded P\ by(rule Case) +next + case(cPar1 \ \\<^sub>Q P M N P' A\<^sub>Q Q x y) + from \x \ P \ Q\ have "x \ P" and "x \ Q" by simp+ + from \y \ P \ Q\ have "y \ P" and "y \ Q" by simp+ + from \x \ P\ \y \ P\ \\x y. \x \ P; y \ P\ \ ([(x, y)] \ (\ \ \\<^sub>Q)) \ P \\([(x, y)] \ M)\N\ \ P'\ + have "([(x, y)] \ \) \ ([(x, y)] \ \\<^sub>Q) \ P \\([(x, y)] \ M)\N\ \ P'" + by(simp add: eqvts) + + moreover from \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ have "([(x, y)] \ (extractFrame Q)) = ([(x, y)] \ \A\<^sub>Q, \\<^sub>Q\)" + by simp + with \A\<^sub>Q \* x\ \x \ Q\ \A\<^sub>Q \* y\ \y \ Q\ have "\A\<^sub>Q, ([(x, y)] \ \\<^sub>Q)\ = extractFrame Q" + by(simp add: eqvts) + moreover from \A\<^sub>Q \* \\ have "([(x, y)] \ A\<^sub>Q) \* ([(x, y)] \ \)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>Q \* x\ \A\<^sub>Q \* y\ have "A\<^sub>Q \* ([(x, y)] \ \)" by simp + moreover from \A\<^sub>Q \* M\ have "([(x, y)] \ A\<^sub>Q) \* ([(x, y)] \ M)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>Q \* x\ \A\<^sub>Q \* y\ have "A\<^sub>Q \* ([(x, y)] \ M)" by simp + ultimately show ?case using \A\<^sub>Q \* P\ \A\<^sub>Q \* N\ + by(force intro!: Par1) +next + case(cPar2 \ \\<^sub>P Q M N Q' A\<^sub>P P x y) + from \x \ P \ Q\ have "x \ P" and "x \ Q" by simp+ + from \y \ P \ Q\ have "y \ P" and "y \ Q" by simp+ + from \x \ Q\ \y \ Q\ \\x y. \x \ Q; y \ Q\ \ ([(x, y)] \ (\ \ \\<^sub>P)) \ Q \\([(x, y)] \ M)\N\ \ Q'\ + have "([(x, y)] \ \) \ ([(x, y)] \ \\<^sub>P) \ Q \\([(x, y)] \ M)\N\ \ Q'" + by(simp add: eqvts) + + moreover from \extractFrame P = \A\<^sub>P, \\<^sub>P\\ have "([(x, y)] \ (extractFrame P)) = ([(x, y)] \ \A\<^sub>P, \\<^sub>P\)" + by simp + with \A\<^sub>P \* x\ \x \ P\ \A\<^sub>P \* y\ \y \ P\ have "\A\<^sub>P, ([(x, y)] \ \\<^sub>P)\ = extractFrame P" + by(simp add: eqvts) + moreover from \A\<^sub>P \* \\ have "([(x, y)] \ A\<^sub>P) \* ([(x, y)] \ \)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>P \* x\ \A\<^sub>P \* y\ have "A\<^sub>P \* ([(x, y)] \ \)" by simp + moreover from \A\<^sub>P \* M\ have "([(x, y)] \ A\<^sub>P) \* ([(x, y)] \ M)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>P \* x\ \A\<^sub>P \* y\ have "A\<^sub>P \* ([(x, y)] \ M)" by simp + ultimately show ?case using \A\<^sub>P \* Q\ \A\<^sub>P \* N\ + by(force intro!: Par2) +next + case (cBrMerge \ \\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q Q' A\<^sub>Q x y) + from \x \ P \ Q\ have "x \ P" and "x \ Q" by simp+ + from \y \ P \ Q\ have "y \ P" and "y \ Q" by simp+ + + from \extractFrame P = \A\<^sub>P, \\<^sub>P\\ have "([(x, y)] \ (extractFrame P)) = ([(x, y)] \ \A\<^sub>P, \\<^sub>P\)" + by simp + with \A\<^sub>P \* x\ \x \ P\ \A\<^sub>P \* y\ \y \ P\ have "extractFrame P = \A\<^sub>P, ([(x, y)] \ \\<^sub>P)\" + by(simp add: eqvts) + + from \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ have "([(x, y)] \ (extractFrame Q)) = ([(x, y)] \ \A\<^sub>Q, \\<^sub>Q\)" + by simp + with \A\<^sub>Q \* x\ \x \ Q\ \A\<^sub>Q \* y\ \y \ Q\ have "extractFrame Q = \A\<^sub>Q, ([(x, y)] \ \\<^sub>Q)\" + by(simp add: eqvts) + + from \x \ P\ \y \ P\ \\x y. \x \ P; y \ P\ \ ([(x, y)] \ (\ \ \\<^sub>Q)) \ P \\([(x, y)] \ M)\N\ \ P'\ + have "([(x, y)] \ \) \ ([(x, y)] \ \\<^sub>Q) \ P \\([(x, y)] \ M)\N\ \ P'" + by(simp add: eqvts) + moreover from \x \ Q\ \y \ Q\ \\x y. \x \ Q; y \ Q\ \ ([(x, y)] \ (\ \ \\<^sub>P)) \ Q \\([(x, y)] \ M)\N\ \ Q'\ + have "([(x, y)] \ \) \ ([(x, y)] \ \\<^sub>P) \ Q \\([(x, y)] \ M)\N\ \ Q'" + by(simp add: eqvts) + moreover note \extractFrame P = \A\<^sub>P, ([(x, y)] \ \\<^sub>P)\\ \extractFrame Q = \A\<^sub>Q, ([(x, y)] \ \\<^sub>Q)\\ + + moreover from \A\<^sub>P \* \\ have "([(x, y)] \ A\<^sub>P) \* ([(x, y)] \ \)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>P \* x\ \A\<^sub>P \* y\ have "A\<^sub>P \* ([(x, y)] \ \)" by simp + moreover from \A\<^sub>P \* \\<^sub>Q\ have "([(x, y)] \ A\<^sub>P) \* ([(x, y)] \ \\<^sub>Q)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>P \* x\ \A\<^sub>P \* y\ have "A\<^sub>P \* ([(x, y)] \ \\<^sub>Q)" by simp + moreover from \A\<^sub>P \* M\ have "([(x, y)] \ A\<^sub>P) \* ([(x, y)] \ M)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>P \* x\ \A\<^sub>P \* y\ have "A\<^sub>P \* ([(x, y)] \ M)" by simp + + moreover from \A\<^sub>Q \* \\ have "([(x, y)] \ A\<^sub>Q) \* ([(x, y)] \ \)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>Q \* x\ \A\<^sub>Q \* y\ have "A\<^sub>Q \* ([(x, y)] \ \)" by simp + moreover from \A\<^sub>Q \* \\<^sub>P\ have "([(x, y)] \ A\<^sub>Q) \* ([(x, y)] \ \\<^sub>P)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>Q \* x\ \A\<^sub>Q \* y\ have "A\<^sub>Q \* ([(x, y)] \ \\<^sub>P)" by simp + moreover from \A\<^sub>Q \* M\ have "([(x, y)] \ A\<^sub>Q) \* ([(x, y)] \ M)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>Q \* x\ \A\<^sub>Q \* y\ have "A\<^sub>Q \* ([(x, y)] \ M)" by simp + + moreover note \distinct A\<^sub>P\ \distinct A\<^sub>Q\ + \A\<^sub>P \* P\ \A\<^sub>P \* N\ \A\<^sub>P \* P'\ \A\<^sub>P \* Q\ \A\<^sub>P \* Q'\ \A\<^sub>P \* A\<^sub>Q\ + \A\<^sub>Q \* P\ \A\<^sub>Q \* N\ \A\<^sub>Q \* P'\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* Q'\ + ultimately show ?case + by(force intro!: semantics.cBrMerge) +next + case(cScope \ P M N P' z x y) + from \x \ \\z\P\ \z \ x\ have "x \ P" by(simp add: abs_fresh) + from \y \ \\z\P\ \z \ y\ have "y \ P" by(simp add: abs_fresh) + from \x \ P\ \y \ P\ \\x y. \x \ P; y \ P\ \ ([(x, y)] \ \) \ P \\([(x, y)] \ M)\N\ \ P'\ + have "([(x, y)] \ \) \ P \\([(x, y)] \ M)\N\ \ P'" by simp + moreover with \z \ \\ have "([(x, y)] \ z) \ [(x, y)] \ \" + by(simp add: pt_fresh_bij[OF pt_name_inst, OF at_name_inst]) + with \z \ x\ \z \ y\ have "z \ [(x, y)] \ \" by simp + moreover with \z \ M\ have "([(x, y)] \ z) \ [(x, y)] \ M" + by(simp add: pt_fresh_bij[OF pt_name_inst, OF at_name_inst]) + with \z \ x\ \z \ y\ have "z \ [(x, y)] \ M" by simp + ultimately show ?case using \z \ N\ + by(force intro: Scope) +next + case(cBang \ P M N P' x y) + then show ?case by(force intro: Bang) +qed + +lemma inputPermFrameSubject: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and M :: 'a + and N :: 'a + and P' :: "('a, 'b, 'c) psi" + and p :: "name prm" + and Xs :: "name set" + and Ys :: "name set" + +assumes "\ \ P \M\N\ \ P'" + and S: "set p \ Xs \ Ys" + and "Xs \* P" + and "Ys \* P" + +shows "(p \ \) \ P \ (p \ M)\N\ \ P'" + using S +proof(induct p) + case Nil + from \\ \ P \M\N\ \ P'\ + show ?case by simp +next + case(Cons a p) + from \set(a#p) \ Xs \ Ys\ have "set p \ Xs \ Ys" by auto + with \set p \ Xs \ Ys \ (p \ \) \ P \ (p \ M)\N\ \ P'\ + have Trans: "(p \ \) \ P \ (p \ M)\N\ \ P'" by simp + from \set(a#p) \ Xs \ Ys\ show ?case + proof(cases a) + case (Pair x y) + then have "x \ Xs" and "y \ Ys" + using \set(a#p) \ Xs \ Ys\ by auto + with \Xs \* P\ \Ys \* P\ have "x \ P" and "y \ P" + by(auto simp add: fresh_star_def) + with Trans have "([(x, y)] \ p \ \) \ P \ ([(x, y)] \ p \ M)\N\ \ P'" + by(rule inputSwapFrameSubject) + then show ?thesis + using Pair by simp + qed +qed + +lemma brinputPermFrameSubject: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and M :: 'a + and N :: 'a + and P' :: "('a, 'b, 'c) psi" + and p :: "name prm" + and Xs :: "name set" + and Ys :: "name set" + +assumes "\ \ P \\M\N\ \ P'" + and S: "set p \ Xs \ Ys" + and "Xs \* P" + and "Ys \* P" + +shows "(p \ \) \ P \ \(p \ M)\N\ \ P'" + using S +proof(induct p) + case Nil + from \\ \ P \\M\N\ \ P'\ + show ?case by simp +next + case(Cons a p) + from \set(a#p) \ Xs \ Ys\ have "set p \ Xs \ Ys" by auto + with \set p \ Xs \ Ys \ (p \ \) \ P \ \(p \ M)\N\ \ P'\ + have Trans: "(p \ \) \ P \ \(p \ M)\N\ \ P'" by simp + from \set(a#p) \ Xs \ Ys\ show ?case + proof(cases a) + case (Pair x y) + then have "x \ Xs" and "y \ Ys" + using \set(a#p) \ Xs \ Ys\ by auto + with \Xs \* P\ \Ys \* P\ have "x \ P" and "y \ P" + by(auto simp add: fresh_star_def) + with Trans have "([(x, y)] \ p \ \) \ P \ \([(x, y)] \ p \ M)\N\ \ P'" + by(rule brinputSwapFrameSubject) + then show ?thesis + using Pair by simp + qed +qed + +lemma inputSwapSubject: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and M :: 'a + and N :: 'a + and P' :: "('a, 'b, 'c) psi" + and x :: name + and y :: name + +assumes "\ \ P \M\N\ \ P'" + and "x \ P" + and "y \ P" + and "x \ \" + and "y \ \" + +shows "\ \ P \ ([(x, y)] \ M)\N\ \ P'" +proof - + from \\ \ P \M\N\ \ P'\ \x \ P\ \y \ P\ + have "([(x, y)] \ \) \ P \([(x, y)] \ M)\N\ \ P'" + by(rule inputSwapFrameSubject) + with \x \ \\ \y \ \\ show ?thesis + by simp +qed + +lemma brinputSwapSubject: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and M :: 'a + and N :: 'a + and P' :: "('a, 'b, 'c) psi" + and x :: name + and y :: name + +assumes "\ \ P \\M\N\ \ P'" + and "x \ P" + and "y \ P" + and "x \ \" + and "y \ \" + +shows "\ \ P \ \([(x, y)] \ M)\N\ \ P'" +proof - + from \\ \ P \\M\N\ \ P'\ \x \ P\ \y \ P\ + have "([(x, y)] \ \) \ P \\([(x, y)] \ M)\N\ \ P'" + by(rule brinputSwapFrameSubject) + with \x \ \\ \y \ \\ show ?thesis + by simp +qed + +lemma inputPermSubject: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and M :: 'a + and N :: 'a + and P' :: "('a, 'b, 'c) psi" + and p :: "name prm" + and Xs :: "name set" + and Ys :: "name set" + +assumes "\ \ P \M\N\ \ P'" + and S: "set p \ Xs \ Ys" + and "Xs \* P" + and "Ys \* P" + and "Xs \* \" + and "Ys \* \" + +shows "\ \ P \ (p \ M)\N\ \ P'" +proof - + from \\ \ P \M\N\ \ P'\ S \Xs \* P\ \Ys \* P\ + have "(p \ \) \ P \(p \ M)\N\ \ P'" + by(rule inputPermFrameSubject) + with \Xs \* \\ \Ys \* \\ S show ?thesis + by simp +qed + +lemma brinputPermSubject: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and M :: 'a + and N :: 'a + and P' :: "('a, 'b, 'c) psi" + and p :: "name prm" + and Xs :: "name set" + and Ys :: "name set" + +assumes "\ \ P \\M\N\ \ P'" + and S: "set p \ Xs \ Ys" + and "Xs \* P" + and "Ys \* P" + and "Xs \* \" + and "Ys \* \" + +shows "\ \ P \ \(p \ M)\N\ \ P'" +proof - + from \\ \ P \\M\N\ \ P'\ S \Xs \* P\ \Ys \* P\ + have "(p \ \) \ P \\(p \ M)\N\ \ P'" + by(rule brinputPermFrameSubject) + with \Xs \* \\ \Ys \* \\ S show ?thesis + by simp +qed + +lemma inputSwapFrame: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and M :: 'a + and N :: 'a + and P' :: "('a, 'b, 'c) psi" + and x :: name + and y :: name + +assumes "\ \ P \M\N\ \ P'" + and "x \ P" + and "y \ P" + and "x \ M" + and "y \ M" + +shows "([(x, y)] \ \) \ P \ M\N\ \ P'" +proof - + from \\ \ P \M\N\ \ P'\ \x \ P\ \y \ P\ + have "([(x, y)] \ \) \ P \([(x, y)] \ M)\N\ \ P'" + by(rule inputSwapFrameSubject) + with \x \ M\ \y \ M\ show ?thesis + by simp +qed + +lemma brinputSwapFrame: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and M :: 'a + and N :: 'a + and P' :: "('a, 'b, 'c) psi" + and x :: name + and y :: name + +assumes "\ \ P \\M\N\ \ P'" + and "x \ P" + and "y \ P" + and "x \ M" + and "y \ M" + +shows "([(x, y)] \ \) \ P \ \M\N\ \ P'" +proof - + from \\ \ P \\M\N\ \ P'\ \x \ P\ \y \ P\ + have "([(x, y)] \ \) \ P \\([(x, y)] \ M)\N\ \ P'" + by(rule brinputSwapFrameSubject) + with \x \ M\ \y \ M\ show ?thesis + by simp +qed + +lemma inputPermFrame: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and M :: 'a + and N :: 'a + and P' :: "('a, 'b, 'c) psi" + and p :: "name prm" + and Xs :: "name set" + and Ys :: "name set" + +assumes "\ \ P \M\N\ \ P'" + and S: "set p \ Xs \ Ys" + and "Xs \* P" + and "Ys \* P" + and "Xs \* M" + and "Ys \* M" + +shows "(p \ \) \ P \ M\N\ \ P'" +proof - + from \\ \ P \M\N\ \ P'\ S \Xs \* P\ \Ys \* P\ + have "(p \ \) \ P \(p \ M)\N\ \ P'" + by(rule inputPermFrameSubject) + with \Xs \* M\ \Ys \* M\ S show ?thesis + by simp +qed + +lemma brinputPermFrame: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and M :: 'a + and N :: 'a + and P' :: "('a, 'b, 'c) psi" + and p :: "name prm" + and Xs :: "name set" + and Ys :: "name set" + +assumes "\ \ P \\M\N\ \ P'" + and S: "set p \ Xs \ Ys" + and "Xs \* P" + and "Ys \* P" + and "Xs \* M" + and "Ys \* M" + +shows "(p \ \) \ P \ \M\N\ \ P'" +proof - + from \\ \ P \\M\N\ \ P'\ S \Xs \* P\ \Ys \* P\ + have "(p \ \) \ P \\(p \ M)\N\ \ P'" + by(rule brinputPermFrameSubject) + with \Xs \* M\ \Ys \* M\ S show ?thesis + by simp +qed + +lemma inputAlpha: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and M :: 'a + and N :: 'a + and P' :: "('a, 'b, 'c) psi" + and p :: "name prm" + and xvec :: "name list" + +assumes "\ \ P \M\N\ \ P'" + and "set p \ (set xvec) \ (set (p \ xvec))" + and "distinctPerm p" + and "xvec \* P" + and "(p \ xvec) \* P" + +shows "\ \ P \M\(p \ N)\ \ (p \ P')" +proof - + from \\ \ P \M\N\ \ P'\ \set p \ (set xvec) \ (set (p \ xvec))\ \xvec \* P\ \(p \ xvec) \* P\ + have "(p \ \) \ P \(p \ M)\N\ \ P'" by - (rule inputPermFrameSubject, auto) + then have "(p \ p \ \) \ (p \ P) \(p \ ((p \ M)\N\ \ P'))" by(rule eqvts) + with \distinctPerm p\ \xvec \* P\ \(p \ xvec) \* P\ \set p \ (set xvec) \ (set (p \ xvec))\ + show ?thesis by(simp add: eqvts) +qed + +lemma brinputAlpha: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and M :: 'a + and N :: 'a + and P' :: "('a, 'b, 'c) psi" + and p :: "name prm" + and xvec :: "name list" + +assumes "\ \ P \\M\N\ \ P'" + and "set p \ (set xvec) \ (set (p \ xvec))" + and "distinctPerm p" + and "xvec \* P" + and "(p \ xvec) \* P" + +shows "\ \ P \\M\(p \ N)\ \ (p \ P')" +proof - + from \\ \ P \\M\N\ \ P'\ \set p \ (set xvec) \ (set (p \ xvec))\ \xvec \* P\ \(p \ xvec) \* P\ + have "(p \ \) \ P \\(p \ M)\N\ \ P'" by - (rule brinputPermFrameSubject, auto) + then have "(p \ p \ \) \ (p \ P) \ (p \ (\(p \ M)\N\ \ P'))" by(rule eqvts) + with \distinctPerm p\ \xvec \* P\ \(p \ xvec) \* P\ \set p \ (set xvec) \ (set (p \ xvec))\ + show ?thesis by(simp add: eqvts) +qed + +lemma frameFresh[dest]: + fixes x :: name + and A\<^sub>F :: "name list" + and \\<^sub>F :: 'b + +assumes "x \ A\<^sub>F" + and "x \ \A\<^sub>F, \\<^sub>F\" + +shows "x \ \\<^sub>F" + using assms + by(simp add: frameResChainFresh) (simp add: fresh_def name_list_supp) + +lemma outputSwapFrameSubject: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and M :: 'a + and xvec :: "name list" + and N :: 'a + and x :: name + and y :: name + +assumes "\ \ P \M\\*xvec\\N\ \ P'" + and "xvec \* M" + and "x \ P" + and "y \ P" + +shows "([(x, y)] \ \) \ P \([(x, y)] \ M)\\*xvec\\N\ \ P'" + using assms +proof(nominal_induct avoiding: x y rule: outputInduct') + case cAlpha + then show ?case by(simp add: create_residual.simps boundOutputChainAlpha'') +next + case(cOutput \ M K N P x y) + from \x \ M\N\.P\ have "x \ M" by simp + from \y \ M\N\.P\ have "y \ M" by simp + from \\ \ M \ K\ have "([(x, y)] \ \) \ ([(x, y)] \ M) \ ([(x, y)] \ K)" + by(rule chanEqClosed) + with \x \ M\ \y \ M\ have "([(x, y)] \ \) \ M \ ([(x, y)] \ K)" + by(simp) + then show ?case by(rule Output) +next + case(cCase \ P M xvec N P' \ Cs x y) + from \x \ Cases Cs\ \y \ Cases Cs\ \(\, P) \ set Cs\ have "x \ \" and "x \ P" and "y \ \" and "y \ P" + by(auto dest: memFresh) + from \x \ P\ \y \ P\ have "([(x ,y)] \ \) \ P \([(x, y)] \ M)\\*xvec\\N\ \ P'" by(rule cCase) + moreover note \(\, P) \ set Cs\ + moreover from \\ \ \\ have "([(x, y)] \ \) \ ([(x, y)] \ \)" by(rule statClosed) + with \x \ \\ \y \ \\ have "([(x, y)] \ \) \ \" by simp + ultimately show ?case using \guarded P\ by(rule Case) +next + case(cPar1 \ \\<^sub>Q P M xvec N P' A\<^sub>Q Q x y) + from \x \ P \ Q\ have "x \ P" and "x \ Q" by simp+ + from \y \ P \ Q\ have "y \ P" and "y \ Q" by simp+ + from \x \ P\ \y \ P\ \\x y. \x \ P; y \ P\ \ ([(x, y)] \ (\ \ \\<^sub>Q)) \ P \([(x, y)] \ M)\\*xvec\\N\ \ P'\ + have "([(x, y)] \ \) \ ([(x, y)] \ \\<^sub>Q) \ P \([(x, y)] \ M)\\*xvec\\N\ \ P'" + by(simp add: eqvts) + + moreover from \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ have "([(x, y)] \ \A\<^sub>Q, \\<^sub>Q\) = ([(x, y)] \ (extractFrame Q))" + by simp + with \A\<^sub>Q \* x\ \x \ Q\ \A\<^sub>Q \* y\ \y \ Q\ have "\A\<^sub>Q, ([(x, y)] \ \\<^sub>Q)\ = extractFrame Q" + by(simp add: eqvts) + moreover from \A\<^sub>Q \* \\ have "([(x, y)] \ A\<^sub>Q) \* ([(x, y)] \ \)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>Q \* x\ \A\<^sub>Q \* y\ have "A\<^sub>Q \* ([(x, y)] \ \)" by simp + moreover from \A\<^sub>Q \* M\ have "([(x, y)] \ A\<^sub>Q) \* ([(x, y)] \ M)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>Q \* x\ \A\<^sub>Q \* y\ have "A\<^sub>Q \* ([(x, y)] \ M)" by simp + ultimately show ?case using \A\<^sub>Q \* P\ \A\<^sub>Q \* N\ \xvec \* Q\ \A\<^sub>Q \* xvec\ + by(force intro: Par1) +next + case(cPar2 \ \\<^sub>P Q M xvec N Q' A\<^sub>P P x y) + from \x \ P \ Q\ have "x \ P" and "x \ Q" by simp+ + from \y \ P \ Q\ have "y \ P" and "y \ Q" by simp+ + from \x \ Q\ \y \ Q\ \\x y. \x \ Q; y \ Q\ \ ([(x, y)] \ (\ \ \\<^sub>P)) \ Q \([(x, y)] \ M)\\*xvec\\N\ \ Q'\ + have "([(x, y)] \ \) \ ([(x, y)] \ \\<^sub>P) \ Q \([(x, y)] \ M)\\*xvec\\N\ \ Q'" + by(simp add: eqvts) + + moreover from \extractFrame P = \A\<^sub>P, \\<^sub>P\\ have "([(x, y)] \ \A\<^sub>P, \\<^sub>P\) = ([(x, y)] \ (extractFrame P))" + by simp + with \A\<^sub>P \* x\ \x \ P\ \A\<^sub>P \* y\ \y \ P\ have "\A\<^sub>P, ([(x, y)] \ \\<^sub>P)\ = extractFrame P" + by(simp add: eqvts) + moreover from \A\<^sub>P \* \\ have "([(x, y)] \ A\<^sub>P) \* ([(x, y)] \ \)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>P \* x\ \A\<^sub>P \* y\ have "A\<^sub>P \* ([(x, y)] \ \)" by simp + moreover from \A\<^sub>P \* M\ have "([(x, y)] \ A\<^sub>P) \* ([(x, y)] \ M)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>P \* x\ \A\<^sub>P \* y\ have "A\<^sub>P \* ([(x, y)] \ M)" by simp + ultimately show ?case using \A\<^sub>P \* Q\ \A\<^sub>P \* N\ \xvec \* P\ \A\<^sub>P \* xvec\ + by(force intro: Par2) +next + case(cOpen \ P M xvec yvec N P' z x y) + from \x \ \\z\P\ \z \ x\ have "x \ P" by(simp add: abs_fresh) + from \y \ \\z\P\ \z \ y\ have "y \ P" by(simp add: abs_fresh) + from \x \ P\ \y \ P\ \\x y. \x \ P; y \ P\ \ ([(x, y)] \ \) \ P \([(x, y)] \ M)\\*(xvec@yvec)\\N\ \ P'\ + have "([(x, y)] \ \) \ P \([(x, y)] \ M)\\*(xvec@yvec)\\N\ \ P'" by simp + moreover with \z \ \\ have "([(x, y)] \ z) \ [(x, y)] \ \" + by(simp add: pt_fresh_bij[OF pt_name_inst, OF at_name_inst]) + with \z \ x\ \z \ y\ have "z \ [(x, y)] \ \" by simp + moreover with \z \ M\ have "([(x, y)] \ z) \ [(x, y)] \ M" + by(simp add: pt_fresh_bij[OF pt_name_inst, OF at_name_inst]) + with \z \ x\ \z \ y\ have "z \ [(x, y)] \ M" by simp + ultimately show ?case using \z \ supp N\ \z \ xvec\ \z \ yvec\ + by(force intro!: Open) +next + case(cScope \ P M xvec N P' z x y) + from \x \ \\z\P\ \z \ x\ have "x \ P" by(simp add: abs_fresh) + from \y \ \\z\P\ \z \ y\ have "y \ P" by(simp add: abs_fresh) + from \x \ P\ \y \ P\ \\x y. \x \ P; y \ P\ \ ([(x, y)] \ \) \ P \([(x, y)] \ M)\\*xvec\\N\ \ P'\ + have "([(x, y)] \ \) \ P \([(x, y)] \ M)\\*xvec\\N\ \ P'" by simp + moreover with \z \ \\ have "([(x, y)] \ z) \ [(x, y)] \ \" + by(simp add: pt_fresh_bij[OF pt_name_inst, OF at_name_inst]) + with \z \ x\ \z \ y\ have "z \ [(x, y)] \ \" by simp + moreover with \z \ M\ have "([(x, y)] \ z) \ [(x, y)] \ M" + by(simp add: pt_fresh_bij[OF pt_name_inst, OF at_name_inst]) + with \z \ x\ \z \ y\ have "z \ [(x, y)] \ M" by simp + ultimately show ?case using \z \ N\ \z \ xvec\ + by(force intro!: Scope) +next + case(cBang \ P M B x y) + then show ?case by(force intro: Bang) +qed + +lemma broutputSwapFrameSubject: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and M :: 'a + and xvec :: "name list" + and N :: 'a + and x :: name + and y :: name + +assumes "\ \ P \\M\\*xvec\\N\ \ P'" + and "xvec \* M" + and "x \ P" + and "y \ P" + +shows "([(x, y)] \ \) \ P \\([(x, y)] \ M)\\*xvec\\N\ \ P'" + using assms +proof(nominal_induct avoiding: x y rule: brOutputInduct') + case cAlpha + then show ?case by(simp add: create_residual.simps boundOutputChainAlpha'') +next + case(cBrOutput \ M K N P x y) + from \x \ M\N\.P\ have "x \ M" by simp + from \y \ M\N\.P\ have "y \ M" by simp + from \\ \ M \ K\ have "([(x, y)] \ \) \ ([(x, y)] \ M) \ ([(x, y)] \ K)" + by(rule chanOutConClosed) + with \x \ M\ \y \ M\ have "([(x, y)] \ \) \ M \ ([(x, y)] \ K)" + by(simp) + then show ?case by(rule BrOutput) +next + case(cCase \ P M xvec N P' \ Cs x y) + from \x \ Cases Cs\ \y \ Cases Cs\ \(\, P) \ set Cs\ have "x \ \" and "x \ P" and "y \ \" and "y \ P" + by(auto dest: memFresh) + from \x \ P\ \y \ P\ have "([(x ,y)] \ \) \ P \\([(x, y)] \ M)\\*xvec\\N\ \ P'" by(rule cCase) + moreover note \(\, P) \ set Cs\ + moreover from \\ \ \\ have "([(x, y)] \ \) \ ([(x, y)] \ \)" by(rule statClosed) + with \x \ \\ \y \ \\ have "([(x, y)] \ \) \ \" by simp + ultimately show ?case using \guarded P\ by(rule Case) +next + case(cPar1 \ \\<^sub>Q P M xvec N P' A\<^sub>Q Q x y) + from \x \ P \ Q\ have "x \ P" and "x \ Q" by simp+ + from \y \ P \ Q\ have "y \ P" and "y \ Q" by simp+ + from \x \ P\ \y \ P\ \\x y. \x \ P; y \ P\ \ ([(x, y)] \ (\ \ \\<^sub>Q)) \ P \\([(x, y)] \ M)\\*xvec\\N\ \ P'\ + have "([(x, y)] \ \) \ ([(x, y)] \ \\<^sub>Q) \ P \\([(x, y)] \ M)\\*xvec\\N\ \ P'" + by(simp add: eqvts) + + moreover from \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ have "([(x, y)] \ \A\<^sub>Q, \\<^sub>Q\) = ([(x, y)] \ (extractFrame Q))" + by simp + with \A\<^sub>Q \* x\ \x \ Q\ \A\<^sub>Q \* y\ \y \ Q\ have "\A\<^sub>Q, ([(x, y)] \ \\<^sub>Q)\ = extractFrame Q" + by(simp add: eqvts) + moreover from \A\<^sub>Q \* \\ have "([(x, y)] \ A\<^sub>Q) \* ([(x, y)] \ \)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>Q \* x\ \A\<^sub>Q \* y\ have "A\<^sub>Q \* ([(x, y)] \ \)" by simp + moreover from \A\<^sub>Q \* M\ have "([(x, y)] \ A\<^sub>Q) \* ([(x, y)] \ M)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>Q \* x\ \A\<^sub>Q \* y\ have "A\<^sub>Q \* ([(x, y)] \ M)" by simp + ultimately show ?case using \A\<^sub>Q \* P\ \A\<^sub>Q \* N\ \xvec \* Q\ \A\<^sub>Q \* xvec\ + by(force intro: Par1) +next + case(cPar2 \ \\<^sub>P Q M xvec N Q' A\<^sub>P P x y) + from \x \ P \ Q\ have "x \ P" and "x \ Q" by simp+ + from \y \ P \ Q\ have "y \ P" and "y \ Q" by simp+ + from \x \ Q\ \y \ Q\ \\x y. \x \ Q; y \ Q\ \ ([(x, y)] \ (\ \ \\<^sub>P)) \ Q \\([(x, y)] \ M)\\*xvec\\N\ \ Q'\ + have "([(x, y)] \ \) \ ([(x, y)] \ \\<^sub>P) \ Q \\([(x, y)] \ M)\\*xvec\\N\ \ Q'" + by(simp add: eqvts) + + moreover from \extractFrame P = \A\<^sub>P, \\<^sub>P\\ have "([(x, y)] \ \A\<^sub>P, \\<^sub>P\) = ([(x, y)] \ (extractFrame P))" + by simp + with \A\<^sub>P \* x\ \x \ P\ \A\<^sub>P \* y\ \y \ P\ have "\A\<^sub>P, ([(x, y)] \ \\<^sub>P)\ = extractFrame P" + by(simp add: eqvts) + moreover from \A\<^sub>P \* \\ have "([(x, y)] \ A\<^sub>P) \* ([(x, y)] \ \)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>P \* x\ \A\<^sub>P \* y\ have "A\<^sub>P \* ([(x, y)] \ \)" by simp + moreover from \A\<^sub>P \* M\ have "([(x, y)] \ A\<^sub>P) \* ([(x, y)] \ M)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>P \* x\ \A\<^sub>P \* y\ have "A\<^sub>P \* ([(x, y)] \ M)" by simp + ultimately show ?case using \A\<^sub>P \* Q\ \A\<^sub>P \* N\ \xvec \* P\ \A\<^sub>P \* xvec\ + by(force intro: Par2) +next + case(cBrComm1 \ \\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q xvec Q' A\<^sub>Q x y) + from \x \ P \ Q\ have "x \ P" and "x \ Q" by simp+ + from \y \ P \ Q\ have "y \ P" and "y \ Q" by simp+ + + from \\ \ \\<^sub>Q \ P \ \M\N\ \ P'\ \x \ P\ \y \ P\ + have "([(x, y)] \ (\ \ \\<^sub>Q)) \ P \\([(x, y)] \ M)\N\ \ P'" + by(rule brinputSwapFrameSubject) + then have permIn: "(([(x, y)] \ \) \ ([(x, y)] \ \\<^sub>Q)) \ P \\([(x, y)] \ M)\N\ \ P'" + by(simp add: eqvts) + moreover from \x \ Q\ \y \ Q\ \\x y. \x \ Q; y \ Q\ \ ([(x, y)] \ (\ \ \\<^sub>P)) \ Q \\([(x, y)] \ M)\\*xvec\\N\ \ Q'\ + have permOut: "([(x, y)] \ \) \ ([(x, y)] \ \\<^sub>P) \ Q \\([(x, y)] \ M)\\*xvec\\N\ \ Q'" + by(simp add: eqvts) + + moreover from \extractFrame P = \A\<^sub>P, \\<^sub>P\\ have "([(x, y)] \ \A\<^sub>P, \\<^sub>P\) = ([(x, y)] \ (extractFrame P))" + by simp + with \A\<^sub>P \* x\ \x \ P\ \A\<^sub>P \* y\ \y \ P\ have "extractFrame P = \A\<^sub>P, ([(x, y)] \ \\<^sub>P)\" + by(simp add: eqvts) + moreover from \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ have "([(x, y)] \ \A\<^sub>Q, \\<^sub>Q\) = ([(x, y)] \ (extractFrame Q))" + by simp + with \A\<^sub>Q \* x\ \x \ Q\ \A\<^sub>Q \* y\ \y \ Q\ have "extractFrame Q = \A\<^sub>Q, ([(x, y)] \ \\<^sub>Q)\" + by(simp add: eqvts) + moreover from \A\<^sub>P \* \\ have "([(x, y)] \ A\<^sub>P) \* ([(x, y)] \ \)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>P \* x\ \A\<^sub>P \* y\ have "A\<^sub>P \* ([(x, y)] \ \)" by simp + moreover from \A\<^sub>Q \* \\ have "([(x, y)] \ A\<^sub>Q) \* ([(x, y)] \ \)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>Q \* x\ \A\<^sub>Q \* y\ have "A\<^sub>Q \* ([(x, y)] \ \)" by simp + moreover from \A\<^sub>P \* \\<^sub>Q\ have "([(x, y)] \ A\<^sub>P) \* ([(x, y)] \ \\<^sub>Q)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>P \* x\ \A\<^sub>P \* y\ have "A\<^sub>P \* ([(x, y)] \ \\<^sub>Q)" by simp + moreover from \A\<^sub>Q \* \\<^sub>P\ have "([(x, y)] \ A\<^sub>Q) \* ([(x, y)] \ \\<^sub>P)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>Q \* x\ \A\<^sub>Q \* y\ have "A\<^sub>Q \* ([(x, y)] \ \\<^sub>P)" by simp + moreover from \A\<^sub>P \* M\ have "([(x, y)] \ A\<^sub>P) \* ([(x, y)] \ M)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>P \* x\ \A\<^sub>P \* y\ have "A\<^sub>P \* ([(x, y)] \ M)" by simp + moreover from \A\<^sub>Q \* M\ have "([(x, y)] \ A\<^sub>Q) \* ([(x, y)] \ M)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>Q \* x\ \A\<^sub>Q \* y\ have "A\<^sub>Q \* ([(x, y)] \ M)" by simp + moreover from \xvec \* \\ have "([(x, y)] \ xvec) \* ([(x, y)] \ \)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \xvec \* x\ \xvec \* y\ have "xvec \* ([(x, y)] \ \)" by simp + moreover from \xvec \* \\<^sub>Q\ have "([(x, y)] \ xvec) \* ([(x, y)] \ \\<^sub>Q)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \xvec \* x\ \xvec \* y\ have "xvec \* ([(x, y)] \ \\<^sub>Q)" by simp + moreover from \xvec \* \\<^sub>P\ have "([(x, y)] \ xvec) \* ([(x, y)] \ \\<^sub>P)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \xvec \* x\ \xvec \* y\ have "xvec \* ([(x, y)] \ \\<^sub>P)" by simp + moreover from \xvec \* M\ have "([(x, y)] \ xvec) \* ([(x, y)] \ M)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \xvec \* x\ \xvec \* y\ have "xvec \* ([(x, y)] \ M)" by simp + + moreover note \distinct A\<^sub>P\ \distinct A\<^sub>Q\ \A\<^sub>P \* P\ \A\<^sub>P \* N\ \A\<^sub>P \* P'\ \A\<^sub>P \* Q\ \A\<^sub>P \* Q'\ \A\<^sub>P \* A\<^sub>Q\ + \A\<^sub>P \* xvec\ \A\<^sub>Q \* P\ \A\<^sub>Q \* N\ \A\<^sub>Q \* P'\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* Q'\ \A\<^sub>Q \* xvec\ + \distinct xvec\ \xvec \* P\ \xvec \* Q\ + ultimately show ?case + by(simp add: semantics.cBrComm1) +next + case(cBrComm2 \ \\<^sub>Q P M xvec N P' A\<^sub>P \\<^sub>P Q Q' A\<^sub>Q x y) + from \x \ P \ Q\ have "x \ P" and "x \ Q" by simp+ + from \y \ P \ Q\ have "y \ P" and "y \ Q" by simp+ + + from \\ \ \\<^sub>P \ Q \ \M\N\ \ Q'\ \x \ Q\ \y \ Q\ + have "([(x, y)] \ (\ \ \\<^sub>P)) \ Q \\([(x, y)] \ M)\N\ \ Q'" + by(rule brinputSwapFrameSubject) + then have permIn: "(([(x, y)] \ \) \ ([(x, y)] \ \\<^sub>P)) \ Q \\([(x, y)] \ M)\N\ \ Q'" + by(simp add: eqvts) + moreover from \x \ P\ \y \ P\ \\x y. \x \ P; y \ P\ \ ([(x, y)] \ (\ \ \\<^sub>Q)) \ P \\([(x, y)] \ M)\\*xvec\\N\ \ P'\ + have permOut: "([(x, y)] \ \) \ ([(x, y)] \ \\<^sub>Q) \ P \\([(x, y)] \ M)\\*xvec\\N\ \ P'" + by(simp add: eqvts) + + moreover from \extractFrame P = \A\<^sub>P, \\<^sub>P\\ have "([(x, y)] \ \A\<^sub>P, \\<^sub>P\) = ([(x, y)] \ (extractFrame P))" + by simp + with \A\<^sub>P \* x\ \x \ P\ \A\<^sub>P \* y\ \y \ P\ have "extractFrame P = \A\<^sub>P, ([(x, y)] \ \\<^sub>P)\" + by(simp add: eqvts) + moreover from \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ have "([(x, y)] \ \A\<^sub>Q, \\<^sub>Q\) = ([(x, y)] \ (extractFrame Q))" + by simp + with \A\<^sub>Q \* x\ \x \ Q\ \A\<^sub>Q \* y\ \y \ Q\ have "extractFrame Q = \A\<^sub>Q, ([(x, y)] \ \\<^sub>Q)\" + by(simp add: eqvts) + moreover from \A\<^sub>P \* \\ have "([(x, y)] \ A\<^sub>P) \* ([(x, y)] \ \)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>P \* x\ \A\<^sub>P \* y\ have "A\<^sub>P \* ([(x, y)] \ \)" by simp + moreover from \A\<^sub>Q \* \\ have "([(x, y)] \ A\<^sub>Q) \* ([(x, y)] \ \)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>Q \* x\ \A\<^sub>Q \* y\ have "A\<^sub>Q \* ([(x, y)] \ \)" by simp + moreover from \A\<^sub>P \* \\<^sub>Q\ have "([(x, y)] \ A\<^sub>P) \* ([(x, y)] \ \\<^sub>Q)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>P \* x\ \A\<^sub>P \* y\ have "A\<^sub>P \* ([(x, y)] \ \\<^sub>Q)" by simp + moreover from \A\<^sub>Q \* \\<^sub>P\ have "([(x, y)] \ A\<^sub>Q) \* ([(x, y)] \ \\<^sub>P)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>Q \* x\ \A\<^sub>Q \* y\ have "A\<^sub>Q \* ([(x, y)] \ \\<^sub>P)" by simp + moreover from \A\<^sub>P \* M\ have "([(x, y)] \ A\<^sub>P) \* ([(x, y)] \ M)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>P \* x\ \A\<^sub>P \* y\ have "A\<^sub>P \* ([(x, y)] \ M)" by simp + moreover from \A\<^sub>Q \* M\ have "([(x, y)] \ A\<^sub>Q) \* ([(x, y)] \ M)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>Q \* x\ \A\<^sub>Q \* y\ have "A\<^sub>Q \* ([(x, y)] \ M)" by simp + moreover from \xvec \* \\ have "([(x, y)] \ xvec) \* ([(x, y)] \ \)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \xvec \* x\ \xvec \* y\ have "xvec \* ([(x, y)] \ \)" by simp + moreover from \xvec \* \\<^sub>Q\ have "([(x, y)] \ xvec) \* ([(x, y)] \ \\<^sub>Q)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \xvec \* x\ \xvec \* y\ have "xvec \* ([(x, y)] \ \\<^sub>Q)" by simp + moreover from \xvec \* \\<^sub>P\ have "([(x, y)] \ xvec) \* ([(x, y)] \ \\<^sub>P)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \xvec \* x\ \xvec \* y\ have "xvec \* ([(x, y)] \ \\<^sub>P)" by simp + moreover from \xvec \* M\ have "([(x, y)] \ xvec) \* ([(x, y)] \ M)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \xvec \* x\ \xvec \* y\ have "xvec \* ([(x, y)] \ M)" by simp + + moreover note \distinct A\<^sub>P\ \distinct A\<^sub>Q\ \A\<^sub>P \* P\ \A\<^sub>P \* N\ \A\<^sub>P \* P'\ \A\<^sub>P \* Q\ \A\<^sub>P \* Q'\ \A\<^sub>P \* A\<^sub>Q\ + \A\<^sub>P \* xvec\ \A\<^sub>Q \* P\ \A\<^sub>Q \* N\ \A\<^sub>Q \* P'\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* Q'\ \A\<^sub>Q \* xvec\ + \distinct xvec\ \xvec \* P\ \xvec \* Q\ + ultimately show ?case + by(simp add: semantics.cBrComm2) +next + case(cBrOpen \ P M xvec yvec N P' z x y) + from \x \ \\z\P\ \z \ x\ have "x \ P" by(simp add: abs_fresh) + from \y \ \\z\P\ \z \ y\ have "y \ P" by(simp add: abs_fresh) + from \x \ P\ \y \ P\ \\x y. \x \ P; y \ P\ \ ([(x, y)] \ \) \ P \\([(x, y)] \ M)\\*(xvec@yvec)\\N\ \ P'\ + have "([(x, y)] \ \) \ P \\([(x, y)] \ M)\\*(xvec@yvec)\\N\ \ P'" by simp + moreover with \z \ \\ have "([(x, y)] \ z) \ [(x, y)] \ \" + by(simp add: pt_fresh_bij[OF pt_name_inst, OF at_name_inst]) + with \z \ x\ \z \ y\ have "z \ [(x, y)] \ \" by simp + moreover with \z \ M\ have "([(x, y)] \ z) \ [(x, y)] \ M" + by(simp add: pt_fresh_bij[OF pt_name_inst, OF at_name_inst]) + with \z \ x\ \z \ y\ have "z \ [(x, y)] \ M" by simp + ultimately show ?case using \z \ supp N\ \z \ xvec\ \z \ yvec\ + by(force intro: BrOpen) +next + case(cScope \ P M xvec N P' z x y) + from \x \ \\z\P\ \z \ x\ have "x \ P" by(simp add: abs_fresh) + from \y \ \\z\P\ \z \ y\ have "y \ P" by(simp add: abs_fresh) + from \x \ P\ \y \ P\ \\x y. \x \ P; y \ P\ \ ([(x, y)] \ \) \ P \\([(x, y)] \ M)\\*xvec\\N\ \ P'\ + have "([(x, y)] \ \) \ P \\([(x, y)] \ M)\\*xvec\\N\ \ P'" by simp + moreover with \z \ \\ have "([(x, y)] \ z) \ [(x, y)] \ \" + by(simp add: pt_fresh_bij[OF pt_name_inst, OF at_name_inst]) + with \z \ x\ \z \ y\ have "z \ [(x, y)] \ \" by simp + moreover with \z \ M\ have "([(x, y)] \ z) \ [(x, y)] \ M" + by(simp add: pt_fresh_bij[OF pt_name_inst, OF at_name_inst]) + with \z \ x\ \z \ y\ have "z \ [(x, y)] \ M" by simp + ultimately show ?case using \z \ N\ \z \ xvec\ + by(force intro: Scope) +next + case(cBang \ P M B x y) + then show ?case by(force intro: Bang) +qed + +lemma outputPermFrameSubject: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and M :: 'a + and xvec :: "name list" + and N :: 'a + and P' :: "('a, 'b, 'c) psi" + and p :: "name prm" + and yvec :: "name list" + and zvec :: "name list" + +assumes "\ \ P \M\\*xvec\\N\ \ P'" + and S: "set p \ set yvec \ set zvec" + and "yvec \* P" + and "zvec \* P" + +shows "(p \ \) \ P \(p \ M)\\*xvec\\N\ \ P'" +proof - + { + fix xvec N P' Xs YS + assume "\ \ P \M\\*xvec\\N\ \ P'" and "xvec \* M" and "xvec \* yvec" and "xvec \* zvec" + have "(p \ \) \ P \(p \ M)\\*xvec\\N\ \ P'" using S + proof(induct p) + case Nil + from \\ \ P \M\\*xvec\\N\ \ P'\ + show ?case by simp + next + case(Cons a p) + from \set(a#p) \ set yvec \ set zvec\ have "set p \ set yvec \ set zvec" by auto + then have Trans: "(p \ \) \ P \(p \ M)\\*xvec\\N\ \ P'" by(rule Cons) + show ?case + proof(cases a) + case (Pair x y) + note Trans + moreover from \xvec \* yvec\ \xvec \* zvec\ \set p \ set yvec \ set zvec\ \xvec \* M\ have "xvec \* (p \ M)" + by(simp add: freshChainSimps) + moreover have "x \ set yvec" and "y \ set zvec" + using \set (a # p) \ set yvec \ set zvec\ Pair by auto + with \yvec \* P\ \zvec \* P\ have "x \ P" and "y \ P" + by(auto simp add: fresh_star_def) + ultimately have "([(x, y)] \ p \ \) \ P \([(x, y)] \ p \ M)\\*xvec\\N\ \ P'" + by(rule outputSwapFrameSubject) + then show ?thesis + using Pair by simp + qed + qed + } + note Goal = this + obtain q::"name prm" where "(q \ xvec) \* yvec" and "(q \ xvec) \* zvec" and "(q \ xvec) \* xvec" + and "(q \ xvec) \* N" and "(q \ xvec) \* P'" and "(q \ xvec) \* M" + and Sq: "(set q) \ (set xvec) \ (set(q \ xvec))" + by(rule name_list_avoiding[where xvec=xvec and c="(P, xvec, yvec, zvec, N, M, P')"]) auto + with \\ \ P \M\\*xvec\\N\ \ P'\ have "\ \ P \M\\*(q \ xvec)\\(q \ N)\ \ (q \ P')" + by(simp add: boundOutputChainAlpha'' residualInject) + then have "(p \ \) \ P \(p \ M)\\*(q \ xvec)\\(q \ N)\ \ (q \ P')" + using \(q \ xvec) \* M\ \(q \ xvec) \* yvec\ \(q \ xvec) \* zvec\ + by(rule Goal) + with \(q \ xvec) \* N\ \(q \ xvec) \* P'\ Sq show ?thesis + by(simp add: boundOutputChainAlpha'' residualInject) +qed + +lemma broutputPermFrameSubject: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and M :: 'a + and xvec :: "name list" + and N :: 'a + and P' :: "('a, 'b, 'c) psi" + and p :: "name prm" + and yvec :: "name list" + and zvec :: "name list" + +assumes "\ \ P \\M\\*xvec\\N\ \ P'" + and S: "set p \ set yvec \ set zvec" + and "yvec \* P" + and "zvec \* P" + +shows "(p \ \) \ P \\(p \ M)\\*xvec\\N\ \ P'" +proof - + { + fix xvec N P' Xs YS + assume "\ \ P \\M\\*xvec\\N\ \ P'" and "xvec \* M" and "xvec \* yvec" and "xvec \* zvec" + have "(p \ \) \ P \\(p \ M)\\*xvec\\N\ \ P'" using S + proof(induct p) + case Nil + from \\ \ P \\M\\*xvec\\N\ \ P'\ + show ?case by simp + next + case(Cons a p) + from \set(a#p) \ set yvec \ set zvec\ have "set p \ set yvec \ set zvec" by auto + then have Trans: "(p \ \) \ P \\(p \ M)\\*xvec\\N\ \ P'" by(rule Cons) + show ?case + proof(cases a) + case (Pair x y) + note Trans + moreover from \xvec \* yvec\ \xvec \* zvec\ \set p \ set yvec \ set zvec\ \xvec \* M\ have "xvec \* (p \ M)" + by(simp add: freshChainSimps) + moreover have "x \ set yvec" and "y \ set zvec" + using \set (a # p) \ set yvec \ set zvec\ Pair by auto + with \yvec \* P\ \zvec \* P\ have "x \ P" and "y \ P" + by(auto simp add: fresh_star_def) + ultimately have "([(x, y)] \ p \ \) \ P \\([(x, y)] \ p \ M)\\*xvec\\N\ \ P'" + by(rule broutputSwapFrameSubject) + then show ?thesis + using Pair by simp + qed + qed + } + note Goal = this + obtain q::"name prm" where "(q \ xvec) \* yvec" and "(q \ xvec) \* zvec" and "(q \ xvec) \* xvec" + and "(q \ xvec) \* N" and "(q \ xvec) \* P'" and "(q \ xvec) \* M" + and Sq: "(set q) \ (set xvec) \ (set(q \ xvec))" + by(rule name_list_avoiding[where xvec=xvec and c="(P, xvec, yvec, zvec, N, M, P')"]) auto + with \\ \ P \\M\\*xvec\\N\ \ P'\ have "\ \ P \\M\\*(q \ xvec)\\(q \ N)\ \ (q \ P')" + by(simp add: boundOutputChainAlpha'' residualInject) + then have "(p \ \) \ P \\(p \ M)\\*(q \ xvec)\\(q \ N)\ \ (q \ P')" + using \(q \ xvec) \* M\ \(q \ xvec) \* yvec\ \(q \ xvec) \* zvec\ + by(rule Goal) + with \(q \ xvec) \* N\ \(q \ xvec) \* P'\ Sq show ?thesis + by(simp add: boundOutputChainAlpha'' residualInject) +qed + +lemma outputSwapSubject: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and M :: 'a + and B :: "('a, 'b, 'c) boundOutput" + and x :: name + and y :: name + +assumes "\ \ P \M\\*xvec\\N\ \ P'" + and "xvec \* M" + and "x \ P" + and "y \ P" + and "x \ \" + and "y \ \" + +shows "\ \ P \([(x, y)] \ M)\\*xvec\\N\ \ P'" +proof - + from \\ \ P \M\\*xvec\\N\ \ P'\ \xvec \* M\ \x \ P\ \y \ P\ + have "([(x, y)] \ \) \ P \([(x, y)] \ M)\\*xvec\\N\ \ P'" + by(rule outputSwapFrameSubject) + with \x \ \\ \y \ \\ show ?thesis + by simp +qed + +lemma broutputSwapSubject: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and M :: 'a + and B :: "('a, 'b, 'c) boundOutput" + and x :: name + and y :: name + +assumes "\ \ P \\M\\*xvec\\N\ \ P'" + and "xvec \* M" + and "x \ P" + and "y \ P" + and "x \ \" + and "y \ \" + +shows "\ \ P \\([(x, y)] \ M)\\*xvec\\N\ \ P'" +proof - + from \\ \ P \\M\\*xvec\\N\ \ P'\ \xvec \* M\ \x \ P\ \y \ P\ + have "([(x, y)] \ \) \ P \\([(x, y)] \ M)\\*xvec\\N\ \ P'" + by(rule broutputSwapFrameSubject) + with \x \ \\ \y \ \\ show ?thesis + by simp +qed + +lemma outputPermSubject: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and M :: 'a + and B :: "('a, 'b, 'c) boundOutput" + and p :: "name prm" + and yvec :: "name list" + and zvec :: "name list" + +assumes "\ \ P \M\\*xvec\\N\ \ P'" + and S: "set p \ set yvec \ set zvec" + and "yvec \* P" + and "zvec \* P" + and "yvec \* \" + and "zvec \* \" + +shows "\ \ P \(p \ M)\\*xvec\\N\ \ P'" +proof - + from assms have "(p \ \) \ P \(p \ M)\\*xvec\\N\ \ P'" + by(metis outputPermFrameSubject) + with S \yvec \* \\ \zvec \* \\ show ?thesis + by simp +qed + +lemma broutputPermSubject: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and M :: 'a + and B :: "('a, 'b, 'c) boundOutput" + and p :: "name prm" + and yvec :: "name list" + and zvec :: "name list" + +assumes "\ \ P \\M\\*xvec\\N\ \ P'" + and S: "set p \ set yvec \ set zvec" + and "yvec \* P" + and "zvec \* P" + and "yvec \* \" + and "zvec \* \" + +shows "\ \ P \\(p \ M)\\*xvec\\N\ \ P'" +proof - + from assms have "(p \ \) \ P \\(p \ M)\\*xvec\\N\ \ P'" + by(metis broutputPermFrameSubject) + with S \yvec \* \\ \zvec \* \\ show ?thesis + by simp +qed + +lemma outputSwapFrame: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and M :: 'a + and B :: "('a, 'b, 'c) boundOutput" + and x :: name + and y :: name + +assumes "\ \ P \M\\*xvec\\N\ \ P'" + and "xvec \* M" + and "x \ P" + and "y \ P" + and "x \ M" + and "y \ M" + +shows "([(x, y)] \ \) \ P \M\\*xvec\\N\ \ P'" +proof - + from \\ \ P \M\\*xvec\\N\ \ P'\ \xvec \* M\ \x \ P\ \y \ P\ + have "([(x, y)] \ \) \ P \([(x, y)] \ M)\\*xvec\\N\ \ P'" + by(rule outputSwapFrameSubject) + with \x \ M\ \y \ M\ show ?thesis + by simp +qed + +lemma broutputSwapFrame: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and M :: 'a + and B :: "('a, 'b, 'c) boundOutput" + and x :: name + and y :: name + +assumes "\ \ P \\M\\*xvec\\N\ \ P'" + and "xvec \* M" + and "x \ P" + and "y \ P" + and "x \ M" + and "y \ M" + +shows "([(x, y)] \ \) \ P \\M\\*xvec\\N\ \ P'" +proof - + from \\ \ P \\M\\*xvec\\N\ \ P'\ \xvec \* M\ \x \ P\ \y \ P\ + have "([(x, y)] \ \) \ P \\([(x, y)] \ M)\\*xvec\\N\ \ P'" + by(rule broutputSwapFrameSubject) + with \x \ M\ \y \ M\ show ?thesis + by simp +qed + +lemma outputPermFrame: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and M :: 'a + and B :: "('a, 'b, 'c) boundOutput" + and p :: "name prm" + and yvec :: "name list" + and zvec :: "name list" + +assumes "\ \ P \M\\*xvec\\N\ \ P'" + and S: "set p \ set yvec \ set zvec" + and "yvec \* P" + and "zvec \* P" + and "yvec \* M" + and "zvec \* M" + +shows "(p \ \) \ P \M\\*xvec\\N\ \ P'" +proof - + from assms have "(p \ \) \ P \(p \ M)\\*xvec\\N\ \ P'" + by(metis outputPermFrameSubject) + with S \yvec \* M\ \zvec \* M\ show ?thesis + by simp +qed + +lemma broutputPermFrame: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and M :: 'a + and B :: "('a, 'b, 'c) boundOutput" + and p :: "name prm" + and yvec :: "name list" + and zvec :: "name list" + +assumes "\ \ P \\M\\*xvec\\N\ \ P'" + and S: "set p \ set yvec \ set zvec" + and "yvec \* P" + and "zvec \* P" + and "yvec \* M" + and "zvec \* M" + +shows "(p \ \) \ P \\M\\*xvec\\N\ \ P'" +proof - + from assms have "(p \ \) \ P \\(p \ M)\\*xvec\\N\ \ P'" + by(metis broutputPermFrameSubject) + with S \yvec \* M\ \zvec \* M\ show ?thesis + by simp +qed + +lemma Comm1: + fixes \ :: 'b + and \\<^sub>Q :: 'b + and P :: "('a, 'b, 'c) psi" + and M :: 'a + and N :: 'a + and P' :: "('a, 'b, 'c) psi" + and A\<^sub>P :: "name list" + and \\<^sub>P :: 'b + and Q :: "('a, 'b, 'c) psi" + and K :: 'a + and xvec :: "name list" + and Q' :: "('a, 'b, 'c) psi" + and A\<^sub>Q :: "name list" + +assumes "\ \ \\<^sub>Q \ P \M\N\ \ P'" + and "extractFrame P = \A\<^sub>P, \\<^sub>P\" + and "\ \ \\<^sub>P \ Q \K\\*xvec\\N\ \ Q'" + and "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" + and "\ \ \\<^sub>P \ \\<^sub>Q \ M \ K" + and "A\<^sub>P \* \" + and "A\<^sub>P \* P" + and "A\<^sub>P \* Q" + and "A\<^sub>P \* M" + and "A\<^sub>P \* A\<^sub>Q" + and "A\<^sub>Q \* \" + and "A\<^sub>Q \* P" + and "A\<^sub>Q \* Q" + and "A\<^sub>Q \* K" + and "xvec \* P" + +shows "\ \ P \ Q \\ \ \\*xvec\(P' \ Q')" +proof - + { + fix \ :: 'b + and \\<^sub>Q :: 'b + and P :: "('a, 'b, 'c) psi" + and M :: 'a + and N :: 'a + and P' :: "('a, 'b, 'c) psi" + and A\<^sub>P :: "name list" + and \\<^sub>P :: 'b + and Q :: "('a, 'b, 'c) psi" + and K :: 'a + and xvec :: "name list" + and Q' :: "('a, 'b, 'c) psi" + and A\<^sub>Q :: "name list" + + assume "\ \ \\<^sub>Q \ P \M\N\ \ P'" + and "extractFrame P = \A\<^sub>P, \\<^sub>P\" + and "distinct A\<^sub>P" + and "\ \ \\<^sub>P \ Q \K\\*xvec\\N\ \ Q'" + and "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" + and "distinct A\<^sub>Q" + and "\ \ \\<^sub>P \ \\<^sub>Q \ M \ K" + and "A\<^sub>P \* \" + and "A\<^sub>P \* P" + and "A\<^sub>P \* Q" + and "A\<^sub>P \* M" + and "A\<^sub>P \* A\<^sub>Q" + and "A\<^sub>Q \* \" + and "A\<^sub>Q \* P" + and "A\<^sub>Q \* Q" + and "A\<^sub>Q \* K" + and "xvec \* P" + + have "\ \ P \ Q \\ \ \\*xvec\(P' \ Q')" + proof - + + obtain r::"name prm" where "(r \ xvec) \* \" and "(r \ xvec) \* P" and "(r \ xvec) \* Q" and "(r \ xvec) \* M" + and "(r \ xvec) \* K" and "(r \ xvec) \* N" and "(r \ xvec) \* A\<^sub>P" and "(r \ xvec) \* A\<^sub>Q" + and "(r \ xvec) \* P'" and "(r \ xvec) \* Q'" and "(r \ xvec) \* \\<^sub>P" and "(r \ xvec) \* \\<^sub>Q" + and Sr: "(set r) \ (set xvec) \ (set(r \ xvec))" and "distinctPerm r" + by(rule name_list_avoiding[where xvec=xvec and c="(\, P, Q, M, K, N, A\<^sub>P, A\<^sub>Q, \\<^sub>P, \\<^sub>Q, P', Q')"]) + (auto simp add: eqvts fresh_star_prod) + obtain q::"name prm" where "(q \ A\<^sub>Q) \* \" and "(q \ A\<^sub>Q) \* P" and "(q \ A\<^sub>Q) \* Q" and "(q \ A\<^sub>Q) \* K" + and "(q \ A\<^sub>Q) \* (r \ N)" and "(q \ A\<^sub>Q) \* (r \ xvec)" and "(q \ A\<^sub>Q) \* (r \ Q')" + and "(q \ A\<^sub>Q) \* (r \ P')" and "(q \ A\<^sub>Q) \* \\<^sub>P" and "(q \ A\<^sub>Q) \* A\<^sub>P" and "(q \ A\<^sub>Q) \* \\<^sub>Q" + and Sq: "set q \ set A\<^sub>Q \ set(q \ A\<^sub>Q)" + by(rule name_list_avoiding[where xvec=A\<^sub>Q and c="(\, P, Q, K, r \ N, r \ xvec, \\<^sub>Q, A\<^sub>P, \\<^sub>P, r \ Q', r \ P')"]) + (auto simp add: eqvts fresh_star_prod) + obtain p::"name prm" where "(p \ A\<^sub>P) \* \" and "(p \ A\<^sub>P) \* P" and "(p \ A\<^sub>P) \* Q" and "(p \ A\<^sub>P) \* M" + and "(p \ A\<^sub>P) \* (r \ N)" and "(p \ A\<^sub>P) \* (r \ xvec)" and "(p \ A\<^sub>P) \* (r \ Q')" + and "(p \ A\<^sub>P) \* (r \ P')" and "(p \ A\<^sub>P) \* \\<^sub>P" and "(p \ A\<^sub>P) \* \\<^sub>Q" and "(p \ A\<^sub>P) \* A\<^sub>Q" + and "(p \ A\<^sub>P) \* (q \ A\<^sub>Q)" and Sp: "(set p) \ (set A\<^sub>P) \ (set(p \ A\<^sub>P))" + by(rule name_list_avoiding[where xvec=A\<^sub>P and c="(\, P, Q, M, r \ N, r \ xvec, A\<^sub>Q, q \ A\<^sub>Q, \\<^sub>Q, \\<^sub>P, r \ Q', r \ P')"]) + (auto simp add: eqvts fresh_star_prod) + have FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" by fact + have FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" by fact + + from \A\<^sub>P \* Q\ FrQ \A\<^sub>P \* A\<^sub>Q\ have "A\<^sub>P \* \\<^sub>Q" + by(force dest: extractFrameFreshChain) + from \A\<^sub>Q \* P\ FrP \A\<^sub>P \* A\<^sub>Q\ have "A\<^sub>Q \* \\<^sub>P" + by(force dest: extractFrameFreshChain) + from \(r \ xvec) \* A\<^sub>P\ \(p \ A\<^sub>P) \* (r \ xvec)\ \(r \ xvec) \* A\<^sub>P\ Sp have "(r \ xvec) \* (p \ A\<^sub>P)" + by(simp add: freshChainSimps) + + from \\ \ \\<^sub>Q \ P \M\N\ \ P'\ Sr \distinctPerm r\ \xvec \* P\ \(r \ xvec) \* P\ + have "\ \ \\<^sub>Q \ P \M\(r \ N)\ \ (r \ P')" + by(rule inputAlpha) + then have "(q \ (\ \ \\<^sub>Q)) \ P \(q \ M)\(r \ N)\ \ (r \ P')" using Sq \A\<^sub>Q \* P\ \(q \ A\<^sub>Q) \* P\ + by - (rule inputPermFrameSubject, (assumption | simp)+) + then have PTrans: "\ \ (q \ \\<^sub>Q) \ P \(q \ M)\(r \ N)\ \ (r \ P')" using Sq \A\<^sub>Q \* \\ \(q \ A\<^sub>Q) \* \\ + by(simp add: eqvts) + + moreover from \extractFrame P = \A\<^sub>P, \\<^sub>P\\ Sp \(p \ A\<^sub>P) \* \\<^sub>P\ + have FrP: "extractFrame P = \(p \ A\<^sub>P), (p \ \\<^sub>P)\" + by(simp add: frameChainAlpha) + moreover from \distinct A\<^sub>P\ have "distinct(p \ A\<^sub>P)" by simp + + moreover from \\ \ \\<^sub>P \ Q \K\\*xvec\\N\ \ Q'\ Sr \(r \ xvec) \* N\ \(r \ xvec) \* Q'\ + have "\ \ \\<^sub>P \ Q \K\\*(r \ xvec)\\(r \ N)\ \ (r \ Q')" + by(simp add: boundOutputChainAlpha'' create_residual.simps) + then have "(p \ (\ \ \\<^sub>P)) \ Q \(p \ K)\\*(r \ xvec)\\(r \ N)\ \ (r \ Q')" using Sp \A\<^sub>P \* Q\ \(p \ A\<^sub>P) \* Q\ \(r \ xvec) \* K\ \(r \ xvec) \* A\<^sub>P\ \(r \ xvec) \* (p \ A\<^sub>P)\ + by(fastforce intro: outputPermFrameSubject) + then have QTrans: "\ \ (p \ \\<^sub>P) \ Q \(p \ K)\\*(r \ xvec)\\(r \ N)\ \ (r \ Q')" using Sp \A\<^sub>P \* \\ \(p \ A\<^sub>P) \* \\ + by(simp add: eqvts) + moreover then have "distinct(r \ xvec)" by(force dest: boundOutputDistinct) + moreover from \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ Sq \(q \ A\<^sub>Q) \* \\<^sub>Q\ + have FrQ: "extractFrame Q = \(q \ A\<^sub>Q), (q \ \\<^sub>Q)\" + by(simp add: frameChainAlpha) + moreover from \distinct A\<^sub>Q\ have "distinct(q \ A\<^sub>Q)" by simp + + moreover from \\ \ \\<^sub>P \ \\<^sub>Q \ M \ K\ have "(p \ q \ (\ \ \\<^sub>P \ \\<^sub>Q)) \ (p \ q \ M) \ (p \ q \ K)" + by(metis chanEqClosed) + with \A\<^sub>P \* \\ \(p \ A\<^sub>P) \* \\ \A\<^sub>Q \* \\ \(q \ A\<^sub>Q) \* \\ \A\<^sub>Q \* \\<^sub>P\ \(q \ A\<^sub>Q) \* \\<^sub>P\ + \A\<^sub>P \* \\<^sub>Q\ \(p \ A\<^sub>P) \* \\<^sub>Q\ \A\<^sub>P \* M\ \(p \ A\<^sub>P) \* M\ \(q \ A\<^sub>Q) \* A\<^sub>P\ \(p \ A\<^sub>P) \* (q \ A\<^sub>Q)\ + \A\<^sub>Q \* K\ \(q \ A\<^sub>Q) \* K\ \A\<^sub>P \* A\<^sub>Q\ \(p \ A\<^sub>P) \* A\<^sub>Q\ Sp Sq + have "\ \ (p \ \\<^sub>P) \ (q \ \\<^sub>Q) \ (q \ M) \ (p \ K)" by(simp add: eqvts freshChainSimps) + moreover note \(p \ A\<^sub>P) \* \\ + moreover from \(p \ A\<^sub>P) \* A\<^sub>Q\ \(p \ A\<^sub>P) \* (q \ A\<^sub>Q)\ \(p \ A\<^sub>P) \* \\<^sub>Q\ Sq have "(p \ A\<^sub>P) \* (q \ \\<^sub>Q)" + by(simp add: freshChainSimps) + moreover note \(p \ A\<^sub>P) \* P\ + moreover from \(p \ A\<^sub>P) \* A\<^sub>Q\ \(p \ A\<^sub>P) \* (q \ A\<^sub>Q)\ \(p \ A\<^sub>P) \* M\ Sq have "(p \ A\<^sub>P) \* (q \ M)" + by(simp add: freshChainSimps) + moreover note \(p \ A\<^sub>P) \* (r \ N)\ \(p \ A\<^sub>P) \* (r \ P')\ \(p \ A\<^sub>P) \* Q\ \(p \ A\<^sub>P) \* (r \ Q')\ \(p \ A\<^sub>P) \* (q \ A\<^sub>Q)\ + \(p \ A\<^sub>P) \* (r \ xvec)\ \(q \ A\<^sub>Q) \* \\ + moreover from \(q \ A\<^sub>Q) \* A\<^sub>P\ \(p \ A\<^sub>P) \* A\<^sub>Q\ \(q \ A\<^sub>Q) \* \\<^sub>P\ \(p \ A\<^sub>P) \* (q \ A\<^sub>Q)\ Sp Sq have "(q \ A\<^sub>Q) \* (p \ \\<^sub>P)" + by(simp add: freshChainSimps) + moreover note \(q \ A\<^sub>Q) \* P\ \(q \ A\<^sub>Q) \* (r \ N)\\(q \ A\<^sub>Q) \* (r \ P')\ \(q \ A\<^sub>Q) \* Q\ + moreover from \(q \ A\<^sub>Q) \* A\<^sub>P\ \(p \ A\<^sub>P) \* A\<^sub>Q\ \(q \ A\<^sub>Q) \* K\ \(p \ A\<^sub>P) \* (q \ A\<^sub>Q)\ Sp Sq have "(q \ A\<^sub>Q) \* (p \ K)" + by(simp add: freshChainSimps) + moreover note \(q \ A\<^sub>Q) \* (r \ Q')\ \(q \ A\<^sub>Q) \* (r \ xvec)\ \(r \ xvec) \* \\ + moreover from \(r \ xvec) \* A\<^sub>P\ \(p \ A\<^sub>P) \* (r \ xvec)\ \(r \ xvec) \* \\<^sub>P\ Sp have "(r \ xvec) \* (p \ \\<^sub>P)" + by(simp add: freshChainSimps) + moreover from \(r \ xvec) \* A\<^sub>Q\ \(q \ A\<^sub>Q) \* (r \ xvec)\ \(r \ xvec) \* \\<^sub>Q\ Sq have "(r \ xvec) \* (q \ \\<^sub>Q)" + by(simp add: freshChainSimps) + moreover note \(r \ xvec) \* P\ + moreover from \(r \ xvec) \* A\<^sub>Q\ \(q \ A\<^sub>Q) \* (r \ xvec)\ \(r \ xvec) \* M\ Sq have "(r \ xvec) \* (q \ M)" + by(simp add: freshChainSimps) + moreover note \(r \ xvec) \* Q\ + moreover from \(r \ xvec) \* A\<^sub>P\ \(p \ A\<^sub>P) \* (r \ xvec)\ \(r \ xvec) \* K\ Sp have "(r \ xvec) \* (p \ K)" + by(simp add: freshChainSimps) + ultimately have "\ \ P \ Q \\ \ \\*(r \ xvec)\((r \ P') \ (r \ Q'))" + by - (rule cComm1) + with \(r \ xvec) \* P'\ \(r \ xvec) \* Q'\ Sr + show ?thesis + by(subst resChainAlpha) auto + qed + } + note Goal = this + note \\ \ \\<^sub>Q \ P \M\N\ \ P'\ \\ \ \\<^sub>P \ Q \K\\*xvec\\N\ \ Q'\ \\ \ \\<^sub>P \ \\<^sub>Q \ M \ K\ + moreover from \extractFrame P = \A\<^sub>P, \\<^sub>P\\ \A\<^sub>P \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* Q\ \A\<^sub>P \* M\ \A\<^sub>P \* A\<^sub>Q\ + obtain A\<^sub>P' where "extractFrame P = \A\<^sub>P', \\<^sub>P\" and "distinct A\<^sub>P'" and "A\<^sub>P' \* \" and "A\<^sub>P' \* P" and "A\<^sub>P' \* Q" and "A\<^sub>P' \* M" and "A\<^sub>P' \* A\<^sub>Q" + by - (rule distinctFrame[where C="(\, P, Q, M, A\<^sub>Q)"], auto) + moreover from \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ \A\<^sub>Q \* \\ \A\<^sub>Q \* P\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* K\ \A\<^sub>P' \* A\<^sub>Q\ + obtain A\<^sub>Q' where "extractFrame Q = \A\<^sub>Q', \\<^sub>Q\" and "distinct A\<^sub>Q'" and "A\<^sub>Q' \* \" and "A\<^sub>Q' \* P" and "A\<^sub>Q' \* Q" and "A\<^sub>Q' \* K" and "A\<^sub>P' \* A\<^sub>Q'" + by - (rule distinctFrame[where C="(\, P, Q, K, A\<^sub>P')"], auto) + ultimately show ?thesis using \xvec \* P\ + by(metis Goal) +qed + +lemma Comm2: + fixes \ :: 'b + and \\<^sub>Q :: 'b + and P :: "('a, 'b, 'c) psi" + and M :: 'a + and xvec :: "name list" + and N :: 'a + and P' :: "('a, 'b, 'c) psi" + and A\<^sub>P :: "name list" + and \\<^sub>P :: 'b + and Q :: "('a, 'b, 'c) psi" + and K :: 'a + and Q' :: "('a, 'b, 'c) psi" + and A\<^sub>Q :: "name list" + +assumes "\ \ \\<^sub>Q \ P \M\\*xvec\\N\ \ P'" + and "extractFrame P = \A\<^sub>P, \\<^sub>P\" + and "\ \ \\<^sub>P \ Q \K\N\ \ Q'" + and "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" + and "\ \ \\<^sub>P \ \\<^sub>Q \ M \ K" + and "A\<^sub>P \* \" + and "A\<^sub>P \* P" + and "A\<^sub>P \* Q" + and "A\<^sub>P \* M" + and "A\<^sub>P \* A\<^sub>Q" + and "A\<^sub>Q \* \" + and "A\<^sub>Q \* P" + and "A\<^sub>Q \* Q" + and "A\<^sub>Q \* K" + and "xvec \* Q" + +shows "\ \ P \ Q \\ \ \\*xvec\(P' \ Q')" +proof - + { + fix \ :: 'b + and \\<^sub>Q :: 'b + and P :: "('a, 'b, 'c) psi" + and M :: 'a + and xvec :: "name list" + and N :: 'a + and P' :: "('a, 'b, 'c) psi" + and A\<^sub>P :: "name list" + and \\<^sub>P :: 'b + and Q :: "('a, 'b, 'c) psi" + and K :: 'a + and Q' :: "('a, 'b, 'c) psi" + and A\<^sub>Q :: "name list" + + assume "\ \ \\<^sub>Q \ P \M\\*xvec\\N\ \ P'" + and "extractFrame P = \A\<^sub>P, \\<^sub>P\" + and "distinct A\<^sub>P" + and "\ \ \\<^sub>P \ Q \K\N\ \ Q'" + and "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" + and "distinct A\<^sub>Q" + and "\ \ \\<^sub>P \ \\<^sub>Q \ M \ K" + and "A\<^sub>P \* \" + and "A\<^sub>P \* P" + and "A\<^sub>P \* Q" + and "A\<^sub>P \* M" + and "A\<^sub>P \* A\<^sub>Q" + and "A\<^sub>Q \* \" + and "A\<^sub>Q \* P" + and "A\<^sub>Q \* Q" + and "A\<^sub>Q \* K" + and "xvec \* Q" + + have "\ \ P \ Q \\ \ \\*xvec\(P' \ Q')" + proof - + + obtain r::"name prm" where "(r \ xvec) \* \" and "(r \ xvec) \* P" and "(r \ xvec) \* Q" and "(r \ xvec) \* M" + and "(r \ xvec) \* K" and "(r \ xvec) \* N" and "(r \ xvec) \* A\<^sub>P" and "(r \ xvec) \* A\<^sub>Q" + and "(r \ xvec) \* P'" and "(r \ xvec) \* Q'" and "(r \ xvec) \* \\<^sub>P" and "(r \ xvec) \* \\<^sub>Q" + and Sr: "(set r) \ (set xvec) \ (set(r \ xvec))" and "distinctPerm r" + by(rule name_list_avoiding[where xvec=xvec and c="(\, P, Q, M, K, N, A\<^sub>P, A\<^sub>Q, \\<^sub>P, \\<^sub>Q, P', Q')"]) + (auto simp add: eqvts fresh_star_prod) + obtain q::"name prm" where "(q \ A\<^sub>Q) \* \" and "(q \ A\<^sub>Q) \* P" and "(q \ A\<^sub>Q) \* Q" and "(q \ A\<^sub>Q) \* K" + and "(q \ A\<^sub>Q) \* (r \ N)" and "(q \ A\<^sub>Q) \* (r \ xvec)" and "(q \ A\<^sub>Q) \* (r \ Q')" + and "(q \ A\<^sub>Q) \* (r \ P')" and "(q \ A\<^sub>Q) \* \\<^sub>P" and "(q \ A\<^sub>Q) \* A\<^sub>P" and "(q \ A\<^sub>Q) \* \\<^sub>Q" + and Sq: "set q \ set A\<^sub>Q \ set(q \ A\<^sub>Q)" + by(rule name_list_avoiding[where xvec=A\<^sub>Q and c="(\, P, Q, K, r \ N, r \ xvec, \\<^sub>Q, A\<^sub>P, \\<^sub>P, r \ Q', r \ P')"]) + (auto simp add: eqvts fresh_star_prod) + obtain p::"name prm" where "(p \ A\<^sub>P) \* \" and "(p \ A\<^sub>P) \* P" and "(p \ A\<^sub>P) \* Q" and "(p \ A\<^sub>P) \* M" + and "(p \ A\<^sub>P) \* (r \ N)" and "(p \ A\<^sub>P) \* (r \ xvec)" and "(p \ A\<^sub>P) \* (r \ Q')" + and "(p \ A\<^sub>P) \* (r \ P')" and "(p \ A\<^sub>P) \* \\<^sub>P" and "(p \ A\<^sub>P) \* \\<^sub>Q" and "(p \ A\<^sub>P) \* A\<^sub>Q" + and "(p \ A\<^sub>P) \* (q \ A\<^sub>Q)" and Sp: "(set p) \ (set A\<^sub>P) \ (set(p \ A\<^sub>P))" + by(rule name_list_avoiding[where xvec=A\<^sub>P and c="(\, P, Q, M, r \ N, r \ xvec, A\<^sub>Q, q \ A\<^sub>Q, \\<^sub>Q, \\<^sub>P, r \ Q', r \ P')"]) + (auto simp add: eqvts fresh_star_prod) + + have FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" by fact + have FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" by fact + + from \A\<^sub>P \* Q\ FrQ \A\<^sub>P \* A\<^sub>Q\ have "A\<^sub>P \* \\<^sub>Q" + by(auto dest: extractFrameFreshChain) + from \A\<^sub>Q \* P\ FrP \A\<^sub>P \* A\<^sub>Q\ have "A\<^sub>Q \* \\<^sub>P" + by(auto dest: extractFrameFreshChain) + + from \(r \ xvec) \* A\<^sub>Q\ \(q \ A\<^sub>Q) \* (r \ xvec)\ \(r \ xvec) \* A\<^sub>Q\ Sq have "(r \ xvec) \* (q \ A\<^sub>Q)" + by(simp add: freshChainSimps) + + from \\ \ \\<^sub>Q \ P \M\\*xvec\\N\ \ P'\ Sr \(r \ xvec) \* N\ \(r \ xvec) \* P'\ + have "\ \ \\<^sub>Q \ P \M\\*(r \ xvec)\\(r \ N)\ \ (r \ P')" + by(simp add: boundOutputChainAlpha'' create_residual.simps) + then have "(q \ (\ \ \\<^sub>Q)) \ P \(q \ M)\\*(r \ xvec)\\(r \ N)\ \ (r \ P')" using Sq \A\<^sub>Q \* P\ \(q \ A\<^sub>Q) \* P\ \(r \ xvec) \* M\ \(r \ xvec) \* A\<^sub>Q\ \(r \ xvec) \* (q \ A\<^sub>Q)\ + by(fastforce intro: outputPermFrameSubject) + then have PTrans: "\ \ (q \ \\<^sub>Q) \ P \(q \ M)\\*(r \ xvec)\\(r \ N)\ \ (r \ P')" using Sq \A\<^sub>Q \* \\ \(q \ A\<^sub>Q) \* \\ + by(simp add: eqvts) + moreover then have "distinct(r \ xvec)" by(force dest: boundOutputDistinct) + + moreover from \extractFrame P = \A\<^sub>P, \\<^sub>P\\ Sp \(p \ A\<^sub>P) \* \\<^sub>P\ + have FrP: "extractFrame P = \(p \ A\<^sub>P), (p \ \\<^sub>P)\" + by(simp add: frameChainAlpha) + moreover from \distinct A\<^sub>P\ have "distinct(p \ A\<^sub>P)" by simp + + moreover from \\ \ \\<^sub>P \ Q \K\N\ \ Q'\ Sr \distinctPerm r\ \xvec \* Q\ \(r \ xvec) \* Q\ + have "\ \ \\<^sub>P \ Q \K\(r \ N)\ \ (r \ Q')" + by(rule inputAlpha) + then have "(p \ (\ \ \\<^sub>P)) \ Q \(p \ K)\(r \ N)\ \ (r \ Q')" using Sp \A\<^sub>P \* Q\ \(p \ A\<^sub>P) \* Q\ + by - (rule inputPermFrameSubject, (assumption | simp)+) + then have QTrans: "\ \ (p \ \\<^sub>P) \ Q \(p \ K)\(r \ N)\ \ (r \ Q')" using Sp \A\<^sub>P \* \\ \(p \ A\<^sub>P) \* \\ + by(simp add: eqvts) + + moreover from \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ Sq \(q \ A\<^sub>Q) \* \\<^sub>Q\ + have FrQ: "extractFrame Q = \(q \ A\<^sub>Q), (q \ \\<^sub>Q)\" + by(simp add: frameChainAlpha) + moreover from \distinct A\<^sub>Q\ have "distinct(q \ A\<^sub>Q)" by simp + + moreover from \\ \ \\<^sub>P \ \\<^sub>Q \ M \ K\ have "(p \ q \ (\ \ \\<^sub>P \ \\<^sub>Q)) \ (p \ q \ M) \ (p \ q \ K)" + by(metis chanEqClosed) + with \A\<^sub>P \* \\ \(p \ A\<^sub>P) \* \\ \A\<^sub>Q \* \\ \(q \ A\<^sub>Q) \* \\ \A\<^sub>Q \* \\<^sub>P\ \(q \ A\<^sub>Q) \* \\<^sub>P\ + \A\<^sub>P \* \\<^sub>Q\ \(p \ A\<^sub>P) \* \\<^sub>Q\ \A\<^sub>P \* M\ \(p \ A\<^sub>P) \* M\ \(q \ A\<^sub>Q) \* A\<^sub>P\ \(p \ A\<^sub>P) \* (q \ A\<^sub>Q)\ + \A\<^sub>Q \* K\ \(q \ A\<^sub>Q) \* K\ \A\<^sub>P \* A\<^sub>Q\ \(p \ A\<^sub>P) \* A\<^sub>Q\ Sp Sq + have "\ \ (p \ \\<^sub>P) \ (q \ \\<^sub>Q) \ (q \ M) \ (p \ K)" + by(simp add: eqvts freshChainSimps) + moreover note \(p \ A\<^sub>P) \* \\ + moreover from \(p \ A\<^sub>P) \* A\<^sub>Q\ \(p \ A\<^sub>P) \* (q \ A\<^sub>Q)\ \(p \ A\<^sub>P) \* \\<^sub>Q\ Sq have "(p \ A\<^sub>P) \* (q \ \\<^sub>Q)" + by(simp add: freshChainSimps) + moreover note \(p \ A\<^sub>P) \* P\ + moreover from \(p \ A\<^sub>P) \* A\<^sub>Q\ \(p \ A\<^sub>P) \* (q \ A\<^sub>Q)\ \(p \ A\<^sub>P) \* M\ Sq have "(p \ A\<^sub>P) \* (q \ M)" + by(simp add: freshChainSimps) + moreover note \(p \ A\<^sub>P) \* (r \ N)\ \(p \ A\<^sub>P) \* (r \ P')\ \(p \ A\<^sub>P) \* Q\ \(p \ A\<^sub>P) \* (r \ Q')\ \(p \ A\<^sub>P) \* (q \ A\<^sub>Q)\ + \(p \ A\<^sub>P) \* (r \ xvec)\ \(q \ A\<^sub>Q) \* \\ + moreover from \(q \ A\<^sub>Q) \* A\<^sub>P\ \(p \ A\<^sub>P) \* A\<^sub>Q\ \(q \ A\<^sub>Q) \* \\<^sub>P\ \(p \ A\<^sub>P) \* (q \ A\<^sub>Q)\ Sp Sq have "(q \ A\<^sub>Q) \* (p \ \\<^sub>P)" + by(simp add: freshChainSimps) + moreover note \(q \ A\<^sub>Q) \* P\ \(q \ A\<^sub>Q) \* (r \ N)\\(q \ A\<^sub>Q) \* (r \ P')\ \(q \ A\<^sub>Q) \* Q\ + moreover from \(q \ A\<^sub>Q) \* A\<^sub>P\ \(p \ A\<^sub>P) \* A\<^sub>Q\ \(q \ A\<^sub>Q) \* K\ \(p \ A\<^sub>P) \* (q \ A\<^sub>Q)\ Sp Sq have "(q \ A\<^sub>Q) \* (p \ K)" + by(simp add: freshChainSimps) + moreover note \(q \ A\<^sub>Q) \* (r \ Q')\ \(q \ A\<^sub>Q) \* (r \ xvec)\ \(r \ xvec) \* \\ + moreover from \(r \ xvec) \* A\<^sub>P\ \(p \ A\<^sub>P) \* (r \ xvec)\ \(r \ xvec) \* \\<^sub>P\ Sp have "(r \ xvec) \* (p \ \\<^sub>P)" + by(simp add: freshChainSimps) + moreover from \(r \ xvec) \* A\<^sub>Q\ \(q \ A\<^sub>Q) \* (r \ xvec)\ \(r \ xvec) \* \\<^sub>Q\ Sq have "(r \ xvec) \* (q \ \\<^sub>Q)" + by(simp add: freshChainSimps) + moreover note \(r \ xvec) \* P\ + moreover from \(r \ xvec) \* A\<^sub>Q\ \(q \ A\<^sub>Q) \* (r \ xvec)\ \(r \ xvec) \* M\ Sq have "(r \ xvec) \* (q \ M)" + by(simp add: freshChainSimps) + moreover note \(r \ xvec) \* Q\ + moreover from \(r \ xvec) \* A\<^sub>P\ \(p \ A\<^sub>P) \* (r \ xvec)\ \(r \ xvec) \* K\ Sp have "(r \ xvec) \* (p \ K)" + by(simp add: freshChainSimps) + ultimately have "\ \ P \ Q \\ \ \\*(r \ xvec)\((r \ P') \ (r \ Q'))" + by - (rule cComm2) + with \(r \ xvec) \* P'\ \(r \ xvec) \* Q'\ Sr + show ?thesis + by(subst resChainAlpha) auto + qed + } + note Goal = this + note \\ \ \\<^sub>Q \ P \M\\*xvec\\N\ \ P'\ \\ \ \\<^sub>P \ Q \K\N\ \ Q'\ \\ \ \\<^sub>P \ \\<^sub>Q \ M \ K\ + moreover from \extractFrame P = \A\<^sub>P, \\<^sub>P\\ \A\<^sub>P \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* Q\ \A\<^sub>P \* M\ \A\<^sub>P \* A\<^sub>Q\ + obtain A\<^sub>P' where "extractFrame P = \A\<^sub>P', \\<^sub>P\" and "distinct A\<^sub>P'" and "A\<^sub>P' \* \" and "A\<^sub>P' \* P" and "A\<^sub>P' \* Q" and "A\<^sub>P' \* M" and "A\<^sub>P' \* A\<^sub>Q" + by - (rule distinctFrame[where C="(\, P, Q, M, A\<^sub>Q)"], auto) + moreover from \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ \A\<^sub>Q \* \\ \A\<^sub>Q \* P\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* K\ \A\<^sub>P' \* A\<^sub>Q\ + obtain A\<^sub>Q' where "extractFrame Q = \A\<^sub>Q', \\<^sub>Q\" and "distinct A\<^sub>Q'" and "A\<^sub>Q' \* \" and "A\<^sub>Q' \* P" and "A\<^sub>Q' \* Q" and "A\<^sub>Q' \* K" and "A\<^sub>P' \* A\<^sub>Q'" + by - (rule distinctFrame[where C="(\, P, Q, K, A\<^sub>P')"], auto) + ultimately show ?thesis using \xvec \* Q\ + by(metis Goal) +qed + +lemma BrMerge: + fixes \ :: 'b + and \\<^sub>Q :: 'b + and P :: "('a, 'b, 'c) psi" + and M :: 'a + and N :: 'a + and P' :: "('a, 'b, 'c) psi" + and A\<^sub>P :: "name list" + and \\<^sub>P :: 'b + and Q :: "('a, 'b, 'c) psi" + and Q' :: "('a, 'b, 'c) psi" + and A\<^sub>Q :: "name list" + +assumes "\ \ \\<^sub>Q \ P \\M\N\ \ P'" + and "extractFrame P = \A\<^sub>P, \\<^sub>P\" + and "\ \ \\<^sub>P \ Q \\M\N\ \ Q'" + and "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" + and "A\<^sub>P \* \" + and "A\<^sub>P \* P" + and "A\<^sub>P \* Q" + and "A\<^sub>P \* M" + and "A\<^sub>P \* A\<^sub>Q" + and "A\<^sub>Q \* \" + and "A\<^sub>Q \* P" + and "A\<^sub>Q \* Q" + and "A\<^sub>Q \* M" + +shows "\ \ P \ Q \\M\N\ \ (P' \ Q')" +proof - + { + fix \ :: 'b + and \\<^sub>Q :: 'b + and P :: "('a, 'b, 'c) psi" + and M :: 'a + and N :: 'a + and P' :: "('a, 'b, 'c) psi" + and A\<^sub>P :: "name list" + and \\<^sub>P :: 'b + and Q :: "('a, 'b, 'c) psi" + and Q' :: "('a, 'b, 'c) psi" + and A\<^sub>Q :: "name list" + and svec :: "name list" + + assume "\ \ \\<^sub>Q \ P \\M\N\ \ P'" + and "extractFrame P = \A\<^sub>P, \\<^sub>P\" + and "distinct A\<^sub>P" + and "\ \ \\<^sub>P \ Q \\M\N\ \ Q'" + and "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" + and "distinct A\<^sub>Q" + and "A\<^sub>P \* \" + and "A\<^sub>P \* P" + and "A\<^sub>P \* Q" + and "A\<^sub>P \* M" + and "A\<^sub>P \* A\<^sub>Q" + and "A\<^sub>Q \* \" + and "A\<^sub>Q \* P" + and "A\<^sub>Q \* Q" + and "A\<^sub>Q \* M" + + have "\ \ P \ Q \\M\N\ \ (P' \ Q')" + proof - + obtain q::"name prm" where "(q \ (A\<^sub>Q::name list)) \* \" and "(q \ A\<^sub>Q) \* P" + and "(q \ A\<^sub>Q) \* Q" and "(q \ A\<^sub>Q) \* M" + and "(q \ A\<^sub>Q) \* \\<^sub>P" and "(q \ A\<^sub>Q) \* A\<^sub>P" and "(q \ A\<^sub>Q) \* \\<^sub>Q" + and "(q \ A\<^sub>Q) \* N" and "(q \ A\<^sub>Q) \* P'" and "(q \ A\<^sub>Q) \* Q'" + and Sq: "set q \ set A\<^sub>Q \ set(q \ A\<^sub>Q)" + and "distinctPerm q" + by(rule name_list_avoiding[where c="(\, P, M, N, P', Q', Q, \\<^sub>Q, A\<^sub>P, \\<^sub>P)"]) + (auto simp add: eqvts fresh_star_prod) + obtain p::"name prm" where "(p \ A\<^sub>P) \* \" and "(p \ A\<^sub>P) \* P" + and "(p \ A\<^sub>P) \* Q" and "(p \ A\<^sub>P) \* M" + and "(p \ A\<^sub>P) \* \\<^sub>P" and "(p \ A\<^sub>P) \* \\<^sub>Q" and "(p \ A\<^sub>P) \* A\<^sub>Q" + and "(p \ A\<^sub>P) \* N" and "(p \ A\<^sub>P) \* P'" and "(p \ A\<^sub>P) \* Q'" + and Sp: "set p \ set A\<^sub>P \ set(p \ A\<^sub>P)" + and "(p \ A\<^sub>P) \* (q \ A\<^sub>Q)" + by(rule name_list_avoiding[where c="(\, P, N, P', Q', Q, M, A\<^sub>Q, q \ A\<^sub>Q, \\<^sub>Q, \\<^sub>P)"]) + (auto simp add: eqvts fresh_star_prod) + + have FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" by fact + have FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" by fact + + from \A\<^sub>P \* Q\ FrQ \A\<^sub>P \* A\<^sub>Q\ have "A\<^sub>P \* \\<^sub>Q" + by(force dest: extractFrameFreshChain) + from \A\<^sub>Q \* P\ FrP \A\<^sub>P \* A\<^sub>Q\ have "A\<^sub>Q \* \\<^sub>P" + by(force dest: extractFrameFreshChain) + + from Sq \A\<^sub>Q \* M\ \(q \ A\<^sub>Q) \* M\ + have "(q \ M) = M" + by simp + from Sp \A\<^sub>P \* M\ \(p \ A\<^sub>P) \* M\ + have "(p \ M) = M" + by simp + + from \distinct A\<^sub>P\ have "distinct(p \ A\<^sub>P)" by simp + moreover from \distinct A\<^sub>Q\ have "distinct(q \ A\<^sub>Q)" by simp + moreover from \extractFrame P = \A\<^sub>P, \\<^sub>P\\ Sp \(p \ A\<^sub>P) \* \\<^sub>P\ + have FrP: "extractFrame P = \(p \ A\<^sub>P), (p \ \\<^sub>P)\" + by(simp add: frameChainAlpha) + moreover from \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ Sq \(q \ A\<^sub>Q) \* \\<^sub>Q\ + have FrQ: "extractFrame Q = \(q \ A\<^sub>Q), (q \ \\<^sub>Q)\" + by(simp add: frameChainAlpha) + + moreover have "(q \ (\ \ \\<^sub>Q)) \ P \\(q \ M)\N\ \ P'" using Sq \A\<^sub>Q \* P\ \(q \ A\<^sub>Q) \* P\ \\ \ \\<^sub>Q \ P \\M\N\ \ P'\ + by - (rule brinputPermFrameSubject, (assumption | simp)+) + then have "(\ \ (q \ \\<^sub>Q)) \ P \\(q \ M)\N\ \ P'" using Sq \A\<^sub>Q \* \\ \(q \ A\<^sub>Q) \* \\ + by(simp add: eqvts) + with \(q \ M) = M\ have PTrans: "(\ \ (q \ \\<^sub>Q)) \ P \\M\N\ \ P'" + by simp + + moreover have "(p \ (\ \ \\<^sub>P)) \ Q \\(p \ M)\N\ \ Q'" using Sp \A\<^sub>P \* Q\ \(p \ A\<^sub>P) \* Q\ \\ \ \\<^sub>P \ Q \\M\N\ \ Q'\ + by - (rule brinputPermFrameSubject, (assumption | simp)+) + then have "(\ \ (p \ \\<^sub>P)) \ Q \\(p \ M)\N\ \ Q'" using Sp \A\<^sub>P \* \\ \(p \ A\<^sub>P) \* \\ + by(simp add: eqvts) + with \(p \ M) = M\ have PTrans: "(\ \ (p \ \\<^sub>P)) \ Q \\M\N\ \ Q'" + by simp + moreover from \(p \ A\<^sub>P) \* A\<^sub>Q\ \(p \ A\<^sub>P) \* (q \ A\<^sub>Q)\ \(p \ A\<^sub>P) \* \\<^sub>Q\ Sq have "(p \ A\<^sub>P) \* (q \ \\<^sub>Q)" + by(simp add: freshChainSimps) + moreover from \(q \ A\<^sub>Q) \* A\<^sub>P\ \(p \ A\<^sub>P) \* A\<^sub>Q\ \(q \ A\<^sub>Q) \* \\<^sub>P\ \(p \ A\<^sub>P) \* (q \ A\<^sub>Q)\ Sp Sq have "(q \ A\<^sub>Q) \* (p \ \\<^sub>P)" + by(simp add: freshChainSimps) + + moreover note + \(p \ A\<^sub>P) \* \\ \(p \ A\<^sub>P) \* M\ + \(p \ A\<^sub>P) \* P\ \(p \ A\<^sub>P) \* N\ \(p \ A\<^sub>P) \* P'\ + \(p \ A\<^sub>P) \* Q\ \(p \ A\<^sub>P) \* Q'\ \(p \ A\<^sub>P) \* (q \ A\<^sub>Q)\ + \(q \ A\<^sub>Q) \* \\ \(q \ A\<^sub>Q) \* M\ + \(q \ A\<^sub>Q) \* P\ \(q \ A\<^sub>Q) \* N\ \(q \ A\<^sub>Q) \* P'\ + \(q \ A\<^sub>Q) \* Q\ \(q \ A\<^sub>Q) \* Q'\ + ultimately show ?thesis + by(simp add: cBrMerge) + qed + } + note Goal = this + + note \\ \ \\<^sub>Q \ P \\M\N\ \ P'\ \\ \ \\<^sub>P \ Q \\M\N\ \ Q'\ + moreover from \extractFrame P = \A\<^sub>P, \\<^sub>P\\ \A\<^sub>P \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* Q\ \A\<^sub>P \* M\ \A\<^sub>P \* A\<^sub>Q\ + obtain A\<^sub>P' where "extractFrame P = \A\<^sub>P', \\<^sub>P\" and "distinct A\<^sub>P'" and "A\<^sub>P' \* \" and "A\<^sub>P' \* P" and "A\<^sub>P' \* Q" and "A\<^sub>P' \* M" and "A\<^sub>P' \* A\<^sub>Q" + by - (rule distinctFrame[where C="(\, P, Q, M, A\<^sub>Q)"], auto) + moreover from \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ \A\<^sub>Q \* \\ \A\<^sub>Q \* P\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* M\ \A\<^sub>P' \* A\<^sub>Q\ + obtain A\<^sub>Q' where "extractFrame Q = \A\<^sub>Q', \\<^sub>Q\" and "distinct A\<^sub>Q'" and "A\<^sub>Q' \* \" and "A\<^sub>Q' \* P" and "A\<^sub>Q' \* Q" and "A\<^sub>Q' \* M" and "A\<^sub>P' \* A\<^sub>Q'" + by - (rule distinctFrame[where C="(\, P, Q, M, A\<^sub>P')"], auto) + ultimately show ?thesis + by(metis Goal) +qed + +lemma BrComm1: + fixes \ :: 'b + and \\<^sub>Q :: 'b + and P :: "('a, 'b, 'c) psi" + and M :: 'a + and N :: 'a + and P' :: "('a, 'b, 'c) psi" + and A\<^sub>P :: "name list" + and \\<^sub>P :: 'b + and Q :: "('a, 'b, 'c) psi" + and xvec :: "name list" + and Q' :: "('a, 'b, 'c) psi" + and A\<^sub>Q :: "name list" + +assumes "\ \ \\<^sub>Q \ P \\M\N\ \ P'" + and "extractFrame P = \A\<^sub>P, \\<^sub>P\" + and "\ \ \\<^sub>P \ Q \\M\\*xvec\\N\ \ Q'" + and "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" + and "A\<^sub>P \* \" + and "A\<^sub>P \* P" + and "A\<^sub>P \* Q" + and "A\<^sub>P \* M" + and "A\<^sub>P \* A\<^sub>Q" + and "A\<^sub>Q \* \" + and "A\<^sub>Q \* P" + and "A\<^sub>Q \* Q" + and "A\<^sub>Q \* M" + and "xvec \* P" + +shows "\ \ P \ Q \\M\\*xvec\\N\ \ (P' \ Q')" +proof - + { + fix \ :: 'b + and \\<^sub>Q :: 'b + and P :: "('a, 'b, 'c) psi" + and M :: 'a + and N :: 'a + and P' :: "('a, 'b, 'c) psi" + and A\<^sub>P :: "name list" + and \\<^sub>P :: 'b + and Q :: "('a, 'b, 'c) psi" + and xvec :: "name list" + and Q' :: "('a, 'b, 'c) psi" + and A\<^sub>Q :: "name list" + + assume "\ \ \\<^sub>Q \ P \\M\N\ \ P'" + and "extractFrame P = \A\<^sub>P, \\<^sub>P\" + and "distinct A\<^sub>P" + and "\ \ \\<^sub>P \ Q \\M\\*xvec\\N\ \ Q'" + and "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" + and "distinct A\<^sub>Q" + and "A\<^sub>P \* \" + and "A\<^sub>P \* P" + and "A\<^sub>P \* Q" + and "A\<^sub>P \* M" + and "A\<^sub>P \* A\<^sub>Q" + and "A\<^sub>Q \* \" + and "A\<^sub>Q \* P" + and "A\<^sub>Q \* Q" + and "A\<^sub>Q \* M" + and "xvec \* P" + + have "\ \ P \ Q \\M\\*xvec\\N\ \ (P' \ Q')" + proof - + obtain r::"name prm" where "(r \ xvec) \* \" and "(r \ xvec) \* P" and "(r \ xvec) \* Q" and "(r \ xvec) \* M" + and "(r \ xvec) \* N" and "(r \ xvec) \* A\<^sub>P" and "(r \ xvec) \* A\<^sub>Q" + and "(r \ xvec) \* P'" and "(r \ xvec) \* Q'" and "(r \ xvec) \* \\<^sub>P" and "(r \ xvec) \* \\<^sub>Q" + and Sr: "(set r) \ (set xvec) \ (set(r \ xvec))" and "distinctPerm r" + by(rule name_list_avoiding[where xvec=xvec and c="(\, P, Q, M, N, A\<^sub>P, A\<^sub>Q, \\<^sub>P, \\<^sub>Q, P', Q')"]) + (auto simp add: eqvts fresh_star_prod) + obtain q::"name prm" where "(q \ A\<^sub>Q) \* \" and "(q \ A\<^sub>Q) \* P" and "(q \ A\<^sub>Q) \* Q" and "(q \ A\<^sub>Q) \* M" + and "(q \ A\<^sub>Q) \* (r \ N)" and "(q \ A\<^sub>Q) \* (r \ xvec)" and "(q \ A\<^sub>Q) \* (r \ Q')" + and "(q \ A\<^sub>Q) \* (r \ P')" and "(q \ A\<^sub>Q) \* \\<^sub>P" and "(q \ A\<^sub>Q) \* A\<^sub>P" and "(q \ A\<^sub>Q) \* \\<^sub>Q" + and Sq: "set q \ set A\<^sub>Q \ set(q \ A\<^sub>Q)" + by(rule name_list_avoiding[where xvec=A\<^sub>Q and c="(\, P, Q, M, r \ N, r \ xvec, \\<^sub>Q, A\<^sub>P, \\<^sub>P, r \ Q', r \ P')"]) + (auto simp add: eqvts fresh_star_prod) + obtain p::"name prm" where "(p \ A\<^sub>P) \* \" and "(p \ A\<^sub>P) \* P" and "(p \ A\<^sub>P) \* Q" and "(p \ A\<^sub>P) \* M" + and "(p \ A\<^sub>P) \* (r \ N)" and "(p \ A\<^sub>P) \* (r \ xvec)" and "(p \ A\<^sub>P) \* (r \ Q')" + and "(p \ A\<^sub>P) \* (r \ P')" and "(p \ A\<^sub>P) \* \\<^sub>P" and "(p \ A\<^sub>P) \* \\<^sub>Q" and "(p \ A\<^sub>P) \* A\<^sub>Q" + and "(p \ A\<^sub>P) \* (q \ A\<^sub>Q)" and Sp: "(set p) \ (set A\<^sub>P) \ (set(p \ A\<^sub>P))" + by(rule name_list_avoiding[where xvec=A\<^sub>P and c="(\, P, Q, M, r \ N, r \ xvec, A\<^sub>Q, q \ A\<^sub>Q, \\<^sub>Q, \\<^sub>P, r \ Q', r \ P')"]) + (auto simp add: eqvts fresh_star_prod) + + have FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" by fact + have FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" by fact + + from Sp \A\<^sub>P \* M\ \(p \ A\<^sub>P) \* M\ + have "(p \ M) = M" + by simp + from Sq \A\<^sub>Q \* M\ \(q \ A\<^sub>Q) \* M\ + have "(q \ M) = M" + by simp + + from \A\<^sub>P \* Q\ FrQ \A\<^sub>P \* A\<^sub>Q\ have "A\<^sub>P \* \\<^sub>Q" + by(auto dest: extractFrameFreshChain) + from \A\<^sub>Q \* P\ FrP \A\<^sub>P \* A\<^sub>Q\ have "A\<^sub>Q \* \\<^sub>P" + by(auto dest: extractFrameFreshChain) + from \(r \ xvec) \* A\<^sub>P\ \(p \ A\<^sub>P) \* (r \ xvec)\ \(r \ xvec) \* A\<^sub>P\ Sp have "(r \ xvec) \* (p \ A\<^sub>P)" + by(simp add: freshChainSimps) + + from \\ \ \\<^sub>Q \ P \\M\N\ \ P'\ Sr \distinctPerm r\ \xvec \* P\ \(r \ xvec) \* P\ + have "\ \ \\<^sub>Q \ P \\M\(r \ N)\ \ (r \ P')" + by(rule brinputAlpha) + then have "(q \ (\ \ \\<^sub>Q)) \ P \\(q \ M)\(r \ N)\ \ (r \ P')" using Sq \A\<^sub>Q \* P\ \(q \ A\<^sub>Q) \* P\ + by - (rule brinputPermFrameSubject, (assumption | simp)+) + then have "\ \ (q \ \\<^sub>Q) \ P \\(q \ M)\(r \ N)\ \ (r \ P')" using Sq \A\<^sub>Q \* \\ \(q \ A\<^sub>Q) \* \\ + by(simp add: eqvts) + with \(q \ M) = M\ have PTrans: "\ \ (q \ \\<^sub>Q) \ P \\M\(r \ N)\ \ (r \ P')" by simp + + moreover from \extractFrame P = \A\<^sub>P, \\<^sub>P\\ Sp \(p \ A\<^sub>P) \* \\<^sub>P\ + have FrP: "extractFrame P = \(p \ A\<^sub>P), (p \ \\<^sub>P)\" + by(simp add: frameChainAlpha) + moreover from \distinct A\<^sub>P\ have "distinct(p \ A\<^sub>P)" by simp + + moreover from \\ \ \\<^sub>P \ Q \\M\\*xvec\\N\ \ Q'\ Sr \(r \ xvec) \* N\ \(r \ xvec) \* Q'\ + have "\ \ \\<^sub>P \ Q \\M\\*(r \ xvec)\\(r \ N)\ \ (r \ Q')" + by(simp add: boundOutputChainAlpha'' create_residual.simps) + then have "(p \ (\ \ \\<^sub>P)) \ Q \\(p \ M)\\*(r \ xvec)\\(r \ N)\ \ (r \ Q')" using Sp \A\<^sub>P \* Q\ \(p \ A\<^sub>P) \* Q\ \(r \ xvec) \* M\ \(r \ xvec) \* A\<^sub>P\ \(r \ xvec) \* (p \ A\<^sub>P)\ + by(fastforce intro: broutputPermFrameSubject) + then have "\ \ (p \ \\<^sub>P) \ Q \\(p \ M)\\*(r \ xvec)\\(r \ N)\ \ (r \ Q')" using Sp \A\<^sub>P \* \\ \(p \ A\<^sub>P) \* \\ + by(simp add: eqvts) + with \(p \ M) = M\ have QTrans: "\ \ (p \ \\<^sub>P) \ Q \\M\\*(r \ xvec)\\(r \ N)\ \ (r \ Q')" by simp + moreover then have "distinct(r \ xvec)" by(force dest: boundOutputDistinct) + moreover from \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ Sq \(q \ A\<^sub>Q) \* \\<^sub>Q\ + have FrQ: "extractFrame Q = \(q \ A\<^sub>Q), (q \ \\<^sub>Q)\" + by(simp add: frameChainAlpha) + moreover from \distinct A\<^sub>Q\ have "distinct(q \ A\<^sub>Q)" by simp + + moreover note \(p \ A\<^sub>P) \* \\ + moreover from \(p \ A\<^sub>P) \* A\<^sub>Q\ \(p \ A\<^sub>P) \* (q \ A\<^sub>Q)\ \(p \ A\<^sub>P) \* \\<^sub>Q\ Sq have "(p \ A\<^sub>P) \* (q \ \\<^sub>Q)" + by(simp add: freshChainSimps) + moreover note \(p \ A\<^sub>P) \* P\ \(p \ A\<^sub>P) \* M\ + moreover note \(p \ A\<^sub>P) \* (r \ N)\ \(p \ A\<^sub>P) \* (r \ P')\ \(p \ A\<^sub>P) \* Q\ \(p \ A\<^sub>P) \* (r \ Q')\ \(p \ A\<^sub>P) \* (q \ A\<^sub>Q)\ + \(p \ A\<^sub>P) \* (r \ xvec)\ \(q \ A\<^sub>Q) \* \\ + moreover from \(q \ A\<^sub>Q) \* A\<^sub>P\ \(p \ A\<^sub>P) \* A\<^sub>Q\ \(q \ A\<^sub>Q) \* \\<^sub>P\ \(p \ A\<^sub>P) \* (q \ A\<^sub>Q)\ Sp Sq have "(q \ A\<^sub>Q) \* (p \ \\<^sub>P)" + by(simp add: freshChainSimps) + moreover note \(q \ A\<^sub>Q) \* P\ \(q \ A\<^sub>Q) \* (r \ N)\\(q \ A\<^sub>Q) \* (r \ P')\ \(q \ A\<^sub>Q) \* Q\ \(q \ A\<^sub>Q) \* M\ + moreover note \(q \ A\<^sub>Q) \* (r \ Q')\ \(q \ A\<^sub>Q) \* (r \ xvec)\ \(r \ xvec) \* \\ + moreover from \(r \ xvec) \* A\<^sub>P\ \(p \ A\<^sub>P) \* (r \ xvec)\ \(r \ xvec) \* \\<^sub>P\ Sp have "(r \ xvec) \* (p \ \\<^sub>P)" + by(simp add: freshChainSimps) + moreover from \(r \ xvec) \* A\<^sub>Q\ \(q \ A\<^sub>Q) \* (r \ xvec)\ \(r \ xvec) \* \\<^sub>Q\ Sq have "(r \ xvec) \* (q \ \\<^sub>Q)" + by(simp add: freshChainSimps) + moreover note \(r \ xvec) \* P\ \(r \ xvec) \* M\ + moreover note \(r \ xvec) \* Q\ \(r \ xvec) \* M\ + ultimately have "\ \ P \ Q \\M\\*(r \ xvec)\\(r \ N)\ \ ((r \ P') \ (r \ Q'))" + by - (rule cBrComm1) + then have permuted: "\ \ P \ Q \\M\\*(r \ xvec)\\(r \ N)\ \ (r \ (P' \ Q'))" by simp + note \(r \ xvec) \* N\ + moreover from \(r \ xvec) \* P'\ \(r \ xvec) \* Q'\ + have "(r \ xvec) \* (P' \ Q')" by simp + moreover note Sr + moreover have "set xvec \ set xvec" by simp + ultimately have "\\*xvec\N \' (P' \ Q') = \\*(r \ xvec)\(r \ N) \' (r \ (P' \ Q'))" + by(rule boundOutputChainAlpha'') + then have "\M\\*xvec\\N\ \ (P' \ Q') = \M\\*(r \ xvec)\\(r \ N)\ \ (r \ (P' \ Q'))" + by(simp only: create_residual.simps) + with permuted show ?thesis + by simp + qed + } + note Goal = this + + note \\ \ \\<^sub>Q \ P \\M\N\ \ P'\ \\ \ \\<^sub>P \ Q \\M\\*xvec\\N\ \ Q'\ + moreover from \extractFrame P = \A\<^sub>P, \\<^sub>P\\ \A\<^sub>P \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* Q\ \A\<^sub>P \* M\ \A\<^sub>P \* A\<^sub>Q\ + obtain A\<^sub>P' where "extractFrame P = \A\<^sub>P', \\<^sub>P\" and "distinct A\<^sub>P'" and "A\<^sub>P' \* \" and "A\<^sub>P' \* P" and "A\<^sub>P' \* Q" and "A\<^sub>P' \* M" and "A\<^sub>P' \* A\<^sub>Q" + by - (rule distinctFrame[where C="(\, P, Q, M, A\<^sub>Q)"], auto) + moreover from \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ \A\<^sub>Q \* \\ \A\<^sub>Q \* P\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* M\ \A\<^sub>P' \* A\<^sub>Q\ + obtain A\<^sub>Q' where "extractFrame Q = \A\<^sub>Q', \\<^sub>Q\" and "distinct A\<^sub>Q'" and "A\<^sub>Q' \* \" and "A\<^sub>Q' \* P" and "A\<^sub>Q' \* Q" and "A\<^sub>Q' \* M" and "A\<^sub>P' \* A\<^sub>Q'" + by - (rule distinctFrame[where C="(\, P, Q, M, A\<^sub>P')"], auto) + ultimately show ?thesis using \xvec \* P\ + by(metis Goal) +qed + +lemma BrComm2: + fixes \ :: 'b + and \\<^sub>Q :: 'b + and P :: "('a, 'b, 'c) psi" + and M :: 'a + and xvec :: "name list" + and N :: 'a + and P' :: "('a, 'b, 'c) psi" + and A\<^sub>P :: "name list" + and \\<^sub>P :: 'b + and Q :: "('a, 'b, 'c) psi" + and Q' :: "('a, 'b, 'c) psi" + and A\<^sub>Q :: "name list" + +assumes "\ \ \\<^sub>Q \ P \\M\\*xvec\\N\ \ P'" + and "extractFrame P = \A\<^sub>P, \\<^sub>P\" + and "\ \ \\<^sub>P \ Q \\M\N\ \ Q'" + and "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" + and "A\<^sub>P \* \" + and "A\<^sub>P \* P" + and "A\<^sub>P \* Q" + and "A\<^sub>P \* M" + and "A\<^sub>P \* A\<^sub>Q" + and "A\<^sub>Q \* \" + and "A\<^sub>Q \* P" + and "A\<^sub>Q \* Q" + and "A\<^sub>Q \* M" + and "xvec \* Q" + +shows "\ \ P \ Q \\M\\*xvec\\N\ \ (P' \ Q')" +proof - + { + fix \ :: 'b + and \\<^sub>Q :: 'b + and P :: "('a, 'b, 'c) psi" + and M :: 'a + and xvec :: "name list" + and N :: 'a + and P' :: "('a, 'b, 'c) psi" + and A\<^sub>P :: "name list" + and \\<^sub>P :: 'b + and Q :: "('a, 'b, 'c) psi" + and Q' :: "('a, 'b, 'c) psi" + and A\<^sub>Q :: "name list" + + assume "\ \ \\<^sub>Q \ P \\M\\*xvec\\N\ \ P'" + and "extractFrame P = \A\<^sub>P, \\<^sub>P\" + and "distinct A\<^sub>P" + and "\ \ \\<^sub>P \ Q \\M\N\ \ Q'" + and "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" + and "distinct A\<^sub>Q" + and "A\<^sub>P \* \" + and "A\<^sub>P \* P" + and "A\<^sub>P \* Q" + and "A\<^sub>P \* M" + and "A\<^sub>P \* A\<^sub>Q" + and "A\<^sub>Q \* \" + and "A\<^sub>Q \* P" + and "A\<^sub>Q \* Q" + and "A\<^sub>Q \* M" + and "xvec \* Q" + + have "\ \ P \ Q \\M\\*xvec\\N\ \ (P' \ Q')" + proof - + obtain r::"name prm" where "(r \ xvec) \* \" and "(r \ xvec) \* P" and "(r \ xvec) \* Q" and "(r \ xvec) \* M" + and "(r \ xvec) \* M" and "(r \ xvec) \* N" and "(r \ xvec) \* A\<^sub>P" and "(r \ xvec) \* A\<^sub>Q" + and "(r \ xvec) \* P'" and "(r \ xvec) \* Q'" and "(r \ xvec) \* \\<^sub>P" and "(r \ xvec) \* \\<^sub>Q" + and Sr: "(set r) \ (set xvec) \ (set(r \ xvec))" and "distinctPerm r" + by(rule name_list_avoiding[where xvec=xvec and c="(\, P, Q, M, N, A\<^sub>P, A\<^sub>Q, \\<^sub>P, \\<^sub>Q, P', Q')"]) + (auto simp add: eqvts fresh_star_prod) + obtain q::"name prm" where "(q \ A\<^sub>Q) \* \" and "(q \ A\<^sub>Q) \* P" and "(q \ A\<^sub>Q) \* Q" and "(q \ A\<^sub>Q) \* M" + and "(q \ A\<^sub>Q) \* (r \ N)" and "(q \ A\<^sub>Q) \* (r \ xvec)" and "(q \ A\<^sub>Q) \* (r \ Q')" + and "(q \ A\<^sub>Q) \* (r \ P')" and "(q \ A\<^sub>Q) \* \\<^sub>P" and "(q \ A\<^sub>Q) \* A\<^sub>P" and "(q \ A\<^sub>Q) \* \\<^sub>Q" + and Sq: "set q \ set A\<^sub>Q \ set(q \ A\<^sub>Q)" + by(rule name_list_avoiding[where xvec=A\<^sub>Q and c="(\, P, Q, M, r \ N, r \ xvec, \\<^sub>Q, A\<^sub>P, \\<^sub>P, r \ Q', r \ P')"]) + (auto simp add: eqvts fresh_star_prod) + obtain p::"name prm" where "(p \ A\<^sub>P) \* \" and "(p \ A\<^sub>P) \* P" and "(p \ A\<^sub>P) \* Q" and "(p \ A\<^sub>P) \* M" + and "(p \ A\<^sub>P) \* (r \ N)" and "(p \ A\<^sub>P) \* (r \ xvec)" and "(p \ A\<^sub>P) \* (r \ Q')" + and "(p \ A\<^sub>P) \* (r \ P')" and "(p \ A\<^sub>P) \* \\<^sub>P" and "(p \ A\<^sub>P) \* \\<^sub>Q" and "(p \ A\<^sub>P) \* A\<^sub>Q" + and "(p \ A\<^sub>P) \* (q \ A\<^sub>Q)" and Sp: "(set p) \ (set A\<^sub>P) \ (set(p \ A\<^sub>P))" + by(rule name_list_avoiding[where xvec=A\<^sub>P and c="(\, P, Q, M, r \ N, r \ xvec, A\<^sub>Q, q \ A\<^sub>Q, \\<^sub>Q, \\<^sub>P, r \ Q', r \ P')"]) + (auto simp add: eqvts fresh_star_prod) + + have FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" by fact + have FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" by fact + + from Sp \A\<^sub>P \* M\ \(p \ A\<^sub>P) \* M\ + have "(p \ M) = M" + by simp + from Sq \A\<^sub>Q \* M\ \(q \ A\<^sub>Q) \* M\ + have "(q \ M) = M" + by simp + + from \A\<^sub>P \* Q\ FrQ \A\<^sub>P \* A\<^sub>Q\ have "A\<^sub>P \* \\<^sub>Q" + by(auto dest: extractFrameFreshChain) + from \A\<^sub>Q \* P\ FrP \A\<^sub>P \* A\<^sub>Q\ have "A\<^sub>Q \* \\<^sub>P" + by(auto dest: extractFrameFreshChain) + from \(r \ xvec) \* A\<^sub>Q\ \(q \ A\<^sub>Q) \* (r \ xvec)\ \(r \ xvec) \* A\<^sub>Q\ Sq have "(r \ xvec) \* (q \ A\<^sub>Q)" + by(simp add: freshChainSimps) + + from \\ \ \\<^sub>Q \ P \\M\\*xvec\\N\ \ P'\ Sr \(r \ xvec) \* N\ \(r \ xvec) \* P'\ + have "\ \ \\<^sub>Q \ P \\M\\*(r \ xvec)\\(r \ N)\ \ (r \ P')" + by(simp add: boundOutputChainAlpha'' create_residual.simps) + then have "(q \ (\ \ \\<^sub>Q)) \ P \\(q \ M)\\*(r \ xvec)\\(r \ N)\ \ (r \ P')" using Sq \A\<^sub>Q \* P\ \(q \ A\<^sub>Q) \* P\ \(r \ xvec) \* M\ \(r \ xvec) \* A\<^sub>Q\ \(r \ xvec) \* (q \ A\<^sub>Q)\ + by(fastforce intro: broutputPermFrameSubject) + then have "\ \ (q \ \\<^sub>Q) \ P \\(q \ M)\\*(r \ xvec)\\(r \ N)\ \ (r \ P')" using Sq \A\<^sub>Q \* \\ \(q \ A\<^sub>Q) \* \\ + by(simp add: eqvts) + with \(q \ M) = M\ have PTrans: "\ \ (q \ \\<^sub>Q) \ P \\M\\*(r \ xvec)\\(r \ N)\ \ (r \ P')" by simp + moreover then have "distinct(r \ xvec)" by(force dest: boundOutputDistinct) + + moreover from \extractFrame P = \A\<^sub>P, \\<^sub>P\\ Sp \(p \ A\<^sub>P) \* \\<^sub>P\ + have FrP: "extractFrame P = \(p \ A\<^sub>P), (p \ \\<^sub>P)\" + by(simp add: frameChainAlpha) + moreover from \distinct A\<^sub>P\ have "distinct(p \ A\<^sub>P)" by simp + + moreover from \\ \ \\<^sub>P \ Q \\M\N\ \ Q'\ Sr \distinctPerm r\ \xvec \* Q\ \(r \ xvec) \* Q\ + have "\ \ \\<^sub>P \ Q \\M\(r \ N)\ \ (r \ Q')" + by(rule brinputAlpha) + then have "(p \ (\ \ \\<^sub>P)) \ Q \\(p \ M)\(r \ N)\ \ (r \ Q')" using Sp \A\<^sub>P \* Q\ \(p \ A\<^sub>P) \* Q\ + by - (rule brinputPermFrameSubject, (assumption | simp)+) + then have QTrans: "\ \ (p \ \\<^sub>P) \ Q \\(p \ M)\(r \ N)\ \ (r \ Q')" using Sp \A\<^sub>P \* \\ \(p \ A\<^sub>P) \* \\ + by(simp add: eqvts) + with \(p \ M) = M\ have "\ \ (p \ \\<^sub>P) \ Q \\M\(r \ N)\ \ (r \ Q')" by simp + + moreover from \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ Sq \(q \ A\<^sub>Q) \* \\<^sub>Q\ + have FrQ: "extractFrame Q = \(q \ A\<^sub>Q), (q \ \\<^sub>Q)\" + by(simp add: frameChainAlpha) + moreover from \distinct A\<^sub>Q\ have "distinct(q \ A\<^sub>Q)" by simp + + moreover note \(p \ A\<^sub>P) \* \\ + moreover from \(p \ A\<^sub>P) \* A\<^sub>Q\ \(p \ A\<^sub>P) \* (q \ A\<^sub>Q)\ \(p \ A\<^sub>P) \* \\<^sub>Q\ Sq have "(p \ A\<^sub>P) \* (q \ \\<^sub>Q)" + by(simp add: freshChainSimps) + moreover note \(p \ A\<^sub>P) \* P\ \(p \ A\<^sub>P) \* M\ + moreover note \(p \ A\<^sub>P) \* (r \ N)\ \(p \ A\<^sub>P) \* (r \ P')\ \(p \ A\<^sub>P) \* Q\ \(p \ A\<^sub>P) \* (r \ Q')\ \(p \ A\<^sub>P) \* (q \ A\<^sub>Q)\ + \(p \ A\<^sub>P) \* (r \ xvec)\ \(q \ A\<^sub>Q) \* \\ + moreover from \(q \ A\<^sub>Q) \* A\<^sub>P\ \(p \ A\<^sub>P) \* A\<^sub>Q\ \(q \ A\<^sub>Q) \* \\<^sub>P\ \(p \ A\<^sub>P) \* (q \ A\<^sub>Q)\ Sp Sq have "(q \ A\<^sub>Q) \* (p \ \\<^sub>P)" + by(simp add: freshChainSimps) + moreover note \(q \ A\<^sub>Q) \* P\ \(q \ A\<^sub>Q) \* (r \ N)\\(q \ A\<^sub>Q) \* (r \ P')\ \(q \ A\<^sub>Q) \* Q\ \(q \ A\<^sub>Q) \* M\ + moreover note \(q \ A\<^sub>Q) \* (r \ Q')\ \(q \ A\<^sub>Q) \* (r \ xvec)\ \(r \ xvec) \* \\ + moreover from \(r \ xvec) \* A\<^sub>P\ \(p \ A\<^sub>P) \* (r \ xvec)\ \(r \ xvec) \* \\<^sub>P\ Sp have "(r \ xvec) \* (p \ \\<^sub>P)" + by(simp add: freshChainSimps) + moreover from \(r \ xvec) \* A\<^sub>Q\ \(q \ A\<^sub>Q) \* (r \ xvec)\ \(r \ xvec) \* \\<^sub>Q\ Sq have "(r \ xvec) \* (q \ \\<^sub>Q)" + by(simp add: freshChainSimps) + moreover note \(r \ xvec) \* P\ \(r \ xvec) \* M\ + moreover note \(r \ xvec) \* Q\ \(r \ xvec) \* M\ + ultimately have "\ \ P \ Q \\M\\*(r \ xvec)\\(r \ N)\ \ ((r \ P') \ (r \ Q'))" + by - (rule cBrComm2) + then have permuted: "\ \ P \ Q \\M\\*(r \ xvec)\\(r \ N)\ \ (r \ (P' \ Q'))" by simp + note \(r \ xvec) \* N\ + moreover from \(r \ xvec) \* P'\ \(r \ xvec) \* Q'\ + have "(r \ xvec) \* (P' \ Q')" by simp + moreover note Sr + moreover have "set xvec \ set xvec" by simp + ultimately have "\\*xvec\N \' (P' \ Q') = \\*(r \ xvec)\(r \ N) \' (r \ (P' \ Q'))" + by(rule boundOutputChainAlpha'') + then have "\M\\*xvec\\N\ \ (P' \ Q') = \M\\*(r \ xvec)\\(r \ N)\ \ (r \ (P' \ Q'))" + by(simp only: create_residual.simps) + with permuted show ?thesis + by simp + qed + } + note Goal = this + + note \\ \ \\<^sub>Q \ P \\M\\*xvec\\N\ \ P'\ \\ \ \\<^sub>P \ Q \\M\N\ \ Q'\ + moreover from \extractFrame P = \A\<^sub>P, \\<^sub>P\\ \A\<^sub>P \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* Q\ \A\<^sub>P \* M\ \A\<^sub>P \* A\<^sub>Q\ + obtain A\<^sub>P' where "extractFrame P = \A\<^sub>P', \\<^sub>P\" and "distinct A\<^sub>P'" and "A\<^sub>P' \* \" and "A\<^sub>P' \* P" and "A\<^sub>P' \* Q" and "A\<^sub>P' \* M" and "A\<^sub>P' \* A\<^sub>Q" + by - (rule distinctFrame[where C="(\, P, Q, M, A\<^sub>Q)"], auto) + moreover from \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ \A\<^sub>Q \* \\ \A\<^sub>Q \* P\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* M\ \A\<^sub>P' \* A\<^sub>Q\ + obtain A\<^sub>Q' where "extractFrame Q = \A\<^sub>Q', \\<^sub>Q\" and "distinct A\<^sub>Q'" and "A\<^sub>Q' \* \" and "A\<^sub>Q' \* P" and "A\<^sub>Q' \* Q" and "A\<^sub>Q' \* M" and "A\<^sub>P' \* A\<^sub>Q'" + by - (rule distinctFrame[where C="(\, P, Q, M, A\<^sub>P')"], auto) + ultimately show ?thesis using \xvec \* Q\ + by(metis Goal) +qed + +lemma BrClose: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and M :: 'a + and xvec :: "name list" + and N :: 'a + and P' :: "('a, 'b, 'c) psi" + and x :: name + +assumes "\ \ P \ \M\\*xvec\\N\ \ P'" + and "x \ supp M" + and "x \ \" + +shows "\ \ \\x\P \ \ \ \\x\(\\*xvec\P')" +proof - + obtain p where xvecFreshPsi: "((p::name prm) \ (xvec::name list)) \* \" + and xvecFreshM: "(p \ xvec) \* M" + and xvecFreshN: "(p \ xvec) \* N" + and xvecFreshP: "(p \ xvec) \* P" + and xvecFreshP: "(p \ xvec) \* P'" + and xvecFrechx: "(p \ xvec) \* x" + and S: "(set p) \ (set xvec) \ (set(p \ xvec))" + and dp: "distinctPerm p" + by(rule name_list_avoiding[where xvec=xvec and c="(\, M, N, P, P', x)"]) + (auto simp add: eqvts fresh_star_prod) + + obtain y::name where "y \ P" and "y \ xvec" and "y \ x" and "y \ N" + and "y \ (p \ xvec)" and "y \ (p \ P')" + and "y \ M" and "y \ \" and "y \ P'" + by(generate_fresh "name") (auto simp add: freshChainSimps) + + from \y \ (p \ xvec)\ \y \ (p \ P')\ + have yFreshRes: "y \ (\\*(p \ xvec)\(p \ P'))" + by(simp add: resChainFresh) + + from \\ \ P \ \M\\*xvec\\N\ \ P'\ S + \(p \ xvec) \* N\ \(p \ xvec) \* P'\ + have "\ \ P \ \M\\*(p \ xvec)\\(p \ N)\ \ (p \ P')" + by(simp add: alphaOutputResidual) + + then have "[(x, y)] \ (\ \ P \ \M\\*(p \ xvec)\\(p \ N)\ \ (p \ P'))" + by simp + with \(p \ xvec) \* x\ \y \ (p \ xvec)\ + \x \ \\ \y \ \\ + have pretrans: "\ \ ([(x, y)] \ P) \ \([(x, y)] \ M)\\*(p \ xvec)\\([(x, y)] \ (p \ N))\ \ ([(x, y)] \ (p \ P'))" + by(simp add: eqvts) + + moreover from \x \ supp M\ \y \ M\ + have "y \ supp ([(x, y)] \ M)" + by (metis fresh_bij fresh_def swap_simps) + + moreover from pretrans + have "distinct (p \ xvec)" + by(force dest: boundOutputDistinct) + + moreover note \(p \ xvec) \* \\ + moreover from \(p \ xvec) \* P\ \(p \ xvec) \* x\ \y \ (p \ xvec)\ + have "(p \ xvec) \* ([(x, y)] \ P)" by simp + moreover from \(p \ xvec) \* M\ \(p \ xvec) \* x\ \y \ (p \ xvec)\ + have "(p \ xvec) \* ([(x, y)] \ M)" by simp + moreover note \y \ \\ \y \ (p \ xvec)\ + + ultimately have "\ \ \\y\([(x, y)] \ P) \ \ \ \\y\(\\*(p \ xvec)\([(x, y)] \ (p \ P')))" + by(rule cBrClose) + with \(p \ xvec) \* x\ \y \ (p \ xvec)\ + have "\ \ \\y\([(x, y)] \ P) \ \ \ \\y\([(x, y)] \ (\\*(p \ xvec)\(p \ P')))" + by(simp add: eqvts) + with yFreshRes \y \ P\ + have "\ \ \\x\P \ \ \ \\x\(\\*(p \ xvec)\(p \ P'))" + by(simp add: alphaRes) + + with \(p \ xvec) \* P'\ S + show ?thesis + by(simp add: resChainAlpha) +qed + +lemma semanticsCasesAux[consumes 1, case_names cInput cBrInput cOutput cBrOutput cCase cPar1 cPar2 cComm1 cComm2 cBrMerge cBrComm1 cBrComm2 cBrClose cOpen cBrOpen cScope cBang]: + fixes cP :: "('a, 'b, 'c) psi" + and cRs :: "('a, 'b, 'c) residual" + and C :: "'f::fs_name" + and x :: name + +assumes "\ \ cP \ cRs" + and rInput: "\M K xvec N Tvec P. \cP = M\\*xvec N\.P; cRs = K\(N[xvec::=Tvec])\ \ P[xvec::=Tvec]; + \ \ M \ K; distinct xvec; set xvec \ supp N; length xvec=length Tvec; + xvec \* Tvec; xvec \* \; xvec \* M; xvec \* K; xvec \* C\ \ Prop" + and rBrInput: "\M K xvec N Tvec P. \cP = M\\*xvec N\.P; cRs = \K\(N[xvec::=Tvec])\ \ P[xvec::=Tvec]; + \ \ K \ M; distinct xvec; set xvec \ supp N; length xvec=length Tvec; + xvec \* Tvec; xvec \* \; xvec \* M; xvec \* K; xvec \* C\ \ Prop" + and rOutput: "\M K N P. \cP = M\N\.P; cRs = K\N\ \ P; \ \ M \ K\ \ Prop" + and rBrOutput: "\M K N P. \cP = M\N\.P; cRs = \K\N\ \ P; \ \ M \ K\ \ Prop" + and rCase: "\Cs P \. \cP = Cases Cs; \ \ P \ cRs; (\, P) \ set Cs; \ \ \; guarded P\ \ Prop" + +and rPar1: "\\\<^sub>Q P \ P' Q A\<^sub>Q. \cP = P \ Q; cRs = \ \ (P' \ Q); + (\ \ \\<^sub>Q) \ P \ (\ \ P'); extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + A\<^sub>Q \* P; A\<^sub>Q \* Q; A\<^sub>Q \* \; A\<^sub>Q \* \; A\<^sub>Q \* C; A\<^sub>Q \* P'; bn \ \* \; bn \ \* \\<^sub>Q; + bn \ \* Q; bn \ \* P; bn \ \* subject \; bn \ \* C; distinct(bn \)\ \ + Prop" +and rPar2: "\\\<^sub>P Q \ Q' P A\<^sub>P. \cP = P \ Q; cRs = \ \ (P \ Q'); + (\ \ \\<^sub>P) \ Q \\ \ Q'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + A\<^sub>P \* P; A\<^sub>P \* Q; A\<^sub>P \* \; A\<^sub>P \* \; A\<^sub>P \* C; + A\<^sub>P \* Q'; bn \ \* \; bn \ \* \\<^sub>P; bn \ \* P; bn \ \* Q; bn \ \* subject \; bn \ \* C; distinct(bn \)\ \ Prop" +and rComm1: "\\\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q K xvec Q' A\<^sub>Q. + \cP = P \ Q; cRs = \ \ \\*xvec\P' \ Q'; + \ \ \\<^sub>Q \ P \M\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \ \ \\<^sub>P \ Q \ K\\*xvec\\N\ \ Q'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + \ \ \\<^sub>P \ \\<^sub>Q \ M \ K; A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* N; + A\<^sub>P \* P'; A\<^sub>P \* Q; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* xvec; A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; + A\<^sub>Q \* P; A\<^sub>Q \* K; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* Q; A\<^sub>Q \* Q'; A\<^sub>Q \* xvec; + xvec \* \; xvec \* \\<^sub>P; xvec \* \\<^sub>Q; xvec \* P; xvec \* M; xvec \* Q; + xvec \* K; A\<^sub>P \* C; A\<^sub>Q \* C; xvec \* C; distinct xvec\ \ Prop" +and rComm2: "\\\<^sub>Q P M xvec N P' A\<^sub>P \\<^sub>P Q K Q' A\<^sub>Q. + \cP = P \ Q; cRs = \ \ \\*xvec\P' \ Q'; + \ \ \\<^sub>Q \ P \M\\*xvec\\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \ \ \\<^sub>P \ Q \ K\N\ \ Q'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + \ \ \\<^sub>P \ \\<^sub>Q \ M \ K; A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* N; + A\<^sub>P \* P'; A\<^sub>P \* Q; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* xvec; A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; + A\<^sub>Q \* P; A\<^sub>Q \* K; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* Q; A\<^sub>Q \* Q'; A\<^sub>Q \* xvec; + xvec \* \; xvec \* \\<^sub>P; xvec \* \\<^sub>Q; xvec \* P; xvec \* M; xvec \* Q; + xvec \* K; A\<^sub>P \* C; A\<^sub>Q \* C; xvec \* C; distinct xvec\ \ Prop" +and rBrMerge: "\\\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q Q' A\<^sub>Q. + \cP = (P \ Q); cRs = \M\N\ \ (P' \ Q'); + \ \ \\<^sub>Q \ P \ \M\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \ \ \\<^sub>P \ Q \ \M\N\ \ Q'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* N; A\<^sub>P \* P'; + A\<^sub>P \* Q; A\<^sub>P \* Q'; A\<^sub>P \* M; A\<^sub>P \* A\<^sub>Q; + A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; A\<^sub>Q \* P; A\<^sub>Q \* N; A\<^sub>Q \* P'; + A\<^sub>Q \* Q; A\<^sub>Q \* Q'; A\<^sub>Q \* M; A\<^sub>P \* C; A\<^sub>Q \* C\ \ Prop" +and rBrComm1: "\\\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q xvec Q' A\<^sub>Q. + \cP = P \ Q; cRs = \M\\*xvec\\N\ \ (P' \ Q'); + \ \ \\<^sub>Q \ P \\M\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \ \ \\<^sub>P \ Q \ \M\\*xvec\\N\ \ Q'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* N; + A\<^sub>P \* P'; A\<^sub>P \* Q; A\<^sub>P \* Q'; A\<^sub>P \* M; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* xvec; A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; + A\<^sub>Q \* P; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* M; A\<^sub>Q \* Q; A\<^sub>Q \* Q'; A\<^sub>Q \* xvec; + xvec \* \; xvec \* \\<^sub>P; xvec \* \\<^sub>Q; xvec \* P; xvec \* Q; xvec \* M; + A\<^sub>P \* C; A\<^sub>Q \* C; xvec \* C; distinct xvec\ \ Prop" +and rBrComm2: "\\\<^sub>Q P M xvec N P' A\<^sub>P \\<^sub>P Q Q' A\<^sub>Q. + \cP = P \ Q; cRs = \M\\*xvec\\N\ \ (P' \ Q'); + \ \ \\<^sub>Q \ P \\M\\*xvec\\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \ \ \\<^sub>P \ Q \ \M\N\ \ Q'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* N; + A\<^sub>P \* P'; A\<^sub>P \* Q; A\<^sub>P \* Q'; A\<^sub>P \* M; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* xvec; A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; + A\<^sub>Q \* P; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* M; A\<^sub>Q \* Q; A\<^sub>Q \* Q'; A\<^sub>Q \* xvec; + xvec \* \; xvec \* \\<^sub>P; xvec \* \\<^sub>Q; xvec \* P; xvec \* Q; xvec \* M; + A\<^sub>P \* C; A\<^sub>Q \* C; xvec \* C; distinct xvec\ \ Prop" +and rBrClose: "\P M xvec N P' x. + \cP = (\\x\P); cRs = \ \ \\x\(\\*xvec\P'); + x \ supp M; + \ \ P \ \M\\*xvec\\N\ \ P'; + distinct xvec; xvec \* \; xvec \* P; + xvec \* M; + x \ \; x \ xvec; + xvec \* C; x \ C\ \ Prop" +and rOpen: "\P M xvec yvec N P' x. + \cP = \\x\P; cRs = M\\*(xvec@x#yvec)\\N\ \ P'; + \ \ P \ M\\*(xvec@yvec)\\N\ \ P'; x \ supp N; x \ xvec; x \ yvec; x \ M; x \ \; distinct xvec; distinct yvec; + xvec \* \; xvec \* P; xvec \* M; xvec \* yvec; yvec \* \; yvec \* P; yvec \* M; xvec \* C; x \ C; yvec \* C\ \ + Prop" +and rBrOpen: "\P M xvec yvec N P' x. + \cP = \\x\P; cRs = \M\\*(xvec@x#yvec)\\N\ \ P'; + \ \ P \ \M\\*(xvec@yvec)\\N\ \ P'; x \ supp N; x \ xvec; x \ yvec; x \ M; x \ \; distinct xvec; distinct yvec; + xvec \* \; xvec \* P; xvec \* M; xvec \* yvec; yvec \* \; yvec \* P; yvec \* M; xvec \* C; x \ C; yvec \* C\ \ + Prop" +and rScope: "\P \ P' x. \cP = \\x\P; cRs = \ \ \\x\P'; + \ \ P \\ \ P'; x \ \; x \ \; x \ C; bn \ \* \; bn \ \* P; bn \ \* subject \; bn \ \* C; distinct(bn \)\ \ Prop" +and rBang: "\P. \cP = !P; \ \ P \ !P \ cRs; guarded P\ \ Prop" +shows Prop + using \\ \ cP \ cRs\ +proof(cases rule: semantics.cases) + case(cInput M K xvec N Tvec P) + obtain p::"name prm" where "(p \ xvec) \* \" and "(p \ xvec) \* M" and "(p \ xvec) \* N" and "(p \ xvec) \* K" + and "(p \ xvec) \* Tvec" and "(p \ xvec) \* P" and "(p \ xvec) \* C" + and S: "(set p) \ (set xvec) \ (set(p \ xvec))" and "distinctPerm p" + by(rule name_list_avoiding[where xvec=xvec and c="(\, M, K, N, P, C, Tvec)"]) + (auto simp add: eqvts fresh_star_prod) + from \cP = M\\*xvec N\.P\ \(p \ xvec) \* N\ \(p \ xvec) \* P\ S + have "cP = M\\*(p \ xvec) (p \ N)\.(p \ P)" + by(simp add: inputChainAlpha') + moreover from \cRs = K\(N[xvec::=Tvec])\ \ P[xvec::=Tvec]\ \(p \ xvec) \* N\ \(p \ xvec) \* P\ S \length xvec = length Tvec\ \distinctPerm p\ + have "cRs = K\((p \ N)[(p \ xvec)::=Tvec])\ \ (p \ P)[(p \ xvec)::=Tvec]" + by(auto simp add: substTerm.renaming renaming residualInject) + + moreover note \\ \ M \ K\ + moreover from \distinct xvec\ have "distinct(p \ xvec)" + by simp + moreover from \(set xvec) \ (supp N)\ have "(p \ (set xvec)) \ (p \ (supp N))" + by simp + then have "set(p \ xvec) \ supp(p \ N)" + by(simp add: eqvts) + moreover from \length xvec = length Tvec\ have "length(p \ xvec) = length Tvec" + by simp + ultimately show ?thesis using \(p \ xvec) \* Tvec\ \(p \ xvec) \* \\ \(p \ xvec) \* M\ \(p \ xvec) \* K\ + \(p \ xvec) \* C\ + by(rule rInput) +next + case(cBrInput K M xvec N Tvec P) + obtain p::"name prm" where "(p \ xvec) \* \" and "(p \ xvec) \* M" and "(p \ xvec) \* N" and "(p \ xvec) \* K" + and "(p \ xvec) \* Tvec" and "(p \ xvec) \* P" and "(p \ xvec) \* C" + and S: "(set p) \ (set xvec) \ (set(p \ xvec))" and "distinctPerm p" + by(rule name_list_avoiding[where xvec=xvec and c="(\, M, K, N, P, C, Tvec)"]) + (auto simp add: eqvts fresh_star_prod) + from \cP = M\\*xvec N\.P\ \(p \ xvec) \* N\ \(p \ xvec) \* P\ S + have "cP = M\\*(p \ xvec) (p \ N)\.(p \ P)" + by(simp add: inputChainAlpha') + moreover from \cRs = \K\(N[xvec::=Tvec])\ \ P[xvec::=Tvec]\ \(p \ xvec) \* N\ \(p \ xvec) \* P\ S \length xvec = length Tvec\ \distinctPerm p\ + have "cRs = \K\((p \ N)[(p \ xvec)::=Tvec])\ \ (p \ P)[(p \ xvec)::=Tvec]" + by(auto simp add: substTerm.renaming renaming residualInject) + + moreover note \\ \ K \ M\ + moreover from \distinct xvec\ have "distinct(p \ xvec)" + by simp + moreover from \(set xvec) \ (supp N)\ have "(p \ (set xvec)) \ (p \ (supp N))" + by simp + then have "set(p \ xvec) \ supp(p \ N)" + by(simp add: eqvts) + moreover from \length xvec = length Tvec\ have "length(p \ xvec) = length Tvec" + by simp + ultimately show ?thesis using \(p \ xvec) \* Tvec\ \(p \ xvec) \* \\ \(p \ xvec) \* M\ \(p \ xvec) \* K\ + \(p \ xvec) \* C\ + by(simp add: rBrInput) +next + case(Output M K N P) + then show ?thesis by(rule rOutput) +next + case(BrOutput M K N P) + then show ?thesis by(rule rBrOutput) +next + case(Case P \ Cs) + then show ?thesis by(rule rCase) +next + case(cPar1 \\<^sub>Q P \ P' Q A\<^sub>Q) + obtain q::"name prm" where "(bn(q \ \)) \* \" and "(bn(q \ \)) \* P" and "(bn(q \ \)) \* Q" + and "(bn(q \ \)) \* \" and "(bn(q \ \)) \* A\<^sub>Q" and "(bn(q \ \)) \* P'" and "(bn(q \ \)) \* \\<^sub>Q" + and "distinctPerm q" + and "(bn(q \ \)) \* C" and Sq: "(set q) \ set(bn \) \ (set(bn(q \ \)))" + by(rule name_list_avoiding[where xvec="bn \" and c="(\, P, Q, \, A\<^sub>Q, \\<^sub>Q, P', C)"]) (auto simp add: eqvts) + obtain p::"name prm" where "(p \ A\<^sub>Q) \* \" and "(p \ A\<^sub>Q) \* P" and "(p \ A\<^sub>Q) \* Q" + and "(p \ A\<^sub>Q) \* \" and "(p \ A\<^sub>Q) \* (q \ \)" and "(p \ A\<^sub>Q) \* P'" + and "(p \ A\<^sub>Q) \* (q \ P')" and "(p \ A\<^sub>Q) \* \\<^sub>Q" and "(p \ A\<^sub>Q) \* C" + and Sp: "(set p) \ (set A\<^sub>Q) \ (set(p \ A\<^sub>Q))" and "distinctPerm p" + by(rule name_list_avoiding[where xvec=A\<^sub>Q and c="(\, P, Q, \, q \ \, P', (q \ P'), \\<^sub>Q, C)"]) auto + from \A\<^sub>Q \* \\ \bn(q \ \) \* A\<^sub>Q\ Sq \distinctPerm q\ have "A\<^sub>Q \* (q \ \)" + by(subst fresh_star_bij[symmetric, of _ _ q]) (simp add: eqvts) + from \bn \ \* subject \\ \distinctPerm q\ have "bn(q \ \) \* subject(q \ \)" + by(subst fresh_star_bij[symmetric, of _ _ q]) (simp add: eqvts) + from \distinct(bn \)\ \distinctPerm q\ have "distinct(bn(q \ \))" + by(subst distinctClosed[symmetric, of _ q]) (simp add: eqvts) + note \cP = P \ Q\ + + moreover from \cRs = \ \ (P' \ Q)\ \bn \ \* subject \\ \(bn(q \ \)) \* \\ \(bn(q \ \)) \* P'\ \(bn(q \ \)) \* Q\ \bn \ \* Q\ Sq + have "cRs = (q \ \) \ (q \ P') \ Q" + by(force simp add: residualAlpha) + moreover from \\ \ \\<^sub>Q \ P \\ \ P'\ \bn \ \* subject \\ \(bn(q \ \)) \* \\ \(bn(q \ \)) \* P'\ Sq + have Trans: "\ \ \\<^sub>Q \ P \(q \ \) \ (q \ P')" + by(force simp add: residualAlpha) + then have "A\<^sub>Q \* (q \ P')" using \bn(q \ \) \* subject(q \ \)\ \distinct(bn(q \ \))\ \A\<^sub>Q \* P\ \A\<^sub>Q \* (q \ \)\ + by(force dest: freeFreshChainDerivative) + + from Trans have "(p \ (\ \ \\<^sub>Q)) \ (p \ P) \p \ ((q \ \) \ (q \ P'))" + by(rule semantics.eqvt) + with \A\<^sub>Q \* \\ \A\<^sub>Q \* P\ \A\<^sub>Q \* (q \ \)\ \A\<^sub>Q \* (q \ P')\ \(bn(q \ \)) \* A\<^sub>Q\ \(p \ A\<^sub>Q) \* (q \ \)\ + \(p \ A\<^sub>Q) \* \\ \(p \ A\<^sub>Q) \* P\ \(p \ A\<^sub>Q) \* (q \ P')\ Sp + have "\ \ (p \ \\<^sub>Q) \ P \(q \ \) \ (q \ P')" by(simp add: eqvts) + moreover from \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ \(p \ A\<^sub>Q) \* \\<^sub>Q\ Sp have "extractFrame Q = \(p \ A\<^sub>Q), (p \ \\<^sub>Q)\" + by(simp add: frameChainAlpha' eqvts) + moreover from \(bn(q \ \)) \* \\<^sub>Q\ \(bn(q \ \)) \* A\<^sub>Q\ \(p \ A\<^sub>Q) \* (q \ \)\ Sp have "(bn(q \ \)) \* (p \ \\<^sub>Q)" + by(simp add: freshAlphaPerm) + moreover from \distinct A\<^sub>Q\ have "distinct(p \ A\<^sub>Q)" by simp + ultimately show ?thesis + using \(p \ A\<^sub>Q) \* P\ \(p \ A\<^sub>Q) \* Q\ \(p \ A\<^sub>Q) \* \\ \(p \ A\<^sub>Q) \* (q \ \)\ + \(p \ A\<^sub>Q) \* (q \ P')\ \(bn(q \ \)) \* \\ \(bn(q \ \)) \* Q\ \(bn(q \ \)) \* P\ + \(bn(q \ \)) \* C\ \(p \ A\<^sub>Q) \* C\ \bn (q \ \) \* subject (q \ \)\ \distinct(bn(q \ \))\ + by(metis rPar1) +next + case(cPar2 \\<^sub>P Q \ Q' P A\<^sub>P) + obtain q::"name prm" where "(bn(q \ \)) \* \" and "(bn(q \ \)) \* P" and "(bn(q \ \)) \* Q" + and "(bn(q \ \)) \* \" and "(bn(q \ \)) \* A\<^sub>P" and "(bn(q \ \)) \* Q'" and "(bn(q \ \)) \* \\<^sub>P" + and "distinctPerm q" + and "(bn(q \ \)) \* C" and Sq: "(set q) \ set(bn \) \ (set(bn(q \ \)))" + by (rule name_list_avoiding[where xvec="bn \" and c="(\, P, Q, \, A\<^sub>P, \\<^sub>P, Q', C)"]) (auto simp add: eqvts) + obtain p::"name prm" where "(p \ A\<^sub>P) \* \" and "(p \ A\<^sub>P) \* P" and "(p \ A\<^sub>P) \* Q" + and "(p \ A\<^sub>P) \* \" and "(p \ A\<^sub>P) \* (q \ \)" and "(p \ A\<^sub>P) \* Q'" + and "(p \ A\<^sub>P) \* (q \ Q')" and "(p \ A\<^sub>P) \* \\<^sub>P" and "(p \ A\<^sub>P) \* C" + and Sp: "(set p) \ (set A\<^sub>P) \ (set(p \ A\<^sub>P))" and "distinctPerm p" + by(rule name_list_avoiding[where xvec=A\<^sub>P and c="(\, P, Q, \, q \ \, Q', (q \ Q'), \\<^sub>P, C)"]) auto + from \A\<^sub>P \* \\ \bn(q \ \) \* A\<^sub>P\ Sq \distinctPerm q\ have "A\<^sub>P \* (q \ \)" + by(subst fresh_star_bij[symmetric, of _ _ q]) (simp add: eqvts) + from \bn \ \* subject \\ \distinctPerm q\ have "bn(q \ \) \* subject(q \ \)" + by(subst fresh_star_bij[symmetric, of _ _ q]) (simp add: eqvts) + from \distinct(bn \)\ \distinctPerm q\ have "distinct(bn(q \ \))" + by(subst distinctClosed[symmetric, of _ q]) (simp add: eqvts) + note \cP = P \ Q\ + + moreover from \cRs = \ \ (P \ Q')\ \bn \ \* subject \\ \(bn(q \ \)) \* \\ \(bn(q \ \)) \* Q'\ \(bn(q \ \)) \* P\ \bn \ \* P\ Sq + have "cRs = (q \ \) \ P \ (q \ Q')" + by(force simp add: residualAlpha) + moreover from \\ \ \\<^sub>P \ Q \\ \ Q'\ \bn \ \* subject \\ \(bn(q \ \)) \* \\ \(bn(q \ \)) \* Q'\ Sq + have Trans: "\ \ \\<^sub>P \ Q \(q \ \) \ (q \ Q')" + by(force simp add: residualAlpha) + then have "A\<^sub>P \* (q \ Q')" using \bn(q \ \) \* subject(q \ \)\ \distinct(bn(q \ \))\ \A\<^sub>P \* Q\ \A\<^sub>P \* (q \ \)\ + by(auto dest: freeFreshChainDerivative) + + from Trans have "(p \ (\ \ \\<^sub>P)) \ (p \ Q) \p \ ((q \ \) \ (q \ Q'))" + by(rule semantics.eqvt) + with \A\<^sub>P \* \\ \A\<^sub>P \* Q\ \A\<^sub>P \* (q \ \)\ \A\<^sub>P \* (q \ Q')\ \(bn(q \ \)) \* A\<^sub>P\ \(p \ A\<^sub>P) \* (q \ \)\ + \(p \ A\<^sub>P) \* \\ \(p \ A\<^sub>P) \* Q\ \(p \ A\<^sub>P) \* (q \ Q')\ Sp + have "\ \ (p \ \\<^sub>P) \ Q \(q \ \) \ (q \ Q')" by(simp add: eqvts) + moreover from \extractFrame P = \A\<^sub>P, \\<^sub>P\\ \(p \ A\<^sub>P) \* \\<^sub>P\ Sp have "extractFrame P = \(p \ A\<^sub>P), (p \ \\<^sub>P)\" + by(simp add: frameChainAlpha' eqvts) + moreover from \(bn(q \ \)) \* \\<^sub>P\ \(bn(q \ \)) \* A\<^sub>P\ \(p \ A\<^sub>P) \* (q \ \)\ Sp have "(bn(q \ \)) \* (p \ \\<^sub>P)" + by(simp add: freshAlphaPerm) + moreover from \distinct A\<^sub>P\ have "distinct(p \ A\<^sub>P)" by simp + ultimately show ?thesis + using \(p \ A\<^sub>P) \* P\ \(p \ A\<^sub>P) \* Q\ \(p \ A\<^sub>P) \* \\ \(p \ A\<^sub>P) \* (q \ \)\ + \(p \ A\<^sub>P) \* (q \ Q')\ \(bn(q \ \)) \* \\ \(bn(q \ \)) \* Q\ \(bn(q \ \)) \* P\ + \(bn(q \ \)) \* C\ \(p \ A\<^sub>P) \* C\ \bn (q \ \) \* subject (q \ \)\ \distinct(bn(q \ \))\ + by(metis rPar2) +next + case(cComm1 \\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q K xvec Q' A\<^sub>Q) + obtain r::"name prm" where "(r \ xvec) \* \" and "(r \ xvec) \* P" and "(r \ xvec) \* Q" and "(r \ xvec) \* M" + and "(r \ xvec) \* K" and "(r \ xvec) \* N" and "(r \ xvec) \* A\<^sub>P" and "(r \ xvec) \* A\<^sub>Q" + and "(r \ xvec) \* P'" and "(r \ xvec) \* Q'" and "(r \ xvec) \* \\<^sub>P" and "(r \ xvec) \* \\<^sub>Q" + and "(r \ xvec) \* C" and Sr: "(set r) \ (set xvec) \ (set(r \ xvec))" and "distinctPerm r" + by(rule name_list_avoiding[where xvec=xvec and c="(\, P, Q, M, K, N, A\<^sub>P, A\<^sub>Q, \\<^sub>P, \\<^sub>Q, P', Q', C)"]) + (auto simp add: eqvts) + + obtain q::"name prm" where "(q \ A\<^sub>Q) \* \" and "(q \ A\<^sub>Q) \* P" and "(q \ A\<^sub>Q) \* Q" and "(q \ A\<^sub>Q) \* K" + and "(q \ A\<^sub>Q) \* N" and "(q \ A\<^sub>Q) \* xvec" and "(q \ A\<^sub>Q) \* Q'" and "(q \ A\<^sub>Q) \* P'" + and "(q \ A\<^sub>Q) \* \\<^sub>P" and "(q \ A\<^sub>Q) \* A\<^sub>P" and "(q \ A\<^sub>Q) \* \\<^sub>Q" and "(q \ A\<^sub>Q) \* (r \ xvec)" + and "(q \ A\<^sub>Q) \* C" and Sq: "(set q) \ (set A\<^sub>Q) \ (set(q \ A\<^sub>Q))" + by(rule name_list_avoiding[where xvec=A\<^sub>Q and c="(\, P, Q, K, N, xvec, r \ xvec, \\<^sub>Q, A\<^sub>P, \\<^sub>P, Q', P', C)"]) clarsimp + + obtain p::"name prm" where "(p \ A\<^sub>P) \* \" and "(p \ A\<^sub>P) \* P" and "(p \ A\<^sub>P) \* Q" and "(p \ A\<^sub>P) \* M" + and "(p \ A\<^sub>P) \* N" and "(p \ A\<^sub>P) \* xvec" and "(p \ A\<^sub>P) \* Q'" and "(p \ A\<^sub>P) \* A\<^sub>Q" + and "(p \ A\<^sub>P) \* P'" and "(p \ A\<^sub>P) \* \\<^sub>P" and "(p \ A\<^sub>P) \* \\<^sub>Q" and "(p \ A\<^sub>P) \* (q \ A\<^sub>Q)" + and "(p \ A\<^sub>P) \* C" and "(p \ A\<^sub>P) \* (r \ xvec)" and Sp: "(set p) \ (set A\<^sub>P) \ (set(p \ A\<^sub>P))" + by(rule name_list_avoiding[where xvec=A\<^sub>P and c="(\, P, Q, M, N, xvec, r \ xvec, A\<^sub>Q, q \ A\<^sub>Q, \\<^sub>Q, \\<^sub>P, Q', P', C)"]) + (auto simp add: eqvts fresh_star_prod) + + have FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" by fact + have FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" by fact + + from \A\<^sub>P \* Q\ FrQ \A\<^sub>P \* A\<^sub>Q\ have "A\<^sub>P \* \\<^sub>Q" + by(auto dest: extractFrameFreshChain) + from \A\<^sub>Q \* P\ FrP \A\<^sub>P \* A\<^sub>Q\ have "A\<^sub>Q \* \\<^sub>P" + by(auto dest: extractFrameFreshChain) + note \cP = P \ Q\ + moreover from \(r \ xvec) \* P'\ \(r \ xvec) \* Q'\ have "(r \ xvec) \* (P' \ Q')" + by simp + with \cRs = \ \ \\*xvec\(P' \ Q')\ \(r \ xvec) \* N\ Sr + have "cRs = \ \ \\*(r \ xvec)\(r \ (P' \ Q'))" by(simp add: resChainAlpha residualInject) + then have "cRs = \ \ \\*(r \ xvec)\((r \ P') \ (r \ Q'))" by simp + + moreover from \\ \ \\<^sub>Q \ P \M\N\ \ P'\ Sr \distinctPerm r\ \xvec \* P\ \(r \ xvec) \* P\ + have "\ \ \\<^sub>Q \ P \M\(r \ N)\ \ (r \ P')" + by(rule inputAlpha) + then have "(q \ (\ \ \\<^sub>Q)) \ P \(q \ M)\(r \ N)\ \ (r \ P')" using Sq \A\<^sub>Q \* P\ \(q \ A\<^sub>Q) \* P\ + by - (rule inputPermFrameSubject, (assumption | simp)+) + then have PTrans: "\ \ (q \ \\<^sub>Q) \ P \(q \ M)\(r \ N)\ \ (r \ P')" using Sq \A\<^sub>Q \* \\ \(q \ A\<^sub>Q) \* \\ + by(simp add: eqvts) + + moreover from \extractFrame P = \A\<^sub>P, \\<^sub>P\\ Sp \(p \ A\<^sub>P) \* \\<^sub>P\ + have FrP: "extractFrame P = \(p \ A\<^sub>P), (p \ \\<^sub>P)\" + by(simp add: frameChainAlpha) + moreover from \distinct A\<^sub>P\ have "distinct(p \ A\<^sub>P)" by simp + + moreover from \\ \ \\<^sub>P \ Q \K\\*xvec\\N\ \ Q'\ Sr \(r \ xvec) \* N\ \(r \ xvec) \* Q'\ + have "\ \ \\<^sub>P \ Q \K\\*(r \ xvec)\\(r \ N)\ \ (r \ Q')" + by(simp add: boundOutputChainAlpha'' residualInject) + then have "(p \ (\ \ \\<^sub>P)) \ Q \(p \ K)\\*(r \ xvec)\\(r \ N)\ \ (r \ Q')" using Sp \A\<^sub>P \* Q\ \(p \ A\<^sub>P) \* Q\ \(r \ xvec) \* K\\(p \ A\<^sub>P) \* (r \ xvec)\ \(r \ xvec) \* A\<^sub>P\ + by(fastforce intro: outputPermFrameSubject) + then have QTrans: "\ \ (p \ \\<^sub>P) \ Q \(p \ K)\\*(r \ xvec)\\(r \ N)\ \ (r \ Q')" using Sp \A\<^sub>P \* \\ \(p \ A\<^sub>P) \* \\ + by(simp add: eqvts) + + moreover from \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ Sq \(q \ A\<^sub>Q) \* \\<^sub>Q\ + have FrQ: "extractFrame Q = \(q \ A\<^sub>Q), (q \ \\<^sub>Q)\" + by(simp add: frameChainAlpha) + moreover from \distinct A\<^sub>Q\ have "distinct(q \ A\<^sub>Q)" by simp + + moreover from \\ \ \\<^sub>P \ \\<^sub>Q \ M \ K\ have "(p \ q \ (\ \ \\<^sub>P \ \\<^sub>Q)) \ (p \ q \ M) \ (p \ q \ K)" + by(metis chanEqClosed) + with \A\<^sub>P \* \\ \(p \ A\<^sub>P) \* \\ \A\<^sub>Q \* \\ \(q \ A\<^sub>Q) \* \\ \A\<^sub>Q \* \\<^sub>P\ \(q \ A\<^sub>Q) \* \\<^sub>P\ + \A\<^sub>P \* \\<^sub>Q\ \(p \ A\<^sub>P) \* \\<^sub>Q\ \A\<^sub>P \* M\ \(p \ A\<^sub>P) \* M\ \(q \ A\<^sub>Q) \* A\<^sub>P\ \(p \ A\<^sub>P) \* (q \ A\<^sub>Q)\ + \A\<^sub>Q \* K\ \(q \ A\<^sub>Q) \* K\ \A\<^sub>P \* A\<^sub>Q\ \(p \ A\<^sub>P) \* A\<^sub>Q\ Sp Sq + have "\ \ (p \ \\<^sub>P) \ (q \ \\<^sub>Q) \ (q \ M) \ (p \ K)" + by(simp add: eqvts freshChainSimps) + moreover note \(p \ A\<^sub>P) \* \\ + moreover from \(p \ A\<^sub>P) \* A\<^sub>Q\ \(p \ A\<^sub>P) \* (q \ A\<^sub>Q)\ \(p \ A\<^sub>P) \* \\<^sub>Q\ Sq have "(p \ A\<^sub>P) \* (q \ \\<^sub>Q)" + by(simp add: freshChainSimps) + moreover note \(p \ A\<^sub>P) \* P\ + moreover from \(p \ A\<^sub>P) \* A\<^sub>Q\ \(p \ A\<^sub>P) \* (q \ A\<^sub>Q)\ \(p \ A\<^sub>P) \* M\ Sq have "(p \ A\<^sub>P) \* (q \ M)" + by(simp add: freshChainSimps) + moreover from \(p \ A\<^sub>P) \* xvec\ \(p \ A\<^sub>P) \* (r \ xvec)\ \(p \ A\<^sub>P) \* N\ Sr have "(p \ A\<^sub>P) \* (r \ N)" + by(simp add: freshChainSimps) + moreover from \(p \ A\<^sub>P) \* xvec\ \(p \ A\<^sub>P) \* (r \ xvec)\ \(p \ A\<^sub>P) \* P'\ Sr have "(p \ A\<^sub>P) \* (r \ P')" + by(simp add: freshChainSimps) + moreover note \(p \ A\<^sub>P) \* Q\ + moreover from \(p \ A\<^sub>P) \* xvec\ \(p \ A\<^sub>P) \* (r \ xvec)\ \(p \ A\<^sub>P) \* Q'\ Sr have "(p \ A\<^sub>P) \* (r \ Q')" + by(simp add: freshChainSimps) + moreover note \(p \ A\<^sub>P) \* (q \ A\<^sub>Q)\ \(p \ A\<^sub>P) \* (r \ xvec)\ \(q \ A\<^sub>Q) \* \\ + moreover from \(q \ A\<^sub>Q) \* A\<^sub>P\ \(p \ A\<^sub>P) \* A\<^sub>Q\ \(q \ A\<^sub>Q) \* \\<^sub>P\ \(p \ A\<^sub>P) \* (q \ A\<^sub>Q)\ Sp Sq have "(q \ A\<^sub>Q) \* (p \ \\<^sub>P)" + by(simp add: freshChainSimps) + moreover note \(q \ A\<^sub>Q) \* P\ + moreover from \(q \ A\<^sub>Q) \* A\<^sub>P\ \(p \ A\<^sub>P) \* A\<^sub>Q\ \(q \ A\<^sub>Q) \* K\ \(p \ A\<^sub>P) \* (q \ A\<^sub>Q)\ Sp Sq have "(q \ A\<^sub>Q) \* (p \ K)" + by(simp add: freshChainSimps) + moreover from \(q \ A\<^sub>Q) \* xvec\ \(q \ A\<^sub>Q) \* (r \ xvec)\ \(q \ A\<^sub>Q) \* N\ Sr have "(q \ A\<^sub>Q) \* (r \ N)" + by(simp add: freshChainSimps) + moreover from \(q \ A\<^sub>Q) \* xvec\ \(q \ A\<^sub>Q) \* (r \ xvec)\ \(q \ A\<^sub>Q) \* P'\ Sr have "(q \ A\<^sub>Q) \* (r \ P')" + by(simp add: freshChainSimps) + moreover note \(q \ A\<^sub>Q) \* Q\ + moreover from \(q \ A\<^sub>Q) \* xvec\ \(q \ A\<^sub>Q) \* (r \ xvec)\ \(q \ A\<^sub>Q) \* Q'\ Sr have "(q \ A\<^sub>Q) \* (r \ Q')" + by(simp add: freshChainSimps) + moreover note \(q \ A\<^sub>Q) \* (r \ xvec)\ \(r \ xvec) \* \\ + moreover from \(r \ xvec) \* A\<^sub>P\ \(p \ A\<^sub>P) \* (r \ xvec)\ \(r \ xvec) \* \\<^sub>P\ Sp have "(r \ xvec) \* (p \ \\<^sub>P)" + by(simp add: freshChainSimps) + moreover from \(r \ xvec) \* A\<^sub>Q\ \(q \ A\<^sub>Q) \* (r \ xvec)\ \(r \ xvec) \* \\<^sub>Q\ Sq have "(r \ xvec) \* (q \ \\<^sub>Q)" + by(simp add: freshChainSimps) + moreover note \(r \ xvec) \* P\ + moreover from \(r \ xvec) \* A\<^sub>Q\ \(q \ A\<^sub>Q) \* (r \ xvec)\ \(r \ xvec) \* M\ Sq have "(r \ xvec) \* (q \ M)" + by(simp add: freshChainSimps) + moreover note \(r \ xvec) \* Q\ + moreover from \(r \ xvec) \* A\<^sub>P\ \(p \ A\<^sub>P) \* (r \ xvec)\ \(r \ xvec) \* K\ Sp have "(r \ xvec) \* (p \ K)" + by(simp add: freshChainSimps) + moreover note \(p \ A\<^sub>P) \* C\ \(q \ A\<^sub>Q) \* C\ \(r \ xvec) \* C\ + moreover from \distinct xvec\ have "distinct(r \ xvec)" by simp + ultimately show ?thesis by(rule rComm1) +next + case(cComm2 \\<^sub>Q P M xvec N P' A\<^sub>P \\<^sub>P Q K Q' A\<^sub>Q) + obtain r::"name prm" where "(r \ xvec) \* \" and "(r \ xvec) \* P" and "(r \ xvec) \* Q" and "(r \ xvec) \* M" + and "(r \ xvec) \* K" and "(r \ xvec) \* N" and "(r \ xvec) \* A\<^sub>P" and "(r \ xvec) \* A\<^sub>Q" + and "(r \ xvec) \* P'" and "(r \ xvec) \* Q'" and "(r \ xvec) \* \\<^sub>P" and "(r \ xvec) \* \\<^sub>Q" + and "(r \ xvec) \* C" and Sr: "(set r) \ (set xvec) \ (set(r \ xvec))" and "distinctPerm r" + by(rule name_list_avoiding[where xvec=xvec and c="(\, P, Q, M, K, N, A\<^sub>P, A\<^sub>Q, \\<^sub>P, \\<^sub>Q, P', Q', C)"]) + (auto simp add: eqvts) + + obtain q::"name prm" where "(q \ A\<^sub>Q) \* \" and "(q \ A\<^sub>Q) \* P" and "(q \ A\<^sub>Q) \* Q" and "(q \ A\<^sub>Q) \* K" + and "(q \ A\<^sub>Q) \* N" and "(q \ A\<^sub>Q) \* xvec" and "(q \ A\<^sub>Q) \* Q'" and "(q \ A\<^sub>Q) \* P'" + and "(q \ A\<^sub>Q) \* \\<^sub>P" and "(q \ A\<^sub>Q) \* A\<^sub>P" and "(q \ A\<^sub>Q) \* \\<^sub>Q" and "(q \ A\<^sub>Q) \* (r \ xvec)" + and "(q \ A\<^sub>Q) \* C" and Sq: "(set q) \ (set A\<^sub>Q) \ (set(q \ A\<^sub>Q))" + by(rule name_list_avoiding[where xvec=A\<^sub>Q and c="(\, P, Q, K, N, xvec, r \ xvec, \\<^sub>Q, A\<^sub>P, \\<^sub>P, Q', P', C)"]) clarsimp + + obtain p::"name prm" where "(p \ A\<^sub>P) \* \" and "(p \ A\<^sub>P) \* P" and "(p \ A\<^sub>P) \* Q" and "(p \ A\<^sub>P) \* M" + and "(p \ A\<^sub>P) \* N" and "(p \ A\<^sub>P) \* xvec" and "(p \ A\<^sub>P) \* Q'" and "(p \ A\<^sub>P) \* A\<^sub>Q" + and "(p \ A\<^sub>P) \* P'" and "(p \ A\<^sub>P) \* \\<^sub>P" and "(p \ A\<^sub>P) \* \\<^sub>Q" and "(p \ A\<^sub>P) \* (q \ A\<^sub>Q)" + and "(p \ A\<^sub>P) \* C" and "(p \ A\<^sub>P) \* (r \ xvec)" and Sp: "(set p) \ (set A\<^sub>P) \ (set(p \ A\<^sub>P))" + by(rule name_list_avoiding[where xvec=A\<^sub>P and c="(\, P, Q, M, N, xvec, r \ xvec, A\<^sub>Q, q \ A\<^sub>Q, \\<^sub>Q, \\<^sub>P, Q', P', C)"]) + (auto simp add: eqvts fresh_star_prod) + + have FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" by fact + have FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" by fact + + from \A\<^sub>P \* Q\ FrQ \A\<^sub>P \* A\<^sub>Q\ have "A\<^sub>P \* \\<^sub>Q" + by(auto dest: extractFrameFreshChain) + from \A\<^sub>Q \* P\ FrP \A\<^sub>P \* A\<^sub>Q\ have "A\<^sub>Q \* \\<^sub>P" + by(auto dest: extractFrameFreshChain) + + note \cP = P \ Q\ + moreover from \(r \ xvec) \* P'\ \(r \ xvec) \* Q'\ have "(r \ xvec) \* (P' \ Q')" + by simp + with \cRs = \ \ \\*xvec\(P' \ Q')\ \(r \ xvec) \* N\ Sr + have "cRs = \ \ \\*(r \ xvec)\(r \ (P' \ Q'))" by(simp add: resChainAlpha residualInject) + then have "cRs = \ \ \\*(r \ xvec)\((r \ P') \ (r \ Q'))" + by simp + + moreover from \\ \ \\<^sub>Q \ P \M\\*xvec\\N\ \ P'\ Sr \(r \ xvec) \* N\ \(r \ xvec) \* P'\ + have "\ \ \\<^sub>Q \ P \M\\*(r \ xvec)\\(r \ N)\ \ (r \ P')" by(simp add: boundOutputChainAlpha'' residualInject) + then have "(q \ (\ \ \\<^sub>Q)) \ P \(q \ M)\\*(r \ xvec)\\(r \ N)\ \ (r \ P')" using Sq \A\<^sub>Q \* P\ \(q \ A\<^sub>Q) \* P\ \(r \ xvec) \* M\ \(r \ xvec) \* A\<^sub>Q\ \(q \ A\<^sub>Q) \* (r \ xvec)\ + by(fastforce intro: outputPermFrameSubject) + then have PTrans: "\ \ (q \ \\<^sub>Q) \ P \(q \ M)\\*(r \ xvec)\\(r \ N)\ \ (r \ P')" using Sq \A\<^sub>Q \* \\ \(q \ A\<^sub>Q) \* \\ + by(simp add: eqvts) + + moreover from \extractFrame P = \A\<^sub>P, \\<^sub>P\\ Sp \(p \ A\<^sub>P) \* \\<^sub>P\ + have FrP: "extractFrame P = \(p \ A\<^sub>P), (p \ \\<^sub>P)\" + by(simp add: frameChainAlpha) + moreover from \distinct A\<^sub>P\ have "distinct(p \ A\<^sub>P)" by simp + + moreover from \\ \ \\<^sub>P \ Q \K\N\ \ Q'\ Sr \distinctPerm r\ \xvec \* Q\ \(r \ xvec) \* Q\ + have "\ \ \\<^sub>P \ Q \K\(r \ N)\ \ (r \ Q')" by(rule inputAlpha) + then have "(p \ (\ \ \\<^sub>P)) \ Q \(p \ K)\(r \ N)\ \ (r \ Q')" using Sp \A\<^sub>P \* Q\ \(p \ A\<^sub>P) \* Q\ + by - (rule inputPermFrameSubject, (assumption | simp)+) + then have QTrans: "\ \ (p \ \\<^sub>P) \ Q \(p \ K)\(r \ N)\ \ (r \ Q')" using Sp \A\<^sub>P \* \\ \(p \ A\<^sub>P) \* \\ + by(simp add: eqvts) + + moreover from \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ Sq \(q \ A\<^sub>Q) \* \\<^sub>Q\ + have FrQ: "extractFrame Q = \(q \ A\<^sub>Q), (q \ \\<^sub>Q)\" + by(simp add: frameChainAlpha) + moreover from \distinct A\<^sub>Q\ have "distinct(q \ A\<^sub>Q)" by simp + + moreover from \\ \ \\<^sub>P \ \\<^sub>Q \ M \ K\ have "(p \ q \ (\ \ \\<^sub>P \ \\<^sub>Q)) \ (p \ q \ M) \ (p \ q \ K)" + by(metis chanEqClosed) + with \A\<^sub>P \* \\ \(p \ A\<^sub>P) \* \\ \A\<^sub>Q \* \\ \(q \ A\<^sub>Q) \* \\ \A\<^sub>Q \* \\<^sub>P\ \(q \ A\<^sub>Q) \* \\<^sub>P\ + \A\<^sub>P \* \\<^sub>Q\ \(p \ A\<^sub>P) \* \\<^sub>Q\ \A\<^sub>P \* M\ \(p \ A\<^sub>P) \* M\ \(q \ A\<^sub>Q) \* A\<^sub>P\ \(p \ A\<^sub>P) \* (q \ A\<^sub>Q)\ + \A\<^sub>Q \* K\ \(q \ A\<^sub>Q) \* K\ \A\<^sub>P \* A\<^sub>Q\ \(p \ A\<^sub>P) \* A\<^sub>Q\ Sp Sq + have "\ \ (p \ \\<^sub>P) \ (q \ \\<^sub>Q) \ (q \ M) \ (p \ K)" + by(simp add: eqvts freshChainSimps) + moreover note \(p \ A\<^sub>P) \* \\ + moreover from \(p \ A\<^sub>P) \* A\<^sub>Q\ \(p \ A\<^sub>P) \* (q \ A\<^sub>Q)\ \(p \ A\<^sub>P) \* \\<^sub>Q\ Sq have "(p \ A\<^sub>P) \* (q \ \\<^sub>Q)" + by(simp add: freshChainSimps) + moreover note \(p \ A\<^sub>P) \* P\ + moreover from \(p \ A\<^sub>P) \* A\<^sub>Q\ \(p \ A\<^sub>P) \* (q \ A\<^sub>Q)\ \(p \ A\<^sub>P) \* M\ Sq have "(p \ A\<^sub>P) \* (q \ M)" + by(simp add: freshChainSimps) + moreover from \(p \ A\<^sub>P) \* xvec\ \(p \ A\<^sub>P) \* (r \ xvec)\ \(p \ A\<^sub>P) \* N\ Sr have "(p \ A\<^sub>P) \* (r \ N)" + by(simp add: freshChainSimps) + moreover from \(p \ A\<^sub>P) \* xvec\ \(p \ A\<^sub>P) \* (r \ xvec)\ \(p \ A\<^sub>P) \* P'\ Sr have "(p \ A\<^sub>P) \* (r \ P')" + by(simp add: freshChainSimps) + moreover note \(p \ A\<^sub>P) \* Q\ + moreover from \(p \ A\<^sub>P) \* xvec\ \(p \ A\<^sub>P) \* (r \ xvec)\ \(p \ A\<^sub>P) \* Q'\ Sr have "(p \ A\<^sub>P) \* (r \ Q')" + by(simp add: freshChainSimps) + moreover note \(p \ A\<^sub>P) \* (q \ A\<^sub>Q)\ \(p \ A\<^sub>P) \* (r \ xvec)\ \(q \ A\<^sub>Q) \* \\ + moreover from \(q \ A\<^sub>Q) \* A\<^sub>P\ \(p \ A\<^sub>P) \* A\<^sub>Q\ \(q \ A\<^sub>Q) \* \\<^sub>P\ \(p \ A\<^sub>P) \* (q \ A\<^sub>Q)\ Sp Sq have "(q \ A\<^sub>Q) \* (p \ \\<^sub>P)" + by(simp add: freshChainSimps) + moreover note \(q \ A\<^sub>Q) \* P\ + moreover from \(q \ A\<^sub>Q) \* A\<^sub>P\ \(p \ A\<^sub>P) \* A\<^sub>Q\ \(q \ A\<^sub>Q) \* K\ \(p \ A\<^sub>P) \* (q \ A\<^sub>Q)\ Sp Sq have "(q \ A\<^sub>Q) \* (p \ K)" + by(simp add: freshChainSimps) + moreover from \(q \ A\<^sub>Q) \* xvec\ \(q \ A\<^sub>Q) \* (r \ xvec)\ \(q \ A\<^sub>Q) \* N\ Sr have "(q \ A\<^sub>Q) \* (r \ N)" + by(simp add: freshChainSimps) + moreover from \(q \ A\<^sub>Q) \* xvec\ \(q \ A\<^sub>Q) \* (r \ xvec)\ \(q \ A\<^sub>Q) \* P'\ Sr have "(q \ A\<^sub>Q) \* (r \ P')" + by(simp add: freshChainSimps) + moreover note \(q \ A\<^sub>Q) \* Q\ + moreover from \(q \ A\<^sub>Q) \* xvec\ \(q \ A\<^sub>Q) \* (r \ xvec)\ \(q \ A\<^sub>Q) \* Q'\ Sr have "(q \ A\<^sub>Q) \* (r \ Q')" + by(simp add: freshChainSimps) + moreover note \(q \ A\<^sub>Q) \* (r \ xvec)\ \(r \ xvec) \* \\ + moreover from \(r \ xvec) \* A\<^sub>P\ \(p \ A\<^sub>P) \* (r \ xvec)\ \(r \ xvec) \* \\<^sub>P\ Sp have "(r \ xvec) \* (p \ \\<^sub>P)" + by(simp add: freshChainSimps) + moreover from \(r \ xvec) \* A\<^sub>Q\ \(q \ A\<^sub>Q) \* (r \ xvec)\ \(r \ xvec) \* \\<^sub>Q\ Sq have "(r \ xvec) \* (q \ \\<^sub>Q)" + by(simp add: freshChainSimps) + moreover note \(r \ xvec) \* P\ + moreover from \(r \ xvec) \* A\<^sub>Q\ \(q \ A\<^sub>Q) \* (r \ xvec)\ \(r \ xvec) \* M\ Sq have "(r \ xvec) \* (q \ M)" + by(simp add: freshChainSimps) + moreover note \(r \ xvec) \* Q\ + moreover from \(r \ xvec) \* A\<^sub>P\ \(p \ A\<^sub>P) \* (r \ xvec)\ \(r \ xvec) \* K\ Sp have "(r \ xvec) \* (p \ K)" + by(simp add: freshChainSimps) + moreover note \(p \ A\<^sub>P) \* C\ \(q \ A\<^sub>Q) \* C\ \(r \ xvec) \* C\ + moreover from \distinct xvec\ have "distinct(r \ xvec)" by simp + ultimately show ?thesis by(rule rComm2) +next + case(cBrMerge \\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q Q' A\<^sub>Q) + obtain q::"name prm" where "(q \ A\<^sub>Q) \* \" and "(q \ A\<^sub>Q) \* P" + and "(q \ A\<^sub>Q) \* Q" and "(q \ A\<^sub>Q) \* M" + and "(q \ A\<^sub>Q) \* \\<^sub>P" and "(q \ A\<^sub>Q) \* A\<^sub>P" and "(q \ A\<^sub>Q) \* \\<^sub>Q" + and "(q \ A\<^sub>Q) \* N" and "(q \ A\<^sub>Q) \* P'" and "(q \ A\<^sub>Q) \* Q'" + and "(q \ A\<^sub>Q) \* C" + and Sq: "set q \ set A\<^sub>Q \ set(q \ A\<^sub>Q)" + by(rule name_list_avoiding[where c="(\, P, N, M, P', Q', Q, \\<^sub>Q, A\<^sub>P, \\<^sub>P, C)"]) + (auto simp add: eqvts fresh_star_prod emptyFresh) + obtain p::"name prm" where "(p \ A\<^sub>P) \* \" and "(p \ A\<^sub>P) \* P" + and "(p \ A\<^sub>P) \* Q" and "(p \ A\<^sub>P) \* M" + and "(p \ A\<^sub>P) \* \\<^sub>P" and "(p \ A\<^sub>P) \* \\<^sub>Q" and "(p \ A\<^sub>P) \* A\<^sub>Q" + and "(p \ A\<^sub>P) \* N" and "(p \ A\<^sub>P) \* P'" and "(p \ A\<^sub>P) \* Q'" + and "(p \ A\<^sub>P) \* C" + and Sp: "set p \ set A\<^sub>P \ set(p \ A\<^sub>P)" + and "(p \ A\<^sub>P) \* (q \ A\<^sub>Q)" + by(rule name_list_avoiding[where c="(\, P, N, P', Q', Q, M, A\<^sub>Q, q \ A\<^sub>Q, \\<^sub>Q, \\<^sub>P, C)"]) + (auto simp add: eqvts fresh_star_prod) + + have FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" by fact + have FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" by fact + + from \A\<^sub>P \* Q\ FrQ \A\<^sub>P \* A\<^sub>Q\ have "A\<^sub>P \* \\<^sub>Q" + by(auto dest: extractFrameFreshChain) + from \A\<^sub>Q \* P\ FrP \A\<^sub>P \* A\<^sub>Q\ have "A\<^sub>Q \* \\<^sub>P" + by(auto dest: extractFrameFreshChain) + + from Sp \A\<^sub>P \* M\ \(p \ A\<^sub>P) \* M\ + have "(p \ M) = M" + by simp + from Sq \A\<^sub>Q \* M\ \(q \ A\<^sub>Q) \* M\ + have "(q \ M) = M" + by simp + + note \cP = P \ Q\ \cRs = \M\N\ \ (P' \ Q')\ + moreover from \distinct A\<^sub>P\ have "distinct(p \ A\<^sub>P)" by simp + moreover from \distinct A\<^sub>Q\ have "distinct(q \ A\<^sub>Q)" by simp + moreover from \extractFrame P = \A\<^sub>P, \\<^sub>P\\ Sp \(p \ A\<^sub>P) \* \\<^sub>P\ + have FrP: "extractFrame P = \(p \ A\<^sub>P), (p \ \\<^sub>P)\" + by(simp add: frameChainAlpha) + moreover from \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ Sq \(q \ A\<^sub>Q) \* \\<^sub>Q\ + have FrQ: "extractFrame Q = \(q \ A\<^sub>Q), (q \ \\<^sub>Q)\" + by(simp add: frameChainAlpha) + + moreover have "(q \ (\ \ \\<^sub>Q)) \ P \\(q \ M)\N\ \ P'" using Sq \A\<^sub>Q \* P\ \(q \ A\<^sub>Q) \* P\ \\ \ \\<^sub>Q \ P \\M\N\ \ P'\ + by - (rule brinputPermFrameSubject, (assumption | simp)+) + then have "(\ \ (q \ \\<^sub>Q)) \ P \\(q \ M)\N\ \ P'" using Sq \A\<^sub>Q \* \\ \(q \ A\<^sub>Q) \* \\ + by(simp add: eqvts) + with \(q \ M) = M\ have PTrans: "(\ \ (q \ \\<^sub>Q)) \ P \\M\N\ \ P'" + by simp + + moreover have "(p \ (\ \ \\<^sub>P)) \ Q \\(p \ M)\N\ \ Q'" using Sp \A\<^sub>P \* Q\ \(p \ A\<^sub>P) \* Q\ \\ \ \\<^sub>P \ Q \\M\N\ \ Q'\ + by - (rule brinputPermFrameSubject, (assumption | simp)+) + then have "(\ \ (p \ \\<^sub>P)) \ Q \\(p \ M)\N\ \ Q'" using Sp \A\<^sub>P \* \\ \(p \ A\<^sub>P) \* \\ + by(simp add: eqvts) + with \(p \ M) = M\ have PTrans: "(\ \ (p \ \\<^sub>P)) \ Q \\M\N\ \ Q'" + by simp + moreover from \(p \ A\<^sub>P) \* A\<^sub>Q\ \(p \ A\<^sub>P) \* (q \ A\<^sub>Q)\ \(p \ A\<^sub>P) \* \\<^sub>Q\ Sq have "(p \ A\<^sub>P) \* (q \ \\<^sub>Q)" + by(simp add: freshChainSimps) + moreover from \(q \ A\<^sub>Q) \* A\<^sub>P\ \(p \ A\<^sub>P) \* A\<^sub>Q\ \(q \ A\<^sub>Q) \* \\<^sub>P\ \(p \ A\<^sub>P) \* (q \ A\<^sub>Q)\ Sp Sq have "(q \ A\<^sub>Q) \* (p \ \\<^sub>P)" + by(simp add: freshChainSimps) + + moreover note + \(p \ A\<^sub>P) \* \\ + \(p \ A\<^sub>P) \* P\ \(p \ A\<^sub>P) \* N\ \(p \ A\<^sub>P) \* P'\ \(p \ A\<^sub>P) \* M\ + \(p \ A\<^sub>P) \* Q\ \(p \ A\<^sub>P) \* Q'\ \(p \ A\<^sub>P) \* (q \ A\<^sub>Q)\ \(q \ A\<^sub>Q) \* \\ + \(q \ A\<^sub>Q) \* P\ \(q \ A\<^sub>Q) \* N\ \(q \ A\<^sub>Q) \* P'\ \(q \ A\<^sub>Q) \* M\ + \(q \ A\<^sub>Q) \* Q\ \(q \ A\<^sub>Q) \* Q'\ \(p \ A\<^sub>P) \* C\ \(q \ A\<^sub>Q) \* C\ + ultimately show ?thesis + by(auto simp add: rBrMerge) +next + case(cBrComm1 \\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q xvec Q' A\<^sub>Q) + obtain r::"name prm" where "(r \ xvec) \* \" and "(r \ xvec) \* P" and "(r \ xvec) \* Q" and "(r \ xvec) \* M" + and "(r \ xvec) \* N" and "(r \ xvec) \* A\<^sub>P" and "(r \ xvec) \* A\<^sub>Q" + and "(r \ xvec) \* P'" and "(r \ xvec) \* Q'" and "(r \ xvec) \* \\<^sub>P" and "(r \ xvec) \* \\<^sub>Q" + and "(r \ xvec) \* C" and Sr: "(set r) \ (set xvec) \ (set(r \ xvec))" and "distinctPerm r" + by(rule name_list_avoiding[where xvec=xvec and c="(\, P, Q, M, N, A\<^sub>P, A\<^sub>Q, \\<^sub>P, \\<^sub>Q, P', Q', C)"]) + (auto simp add: eqvts) + + obtain q::"name prm" where "(q \ A\<^sub>Q) \* \" and "(q \ A\<^sub>Q) \* P" and "(q \ A\<^sub>Q) \* Q" and "(q \ A\<^sub>Q) \* M" + and "(q \ A\<^sub>Q) \* N" and "(q \ A\<^sub>Q) \* xvec" and "(q \ A\<^sub>Q) \* Q'" and "(q \ A\<^sub>Q) \* P'" + and "(q \ A\<^sub>Q) \* \\<^sub>P" and "(q \ A\<^sub>Q) \* A\<^sub>P" and "(q \ A\<^sub>Q) \* \\<^sub>Q" and "(q \ A\<^sub>Q) \* (r \ xvec)" + and "(q \ A\<^sub>Q) \* C" and Sq: "(set q) \ (set A\<^sub>Q) \ (set(q \ A\<^sub>Q))" + by(rule name_list_avoiding[where xvec=A\<^sub>Q and c="(\, P, Q, M, N, xvec, r \ xvec, \\<^sub>Q, A\<^sub>P, \\<^sub>P, Q', P', C)"]) clarsimp + + obtain p::"name prm" where "(p \ A\<^sub>P) \* \" and "(p \ A\<^sub>P) \* P" and "(p \ A\<^sub>P) \* Q" and "(p \ A\<^sub>P) \* M" + and "(p \ A\<^sub>P) \* N" and "(p \ A\<^sub>P) \* xvec" and "(p \ A\<^sub>P) \* Q'" and "(p \ A\<^sub>P) \* A\<^sub>Q" + and "(p \ A\<^sub>P) \* P'" and "(p \ A\<^sub>P) \* \\<^sub>P" and "(p \ A\<^sub>P) \* \\<^sub>Q" and "(p \ A\<^sub>P) \* (q \ A\<^sub>Q)" + and "(p \ A\<^sub>P) \* C" and "(p \ A\<^sub>P) \* (r \ xvec)" and Sp: "(set p) \ (set A\<^sub>P) \ (set(p \ A\<^sub>P))" + by(rule name_list_avoiding[where xvec=A\<^sub>P and c="(\, P, Q, M, N, xvec, r \ xvec, A\<^sub>Q, q \ A\<^sub>Q, \\<^sub>Q, \\<^sub>P, Q', P', C)"]) + (auto simp add: eqvts fresh_star_prod) + + from Sp \A\<^sub>P \* M\ \(p \ A\<^sub>P) \* M\ + have "(p \ M) = M" + by simp + from Sq \A\<^sub>Q \* M\ \(q \ A\<^sub>Q) \* M\ + have "(q \ M) = M" + by simp + from Sr \xvec \* M\ \(r \ xvec) \* M\ + have "(r \ M) = M" + by simp + + have FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" by fact + have FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" by fact + + from \A\<^sub>P \* Q\ FrQ \A\<^sub>P \* A\<^sub>Q\ have "A\<^sub>P \* \\<^sub>Q" + by(auto dest: extractFrameFreshChain) + from \A\<^sub>Q \* P\ FrP \A\<^sub>P \* A\<^sub>Q\ have "A\<^sub>Q \* \\<^sub>P" + by(auto dest: extractFrameFreshChain) + + note \cP = P \ Q\ + moreover from \(r \ xvec) \* P'\ \(r \ xvec) \* Q'\ have "(r \ xvec) \* (P' \ Q')" + by simp + with \cRs = \M\\*xvec\\N\ \ (P' \ Q')\ \(r \ xvec) \* N\ \(r \ xvec) \* P'\ \(r \ xvec) \* Q'\ \xvec \* M\ \(r \ xvec) \* M\ Sr + have "cRs = (r \ (\M\\*xvec\\N\)) \ (r \ (P' \ Q'))" + by (simp add: residualAlpha) + with \(r \ M) = M\ have "cRs = \M\\*(r \ xvec)\\(r \ N)\ \ ((r \ P') \ (r \ Q'))" by simp + + moreover from \\ \ \\<^sub>Q \ P \\M\N\ \ P'\ Sr \distinctPerm r\ \xvec \* P\ \(r \ xvec) \* P\ + have "\ \ \\<^sub>Q \ P \\M\(r \ N)\ \ (r \ P')" + by(rule brinputAlpha) + then have "(q \ (\ \ \\<^sub>Q)) \ P \\(q \ M)\(r \ N)\ \ (r \ P')" using Sq \A\<^sub>Q \* P\ \(q \ A\<^sub>Q) \* P\ + by - (rule brinputPermFrameSubject, (assumption | simp)+) + then have PTrans: "\ \ (q \ \\<^sub>Q) \ P \\M\(r \ N)\ \ (r \ P')" using Sq \A\<^sub>Q \* \\ \(q \ A\<^sub>Q) \* \\ \(q \ M) = M\ + by(simp add: eqvts) + + moreover from \extractFrame P = \A\<^sub>P, \\<^sub>P\\ Sp \(p \ A\<^sub>P) \* \\<^sub>P\ + have FrP: "extractFrame P = \(p \ A\<^sub>P), (p \ \\<^sub>P)\" + by(simp add: frameChainAlpha) + moreover from \distinct A\<^sub>P\ have "distinct(p \ A\<^sub>P)" by simp + + moreover from \\ \ \\<^sub>P \ Q \\M\\*xvec\\N\ \ Q'\ Sr \(r \ xvec) \* N\ \(r \ xvec) \* Q'\ + have "\ \ \\<^sub>P \ Q \\M\\*(r \ xvec)\\(r \ N)\ \ (r \ Q')" + by(simp add: boundOutputChainAlpha'' residualInject) + then have "(p \ (\ \ \\<^sub>P)) \ Q \\(p \ M)\\*(r \ xvec)\\(r \ N)\ \ (r \ Q')" using Sp \A\<^sub>P \* Q\ \(p \ A\<^sub>P) \* Q\ \(r \ xvec) \* M\ \(p \ A\<^sub>P) \* (r \ xvec)\ \(r \ xvec) \* A\<^sub>P\ + by(fastforce intro: broutputPermFrameSubject) + then have QTrans: "\ \ (p \ \\<^sub>P) \ Q \\M\\*(r \ xvec)\\(r \ N)\ \ (r \ Q')" using Sp \A\<^sub>P \* \\ \(p \ A\<^sub>P) \* \\ \(p \ M) = M\ + by(simp add: eqvts) + + moreover from \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ Sq \(q \ A\<^sub>Q) \* \\<^sub>Q\ + have FrQ: "extractFrame Q = \(q \ A\<^sub>Q), (q \ \\<^sub>Q)\" + by(simp add: frameChainAlpha) + moreover from \distinct A\<^sub>Q\ have "distinct(q \ A\<^sub>Q)" by simp + + moreover note \(p \ A\<^sub>P) \* \\ + moreover from \(p \ A\<^sub>P) \* A\<^sub>Q\ \(p \ A\<^sub>P) \* (q \ A\<^sub>Q)\ \(p \ A\<^sub>P) \* \\<^sub>Q\ Sq have "(p \ A\<^sub>P) \* (q \ \\<^sub>Q)" + by(simp add: freshChainSimps) + moreover note \(p \ A\<^sub>P) \* P\ + moreover from \(p \ A\<^sub>P) \* xvec\ \(p \ A\<^sub>P) \* (r \ xvec)\ \(p \ A\<^sub>P) \* N\ Sr have "(p \ A\<^sub>P) \* (r \ N)" + by(simp add: freshChainSimps) + moreover from \(p \ A\<^sub>P) \* xvec\ \(p \ A\<^sub>P) \* (r \ xvec)\ \(p \ A\<^sub>P) \* P'\ Sr have "(p \ A\<^sub>P) \* (r \ P')" + by(simp add: freshChainSimps) + moreover note \(p \ A\<^sub>P) \* Q\ + moreover from \(p \ A\<^sub>P) \* xvec\ \(p \ A\<^sub>P) \* (r \ xvec)\ \(p \ A\<^sub>P) \* Q'\ Sr have "(p \ A\<^sub>P) \* (r \ Q')" + by(simp add: freshChainSimps) + moreover note \(p \ A\<^sub>P) \* (q \ A\<^sub>Q)\ \(p \ A\<^sub>P) \* (r \ xvec)\ \(q \ A\<^sub>Q) \* \\ + moreover from \(q \ A\<^sub>Q) \* A\<^sub>P\ \(p \ A\<^sub>P) \* A\<^sub>Q\ \(q \ A\<^sub>Q) \* \\<^sub>P\ \(p \ A\<^sub>P) \* (q \ A\<^sub>Q)\ Sp Sq have "(q \ A\<^sub>Q) \* (p \ \\<^sub>P)" + by(simp add: freshChainSimps) + moreover note \(q \ A\<^sub>Q) \* P\ + moreover from \(q \ A\<^sub>Q) \* xvec\ \(q \ A\<^sub>Q) \* (r \ xvec)\ \(q \ A\<^sub>Q) \* N\ Sr have "(q \ A\<^sub>Q) \* (r \ N)" + by(simp add: freshChainSimps) + moreover from \(q \ A\<^sub>Q) \* xvec\ \(q \ A\<^sub>Q) \* (r \ xvec)\ \(q \ A\<^sub>Q) \* P'\ Sr have "(q \ A\<^sub>Q) \* (r \ P')" + by(simp add: freshChainSimps) + moreover note \(q \ A\<^sub>Q) \* Q\ + moreover from \(q \ A\<^sub>Q) \* xvec\ \(q \ A\<^sub>Q) \* (r \ xvec)\ \(q \ A\<^sub>Q) \* Q'\ Sr have "(q \ A\<^sub>Q) \* (r \ Q')" + by(simp add: freshChainSimps) + moreover note \(q \ A\<^sub>Q) \* (r \ xvec)\ \(r \ xvec) \* \\ + moreover from \(r \ xvec) \* A\<^sub>P\ \(p \ A\<^sub>P) \* (r \ xvec)\ \(r \ xvec) \* \\<^sub>P\ Sp have "(r \ xvec) \* (p \ \\<^sub>P)" + by(simp add: freshChainSimps) + moreover from \(r \ xvec) \* A\<^sub>Q\ \(q \ A\<^sub>Q) \* (r \ xvec)\ \(r \ xvec) \* \\<^sub>Q\ Sq have "(r \ xvec) \* (q \ \\<^sub>Q)" + by(simp add: freshChainSimps) + moreover note \(p \ A\<^sub>P) \* M\ \(q \ A\<^sub>Q) \* M\ + moreover note \(r \ xvec) \* P\ + moreover note \(r \ xvec) \* Q\ \(r \ xvec) \* M\ + moreover note \(p \ A\<^sub>P) \* C\ \(q \ A\<^sub>Q) \* C\ \(r \ xvec) \* C\ + moreover from \distinct xvec\ have "distinct(r \ xvec)" by simp + ultimately show ?thesis by(simp add: rBrComm1) +next + case(cBrComm2 \\<^sub>Q P M xvec N P' A\<^sub>P \\<^sub>P Q Q' A\<^sub>Q) + obtain r::"name prm" where "(r \ xvec) \* \" and "(r \ xvec) \* P" and "(r \ xvec) \* Q" and "(r \ xvec) \* M" + and "(r \ xvec) \* N" and "(r \ xvec) \* A\<^sub>P" and "(r \ xvec) \* A\<^sub>Q" + and "(r \ xvec) \* P'" and "(r \ xvec) \* Q'" and "(r \ xvec) \* \\<^sub>P" and "(r \ xvec) \* \\<^sub>Q" + and "(r \ xvec) \* C" and Sr: "(set r) \ (set xvec) \ (set(r \ xvec))" and "distinctPerm r" + by(rule name_list_avoiding[where xvec=xvec and c="(\, P, Q, M, N, A\<^sub>P, A\<^sub>Q, \\<^sub>P, \\<^sub>Q, P', Q', C)"]) + (auto simp add: eqvts) + + obtain q::"name prm" where "(q \ A\<^sub>Q) \* \" and "(q \ A\<^sub>Q) \* P" and "(q \ A\<^sub>Q) \* Q" and "(q \ A\<^sub>Q) \* M" + and "(q \ A\<^sub>Q) \* N" and "(q \ A\<^sub>Q) \* xvec" and "(q \ A\<^sub>Q) \* Q'" and "(q \ A\<^sub>Q) \* P'" + and "(q \ A\<^sub>Q) \* \\<^sub>P" and "(q \ A\<^sub>Q) \* A\<^sub>P" and "(q \ A\<^sub>Q) \* \\<^sub>Q" and "(q \ A\<^sub>Q) \* (r \ xvec)" + and "(q \ A\<^sub>Q) \* C" and Sq: "(set q) \ (set A\<^sub>Q) \ (set(q \ A\<^sub>Q))" + by(rule name_list_avoiding[where xvec=A\<^sub>Q and c="(\, P, Q, N, M, xvec, r \ xvec, \\<^sub>Q, A\<^sub>P, \\<^sub>P, Q', P', C)"]) clarsimp + + obtain p::"name prm" where "(p \ A\<^sub>P) \* \" and "(p \ A\<^sub>P) \* P" and "(p \ A\<^sub>P) \* Q" and "(p \ A\<^sub>P) \* M" + and "(p \ A\<^sub>P) \* N" and "(p \ A\<^sub>P) \* xvec" and "(p \ A\<^sub>P) \* Q'" and "(p \ A\<^sub>P) \* A\<^sub>Q" + and "(p \ A\<^sub>P) \* P'" and "(p \ A\<^sub>P) \* \\<^sub>P" and "(p \ A\<^sub>P) \* \\<^sub>Q" and "(p \ A\<^sub>P) \* (q \ A\<^sub>Q)" + and "(p \ A\<^sub>P) \* C" and "(p \ A\<^sub>P) \* (r \ xvec)" and Sp: "(set p) \ (set A\<^sub>P) \ (set(p \ A\<^sub>P))" + by(rule name_list_avoiding[where xvec=A\<^sub>P and c="(\, P, Q, M, N, xvec, r \ xvec, A\<^sub>Q, q \ A\<^sub>Q, \\<^sub>Q, \\<^sub>P, Q', P', C)"]) + (auto simp add: eqvts fresh_star_prod) + + from Sp \A\<^sub>P \* M\ \(p \ A\<^sub>P) \* M\ + have "(p \ M) = M" + by simp + from Sq \A\<^sub>Q \* M\ \(q \ A\<^sub>Q) \* M\ + have "(q \ M) = M" + by simp + from Sr \xvec \* M\ \(r \ xvec) \* M\ + have "(r \ M) = M" + by simp + + have FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" by fact + have FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" by fact + + from \A\<^sub>P \* Q\ FrQ \A\<^sub>P \* A\<^sub>Q\ have "A\<^sub>P \* \\<^sub>Q" + by(auto dest: extractFrameFreshChain) + from \A\<^sub>Q \* P\ FrP \A\<^sub>P \* A\<^sub>Q\ have "A\<^sub>Q \* \\<^sub>P" + by(auto dest: extractFrameFreshChain) + + note \cP = P \ Q\ + moreover from \(r \ xvec) \* P'\ \(r \ xvec) \* Q'\ have "(r \ xvec) \* (P' \ Q')" + by simp + with \cRs = \M\\*xvec\\N\ \ (P' \ Q')\ \(r \ xvec) \* N\ \(r \ xvec) \* P'\ \(r \ xvec) \* Q'\ \xvec \* M\ \(r \ xvec) \* M\ Sr + have "cRs = (r \ (\M\\*xvec\\N\)) \ (r \ (P' \ Q'))" + by (simp add: residualAlpha) + with \(r \ M) = M\ have "cRs = \M\\*(r \ xvec)\\(r \ N)\ \ ((r \ P') \ (r \ Q'))" by simp + + moreover from \\ \ \\<^sub>Q \ P \\M\\*xvec\\N\ \ P'\ Sr \(r \ xvec) \* N\ \(r \ xvec) \* P'\ + have "\ \ \\<^sub>Q \ P \\M\\*(r \ xvec)\\(r \ N)\ \ (r \ P')" by(simp add: boundOutputChainAlpha'' residualInject) + then have "(q \ (\ \ \\<^sub>Q)) \ P \\(q \ M)\\*(r \ xvec)\\(r \ N)\ \ (r \ P')" using Sq \A\<^sub>Q \* P\ \(q \ A\<^sub>Q) \* P\ \(r \ xvec) \* M\ \(r \ xvec) \* A\<^sub>Q\ \(q \ A\<^sub>Q) \* (r \ xvec)\ + by(fastforce intro: broutputPermFrameSubject) + then have PTrans: "\ \ (q \ \\<^sub>Q) \ P \\M\\*(r \ xvec)\\(r \ N)\ \ (r \ P')" using Sq \A\<^sub>Q \* \\ \(q \ A\<^sub>Q) \* \\ \(q \ M) = M\ + by(simp add: eqvts) + + moreover from \extractFrame P = \A\<^sub>P, \\<^sub>P\\ Sp \(p \ A\<^sub>P) \* \\<^sub>P\ + have FrP: "extractFrame P = \(p \ A\<^sub>P), (p \ \\<^sub>P)\" + by(simp add: frameChainAlpha) + moreover from \distinct A\<^sub>P\ have "distinct(p \ A\<^sub>P)" by simp + + moreover from \\ \ \\<^sub>P \ Q \\M\N\ \ Q'\ Sr \distinctPerm r\ \xvec \* Q\ \(r \ xvec) \* Q\ + have "\ \ \\<^sub>P \ Q \\M\(r \ N)\ \ (r \ Q')" by(rule brinputAlpha) + then have "(p \ (\ \ \\<^sub>P)) \ Q \\(p \ M)\(r \ N)\ \ (r \ Q')" using Sp \A\<^sub>P \* Q\ \(p \ A\<^sub>P) \* Q\ + by - (rule brinputPermFrameSubject, (assumption | simp)+) + then have QTrans: "\ \ (p \ \\<^sub>P) \ Q \\M\(r \ N)\ \ (r \ Q')" using Sp \A\<^sub>P \* \\ \(p \ A\<^sub>P) \* \\ \(p \ M) = M\ + by(simp add: eqvts) + + moreover from \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ Sq \(q \ A\<^sub>Q) \* \\<^sub>Q\ + have FrQ: "extractFrame Q = \(q \ A\<^sub>Q), (q \ \\<^sub>Q)\" + by(simp add: frameChainAlpha) + moreover from \distinct A\<^sub>Q\ have "distinct(q \ A\<^sub>Q)" by simp + + moreover note \(p \ A\<^sub>P) \* \\ + moreover from \(p \ A\<^sub>P) \* A\<^sub>Q\ \(p \ A\<^sub>P) \* (q \ A\<^sub>Q)\ \(p \ A\<^sub>P) \* \\<^sub>Q\ Sq have "(p \ A\<^sub>P) \* (q \ \\<^sub>Q)" + by(simp add: freshChainSimps) + moreover note \(p \ A\<^sub>P) \* P\ + moreover from \(p \ A\<^sub>P) \* xvec\ \(p \ A\<^sub>P) \* (r \ xvec)\ \(p \ A\<^sub>P) \* N\ Sr have "(p \ A\<^sub>P) \* (r \ N)" + by(simp add: freshChainSimps) + moreover from \(p \ A\<^sub>P) \* xvec\ \(p \ A\<^sub>P) \* (r \ xvec)\ \(p \ A\<^sub>P) \* P'\ Sr have "(p \ A\<^sub>P) \* (r \ P')" + by(simp add: freshChainSimps) + moreover note \(p \ A\<^sub>P) \* Q\ + moreover from \(p \ A\<^sub>P) \* xvec\ \(p \ A\<^sub>P) \* (r \ xvec)\ \(p \ A\<^sub>P) \* Q'\ Sr have "(p \ A\<^sub>P) \* (r \ Q')" + by(simp add: freshChainSimps) + moreover note \(p \ A\<^sub>P) \* (q \ A\<^sub>Q)\ \(p \ A\<^sub>P) \* (r \ xvec)\ \(q \ A\<^sub>Q) \* \\ + moreover from \(q \ A\<^sub>Q) \* A\<^sub>P\ \(p \ A\<^sub>P) \* A\<^sub>Q\ \(q \ A\<^sub>Q) \* \\<^sub>P\ \(p \ A\<^sub>P) \* (q \ A\<^sub>Q)\ Sp Sq have "(q \ A\<^sub>Q) \* (p \ \\<^sub>P)" + by(simp add: freshChainSimps) + moreover note \(q \ A\<^sub>Q) \* P\ + moreover from \(q \ A\<^sub>Q) \* xvec\ \(q \ A\<^sub>Q) \* (r \ xvec)\ \(q \ A\<^sub>Q) \* N\ Sr have "(q \ A\<^sub>Q) \* (r \ N)" + by(simp add: freshChainSimps) + moreover from \(q \ A\<^sub>Q) \* xvec\ \(q \ A\<^sub>Q) \* (r \ xvec)\ \(q \ A\<^sub>Q) \* P'\ Sr have "(q \ A\<^sub>Q) \* (r \ P')" + by(simp add: freshChainSimps) + moreover note \(q \ A\<^sub>Q) \* Q\ + moreover from \(q \ A\<^sub>Q) \* xvec\ \(q \ A\<^sub>Q) \* (r \ xvec)\ \(q \ A\<^sub>Q) \* Q'\ Sr have "(q \ A\<^sub>Q) \* (r \ Q')" + by(simp add: freshChainSimps) + moreover note \(q \ A\<^sub>Q) \* (r \ xvec)\ \(r \ xvec) \* \\ + moreover from \(r \ xvec) \* A\<^sub>P\ \(p \ A\<^sub>P) \* (r \ xvec)\ \(r \ xvec) \* \\<^sub>P\ Sp have "(r \ xvec) \* (p \ \\<^sub>P)" + by(simp add: freshChainSimps) + moreover from \(r \ xvec) \* A\<^sub>Q\ \(q \ A\<^sub>Q) \* (r \ xvec)\ \(r \ xvec) \* \\<^sub>Q\ Sq have "(r \ xvec) \* (q \ \\<^sub>Q)" + by(simp add: freshChainSimps) + moreover note \(p \ A\<^sub>P) \* M\ \(q \ A\<^sub>Q) \* M\ + moreover note \(r \ xvec) \* P\ + moreover note \(r \ xvec) \* Q\ \(r \ xvec) \* M\ + moreover note \(p \ A\<^sub>P) \* C\ \(q \ A\<^sub>Q) \* C\ \(r \ xvec) \* C\ + moreover from \distinct xvec\ have "distinct(r \ xvec)" by simp + ultimately show ?thesis by(simp add: rBrComm2) +next + case(cBrClose P M xvec N P' x) + obtain r::"name prm" where "(r \ xvec) \* \" and "(r \ xvec) \* P" and "(r \ xvec) \* M" + and "(r \ xvec) \* N" and "(r \ xvec) \* P'" and "(r \ xvec) \* x" + and "(r \ xvec) \* C" and Sr: "(set r) \ (set xvec) \ (set(r \ xvec))" and "distinctPerm r" + by(rule name_list_avoiding[where xvec=xvec and c="(\, P, M, N, P', x, C)"]) + (auto simp add: eqvts) + obtain y::name where "y \ P" and "y \ C" and "y \ xvec" and "y \ x" and "y \ N" + and "y \ (r \ xvec)" and "y \ r" and "y \ M" and "y \ \" + and "y \ P'" and "y \ (r \ P')" and "y \ (r \ N)" + by(generate_fresh "name") (auto simp add: freshChainSimps) + from \cP = \\x\P\ \y \ P\ have cP_perm: "cP = \\y\([(x, y)] \ P)" by(simp add: alphaRes) + from \cRs = \ \ \\x\(\\*xvec\P')\ \(r \ xvec) \* P'\ Sr + have "cRs = \ \ \\x\(\\*(r \ xvec)\(r \ P'))" by(simp add: resChainAlpha) + moreover from \y \ (r \ P')\ have "y \ (\\*(r \ xvec)\(r \ P'))" by(simp add: resChainFresh) + ultimately have "cRs = \ \ \\y\([(x, y)] \ (\\*(r \ xvec)\(r \ P')))" by(simp add: alphaRes) + with \(r \ xvec) \* x\ \y \ (r \ xvec)\ + have cRs_perm: "cRs = \ \ \\y\(\\*(r \ xvec)\([(x, y)] \ (r \ P')))" by(simp add: eqvts) + + from \x \ xvec\ \(r \ xvec) \* x\ Sr have "r \ x = x" by simp + + from \\ \ P \ \M\\*xvec\\N\ \ P'\ \(r \ xvec) \* N\ \(r \ xvec) \* P'\ + \set r \ set xvec \ set (r \ xvec)\ + have "\ \ P \ \M\\*(r \ xvec)\\(r \ N)\ \ (r \ P')" + by(simp add: boundOutputChainAlpha'' create_residual.simps) + then have "[(x, y)] \ (\ \ P \ \M\\*(r \ xvec)\\(r \ N)\ \ (r \ P'))" + by(simp add: perm_bool) + with \x \ \\ \y \ \\ \(r \ xvec) \* x\ \y \ (r \ xvec)\ + \y \ (r \ N)\ + have trans_perm: "\ \ ([(x, y)] \ P) \ \([(x, y)] \ M)\\*(r \ xvec)\\([(x, y)] \ (r \ N))\ \ ([(x, y)] \ (r \ P'))" + by(auto simp add: eqvts) + + note cP_perm cRs_perm + moreover from \x \ supp M\ have "y \ supp ([(x, y)] \ M)" + by (metis fresh_bij fresh_def swap_simps) + moreover note trans_perm + moreover from \distinct xvec\ \distinctPerm r\ have "distinct (r \ xvec)" + by simp + moreover note \(r \ xvec) \* \\ + moreover from \(r \ xvec) \* x\ \y \ (r \ xvec)\ \(r \ xvec) \* P\ + have "(r \ xvec) \* ([(x, y)] \ P)" by simp + moreover from \(r \ xvec) \* x\ \y \ (r \ xvec)\ \(r \ xvec) \* M\ + have "(r \ xvec) \* ([(x, y)] \ M)" by simp + moreover note \y \ \\ \y \ (r \ xvec)\ + \(r \ xvec) \* C\ \y \ C\ + ultimately show ?thesis + by(rule rBrClose) +next + case(cOpen P M xvec yvec N P' x) + from \\ \ P \ M\\*(xvec@yvec)\\N\ \ P'\ have "distinct(xvec@yvec)" by(force dest: boundOutputDistinct) + then have "xvec \* yvec" by(induct xvec) auto + obtain p where "(p \ yvec) \* \" and "(p \ yvec) \* P" and "(p \ yvec) \* M" + and "(p \ yvec) \* yvec" and "(p \ yvec) \* N" and "(p \ yvec) \* P'" + and "x \ (p \ yvec)" and "(p \ yvec) \* xvec" + and "(p \ yvec) \* C" and Sp: "(set p) \ (set yvec) \ (set(p \ yvec))" + by(rule name_list_avoiding[where xvec=yvec and c="(\, P, M, xvec, yvec, N, P', x, C)"]) + (auto simp add: eqvts fresh_star_prod) + obtain q where "(q \ xvec) \* \" and "(q \ xvec) \* P" and "(q \ xvec) \* M" + and "(q \ xvec) \* xvec" and "(q \ xvec) \* N" and "(q \ xvec) \* P'" + and "x \ (q \ xvec)" and "(q \ xvec) \* yvec" + and "(q \ xvec) \* p" and "(q \ xvec) \* (p \ yvec)" + and "(q \ xvec) \* C" and Sq: "(set q) \ (set xvec) \ (set(q \ xvec))" + by(rule name_list_avoiding[where xvec=xvec and c="(\, P, M, xvec, yvec, p \ yvec, N, P', x, p, C)"]) + (auto simp add: eqvts fresh_star_prod) + obtain y::name where "y \ P" and "y \ C" and "y \ xvec" and "y \ yvec" and "y \ x" and "y \ N" + and "y \ (q \ xvec)" and "y \ (p \ yvec)" and "y \ M" and "y \ \" and "y \ P'" + by(generate_fresh "name") (auto simp add: freshChainSimps) + from \cP = \\x\P\ \y \ P\ have "cP = \\y\([(x, y)] \ P)" by(simp add: alphaRes) + moreover have "cRs = M\\*((q \ xvec)@y#(p \ yvec))\\((q@(x, y)#p) \ N)\ \ ((q@(x, y)#p) \ P')" + proof - + note \cRs = M\\*(xvec@x#yvec)\\N\ \ P'\ + moreover have "\\*(xvec@x#yvec)\N \' P' = \\*xvec\(\\x\(\\*yvec\N \' P'))" by(simp add: boundOutputApp) + moreover from \(p \ yvec) \* N\ \(p \ yvec) \* P'\ Sp have "\ = \\*xvec\(\\x\(\\*(p \ yvec)\(p \ N) \' (p \ P')))" + by(simp add: boundOutputChainAlpha'') + moreover with \y \ N\ \y \ P'\ \y \ (p \ yvec)\ \y \ yvec\ \x \ yvec\ \x \ (p \ yvec)\ Sp + moreover have "\ = \\*xvec\(\\y\(\\*(p \ yvec)\(([(x, y)] \ p \ N) \' ([(x, y)] \ p \ P'))))" + by(subst alphaBoundOutput[where y=y]) (simp add: freshChainSimps eqvts)+ + moreover then have "\ = \\*xvec\(\\y\(\\*(p \ yvec)\((((x, y)#p) \ N) \' (((x, y)#p) \ P'))))" + by simp + moreover from \(q \ xvec) \* N\ \(q \ xvec) \* P'\ \xvec \* yvec\ \(p \ yvec) \* xvec\ \(q \ xvec) \* yvec\ \(q \ xvec) \* (p \ yvec)\ + \y \ xvec\ \y \ (q \ xvec)\ \x \ xvec\ \x \ (q \ xvec)\ Sp Sq + have "\ = \\*(q \ xvec)\(\\y\(\\*(p \ yvec)\((q \ ((x, y)#p) \ N) \' (q \ ((x, y)#p) \ P'))))" + apply(subst boundOutputChainAlpha[where p=q and xvec=xvec and yvec="xvec"]) + defer + apply assumption + apply simp + apply(simp add: eqvts) + apply(simp add: eqvts) + apply(simp add: boundOutputFreshSet(4)) + apply(rule conjI) + apply(simp add: freshChainSimps) + apply(simp add: freshChainSimps) + done + moreover then have "\ = \\*(q \ xvec@y#(p \ yvec))\((q@(x, y)#p) \ N) \' ((q@(x, y)#p) \ P')" + by(simp only: pt2[OF pt_name_inst] boundOutputApp BOresChain.simps) + ultimately show ?thesis + by(simp only: residualInject) + qed + moreover have "\ \ ([(x, y)] \ P) \M\\*((q \ xvec)@(p \ yvec))\\((q@(x, y)#p) \ N)\ \ ((q@(x, y)#p) \ P')" + proof - + note\\ \ P \M\\*(xvec@yvec)\\N\ \ P'\ + moreover from \(p \ yvec) \* N\ \(q \ xvec) \* N\ \xvec \* yvec\ \(q \ xvec) \* yvec\ \(q \ xvec) \* (p \ yvec)\ \(p \ yvec) \* xvec\ Sp Sq + have "((q@p) \ (xvec @ yvec)) \* N" apply(simp only: eqvts) apply(simp only: pt2[OF pt_name_inst]) + by simp + moreover from \(p \ yvec) \* P'\ \(q \ xvec) \* P'\ \xvec \* yvec\ \(q \ xvec) \* yvec\ \(q \ xvec) \* (p \ yvec)\ \(p \ yvec) \* xvec\ Sp Sq + have "((q@p) \ (xvec @ yvec)) \* P'" by(simp del: freshAlphaPerm add: eqvts pt2[OF pt_name_inst]) + moreover from Sp Sq \xvec \* yvec\ \(q \ xvec) \* yvec\ \(q \ xvec) \* (p \ yvec)\ \(p \ yvec) \* xvec\ + have Spq: "set(q@p) \ set(xvec@yvec) \ set((q@p) \ (xvec@yvec))" + by(simp add: pt2[OF pt_name_inst] eqvts) blast + ultimately have "\ \ P \M\\*((q@p) \ (xvec@yvec))\\((q@p) \ N)\ \ ((q@p) \ P')" + apply(simp only: residualInject) + by(erule rev_mp) (subst boundOutputChainAlpha, auto) + with Sp Sq \xvec \* yvec\ \(q \ xvec) \* yvec\ \(q \ xvec) \* (p \ yvec)\ \(p \ yvec) \* xvec\ + have "\ \ P \M\\*((q \ xvec)@(p \ yvec))\\((q@p) \ N)\ \ ((q@p) \ P')" + by(simp add: eqvts pt2[OF pt_name_inst] del: freshAlphaPerm) + then have "([(x, y)] \ \) \ ([(x, y)] \ P) \ [(x, y)] \ (M\\*((q \ xvec)@(p \ yvec))\\((q@p) \ N)\ \ ((q@p) \ P'))" + by(rule semantics.eqvt) + with \x \ \\ \y \ \\ \x \ M\ \y \ M\ \x \ xvec\ \y \ xvec\ \x \ (q \ xvec)\ \y \ (q \ xvec) \\x \ yvec\ \y \ yvec\ \x \ (p \ yvec)\ \y \ (p \ yvec)\ Sp Sq + show ?thesis + apply(simp add: eqvts pt2[OF pt_name_inst]) + by(subst perm_compose[of q], simp)+ + qed + moreover from \x \ supp N\ have "((q@(x, y)#p) \ x) \ ((q@(x, y)#p) \ (supp N))" + by(simp add: pt_set_bij[OF pt_name_inst, OF at_name_inst]) + with \x \ xvec\ \x \ yvec\ \x \ (q \ xvec)\ \x \ (p \ yvec)\ \y \ xvec\ \y \ (q \ xvec)\ Sp Sq + have "y \ supp((q@(x, y)#p)\ N)" by(simp add: pt2[OF pt_name_inst] calc_atm eqvts) + moreover from \distinct xvec\ have "distinct(q \ xvec)" by simp + moreover from \distinct yvec\ have "distinct(p \ yvec)" by simp + moreover note \x \ (q \ xvec)\ \x \ (p \ yvec)\ \x \ M\ \x \ \\ + \(q \ xvec) \* \\ \(q \ xvec) \* P\ \(q \ xvec) \* M\ \(q \ xvec) \* (p \ yvec)\ + \(p \ yvec) \* \\ \(p \ yvec) \* P\ \(p \ yvec) \* M\ \y \ (q \ xvec)\ \y \ (p \ yvec)\ \y \ M\ \y \ C\ \y \ \\ + \(p \ yvec) \* C\ \(q \ xvec) \* C\ + ultimately show Prop by - (rule rOpen, (assumption | simp)+) +next + case(cBrOpen P M xvec yvec N P' x) + from \\ \ P \ \M\\*(xvec@yvec)\\N\ \ P'\ have "distinct(xvec@yvec)" by(force dest: boundOutputDistinct) + then have "xvec \* yvec" by(induct xvec) auto + obtain p where "(p \ yvec) \* \" and "(p \ yvec) \* P" and "(p \ yvec) \* M" + and "(p \ yvec) \* yvec" and "(p \ yvec) \* N" and "(p \ yvec) \* P'" + and "x \ (p \ yvec)" and "(p \ yvec) \* xvec" + and "(p \ yvec) \* C" and Sp: "(set p) \ (set yvec) \ (set(p \ yvec))" + by(rule name_list_avoiding[where xvec=yvec and c="(\, P, M, xvec, yvec, N, P', x, C)"]) + (auto simp add: eqvts fresh_star_prod) + obtain q where "(q \ xvec) \* \" and "(q \ xvec) \* P" and "(q \ xvec) \* M" + and "(q \ xvec) \* xvec" and "(q \ xvec) \* N" and "(q \ xvec) \* P'" + and "x \ (q \ xvec)" and "(q \ xvec) \* yvec" + and "(q \ xvec) \* p" and "(q \ xvec) \* (p \ yvec)" + and "(q \ xvec) \* C" and Sq: "(set q) \ (set xvec) \ (set(q \ xvec))" + by(rule name_list_avoiding[where xvec=xvec and c="(\, P, M, xvec, yvec, p \ yvec, N, P', x, p, C)"]) + (auto simp add: eqvts fresh_star_prod) + obtain y::name where "y \ P" and "y \ C" and "y \ xvec" and "y \ yvec" and "y \ x" and "y \ N" + and "y \ (q \ xvec)" and "y \ (p \ yvec)" and "y \ M" and "y \ \" and "y \ P'" + by(generate_fresh "name") (auto simp add: freshChainSimps) + from \cP = \\x\P\ \y \ P\ have "cP = \\y\([(x, y)] \ P)" by(simp add: alphaRes) + moreover have "cRs = \M\\*((q \ xvec)@y#(p \ yvec))\\((q@(x, y)#p) \ N)\ \ ((q@(x, y)#p) \ P')" + proof - + note \cRs = \M\\*(xvec@x#yvec)\\N\ \ P'\ + moreover have "\\*(xvec@x#yvec)\N \' P' = \\*xvec\(\\x\(\\*yvec\N \' P'))" by(simp add: boundOutputApp) + moreover from \(p \ yvec) \* N\ \(p \ yvec) \* P'\ Sp have "\ = \\*xvec\(\\x\(\\*(p \ yvec)\(p \ N) \' (p \ P')))" + by(simp add: boundOutputChainAlpha'') + moreover with \y \ N\ \y \ P'\ \y \ (p \ yvec)\ \y \ yvec\ \x \ yvec\ \x \ (p \ yvec)\ Sp + moreover have "\ = \\*xvec\(\\y\(\\*(p \ yvec)\(([(x, y)] \ p \ N) \' ([(x, y)] \ p \ P'))))" + by(subst alphaBoundOutput[where y=y]) (simp add: freshChainSimps eqvts)+ + moreover then have "\ = \\*xvec\(\\y\(\\*(p \ yvec)\((((x, y)#p) \ N) \' (((x, y)#p) \ P'))))" + by simp + moreover from \(q \ xvec) \* N\ \(q \ xvec) \* P'\ \xvec \* yvec\ \(p \ yvec) \* xvec\ \(q \ xvec) \* yvec\ \(q \ xvec) \* (p \ yvec)\ + \y \ xvec\ \y \ (q \ xvec)\ \x \ xvec\ \x \ (q \ xvec)\ Sp Sq + have "\ = \\*(q \ xvec)\(\\y\(\\*(p \ yvec)\((q \ ((x, y)#p) \ N) \' (q \ ((x, y)#p) \ P'))))" + apply(subst boundOutputChainAlpha[where p=q and xvec=xvec and yvec="xvec"]) + defer + apply assumption + apply simp + apply(simp add: eqvts) + apply(simp add: eqvts) + apply(simp add: boundOutputFreshSet(4)) + apply(rule conjI) + apply(simp add: freshChainSimps) + apply(simp add: freshChainSimps) + done + moreover then have "\ = \\*(q \ xvec@y#(p \ yvec))\((q@(x, y)#p) \ N) \' ((q@(x, y)#p) \ P')" + by(simp only: pt2[OF pt_name_inst] boundOutputApp BOresChain.simps) + ultimately show ?thesis + by(simp only: residualInject) + qed + moreover have "\ \ ([(x, y)] \ P) \\M\\*((q \ xvec)@(p \ yvec))\\((q@(x, y)#p) \ N)\ \ ((q@(x, y)#p) \ P')" + proof - + note\\ \ P \\M\\*(xvec@yvec)\\N\ \ P'\ + moreover from \(p \ yvec) \* N\ \(q \ xvec) \* N\ \xvec \* yvec\ \(q \ xvec) \* yvec\ \(q \ xvec) \* (p \ yvec)\ \(p \ yvec) \* xvec\ Sp Sq + have "((q@p) \ (xvec @ yvec)) \* N" apply(simp only: eqvts) apply(simp only: pt2[OF pt_name_inst]) + by simp + moreover from \(p \ yvec) \* P'\ \(q \ xvec) \* P'\ \xvec \* yvec\ \(q \ xvec) \* yvec\ \(q \ xvec) \* (p \ yvec)\ \(p \ yvec) \* xvec\ Sp Sq + have "((q@p) \ (xvec @ yvec)) \* P'" by(simp del: freshAlphaPerm add: eqvts pt2[OF pt_name_inst]) + moreover from Sp Sq \xvec \* yvec\ \(q \ xvec) \* yvec\ \(q \ xvec) \* (p \ yvec)\ \(p \ yvec) \* xvec\ + have Spq: "set(q@p) \ set(xvec@yvec) \ set((q@p) \ (xvec@yvec))" + by(simp add: pt2[OF pt_name_inst] eqvts) blast + ultimately have "\ \ P \\M\\*((q@p) \ (xvec@yvec))\\((q@p) \ N)\ \ ((q@p) \ P')" + apply(simp only: residualInject) + by(erule rev_mp) (subst boundOutputChainAlpha, auto) + with Sp Sq \xvec \* yvec\ \(q \ xvec) \* yvec\ \(q \ xvec) \* (p \ yvec)\ \(p \ yvec) \* xvec\ + have "\ \ P \\M\\*((q \ xvec)@(p \ yvec))\\((q@p) \ N)\ \ ((q@p) \ P')" + by(simp add: eqvts pt2[OF pt_name_inst] del: freshAlphaPerm) + then have "([(x, y)] \ \) \ ([(x, y)] \ P) \ [(x, y)] \ (\M\\*((q \ xvec)@(p \ yvec))\\((q@p) \ N)\ \ ((q@p) \ P'))" + by(rule semantics.eqvt) + with \x \ \\ \y \ \\ \x \ M\ \y \ M\ \x \ xvec\ \y \ xvec\ \x \ (q \ xvec)\ \y \ (q \ xvec) \\x \ yvec\ \y \ yvec\ \x \ (p \ yvec)\ \y \ (p \ yvec)\ Sp Sq + show ?thesis + apply(simp add: eqvts pt2[OF pt_name_inst]) + by(subst perm_compose[of q], simp)+ + qed + moreover from \x \ supp N\ have "((q@(x, y)#p) \ x) \ ((q@(x, y)#p) \ (supp N))" + by(simp add: pt_set_bij[OF pt_name_inst, OF at_name_inst]) + with \x \ xvec\ \x \ yvec\ \x \ (q \ xvec)\ \x \ (p \ yvec)\ \y \ xvec\ \y \ (q \ xvec)\ Sp Sq + have "y \ supp((q@(x, y)#p)\ N)" by(simp add: pt2[OF pt_name_inst] calc_atm eqvts) + moreover from \distinct xvec\ have "distinct(q \ xvec)" by simp + moreover from \distinct yvec\ have "distinct(p \ yvec)" by simp + moreover note \x \ (q \ xvec)\ \x \ (p \ yvec)\ \x \ M\ \x \ \\ + \(q \ xvec) \* \\ \(q \ xvec) \* P\ \(q \ xvec) \* M\ \(q \ xvec) \* (p \ yvec)\ + \(p \ yvec) \* \\ \(p \ yvec) \* P\ \(p \ yvec) \* M\ \y \ (q \ xvec)\ \y \ (p \ yvec)\ \y \ M\ \y \ C\ \y \ \\ + \(p \ yvec) \* C\ \(q \ xvec) \* C\ + ultimately show Prop by - (rule rBrOpen, (assumption | simp)+) +next + case(cScope P \ P' x) + obtain p::"name prm" where "(bn(p \ \)) \* \" and "(bn(p \ \)) \* P" + and "(bn(p \ \)) \* \" and "(bn(p \ \)) \* P'" and "x \ bn(p \ \)" + and "distinctPerm p" + and "(bn(p \ \)) \* C" and Sp: "(set p) \ set(bn \) \ (set(bn(p \ \)))" + by(rule name_list_avoiding[where xvec="bn \" and c="(\, P, \, x, P', C)"]) (auto simp add: eqvts) + obtain y::name where "y \ \" and "y \ P" and "y \ (p \ P')" and "y \ (p \ \)" and "y \ C" + by(generate_fresh "name") (auto simp add: freshChainSimps simp del: actionFresh) + from \bn \ \* subject \\ \distinctPerm p\ have "bn(p \ \) \* subject(p \ \)" + by(subst fresh_star_bij[symmetric, of _ _ p]) (simp add: eqvts) + from \distinct(bn \)\ \distinctPerm p\ have "distinct(bn(p \ \))" + by(subst distinctClosed[symmetric, of _ p]) (simp add: eqvts) + from \x \ \\ \x \ (bn(p \ \))\ \distinctPerm p\ Sp have "x \ (p \ \)" + by(subst fresh_bij[symmetric, of _ _ p]) (simp add: eqvts freshChainSimps) + + from \cP = \\x\P\ \y \ P\ have "cP = \\y\([(x, y)] \ P)" by(simp add: alphaRes) + moreover from \cRs = \ \ \\x\P'\ \bn \ \* subject \\ \(bn(p \ \)) \* \\ \x \ bn(p \ \)\ \(bn(p \ \)) \* P'\ \x \ \\ Sp + have "cRs = (p \ \) \ \\x\(p \ P')" + by(force simp add: residualAlpha) + with \y \ (p \ P')\ have "cRs = (p \ \) \ \\y\([(x, y)] \ p \ P')" + by(simp add: alphaRes) + moreover from \\ \ P \\ \ P'\ \bn \ \* subject \\ \(bn(p \ \)) \* \\ \(bn(p \ \)) \* P'\ Sp + have "\ \ P \(p \ \) \ (p \ P')" by(force simp add: residualAlpha) + then have"([(x, y)] \ \) \ ([(x, y)] \ P) \[(x, y)] \ ((p \ \) \ (p \ P'))" + by(rule eqvts) + with \x \ \\ \y \ \\ \y \ (p \ \)\ \x \ (p \ \)\ Sp \distinctPerm p\ + have "\ \ ([(x, y)] \ P) \(p \ \) \ ([(x, y)] \ p \ P')" + by(simp add: eqvts) + moreover from \bn(p \ \) \* P\ \y \ (p \ \)\ \y \ P\ have "bn(p \ \) \* ([(x, y)] \ P)" + by(auto simp add: fresh_star_def fresh_left calc_atm) (simp add: fresh_def name_list_supp) + moreover from \distinct(bn \)\ have "distinct(p \ bn \)" by simp + then have "distinct(bn(p \ \))" by(simp add: eqvts) + ultimately show ?thesis + using \y \ \\ \y \ (p \ \)\ \y \ C\ \bn(p \ \) \* \\ \bn(p \ \) \* subject(p \ \)\ \bn(p \ \) \* C\ + by(metis rScope) +next + case(Bang P) + then show ?thesis by(metis rBang) +qed + +nominal_primrec + inputLength :: "('a::fs_name, 'b::fs_name, 'c::fs_name) psi \ nat" + and inputLength' :: "('a::fs_name, 'b::fs_name, 'c::fs_name) input \ nat" + and inputLength'' :: "('a::fs_name, 'b::fs_name, 'c::fs_name) psiCase \ nat" + +where + "inputLength (\) = 0" +| "inputLength (M\N\.P) = 0" +| "inputLength (M\I) = inputLength' I" +| "inputLength (Case C) = 0" +| "inputLength (P \ Q) = 0" +| "inputLength (\\x\P) = 0" +| "inputLength (\\\) = 0" +| "inputLength (!P) = 0" + +| "inputLength' (Trm M P) = 0" +| "inputLength' (\ y I) = 1 + (inputLength' I)" + +| "inputLength'' (\\<^sub>c) = 0" +| "inputLength'' (\\ \ P C) = 0" + apply(finite_guess)+ + apply(rule TrueI)+ + by(fresh_guess add: fresh_nat)+ + +nominal_primrec boundOutputLength :: "('a, 'b, 'c) boundOutput \ nat" + where + "boundOutputLength (BOut M P) = 0" + | "boundOutputLength (BStep x B) = (boundOutputLength B) + 1" + apply(finite_guess)+ + apply(rule TrueI)+ + by(fresh_guess add: fresh_nat)+ + +nominal_primrec residualLength :: "('a, 'b, 'c) residual \ nat" + where + "residualLength (RIn M N P) = 0" + | "residualLength (RBrIn M N P) = 0" + | "residualLength (ROut M B) = boundOutputLength B" + | "residualLength (RBrOut M B) = boundOutputLength B" + | "residualLength (RTau P) = 0" + by(rule TrueI)+ + +lemma inputLengthProc[simp]: + shows "inputLength(M\\*xvec N\.P) = length xvec" + by(induct xvec) auto + +lemma boundOutputLengthSimp[simp]: + shows "residualLength(M\\*xvec\\N\ \ P) = length xvec" + and "residualLength(\M\\*xvec\\N\ \ P) = length xvec" + by(induct xvec) (auto simp add: residualInject) + +lemma boundOuputLengthSimp2[simp]: + shows "residualLength(\ \ P) = length(bn \)" + by(nominal_induct \ rule: action.strong_induct, auto) (auto simp add: residualInject) + +lemmas [simp del] = inputLength_inputLength'_inputLength''.simps residualLength.simps boundOutputLength.simps + +lemma constructPerm: + fixes xvec :: "name list" + and yvec :: "name list" + +assumes "length xvec = length yvec" + and "xvec \* yvec" + and "distinct xvec" + and "distinct yvec" + +obtains p where "set p \ set xvec \ set(p \ xvec)" and "distinctPerm p" and "yvec = p \ xvec" +proof - + assume "\p. \set p \ set xvec \ set (p \ xvec); distinctPerm p; yvec = p \ xvec\ \ thesis" + moreover obtain n where "n = length xvec" by auto + with assms have "\p. (set p) \ (set xvec) \ set (yvec) \ distinctPerm p \ yvec = p \ xvec" + proof(induct n arbitrary: xvec yvec) + case(0 xvec yvec) + then show ?case by simp + next + case(Suc n xvec yvec) + from \Suc n = length xvec\ + obtain x xvec' where "xvec = x#xvec'" and "length xvec' = n" + by(cases xvec) auto + from \length xvec = length yvec\ \xvec = x # xvec'\ + obtain y yvec' where "length xvec' = length yvec'" and "yvec = y#yvec'" + by(cases yvec) auto + from \xvec = x#xvec'\ \yvec=y#yvec'\ \xvec \* yvec\ + have "x \ y" and "xvec' \* yvec'" and "x \ yvec'" and "y \ xvec'" + by(auto simp add: fresh_list_cons) + from \distinct xvec\ \distinct yvec\ \xvec=x#xvec'\ \yvec=y#yvec'\ have "x \ xvec'" and "y \ yvec'" and "distinct xvec'" and "distinct yvec'" + by simp+ + from \Suc n = length xvec\ \xvec=x#xvec'\ have "n = length xvec'" by simp + with \length xvec' = length yvec'\ \xvec' \* yvec'\ \distinct xvec'\ \distinct yvec'\ + obtain p where S: "set p \ set xvec' \ set yvec'" and "distinctPerm p" and "yvec' = p \ xvec'" + by - (drule Suc,auto) + from S have "set((x, y)#p) \ set(x#xvec') \ set(y#yvec')" by auto + moreover from \x \ xvec'\ \x \ yvec'\ \y \ xvec'\ \y \ yvec'\ S have "x \ p" and "y \ p" + apply(induct p) + by(clarsimp simp add: fresh_list_nil fresh_list_cons fresh_prod name_list_supp; force simp add: fresh_def)+ + + with S \distinctPerm p\ \x \ y\ have "distinctPerm((x, y)#p)" by auto + moreover from \yvec' = p \ xvec'\ \x \ p\ \y \ p\ \x \ xvec'\ \y \ xvec'\ have "(y#yvec') = ((x, y)#p) \ (x#xvec')" + by(simp add: calc_atm freshChainSimps) + ultimately show ?case using \xvec=x#xvec'\ \yvec=y#yvec'\ + by blast + qed + ultimately show ?thesis by blast +qed + +lemma distinctApend[simp]: + fixes xvec :: "name list" + and yvec :: "name list" + +shows "(set xvec \ set yvec = {}) = xvec \* yvec" + by(auto simp add: fresh_star_def name_list_supp fresh_def) + +lemma lengthAux: + fixes xvec :: "name list" + and y :: name + and yvec :: "name list" + +assumes "length xvec = length(y#yvec)" + +obtains z zvec where "xvec = z#zvec" and "length zvec = length yvec" + using assms + by(induct xvec arbitrary: yvec y) auto + +lemma lengthAux2: + fixes xvec :: "name list" + and yvec :: "name list" + and zvec :: "name list" + +assumes "length xvec = length(yvec@y#zvec)" + +obtains xvec1 x xvec2 where "xvec=xvec1@x#xvec2" and "length xvec1 = length yvec" and "length xvec2 = length zvec" +proof - + assume "\xvec1 x xvec2. + \xvec = xvec1 @ x # xvec2; length xvec1 = length yvec; + length xvec2 = length zvec\ + \ thesis" + moreover from assms have "\xvec1 x xvec2. xvec=xvec1@x#xvec2 \ length xvec1 = length yvec \ length xvec2 = length zvec" + apply - + apply(rule exI[where x="take (length yvec) xvec"]) + apply(rule exI[where x="hd(drop (length yvec) xvec)"]) + apply(rule exI[where x="tl(drop (length yvec) xvec)"]) + by auto + ultimately show ?thesis by blast +qed + +lemma semanticsCases[consumes 19, case_names cInput cBrInput cOutput cBrOutput cCase cPar1 cPar2 cComm1 cComm2 cBrMerge cBrComm1 cBrComm2 cBrClose cOpen cBrOpen cScope cBang]: + fixes \ :: 'b + and cP :: "('a, 'b, 'c) psi" + and cRs :: "('a, 'b, 'c) residual" + and C :: "'f::fs_name" + and x1 :: name + and x2 :: name + and x3 :: name + and x4 :: name + and xvec1 :: "name list" + and xvec2 :: "name list" + and xvec3 :: "name list" + and xvec4 :: "name list" + and xvec5 :: "name list" + and xvec6 :: "name list" + and xvec7 :: "name list" + and xvec8 :: "name list" + and xvec9 :: "name list" + +assumes "\ \ cP \cRs" + and "length xvec1 = inputLength cP" and "distinct xvec1" + and "length xvec6 = inputLength cP" and "distinct xvec6" + and "length xvec2 = residualLength cRs" and "distinct xvec2" + and "length xvec3 = residualLength cRs" and "distinct xvec3" + and "length xvec4 = residualLength cRs" and "distinct xvec4" + and "length xvec5 = residualLength cRs" and "distinct xvec5" + and "length xvec7 = residualLength cRs" and "distinct xvec7" + and "length xvec8 = residualLength cRs" and "distinct xvec8" + and "length xvec9 = residualLength cRs" and "distinct xvec9" + and rInput: "\M K N Tvec P. (\xvec1 \* \; xvec1 \* cP; xvec1 \* cRs\ \ cP = M\\*xvec1 N\.P \ cRs = K\(N[xvec1::=Tvec])\ \ P[xvec1::=Tvec] \ + \ \ M \ K \ distinct xvec1 \ set xvec1 \ supp N \ length xvec1=length Tvec \ + xvec1 \* Tvec \ xvec1 \* \ \ xvec1 \* M \ xvec1 \* K) \ Prop" + and rBrInput: "\M K N Tvec P. (\xvec6 \* \; xvec6 \* cP; xvec6 \* cRs\ \ cP = M\\*xvec6 N\.P \ cRs = \K\(N[xvec6::=Tvec])\ \ P[xvec6::=Tvec] \ + \ \ K \ M \ distinct xvec6 \ set xvec6 \ supp N \ length xvec6=length Tvec \ + xvec6 \* Tvec \ xvec6 \* \ \ xvec6 \* M \ xvec6 \* K) \ Prop" + and rOutput: "\M K N P. \cP = M\N\.P; cRs = K\N\ \ P; \ \ M \ K\ \ Prop" + and rBrOutput: "\M K N P. \cP = M\N\.P; cRs = \K\N\ \ P; \ \ M \ K\ \ Prop" + and rCase: "\Cs P \. \cP = Cases Cs; \ \ P \cRs; (\, P) \ set Cs; \ \ \; guarded P\ \ Prop" + and rPar1: "\\\<^sub>Q P \ P' Q A\<^sub>Q. (\xvec2 \* \; xvec2 \* cP; xvec2 \* cRs\ \ + cP = P \ Q \ cRs = \ \ (P' \ Q) \ xvec2 = bn \ \ + \ \ \\<^sub>Q \ P \\ \ P' \ extractFrame Q = \A\<^sub>Q, \\<^sub>Q\ \ distinct A\<^sub>Q \ + A\<^sub>Q \* P \ A\<^sub>Q \* Q \ A\<^sub>Q \* \ \ A\<^sub>Q \* \ \ A\<^sub>Q \* P' \ A\<^sub>Q \* C) \ Prop" + and rPar2: "\\\<^sub>P Q \ Q' P A\<^sub>P. (\xvec3 \* \; xvec3 \* cP; xvec3 \* cRs\ \ + cP = P \ Q \ cRs = \ \ (P \ Q') \ xvec3 = bn \ \ + \ \ \\<^sub>P \ Q \\ \ Q' \ extractFrame P = \A\<^sub>P, \\<^sub>P\ \ distinct A\<^sub>P \ + A\<^sub>P \* P \ A\<^sub>P \* Q \ A\<^sub>P \* \ \ A\<^sub>P \* \ \ A\<^sub>P \* Q' \ A\<^sub>P \* C) \ Prop" + and rComm1: "\\\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q K xvec Q' A\<^sub>Q. + \cP = P \ Q; cRs = \ \ \\*xvec\P' \ Q'; + \ \ \\<^sub>Q \ P \M\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \ \ \\<^sub>P \ Q \ K\\*xvec\\N\ \ Q'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + \ \ \\<^sub>P \ \\<^sub>Q \ M \ K; A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* N; + A\<^sub>P \* P'; A\<^sub>P \* Q; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* xvec; A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; + A\<^sub>Q \* P; A\<^sub>Q \* K; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* Q; A\<^sub>Q \* Q'; A\<^sub>Q \* xvec; + xvec \* \; xvec \* \\<^sub>P; xvec \* \\<^sub>Q; xvec \* P; xvec \* M; xvec \* Q; + xvec \* K; A\<^sub>P \* C; A\<^sub>Q \* C; xvec \* C; distinct xvec\ \ Prop" + and rComm2: "\\\<^sub>Q P M xvec N P' A\<^sub>P \\<^sub>P Q K Q' A\<^sub>Q. + \cP = P \ Q; cRs = \ \ \\*xvec\P' \ Q'; + \ \ \\<^sub>Q \ P \M\\*xvec\\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \ \ \\<^sub>P \ Q \ K\N\ \ Q'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + \ \ \\<^sub>P \ \\<^sub>Q \ M \ K; A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* N; + A\<^sub>P \* P'; A\<^sub>P \* Q; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* xvec; A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; + A\<^sub>Q \* P; A\<^sub>Q \* K; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* Q; A\<^sub>Q \* Q'; A\<^sub>Q \* xvec; + xvec \* \; xvec \* \\<^sub>P; xvec \* \\<^sub>Q; xvec \* P; xvec \* M; xvec \* Q; + xvec \* K; A\<^sub>P \* C; A\<^sub>Q \* C; xvec \* C; distinct xvec\ \ Prop" + and rBrMerge: "\\\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q Q' A\<^sub>Q. + \cP = (P \ Q); cRs = \M\N\ \ (P' \ Q'); + \ \ \\<^sub>Q \ P \ \M\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \ \ \\<^sub>P \ Q \ \M\N\ \ Q'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* N; A\<^sub>P \* P'; + A\<^sub>P \* Q; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* M; A\<^sub>Q \* M; + A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; A\<^sub>Q \* P; A\<^sub>Q \* N; A\<^sub>Q \* P'; + A\<^sub>Q \* Q; A\<^sub>Q \* Q'; A\<^sub>P \* C; A\<^sub>Q \* C\ \ Prop" + and rBrComm1: "\\\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q Q' A\<^sub>Q. + (\xvec7 \* \; xvec7 \* cP; xvec7 \* cRs\ \ + cP = P \ Q \ cRs = \M\\*xvec7\\N\ \ (P' \ Q') \ + \ \ \\<^sub>Q \ P \\M\N\ \ P' \ extractFrame P = \A\<^sub>P, \\<^sub>P\ \ distinct A\<^sub>P \ + \ \ \\<^sub>P \ Q \ \M\\*xvec7\\N\ \ Q' \ extractFrame Q = \A\<^sub>Q, \\<^sub>Q\ \ distinct A\<^sub>Q \ + A\<^sub>P \* \ \ A\<^sub>P \* \\<^sub>Q \ A\<^sub>P \* P \ A\<^sub>P \* N \ + A\<^sub>P \* P'\ A\<^sub>P \* Q \ A\<^sub>P \* Q' \ A\<^sub>P \* A\<^sub>Q \ A\<^sub>P \* xvec7 \ A\<^sub>Q \* \ \ A\<^sub>Q \* \\<^sub>P \ + A\<^sub>Q \* P \ A\<^sub>Q \* N \ A\<^sub>Q \* P' \ A\<^sub>Q \* Q \ A\<^sub>Q \* Q' \ A\<^sub>Q \* xvec7 \ + xvec7 \* \ \ xvec7 \* \\<^sub>P \ xvec7 \* \\<^sub>Q \ xvec7 \* P \ xvec7 \* Q \ + A\<^sub>P \* M \ A\<^sub>Q \* M \ xvec7 \* M \ + A\<^sub>P \* C \ A\<^sub>Q \* C \ distinct xvec7) \ Prop" + and rBrComm2: "\\\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q Q' A\<^sub>Q. + (\xvec8 \* \; xvec8 \* cP; xvec8 \* cRs\ \ + cP = P \ Q \ cRs = \M\\*xvec8\\N\ \ (P' \ Q') \ + \ \ \\<^sub>Q \ P \\M\\*xvec8\\N\ \ P' \ extractFrame P = \A\<^sub>P, \\<^sub>P\ \ distinct A\<^sub>P \ + \ \ \\<^sub>P \ Q \ \M\N\ \ Q' \ extractFrame Q = \A\<^sub>Q, \\<^sub>Q\ \ distinct A\<^sub>Q \ + A\<^sub>P \* \ \ A\<^sub>P \* \\<^sub>Q \ A\<^sub>P \* P \ A\<^sub>P \* N \ + A\<^sub>P \* P' \ A\<^sub>P \* Q \ A\<^sub>P \* Q' \ A\<^sub>P \* A\<^sub>Q \ A\<^sub>P \* xvec8 \ A\<^sub>Q \* \ \ A\<^sub>Q \* \\<^sub>P \ + A\<^sub>Q \* P \ A\<^sub>Q \* N \ A\<^sub>Q \* P' \ A\<^sub>Q \* Q \ A\<^sub>Q \* Q' \ A\<^sub>Q \* xvec8 \ + xvec8 \* \ \ xvec8 \* \\<^sub>P \ xvec8 \* \\<^sub>Q \ xvec8 \* P \ xvec8 \* Q \ + A\<^sub>P \* M \ A\<^sub>Q \* M \ xvec8 \* M \ + A\<^sub>P \* C \ A\<^sub>Q \* C \ distinct xvec8) \ Prop" + and rBrClose: "\P M N xvec P'. + (\x3 \ \; x3 \ cP; x3 \ cRs\ \ + cP = (\\x3\P) \ cRs = \ \ \\x3\(\\*xvec\P') \ + x3 \ supp M \ + \ \ P \ \M\\*xvec\\N\ \ P' \ + distinct xvec \ xvec \* \ \ xvec \* P \ + xvec \* M \ xvec \* C \ + x3 \ \ \ x3 \ xvec) \ Prop" + and rOpen: "\P M xvec y yvec N P'. + (\xvec4 \* \; xvec4 \* cP; xvec4 \* cRs; x1 \ \; x1 \ cP; x1 \ cRs; x1 \ xvec4\ \ + cP = \\x1\P \ cRs = M\\*(xvec@x1#yvec)\\N\ \ P' \ xvec4=xvec@y#yvec \ + \ \ P \M\\*(xvec@yvec)\\N\ \ P' \ x1 \ supp N \ x1 \ xvec \ x1 \ yvec \ + distinct xvec \ distinct yvec \ xvec \* \ \ xvec \* P \ xvec \* M \ xvec \* yvec \ + yvec \* \ \ yvec \* P \ yvec \* M) \ Prop" + and rBrOpen: "\P M xvec y yvec N P'. + (\xvec9 \* \; xvec9 \* cP; xvec9 \* cRs; x4 \ \; x4 \ cP; x4 \ cRs; x4 \ xvec9\ \ + cP = \\x4\P \ cRs = \M\\*(xvec@x4#yvec)\\N\ \ P' \ xvec9=xvec@y#yvec \ + \ \ P \\M\\*(xvec@yvec)\\N\ \ P' \ x4 \ supp N \ x4 \ xvec \ x4 \ yvec \ + distinct xvec \ distinct yvec \ xvec \* \ \ xvec \* P \ xvec \* M \ xvec \* yvec \ + yvec \* \ \ yvec \* P \ yvec \* M) \ Prop" + and rScope: "\P \ P'. (\xvec5 \* \; xvec5 \* cP; xvec5 \* cRs; x2 \ \; x2 \ cP; x2 \ cRs; x2 \ xvec5\ \ + cP = \\x2\P \ cRs = \ \ \\x2\P' \ xvec5 = bn \ \ + \ \ P \\ \ P' \ x2 \ \ \ x2 \ \ \ bn \ \* subject \ \ distinct(bn \)) \ Prop" + and rBang: "\P. \cP = !P; + \ \ P \ !P \cRs; guarded P\ \ Prop" +shows Prop + using \\ \ cP \cRs\ +proof(cases rule: semanticsCasesAux[where C="(xvec1, xvec2, xvec3, xvec4, xvec5, xvec6, xvec7, xvec8, xvec9, x1, x2, x3, x4, cP, cRs, C)"]) + case(cInput M K xvec N Tvec P) + have B: "cP = M\\*xvec N\.P" and C: "cRs = K\(N[xvec::=Tvec])\ \ (P[xvec::=Tvec])" + by fact+ + from \xvec \* (xvec1, xvec2, xvec3, xvec4, xvec5, xvec6, xvec7, xvec8, xvec9, x1, x2, x3, x4, cP, cRs, C)\ have "xvec \* xvec1" by simp + + from \length xvec1 = inputLength cP\ B have "length xvec1 = length xvec" + by simp + then obtain p where S: "set p \ set xvec \ set(p \ xvec)" and "distinctPerm p" and "xvec1 = p \ xvec" + using \xvec \* xvec1\ \distinct xvec\ \distinct xvec1\ + by - (rule constructPerm[where xvec=xvec and yvec=xvec1], auto) + show ?thesis + proof(rule rInput[where M=M and K=K and N = "p \ N" and Tvec=Tvec and P="p \ P"]) + assume "xvec1 \* \" and "xvec1 \* cP" and "xvec1 \* cRs" + from B \xvec \* xvec1\ \xvec1 \* cP\ have "xvec1 \* N" and "xvec1 \* P" + by(auto simp add: fresh_star_def inputChainFresh name_list_supp) (auto simp add: fresh_def) + moreover from \cP = M\\*xvec N\.P\ S \xvec1 \* N\ \xvec1 \* P\ \xvec1 = p \ xvec\ + have "cP = M\\*xvec1 (p \ N)\.(p \ P)" + apply simp + by(subst inputChainAlpha) auto + moreover from \cRs = K\(N[xvec::=Tvec])\ \ P[xvec::=Tvec]\ S \xvec1 \* N\ \xvec1 \* P\ \xvec1 = p \ xvec\ \length xvec = length Tvec\ \distinctPerm p\ + have "cRs = K\((p \ N)[xvec1::=Tvec])\ \ (p \ P)[xvec1::=Tvec]" + by(simp add: renaming substTerm.renaming) + moreover note \\ \ M \ K\ + moreover from \distinct xvec\ \xvec1 = p \ xvec\ have "distinct xvec1" by simp + moreover from \set xvec \ supp N\ have "(p \ set xvec) \ (p \ (supp N))" + by(simp add: eqvts) + with \xvec1 = p \ xvec\ have "set xvec1 \ supp(p \ N)" by(simp add: eqvts) + moreover from \length xvec = length Tvec\ \xvec1 = p \ xvec\ have "length xvec1 = length Tvec" + by simp + + moreover from \xvec1 \* cRs\ C \length xvec = length Tvec\ \distinct xvec\ \set xvec \ supp N\ + have "(set xvec1) \* Tvec" + by - (rule substTerm.subst3Chain[where T=N], auto) + then have "xvec1 \* Tvec" by simp + moreover from \xvec \* Tvec\ have "(p \ xvec) \* (p \ Tvec)" by(simp add: fresh_star_bij) + with S \xvec \* Tvec\ \xvec1 \* Tvec\ \xvec1 = p \ xvec\ have "xvec1 \* Tvec" by simp + moreover note \xvec1 \* \\ + moreover from \xvec \* M\ have "(p \ xvec) \* (p \ M)" by(simp add: fresh_star_bij) + with S \xvec \* M\ \xvec1 \* cP\ B \xvec1 = p \ xvec\ have "xvec1 \* M" by simp + moreover from \xvec \* K\ have "(p \ xvec) \* (p \ K)" by(simp add: fresh_star_bij) + with S \xvec \* K\ \xvec1 \* cRs\ C \xvec1 = p \ xvec\ have "xvec1 \* K" by simp + ultimately show "cP = M\\*xvec1 p \ N\.p \ P \ cRs = K\(p \ N)[xvec1::=Tvec]\ \ (p \ P)[xvec1::=Tvec] \ + \ \ M \ K \ distinct xvec1 \ set xvec1 \ supp (p \ N) \ length xvec1 = length Tvec \ + xvec1 \* Tvec \ xvec1 \* \ \ xvec1 \* M \ xvec1 \* K" + by blast + qed +next + case(cBrInput M K xvec N Tvec P) + have B: "cP = M\\*xvec N\.P" and C: "cRs = \K\(N[xvec::=Tvec])\ \ (P[xvec::=Tvec])" + by fact+ + from \xvec \* (xvec1, xvec2, xvec3, xvec4, xvec5, xvec6, xvec7, xvec8, xvec9, x1, x2, x3, x4, cP, cRs, C)\ have "xvec \* xvec6" by simp + + from \length xvec6 = inputLength cP\ B have "length xvec6 = length xvec" + by simp + then obtain p where S: "set p \ set xvec \ set(p \ xvec)" and "distinctPerm p" and "xvec6 = p \ xvec" + using \xvec \* xvec6\ \distinct xvec\ \distinct xvec6\ + by - (rule constructPerm[where xvec=xvec and yvec=xvec6], auto) + show ?thesis + proof(rule rBrInput[where M=M and K=K and N = "p \ N" and Tvec=Tvec and P="p \ P"]) + assume "xvec6 \* \" and "xvec6 \* cP" and "xvec6 \* cRs" + from B \xvec \* xvec6\ \xvec6 \* cP\ have "xvec6 \* N" and "xvec6 \* P" + by(auto simp add: fresh_star_def inputChainFresh name_list_supp) (auto simp add: fresh_def) + + moreover from \cP = M\\*xvec N\.P\ S \xvec6 \* N\ \xvec6 \* P\ \xvec6 = p \ xvec\ + have "cP = M\\*xvec6 (p \ N)\.(p \ P)" + apply simp + by(subst inputChainAlpha) auto + moreover from \cRs = \K\(N[xvec::=Tvec])\ \ P[xvec::=Tvec]\ S \xvec6 \* N\ \xvec6 \* P\ \xvec6 = p \ xvec\ \length xvec = length Tvec\ \distinctPerm p\ + have "cRs = \K\((p \ N)[xvec6::=Tvec])\ \ (p \ P)[xvec6::=Tvec]" + by(simp add: renaming substTerm.renaming) + moreover note \\ \ K \ M\ + moreover from \distinct xvec\ \xvec6 = p \ xvec\ have "distinct xvec6" by simp + moreover from \set xvec \ supp N\ have "(p \ set xvec) \ (p \ (supp N))" + by(simp add: eqvts) + with \xvec6 = p \ xvec\ have "set xvec6 \ supp(p \ N)" by(simp add: eqvts) + moreover from \length xvec = length Tvec\ \xvec6 = p \ xvec\ have "length xvec6 = length Tvec" + by simp + + moreover from \xvec6 \* cRs\ C \length xvec = length Tvec\ \distinct xvec\ \set xvec \ supp N\ + have "(set xvec6) \* Tvec" + by - (rule substTerm.subst3Chain[where T=N], auto) + then have "xvec6 \* Tvec" by simp + moreover from \xvec \* Tvec\ have "(p \ xvec) \* (p \ Tvec)" by(simp add: fresh_star_bij) + with S \xvec \* Tvec\ \xvec6 \* Tvec\ \xvec6 = p \ xvec\ have "xvec6 \* Tvec" by simp + moreover note \xvec6 \* \\ + moreover from \xvec \* M\ have "(p \ xvec) \* (p \ M)" by(simp add: fresh_star_bij) + with S \xvec \* M\ \xvec6 \* cP\ B \xvec6 = p \ xvec\ have "xvec6 \* M" by simp + moreover from \xvec \* K\ have "(p \ xvec) \* (p \ K)" by(simp add: fresh_star_bij) + with S \xvec \* K\ \xvec6 \* cRs\ C \xvec6 = p \ xvec\ have "xvec6 \* K" by simp + ultimately show "cP = M\\*xvec6 p \ N\.p \ P \ + cRs = \K\(p \ N)[xvec6::=Tvec]\ \ (p \ P)[xvec6::=Tvec] \ + \ \ K \ M \ distinct xvec6 \ set xvec6 \ supp (p \ N) \ length xvec6 = length Tvec \ + xvec6 \* Tvec \ xvec6 \* \ \ xvec6 \* M \ xvec6 \* K" by blast + qed +next + case(cOutput M K N P) + then show ?thesis by(rule rOutput) +next + case(cBrOutput M N P) + then show ?thesis by(rule rBrOutput) +next + case(cCase Cs P \) + then show ?thesis by(rule rCase) +next + case(cPar1 \\<^sub>Q P \ P' Q A\<^sub>Q) + have B: "cP = P \ Q" and C: "cRs = \ \ P' \ Q" + by fact+ + from \bn \ \* (xvec1, xvec2, xvec3, xvec4, xvec5, xvec6, xvec7, xvec8, xvec9, x1, x2, x3, x4, cP, cRs, C)\ have "bn \ \* xvec2" by simp + from \A\<^sub>Q \* (xvec1, xvec2, xvec3, xvec4, xvec5, xvec6, xvec7, xvec8, xvec9, x1, x2, x3, x4, cP, cRs, C)\ have "A\<^sub>Q \* xvec2" and "A\<^sub>Q \* C" by simp+ + + from \length xvec2 = residualLength cRs\ C have "length xvec2 = length(bn \)" + by simp + then obtain p where S: "set p \ set(bn \) \ set(bn(p \ \))" and "distinctPerm p" and "xvec2= bn(p \ \)" + using \bn \ \* xvec2\ \distinct(bn \)\ \distinct xvec2\ + by - (rule constructPerm[where xvec="bn \" and yvec=xvec2], auto simp add: eqvts) + show ?thesis + proof(rule rPar1[where P=P and Q=Q and \="p \ \" and P'="p \ P'" and A\<^sub>Q=A\<^sub>Q and \\<^sub>Q=\\<^sub>Q]) + assume "xvec2 \* \" and "xvec2 \* cP" and "xvec2 \* cRs" + note \cP = P \ Q\ + moreover from C S \bn \ \* xvec2\ \xvec2 \* cRs\ \xvec2 = bn(p \ \)\ \bn \ \* subject \\ \xvec2 \* cP\ \bn \ \* Q\ + have "cRs = (p \ \) \ (p \ P') \ Q" + apply clarsimp + by(subst residualAlpha[where p=p]) auto + moreover note \xvec2 = bn(p \ \)\ + moreover from \\ \ \\<^sub>Q \ P \\ \ P'\ S B C S \bn \ \* xvec2\ \xvec2 \* cRs\ \xvec2 = bn(p \ \)\ \bn \ \* subject \\ \xvec2 \* cP\ + have "\ \ \\<^sub>Q \ P \(p \ \) \ (p \ P')" + by(subst residualAlpha[symmetric]) auto + moreover note \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ \distinct A\<^sub>Q\ \A\<^sub>Q \* P\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* \\ \A\<^sub>Q \* \\ + moreover from \A\<^sub>Q \* \\ \A\<^sub>Q \* xvec2\ S \xvec2 = bn(p \ \)\ \distinctPerm p\ have "A\<^sub>Q \* (p \ \)" + by(subst fresh_star_bij[symmetric, where pi=p]) simp + moreover from \A\<^sub>Q \* P'\ \A\<^sub>Q \* \\ \A\<^sub>Q \* xvec2\ S \xvec2 = bn(p \ \)\ \distinctPerm p\ have "A\<^sub>Q \* (p \ P')" + by(subst fresh_star_bij[symmetric, where pi=p]) simp + moreover note \A\<^sub>Q \* C\ + ultimately show "cP = P \ Q \ cRs = (p \ \) \ (p \ P') \ Q \ xvec2 = bn (p \ \) \ + \ \ \\<^sub>Q \ P \ (p \ \) \ p \ P' \ extractFrame Q = \A\<^sub>Q, \\<^sub>Q\ \ distinct A\<^sub>Q \ A\<^sub>Q \* P \ + A\<^sub>Q \* Q \ A\<^sub>Q \* \ \ A\<^sub>Q \* (p \ \) \ A\<^sub>Q \* (p \ P') \ A\<^sub>Q \* C" by blast + qed +next + case(cPar2 \\<^sub>P Q \ Q' P A\<^sub>P) + have B: "cP = P \ Q" and C: "cRs = \ \ P \ Q'" + by fact+ + from \bn \ \* (xvec1, xvec2, xvec3, xvec4, xvec5, xvec6, xvec7, xvec8, xvec9, x1, x2, x3, x4, cP, cRs, C)\ have "bn \ \* xvec3" by simp + from \A\<^sub>P \* (xvec1, xvec2, xvec3, xvec4, xvec5, xvec6, xvec7, xvec8, xvec9, x1, x2, x3, x4, cP, cRs, C)\ have "A\<^sub>P \* xvec3" and "A\<^sub>P \* C" by simp+ + + from \length xvec3 = residualLength cRs\ C have "length xvec3 = length(bn \)" + by simp + then obtain p where S: "set p \ set(bn \) \ set(bn(p \ \))" and "distinctPerm p" and "xvec3 = bn(p \ \)" + using \bn \ \* xvec3\ \distinct(bn \)\ \distinct xvec3\ + by - (rule constructPerm[where xvec="bn \" and yvec=xvec3], auto simp add: eqvts) + show ?thesis + proof(rule rPar2[where P=P and Q=Q and \="p \ \" and Q'="p \ Q'" and A\<^sub>P=A\<^sub>P and \\<^sub>P=\\<^sub>P]) + assume "xvec3 \* \" and "xvec3 \* cP" and "xvec3 \* cRs" + note \cP = P \ Q\ + moreover from B C S \bn \ \* xvec3\ \xvec3 \* cRs\ \xvec3 = bn(p \ \)\ \bn \ \* subject \\ \xvec3 \* cP\ \bn \ \* P\ + have "cRs = (p \ \) \ P \ (p \ Q')" + apply clarsimp + by(subst residualAlpha[where p=p]) auto + moreover note \xvec3 = bn(p \ \)\ + moreover from \\ \ \\<^sub>P \ Q \\ \ Q'\ S B C S \bn \ \* xvec3\ \xvec3 \* cRs\ \xvec3 = bn(p \ \)\ \bn \ \* subject \\ \xvec3 \* cP\ + have "\ \ \\<^sub>P \ Q \(p \ \) \ (p \ Q')" + by(subst residualAlpha[symmetric]) auto + moreover note \extractFrame P = \A\<^sub>P, \\<^sub>P\\ \distinct A\<^sub>P\ \A\<^sub>P \* P\ \A\<^sub>P \* Q\ \A\<^sub>P \* \\ \A\<^sub>P \* \\ + moreover from \A\<^sub>P \* \\ \A\<^sub>P \* xvec3\ S \xvec3 = bn(p \ \)\ \distinctPerm p\ have "A\<^sub>P \* (p \ \)" + by(subst fresh_star_bij[symmetric, where pi=p]) simp + moreover from \A\<^sub>P \* Q'\ \A\<^sub>P \* \\ \A\<^sub>P \* xvec3\ S \xvec3 = bn(p \ \)\ \distinctPerm p\ have "A\<^sub>P \* (p \ Q')" + by(subst fresh_star_bij[symmetric, where pi=p]) simp + moreover note \A\<^sub>P \* C\ + ultimately show "cP = P \ Q \ cRs = (p \ \) \ P \ (p \ Q') \ xvec3 = bn (p \ \) \ + \ \ \\<^sub>P \ Q \ (p \ \) \ p \ Q' \ + extractFrame P = \A\<^sub>P, \\<^sub>P\ \ distinct A\<^sub>P \ A\<^sub>P \* P \ A\<^sub>P \* Q \ A\<^sub>P \* \ \ A\<^sub>P \* (p \ \) \ + A\<^sub>P \* (p \ Q') \ A\<^sub>P \* C" by blast + qed +next + case(cComm1 \\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q K xvec Q' A\<^sub>Q) + then show ?thesis by - (rule rComm1[where P=P and Q=Q], (assumption | simp)+) +next + case(cComm2 \\<^sub>Q P M xvec N P' A\<^sub>P \\<^sub>P Q K Q' A\<^sub>Q) + then show ?thesis by - (rule rComm2[where P=P and Q=Q], (assumption | simp)+) +next + case(cBrMerge \\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q Q' A\<^sub>Q) + then show ?thesis by - (rule rBrMerge[where P=P and Q=Q], (assumption | simp)+) +next + case(cBrComm1 \\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q xvec Q' A\<^sub>Q) + have B: "cP = P \ Q" and C: "cRs = \M\\*xvec\\N\ \ P' \ Q'" + by fact+ + from \xvec \* (xvec1, xvec2, xvec3, xvec4, xvec5, xvec6, xvec7, xvec8, xvec9, x1, x2, x3, x4, cP, cRs, C)\ have "xvec \* xvec7" and "xvec \* C" by simp+ + from \A\<^sub>P \* (xvec1, xvec2, xvec3, xvec4, xvec5, xvec6, xvec7, xvec8, xvec9, x1, x2, x3, x4, cP, cRs, C)\ have "A\<^sub>P \* xvec7" and "A\<^sub>P \* C" by simp+ + from \A\<^sub>Q \* (xvec1, xvec2, xvec3, xvec4, xvec5, xvec6, xvec7, xvec8, xvec9, x1, x2, x3, x4, cP, cRs, C)\ have "A\<^sub>Q \* xvec7" and "A\<^sub>Q \* C" by simp+ + + from \length xvec7 = residualLength cRs\ C have "length xvec7 = length xvec" + by simp + then obtain p where S: "set p \ set xvec \ set (p \ xvec)" and "distinctPerm p" and "xvec7 = p \ xvec" + using \xvec \* xvec7\ \distinct xvec\ \distinct xvec7\ + by - (rule constructPerm[where xvec=xvec and yvec=xvec7], auto simp add: eqvts) + show ?thesis + proof(rule rBrComm1[where P=P and Q=Q and P'="p \ P'" + and Q'="p \ Q'" and N="p \ N" and A\<^sub>P=A\<^sub>P and \\<^sub>P=\\<^sub>P and A\<^sub>Q=A\<^sub>Q and \\<^sub>Q=\\<^sub>Q and M=M]) + assume "xvec7 \* \" and "xvec7 \* cP" and "xvec7 \* cRs" + from \A\<^sub>Q \* xvec7\ \xvec7 = p \ xvec\ + have "A\<^sub>Q \* (p \ xvec)" by simp + from \A\<^sub>P \* xvec7\ \xvec7 = p \ xvec\ + have "A\<^sub>P \* (p \ xvec)" by simp + from \xvec \* M\ have "(p \ xvec) \* (p \ M)" by(simp add: fresh_star_bij) + with S \xvec \* M\ \xvec7 \* cRs\ C \xvec7 = p \ xvec\ have "xvec7 \* M" by simp + + note \cP = P \ Q\ + moreover from C S \xvec \* xvec7\ \xvec7 \* cRs\ \xvec7 = p \ xvec\ \xvec \* M\ \xvec7 \* cP\ + have "cRs = \M\\*xvec7\\(p \ N)\ \ (p \ P') \ (p \ Q')" + apply clarsimp + by(subst residualAlpha[where p=p]) simp+ + + moreover note \xvec7 = p \ xvec\ + moreover from \\ \ \\<^sub>Q \ P \ \M\N\ \ P'\ S B \distinctPerm p\ \xvec \* xvec7\ \xvec7 = p \ xvec\ \xvec \* P\ \xvec7 \* cP\ + have "\ \ \\<^sub>Q \ P \ \M\(p \ N)\ \ p \ P'" + by(simp add: brinputAlpha) + + moreover from \\ \ \\<^sub>P \ Q \ \M\\*xvec\\N\ \ Q'\ S B C \xvec \* xvec7\ \xvec7 \* cRs\ \xvec7 = p \ xvec\ \xvec \* M\ \xvec7 \* cP\ \xvec7 \* M\ + have "\ \ \\<^sub>P \ Q \ \M\\*xvec7\\(p \ N)\ \ p \ Q'" + by(auto simp add: residualAlpha) + + moreover note + \extractFrame P = \A\<^sub>P, \\<^sub>P\\ \distinct A\<^sub>P\ \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ \distinct A\<^sub>Q\ + \A\<^sub>P \* \\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>P \* P\ \A\<^sub>P \* Q\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* xvec7\ + \A\<^sub>Q \* \\ \A\<^sub>Q \* \\<^sub>P\ \A\<^sub>Q \* P\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* xvec7\ + + moreover from S \A\<^sub>P \* xvec\ \A\<^sub>P \* (p \ xvec)\ \A\<^sub>P \* N\ + have "A\<^sub>P \* (p \ N)" + by(simp add: freshChainSimps) + moreover from S \A\<^sub>P \* xvec\ \A\<^sub>P \* (p \ xvec)\ \A\<^sub>P \* P'\ + have "A\<^sub>P \* (p \ P')" + by(simp add: freshChainSimps) + moreover from S \A\<^sub>P \* xvec\ \A\<^sub>P \* (p \ xvec)\ \A\<^sub>P \* Q'\ + have "A\<^sub>P \* (p \ Q')" + by(simp add: freshChainSimps) + moreover from S \A\<^sub>Q \* xvec\ \A\<^sub>Q \* (p \ xvec)\ \A\<^sub>Q \* N\ + have "A\<^sub>Q \* (p \ N)" + by(simp add: freshChainSimps) + moreover from S \A\<^sub>Q \* xvec\ \A\<^sub>Q \* (p \ xvec)\ \A\<^sub>Q \* P'\ + have "A\<^sub>Q \* (p \ P')" + by(simp add: freshChainSimps) + moreover from S \A\<^sub>Q \* xvec\ \A\<^sub>Q \* (p \ xvec)\ \A\<^sub>Q \* Q'\ + have "A\<^sub>Q \* (p \ Q')" + by(simp add: freshChainSimps) + + moreover note \xvec7 \* \\ + + moreover from \xvec7 \* cP\ \cP = P \ Q\ \extractFrame P = \A\<^sub>P, \\<^sub>P\\ \A\<^sub>P \* xvec7\ + have "xvec7 \* \\<^sub>P" + by simp (metis extractFrameFreshChain freshFrameDest) + + moreover from \xvec7 \* cP\ \cP = P \ Q\ \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ \A\<^sub>Q \* xvec7\ + have "xvec7 \* \\<^sub>Q" + by simp (metis extractFrameFreshChain freshFrameDest) + + moreover from \xvec7 \* cP\ \cP = P \ Q\ have "xvec7 \* P" by simp + moreover from \xvec7 \* cP\ \cP = P \ Q\ have "xvec7 \* Q" by simp + + moreover note \A\<^sub>P \* M\ \A\<^sub>Q \* M\ + + moreover note \xvec7 \* M\ + + moreover note \A\<^sub>P \* C\ \A\<^sub>Q \* C\ \distinct xvec7\ + + ultimately show "cP = P \ Q \ cRs = \M\\*xvec7\\(p \ N)\ \ (p \ P') \ (p \ Q') \ + \ \ \\<^sub>Q \ P \ \M\(p \ N)\ \ p \ P' \ extractFrame P = \A\<^sub>P, \\<^sub>P\ \ distinct A\<^sub>P \ + \ \ \\<^sub>P \ Q \ \M\\*xvec7\\(p \ N)\ \ p \ Q' \ extractFrame Q = \A\<^sub>Q, \\<^sub>Q\ \ distinct A\<^sub>Q \ + A\<^sub>P \* \ \ A\<^sub>P \* \\<^sub>Q \ A\<^sub>P \* P \ A\<^sub>P \* (p \ N) \ A\<^sub>P \* (p \ P') \ A\<^sub>P \* Q \ A\<^sub>P \* (p \ Q') \ + A\<^sub>P \* A\<^sub>Q \ A\<^sub>P \* xvec7 \ A\<^sub>Q \* \ \ A\<^sub>Q \* \\<^sub>P \ A\<^sub>Q \* P \ A\<^sub>Q \* (p \ N) \ A\<^sub>Q \* (p \ P') \ + A\<^sub>Q \* Q \ A\<^sub>Q \* (p \ Q') \ A\<^sub>Q \* xvec7 \ xvec7 \* \ \ xvec7 \* \\<^sub>P \ xvec7 \* \\<^sub>Q \ xvec7 \* P \ + xvec7 \* Q \ A\<^sub>P \* M \ A\<^sub>Q \* M \ xvec7 \* M \ A\<^sub>P \* C \ A\<^sub>Q \* C \ distinct xvec7" by blast + qed +next + case(cBrComm2 \\<^sub>Q P M xvec N P' A\<^sub>P \\<^sub>P Q Q' A\<^sub>Q) + have B: "cP = P \ Q" and C: "cRs = \M\\*xvec\\N\ \ P' \ Q'" + by fact+ + from \xvec \* (xvec1, xvec2, xvec3, xvec4, xvec5, xvec6, xvec7, xvec8, xvec9, x1, x2, x3, x4, cP, cRs, C)\ have "xvec \* xvec8" and "xvec \* C" by simp+ + from \A\<^sub>P \* (xvec1, xvec2, xvec3, xvec4, xvec5, xvec6, xvec7, xvec8, xvec9, x1, x2, x3, x4, cP, cRs, C)\ have "A\<^sub>P \* xvec8" and "A\<^sub>P \* C" by simp+ + from \A\<^sub>Q \* (xvec1, xvec2, xvec3, xvec4, xvec5, xvec6, xvec7, xvec8, xvec9, x1, x2, x3, x4, cP, cRs, C)\ have "A\<^sub>Q \* xvec8" and "A\<^sub>Q \* C" by simp+ + + from \length xvec8 = residualLength cRs\ C have "length xvec8 = length xvec" + by simp + then obtain p where S: "set p \ set xvec \ set (p \ xvec)" and "distinctPerm p" and "xvec8 = p \ xvec" + using \xvec \* xvec8\ \distinct xvec\ \distinct xvec8\ + by - (rule constructPerm[where xvec=xvec and yvec=xvec8], auto simp add: eqvts) + show ?thesis + proof(rule rBrComm2[where P=P and Q=Q and P'="p \ P'" + and Q'="p \ Q'" and N="p \ N" and A\<^sub>P=A\<^sub>P and \\<^sub>P=\\<^sub>P and A\<^sub>Q=A\<^sub>Q and \\<^sub>Q=\\<^sub>Q and M=M]) + assume "xvec8 \* \" and "xvec8 \* cP" and "xvec8 \* cRs" + from \A\<^sub>Q \* xvec8\ \xvec8 = p \ xvec\ + have "A\<^sub>Q \* (p \ xvec)" by simp + from \A\<^sub>P \* xvec8\ \xvec8 = p \ xvec\ + have "A\<^sub>P \* (p \ xvec)" by simp + from \xvec \* M\ have "(p \ xvec) \* (p \ M)" by(simp add: fresh_star_bij) + with S \xvec \* M\ \xvec8 \* cRs\ C \xvec8 = p \ xvec\ have "xvec8 \* M" by simp + + note \cP = P \ Q\ + moreover from C S \xvec \* xvec8\ \xvec8 \* cRs\ \xvec8 = p \ xvec\ \xvec \* M\ \xvec8 \* cP\ + have "cRs = \M\\*xvec8\\(p \ N)\ \ (p \ P') \ (p \ Q')" + apply clarsimp + by(subst residualAlpha[where p=p]) simp+ + + moreover note \xvec8 = p \ xvec\ + moreover from \\ \ \\<^sub>P \ Q \ \M\N\ \ Q'\ S B \distinctPerm p\ \xvec \* xvec8\ \xvec8 = p \ xvec\ \xvec \* Q\ \xvec8 \* cP\ + have "\ \ \\<^sub>P \ Q \ \M\(p \ N)\ \ p \ Q'" + by(simp add: brinputAlpha) + + moreover from \\ \ \\<^sub>Q \ P \ \M\\*xvec\\N\ \ P'\ S B C \xvec \* xvec8\ \xvec8 \* cRs\ \xvec8 = p \ xvec\ \xvec \* M\ \xvec8 \* cP\ \xvec8 \* M\ + have "\ \ \\<^sub>Q \ P \ \M\\*xvec8\\(p \ N)\ \ p \ P'" + by(auto simp add: residualAlpha) + + moreover note + \extractFrame P = \A\<^sub>P, \\<^sub>P\\ \distinct A\<^sub>P\ \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ \distinct A\<^sub>Q\ + \A\<^sub>P \* \\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>P \* P\ \A\<^sub>P \* Q\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* xvec8\ + \A\<^sub>Q \* \\ \A\<^sub>Q \* \\<^sub>P\ \A\<^sub>Q \* P\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* xvec8\ + + moreover from S \A\<^sub>P \* xvec\ \A\<^sub>P \* (p \ xvec)\ \A\<^sub>P \* N\ + have "A\<^sub>P \* (p \ N)" + by(simp add: freshChainSimps) + moreover from S \A\<^sub>P \* xvec\ \A\<^sub>P \* (p \ xvec)\ \A\<^sub>P \* P'\ + have "A\<^sub>P \* (p \ P')" + by(simp add: freshChainSimps) + moreover from S \A\<^sub>P \* xvec\ \A\<^sub>P \* (p \ xvec)\ \A\<^sub>P \* Q'\ + have "A\<^sub>P \* (p \ Q')" + by(simp add: freshChainSimps) + moreover from S \A\<^sub>Q \* xvec\ \A\<^sub>Q \* (p \ xvec)\ \A\<^sub>Q \* N\ + have "A\<^sub>Q \* (p \ N)" + by(simp add: freshChainSimps) + moreover from S \A\<^sub>Q \* xvec\ \A\<^sub>Q \* (p \ xvec)\ \A\<^sub>Q \* P'\ + have "A\<^sub>Q \* (p \ P')" + by(simp add: freshChainSimps) + moreover from S \A\<^sub>Q \* xvec\ \A\<^sub>Q \* (p \ xvec)\ \A\<^sub>Q \* Q'\ + have "A\<^sub>Q \* (p \ Q')" + by(simp add: freshChainSimps) + + moreover note \xvec8 \* \\ + + moreover from \xvec8 \* cP\ \cP = P \ Q\ \extractFrame P = \A\<^sub>P, \\<^sub>P\\ \A\<^sub>P \* xvec8\ + have "xvec8 \* \\<^sub>P" + by simp (metis extractFrameFreshChain freshFrameDest) + + moreover from \xvec8 \* cP\ \cP = P \ Q\ \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ \A\<^sub>Q \* xvec8\ + have "xvec8 \* \\<^sub>Q" + by simp (metis extractFrameFreshChain freshFrameDest) + + moreover from \xvec8 \* cP\ \cP = P \ Q\ have "xvec8 \* P" by simp + moreover from \xvec8 \* cP\ \cP = P \ Q\ have "xvec8 \* Q" by simp + + moreover note \A\<^sub>P \* M\ \A\<^sub>Q \* M\ + + moreover note \xvec8 \* M\ + + moreover note \A\<^sub>P \* C\ \A\<^sub>Q \* C\ \distinct xvec8\ + + ultimately show "cP = P \ Q \ cRs = \M\\*xvec8\\(p \ N)\ \ (p \ P') \ (p \ Q') \ + \ \ \\<^sub>Q \ P \ \M\\*xvec8\\(p \ N)\ \ p \ P' \ extractFrame P = \A\<^sub>P, \\<^sub>P\ \ distinct A\<^sub>P \ + \ \ \\<^sub>P \ Q \ \M\(p \ N)\ \ p \ Q' \ extractFrame Q = \A\<^sub>Q, \\<^sub>Q\ \ distinct A\<^sub>Q \ + A\<^sub>P \* \ \ A\<^sub>P \* \\<^sub>Q \ A\<^sub>P \* P \ A\<^sub>P \* (p \ N) \ A\<^sub>P \* (p \ P') \ A\<^sub>P \* Q \ A\<^sub>P \* (p \ Q') \ + A\<^sub>P \* A\<^sub>Q \ A\<^sub>P \* xvec8 \ A\<^sub>Q \* \ \ A\<^sub>Q \* \\<^sub>P \ A\<^sub>Q \* P \ A\<^sub>Q \* (p \ N) \ A\<^sub>Q \* (p \ P') \ + A\<^sub>Q \* Q \ A\<^sub>Q \* (p \ Q') \ A\<^sub>Q \* xvec8 \ xvec8 \* \ \ xvec8 \* \\<^sub>P \ xvec8 \* \\<^sub>Q \ xvec8 \* P \ + xvec8 \* Q \ A\<^sub>P \* M \ A\<^sub>Q \* M \ xvec8 \* M \ A\<^sub>P \* C \ A\<^sub>Q \* C \ distinct xvec8" by blast + qed +next + case(cBrClose P M xvec N P' x) + have B: "cP = \\x\P" and C: "cRs = \ \ \\x\(\\*xvec\P')" + by fact+ + from \x \ (xvec1, xvec2, xvec3, xvec4, xvec5, xvec6, xvec7, xvec8, xvec9, x1, x2, x3, x4, cP, cRs, C)\ have "x \ cP" and "x \ cRs" and "x \ x3" by simp+ + from \xvec \* (xvec1, xvec2, xvec3, xvec4, xvec5, xvec6, xvec7, xvec8, xvec9, x1, x2, x3, x4, cP, cRs, C)\ have "xvec \* cP" and "xvec \* cRs" and "x3 \ xvec" and "xvec \* x3" by simp+ + + obtain r::"name prm" where "(r \ xvec) \* \" and "(r \ xvec) \* P" and "(r \ xvec) \* M" + and "(r \ xvec) \* N" and "(r \ xvec) \* P'" and "(r \ xvec) \* x" + and "(r \ xvec) \* C" and "(r \ xvec) \* x3" + and "(r \ xvec) \* ([(x, x3)] \ P)" and "(r \ xvec) \* ([(x, x3)] \ M)" + and Sr: "(set r) \ (set xvec) \ (set(r \ xvec))" and "distinctPerm r" + by(rule name_list_avoiding[where xvec=xvec and c="(\, P, M, N, P', x, x3, ([(x, x3)] \ P), ([(x, x3)] \ P), C)"]) + (auto simp add: eqvts) + from \(r \ xvec) \* x3\ have "x3 \ (r \ xvec)" by simp + + show ?thesis + proof(rule rBrClose[where P="[(x, x3)] \ P" and M="[(x, x3)] \ M" and N="[(x, x3)] \ (r \ N)" and xvec="(r \ xvec)" and P'="[(x, x3)] \ (r \ P')"]) + assume "x3 \ \" and "x3 \ cP" and "x3 \ cRs" + with \x3 \ xvec\ have "x3 \ set xvec" by simp + + from \x3 \ cRs\ C have "x3 \ (\ \ \\x\(\\*xvec\P'))" by simp + + then have "x3 \ (\\x\(\\*xvec\P'))" by simp + with \x \ x3\ have "x3 \ (\\*xvec\P')" by(simp add: psi.fresh abs_fresh) + then have "(x3 \ set xvec) \ (x3 \ P')" by(simp add: resChainFresh) + + with \x3 \ set xvec\ have "x3 \ P'" by blast + + with \x3 \ xvec\ \x3 \ (r \ xvec)\ Sr have "x3 \ (r \ P')" + by(simp add: freshChainSimps) + + from \cP = \\x\P\ \x3 \ cP\ \x \ x3\ have cP_perm: "cP = \\x3\([(x, x3)] \ P)" + by(simp add: alphaRes abs_fresh) + from \(r \ xvec) \* P'\ Sr C + have "cRs = \ \ \\x\(\\*(r \ xvec)\(r \ P'))" + by(simp add: resChainAlpha) + moreover from \x3 \ (r \ P')\ + have "x3 \ \\*(r \ xvec)\(r \ P')" + by(simp add: resChainFresh) + ultimately have "cRs = \ \ \\x3\([(x, x3)] \ \\*(r \ xvec)\(r \ P'))" by(simp add: alphaRes) + with \(r \ xvec) \* x\ \(r \ xvec) \* x3\ + have cRs_perm: "cRs = \ \ \\x3\(\\*(r \ xvec)\[(x, x3)] \ (r \ P'))" + by(simp add: eqvts) + + from \x \ supp M\ + have supp_inc_perm: "x3 \ supp ([(x, x3)] \ M)" + by (metis fresh_bij fresh_def swap_simps) + + from \x \ xvec\ \(r \ xvec) \* x\ Sr have "r \ x = x" by simp + + from \\ \ P \ \M\\*xvec\\N\ \ P'\ \(r \ xvec) \* N\ \(r \ xvec) \* P'\ Sr + have "\ \ P \ \M\\*(r \ xvec)\\(r \ N)\ \ (r \ P')" + by(simp add: boundOutputChainAlpha'' create_residual.simps) + then have "[(x, x3)] \ (\ \ P \ \M\\*(r \ xvec)\\(r \ N)\ \ (r \ P'))" + by(simp add: perm_bool) + with \x \ \\ \x3 \ \\ \(r \ xvec) \* x\ \x3 \ (r \ xvec)\ + have trans_perm: "\ \ ([(x, x3)] \ P) \ \([(x, x3)] \ M)\\*(r \ xvec)\\([(x, x3)] \ (r \ N))\ \ ([(x, x3)] \ (r \ P'))" + by(auto simp add: eqvts) + + from \distinctPerm r\ \distinct xvec\ + have distinct_perm: "distinct (r \ xvec)" by simp + + note cP_perm cRs_perm supp_inc_perm trans_perm distinct_perm + \(r \ xvec) \* \\ \(r \ xvec) \* ([(x, x3)] \ P)\ \(r \ xvec) \* ([(x, x3)] \ M)\ + \(r \ xvec) \* C\ \x3 \ \\ \x3 \ (r \ xvec)\ + + then show "cP = \\x3\([(x, x3)] \ P) \ cRs = \ \ \\x3\(\\*r \ xvec\[(x, x3)] \ r \ P') \ + x3 \ supp ([(x, x3)] \ M) \ + \ \ [(x, x3)] \ P \ \([(x, x3)] \ M)\\*(r \ xvec)\\([(x, x3)] \ r \ N)\ \ [(x, x3)] \ r \ P' \ + distinct (r \ xvec) \ (r \ xvec) \* \ \ (r \ xvec) \* ([(x, x3)] \ P) \ + (r \ xvec) \* ([(x, x3)] \ M) \ (r \ xvec) \* C \ x3 \ \ \ x3 \ r \ xvec" + by simp + qed +next + case(cOpen P M xvec yvec N P' x) + have B: "cP = \\x\P" and C: "cRs = M\\*(xvec@x#yvec)\\N\ \ P'" + by fact+ + from \xvec \* (xvec1, xvec2, xvec3, xvec4, xvec5, xvec6, xvec7, xvec8, xvec9, x1, x2, x3, x4, cP, cRs, C)\ have "xvec \* xvec4" and "xvec \* cP" and "xvec \* cRs" and "x1 \ xvec" by simp+ + from \x \ (xvec1, xvec2, xvec3, xvec4, xvec5, xvec6, xvec7, xvec8, xvec9, x1, x2, x3, x4, cP, cRs, C)\ have "x \ xvec4" and "x \ cP" and "x \ cRs" and "x \ x1" by simp+ + from \yvec \* (xvec1, xvec2, xvec3, xvec4, xvec5, xvec6, xvec7, xvec8, xvec9, x1, x2, x3, x4, cP, cRs, C)\ have "yvec \* xvec4" and "yvec \* cP" and "yvec \* cRs" and "x1 \ yvec" by simp+ + + from \xvec \* cRs\ \x \ cRs\ \yvec \* cRs\ C have "(xvec@x#yvec) \* M" by simp + from \xvec \* \\ \x \ \\ \yvec \* \\ have "(xvec@x#yvec) \* \" by simp + from \length xvec4 = residualLength cRs\ C obtain xvec' y yvec' where D: "xvec4 = xvec'@y#yvec'" and "length xvec' = length xvec" and "length yvec' = length yvec" + by(auto intro: lengthAux2) + with \distinct xvec\ \distinct yvec\ \x \ xvec\ \x \ yvec\ \xvec \* yvec\ \xvec \* xvec4\ \yvec \* xvec4\ \x \ xvec4\ \distinct xvec4\ + have "distinct xvec'" and "distinct yvec'" and "xvec' \* yvec'" and "x \ y" and "y \ xvec'" and "y \ yvec'" + and "x \ xvec'" and "x \ yvec'" and "y \ xvec" and "y \ yvec" and "xvec \* xvec'" and "yvec \* yvec'" + by auto + from \length xvec' = length xvec\ \xvec \* xvec'\ \distinct xvec\ \distinct xvec'\ + obtain p where Sp: "set p \ set xvec \ set(p \ xvec)" and "distinctPerm p" and E: "xvec' = p \ xvec" + by(metis constructPerm) + from \length yvec' = length yvec\ \yvec \* yvec'\ \distinct yvec\ \distinct yvec'\ + obtain q where Sq: "set q \ set yvec \ set(q \ yvec)" and "distinctPerm q" and F: "yvec' = q \ yvec" + by(metis constructPerm) + + show ?thesis + proof(rule rOpen[where P="([(x, x1)] \ P)" and xvec="p \ xvec" and y="y" and yvec="q \ yvec" and N="(p@(x1, x)#q) \ N" and P'="(p@(x1, x)#q) \ P'" and M=M]) + assume "xvec4 \* \" and "xvec4 \* cP" and "xvec4 \* cRs" and "x1 \ \" and "x1 \ cP" and "x1 \ cRs" and "x1 \ xvec4" + from \xvec \* xvec4\ \x \ xvec4\ \x1 \ xvec4\ \yvec \* xvec4\ D E F + have "x \ y" and "x1 \ y" and "x1 \ p \ xvec" and "x1 \ q \ yvec" by simp+ + from \xvec4 \* cRs\ \x1 \ cRs\ C have "xvec4 \* M" and "x1 \ M" by simp+ + moreover from \cP = \\x\P\ \x \ cP\ \x \ x1\ have "([(x, x1)] \ cP) = [(x, x1)] \ \\x\P" + by simp + with \x \ cP\ \x1 \ cP\ have "cP = \\x1\([(x, x1)] \ P)" by(simp add: eqvts calc_atm) + moreover from C have "((p@(x1, x)#q) \ cRs) = (p@(x1, x)#q) \ (M\\*(xvec@x#yvec)\\N\ \ P')" by(simp add: fresh_star_bij) + with Sp Sq \xvec4 \* cRs\ D E F \xvec \* cRs\ \x \ cRs\ \yvec \* cRs\ \xvec4 \* M\ \(xvec@x#yvec) \* M\ \xvec \* xvec4\ \x \ xvec4\ \yvec \* xvec4\ \xvec \* yvec\ \x \ xvec\ \x \ yvec\ \y \ xvec'\ \y \ yvec'\ \xvec' \* yvec'\ \x1 \ xvec\ \x1 \ yvec\ \x1 \ y\ \x1 \ xvec4\ \x1 \ cRs\ \x1 \ cRs\ \x \ x1\ \x1 \ M\ + have "cRs = M\\*((p \ xvec)@x1#(q \ yvec))\\((p@(x1, x)#q) \ N)\ \ ((p@(x1, x)#q) \ P')" + by(simp add: eqvts pt2[OF pt_name_inst] calc_atm) + moreover from D E F have "xvec4 = (p \ xvec)@y#(q \ yvec)" by simp + moreover from \\ \ P \M\\*(xvec@yvec)\\N\ \ P'\ have "((p@(x1, x)#q) \ \) \ ((p@(x1, x)#q) \ P) \((p@(x1, x)#q) \ (M\\*(xvec@yvec)\\N\ \ P'))" + by(intro eqvts) + with Sp Sq B C D E F \xvec4 \* \\ \(xvec@x#yvec) \* \\ \xvec4 \* cRs\ \x \ xvec4\ C D \x \ cRs\ \yvec \* cRs\ \xvec4 \* M\ \(xvec@x#yvec) \* M\ \x \ M\ \x1 \ cRs\ \x \ x1\ \x1 \ xvec\ \x1 \ yvec\ \xvec \* xvec4\ \yvec \* xvec4\ \x1 \ xvec4\ \x \ xvec\ \x \ yvec\ \x1 \ \\ \xvec4 \* cP\ \xvec \* P\ \yvec \* P\ \xvec' \* yvec'\ \x1 \ xvec4\ \xvec4 \* cP\ \yvec \* xvec4\ \xvec \* xvec4\ \x \ x1\ \xvec \* yvec\ + have "\ \ ([(x, x1)] \ P) \M\\*((p \ xvec)@(q \ yvec))\\((p@(x1, x)#q) \ N)\ \ ((p@(x1, x)#q) \ P')" + by(simp add: eqvts pt_fresh_bij[OF pt_name_inst, OF at_name_inst] pt2[OF pt_name_inst] name_swap) + + moreover from \x \ supp N\ have "((p@(x1, x)#q) \ x) \ ((p@(x1, x)#q) \ supp N)" + by(simp add: pt_set_bij[OF pt_name_inst, OF at_name_inst]) + then have "x1 \ supp((p@(x1, x)#q) \ N)" + using \x \ xvec\ \x \ yvec\ \x1 \ xvec\ \x1 \ yvec\ \x \ xvec4\ \x1 \ xvec4\ \xvec \* xvec4\ \yvec \* xvec4\ \xvec' \* yvec'\ D E F Sp Sq \x \ x1\ + by(simp add: eqvts pt2[OF pt_name_inst] calc_atm) + moreover from \x1 \ xvec4\ D E F have "x1 \ (p \ xvec)" and "x1 \ (q \ yvec)" by simp+ + moreover from \distinct xvec'\ \distinct yvec'\ E F have "distinct(p \ xvec)" and "distinct(q \ yvec)" by simp+ + moreover from \xvec' \* yvec'\ E F have "(p \ xvec) \* (q \ yvec)" by auto + moreover from \xvec \* \\ have "(p \ xvec) \* (p \ \)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with Sp D E \xvec4 \* \\ \xvec \* \\ have "(p \ xvec) \* \" by(simp add: eqvts) + moreover from \yvec \* \\ have "(p \ yvec) \* (p \ \)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with Sq D F \xvec4 \* \\ \yvec \* \\ have "(q \ yvec) \* \" by(simp add: eqvts) + moreover from \xvec4 \* cP\ \x \ xvec4\ \x1 \ xvec4\ B D E F have "(p \ xvec) \* ([(x, x1)] \ P)" and "(q \ yvec) \* ([(x, x1)] \ P)" + by simp+ + moreover from \xvec4 \* M\ C D E F have "(p \ xvec) \* M" and "(q \ yvec) \* M" by simp+ + ultimately show "cP = \\x1\([(x, x1)] \ P) \ + cRs = M\\*(p \ xvec @ x1 # q \ yvec)\\((p @ (x1, x) # q) \ N)\ \ (p @ (x1, x) # q) \ P' \ + xvec4 = p \ xvec @ y # q \ yvec \ + \ \ [(x, x1)] \ P \ M\\*(p \ xvec @ q \ yvec)\\((p @ (x1, x) # q) \ N)\ \ (p @ (x1, x) # q) \ P' \ + x1 \ supp ((p @ (x1, x) # q) \ N) \ + x1 \ p \ xvec \ + x1 \ q \ yvec \ + distinct (p \ xvec) \ + distinct (q \ yvec) \ + (p \ xvec) \* \ \ + (p \ xvec) \* ([(x, x1)] \ P) \ + (p \ xvec) \* M \ + (p \ xvec) \* (q \ yvec) \ + (q \ yvec) \* \ \ + (q \ yvec) \* ([(x, x1)] \ P) \ + (q \ yvec) \* M" + by blast + qed +next + case(cBrOpen P M xvec yvec N P' x) + have B: "cP = \\x\P" and C: "cRs = \M\\*(xvec@x#yvec)\\N\ \ P'" + by fact+ + from \xvec \* (xvec1, xvec2, xvec3, xvec4, xvec5, xvec6, xvec7, xvec8, xvec9, x1, x2, x3, x4, cP, cRs, C)\ have "xvec \* xvec9" and "xvec \* cP" and "xvec \* cRs" and "x4 \ xvec" by simp+ + from \x \ (xvec1, xvec2, xvec3, xvec4, xvec5, xvec6, xvec7, xvec8, xvec9, x1, x2, x3, x4, cP, cRs, C)\ have "x \ xvec9" and "x \ cP" and "x \ cRs" and "x \ x4" by simp+ + from \yvec \* (xvec1, xvec2, xvec3, xvec4, xvec5, xvec6, xvec7, xvec8, xvec9, x1, x2, x3, x4, cP, cRs, C)\ have "yvec \* xvec9" and "yvec \* cP" and "yvec \* cRs" and "x4 \ yvec" by simp+ + + from \xvec \* cRs\ \x \ cRs\ \yvec \* cRs\ C have "(xvec@x#yvec) \* M" by simp + from \xvec \* \\ \x \ \\ \yvec \* \\ have "(xvec@x#yvec) \* \" by simp + from \length xvec9 = residualLength cRs\ C obtain xvec' y yvec' where D: "xvec9 = xvec'@y#yvec'" and "length xvec' = length xvec" and "length yvec' = length yvec" + by(auto intro: lengthAux2) + with \distinct xvec\ \distinct yvec\ \x \ xvec\ \x \ yvec\ \xvec \* yvec\ \xvec \* xvec9\ \yvec \* xvec9\ \x \ xvec9\ \distinct xvec9\ + have "distinct xvec'" and "distinct yvec'" and "xvec' \* yvec'" and "x \ y" and "y \ xvec'" and "y \ yvec'" + and "x \ xvec'" and "x \ yvec'" and "y \ xvec" and "y \ yvec" and "xvec \* xvec'" and "yvec \* yvec'" + by auto + from \length xvec' = length xvec\ \xvec \* xvec'\ \distinct xvec\ \distinct xvec'\ + obtain p where Sp: "set p \ set xvec \ set(p \ xvec)" and "distinctPerm p" and E: "xvec' = p \ xvec" + by(metis constructPerm) + from \length yvec' = length yvec\ \yvec \* yvec'\ \distinct yvec\ \distinct yvec'\ + obtain q where Sq: "set q \ set yvec \ set(q \ yvec)" and "distinctPerm q" and F: "yvec' = q \ yvec" + by(metis constructPerm) + + show ?thesis + proof(rule rBrOpen[where P="([(x, x4)] \ P)" and xvec="p \ xvec" and y="y" and yvec="q \ yvec" and N="(p@(x4, x)#q) \ N" and P'="(p@(x4, x)#q) \ P'" and M=M]) + assume "xvec9 \* \" and "xvec9 \* cP" and "xvec9 \* cRs" and "x4 \ \" and "x4 \ cP" and "x4 \ cRs" and "x4 \ xvec9" + from \xvec \* xvec9\ \x \ xvec9\ \x4 \ xvec9\ \yvec \* xvec9\ D E F + have "x \ y" and "x4 \ y" and "x4 \ p \ xvec" and "x4 \ q \ yvec" by simp+ + from \xvec9 \* cRs\ \x4 \ cRs\ C have "xvec9 \* M" and "x4 \ M" by simp+ + moreover from \cP = \\x\P\ \x \ cP\ \x \ x4\ have "([(x, x4)] \ cP) = [(x, x4)] \ \\x\P" + by simp + with \x \ cP\ \x4 \ cP\ have "cP = \\x4\([(x, x4)] \ P)" by(simp add: eqvts calc_atm) + moreover from C have "((p@(x4, x)#q) \ cRs) = (p@(x4, x)#q) \ (\M\\*(xvec@x#yvec)\\N\ \ P')" by(simp add: fresh_star_bij) + with Sp Sq \xvec9 \* cRs\ D E F \xvec \* cRs\ \x \ cRs\ \yvec \* cRs\ \xvec9 \* M\ \(xvec@x#yvec) \* M\ \xvec \* xvec9\ \x \ xvec9\ \yvec \* xvec9\ \xvec \* yvec\ \x \ xvec\ \x \ yvec\ \y \ xvec'\ \y \ yvec'\ \xvec' \* yvec'\ \x4 \ xvec\ \x4 \ yvec\ \x4 \ y\ \x4 \ xvec9\ \x4 \ cRs\ \x4 \ cRs\ \x \ x4\ \x4 \ M\ + have "cRs = \M\\*((p \ xvec)@x4#(q \ yvec))\\((p@(x4, x)#q) \ N)\ \ ((p@(x4, x)#q) \ P')" + by(simp add: eqvts pt2[OF pt_name_inst] calc_atm) + moreover from D E F have "xvec9 = (p \ xvec)@y#(q \ yvec)" by simp + moreover from \\ \ P \\M\\*(xvec@yvec)\\N\ \ P'\ have "((p@(x4, x)#q) \ \) \ ((p@(x4, x)#q) \ P) \((p@(x4, x)#q) \ (\M\\*(xvec@yvec)\\N\ \ P'))" + by(intro eqvts) + with Sp Sq B C D E F \xvec9 \* \\ \(xvec@x#yvec) \* \\ \xvec9 \* cRs\ \x \ xvec9\ C D \x \ cRs\ \yvec \* cRs\ \xvec9 \* M\ \(xvec@x#yvec) \* M\ \x \ M\ \x4 \ cRs\ \x \ x4\ \x4 \ xvec\ \x4 \ yvec\ \xvec \* xvec9\ \yvec \* xvec9\ \x4 \ xvec9\ \x \ xvec\ \x \ yvec\ \x4 \ \\ \xvec9 \* cP\ \xvec \* P\ \yvec \* P\ \xvec' \* yvec'\ \x4 \ xvec9\ \xvec9 \* cP\ \yvec \* xvec9\ \xvec \* xvec9\ \x \ x4\ \xvec \* yvec\ + have "\ \ ([(x, x4)] \ P) \\M\\*((p \ xvec)@(q \ yvec))\\((p@(x4, x)#q) \ N)\ \ ((p@(x4, x)#q) \ P')" + by(simp add: eqvts pt_fresh_bij[OF pt_name_inst, OF at_name_inst] pt2[OF pt_name_inst] name_swap) + + moreover from \x \ supp N\ have "((p@(x4, x)#q) \ x) \ ((p@(x4, x)#q) \ supp N)" + by(simp add: pt_set_bij[OF pt_name_inst, OF at_name_inst]) + then have "x4 \ supp((p@(x4, x)#q) \ N)" + using \x \ xvec\ \x \ yvec\ \x4 \ xvec\ \x4 \ yvec\ \x \ xvec9\ \x4 \ xvec9\ \xvec \* xvec9\ \yvec \* xvec9\ \xvec' \* yvec'\ D E F Sp Sq \x \ x4\ + by(simp add: eqvts pt2[OF pt_name_inst] calc_atm) + moreover from \x4 \ xvec9\ D E F have "x4 \ (p \ xvec)" and "x4 \ (q \ yvec)" by simp+ + moreover from \distinct xvec'\ \distinct yvec'\ E F have "distinct(p \ xvec)" and "distinct(q \ yvec)" by simp+ + moreover from \xvec' \* yvec'\ E F have "(p \ xvec) \* (q \ yvec)" by auto + moreover from \xvec \* \\ have "(p \ xvec) \* (p \ \)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with Sp D E \xvec9 \* \\ \xvec \* \\ have "(p \ xvec) \* \" by(simp add: eqvts) + moreover from \yvec \* \\ have "(p \ yvec) \* (p \ \)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with Sq D F \xvec9 \* \\ \yvec \* \\ have "(q \ yvec) \* \" by(simp add: eqvts) + moreover from \xvec9 \* cP\ \x \ xvec9\ \x4 \ xvec9\ B D E F have "(p \ xvec) \* ([(x, x4)] \ P)" and "(q \ yvec) \* ([(x, x4)] \ P)" + by simp+ + moreover from \xvec9 \* M\ C D E F have "(p \ xvec) \* M" and "(q \ yvec) \* M" by simp+ + ultimately show "cP = \\x4\([(x, x4)] \ P) \ + cRs = \M\\*(p \ xvec @ x4 # q \ yvec)\\((p @ (x4, x) # q) \ N)\ \ (p @ (x4, x) # q) \ P' \ + xvec9 = p \ xvec @ y # q \ yvec \ + \ \ [(x, x4)] \ P \ \M\\*(p \ xvec @ q \ yvec)\\((p @ (x4, x) # q) \ N)\ \ (p @ (x4, x) # q) \ P' \ + x4 \ supp ((p @ (x4, x) # q) \ N) \ + x4 \ p \ xvec \ + x4 \ q \ yvec \ + distinct (p \ xvec) \ + distinct (q \ yvec) \ + (p \ xvec) \* \ \ + (p \ xvec) \* ([(x, x4)] \ P) \ + (p \ xvec) \* M \ + (p \ xvec) \* (q \ yvec) \ + (q \ yvec) \* \ \ + (q \ yvec) \* ([(x, x4)] \ P) \ + (q \ yvec) \* M" + by blast + qed +next + case(cScope P \ P' x) + have B: "cP = \\x\P" and C: "cRs = \ \ \\x\P'" + by fact+ + from \bn \ \* (xvec1, xvec2, xvec3, xvec4, xvec5, xvec6, xvec7, xvec8, xvec9, x1, x2, x3, x4, cP, cRs, C)\ have "bn \ \* xvec5" and "x2 \ bn \" by simp+ + from \x \ (xvec1, xvec2, xvec3, xvec4, xvec5, xvec6, xvec7, xvec8, xvec9, x1, x2, x3, x4, cP, cRs, C)\ have "x \ xvec5" and "x \ x2" and "x \ cRs" by simp+ + + from \length xvec5 = residualLength cRs\ C have "length xvec5 = length(bn \)" + by simp + then obtain p where S: "set p \ set(bn \) \ set(bn(p \ \))" and "distinctPerm p" and "xvec5= bn(p \ \)" + using \bn \ \* xvec5\ \distinct(bn \)\ \distinct xvec5\ + by - (rule constructPerm[where xvec="bn \" and yvec=xvec5], auto simp add: eqvts) + show ?thesis + proof(rule rScope[where P="[(x, x2)] \ P" and \="[(x, x2)] \ p \ \" and P'="[(x, x2)] \ p \ P'"]) + assume "xvec5 \* \" and "xvec5 \* cP" and "xvec5 \* cRs" and "x2 \ \" and "x2 \ cP" and "x2 \ cRs" and "x2 \ xvec5" + from \x2 \ cRs\ C \x2 \ bn \\ \x \ x2\ have "x2 \ \" and "x2 \ P'" by(auto simp add: abs_fresh) + moreover from \cP = \\x\P\ \x2 \ cP\ \x \ x2\ have "cP = \\x2\([(x, x2)] \ P)" + by(simp add: alphaRes abs_fresh) + moreover from B C S \bn \ \* xvec5\ \xvec5 \* cRs\ \xvec5 = bn(p \ \)\ \bn \ \* subject \\ \xvec5 \* cP\ \x \ \\ \x \ xvec5\ + have "cRs = (p \ \) \ \\x\(p \ P')" + apply clarsimp + by(subst residualAlpha[where p=p] alphaRes) (auto simp del: actionFresh) + then have "([(x, x2)] \ cRs) = [(x, x2)] \ ((p \ \) \ \\x\(p \ P'))" + by simp + with \x2 \ cRs\ \x \ cRs\ have "cRs = ([(x, x2)] \ p \ \) \ \\x2\([(x, x2)] \ p \ P')" + by(simp add: eqvts calc_atm) + moreover from \xvec5= bn(p \ \)\ have "([(x, x2)] \ xvec5) = ([(x, x2)] \ bn(p \ \))" + by simp + with \x \ xvec5\ \x2 \ xvec5\ have "xvec5 = bn([(x, x2)] \ p \ \)" + by(simp add: eqvts) + moreover from \\ \ P \\ \ P'\ S B C S \bn \ \* xvec5\ \xvec5 \* cRs\ \xvec5 = bn(p \ \)\ \bn \ \* subject \\ \xvec5 \* cP\ \x \ xvec5\ + have "\ \ P \(p \ \) \ (p \ P')" + by(subst residualAlpha[symmetric]) auto + then have "([(x, x2)] \ \) \ ([(x, x2)] \ P) \([(x, x2)] \ ((p \ \) \ (p \ P')))" + by(rule eqvt) + with \x \ \\ \x2 \ \\ have "\ \ ([(x, x2)] \ P) \([(x, x2)] \ p \ \) \ ([(x, x2)] \ p \ P')" + by(simp add: eqvts) + moreover note \x2 \ \\ + moreover from \x \ \\ \x2 \ \\ \x \ xvec5\ \x2 \ xvec5\ S \x \ x2\ \xvec5 = bn(p \ \)\ have "x2 \ [(x, x2)] \ p \ \" + apply(subgoal_tac "x \ p \ x2 \ p") + apply(simp add: perm_compose freshChainSimps del: actionFresh) + by(auto dest: freshAlphaSwap) + moreover from \bn \ \* subject \\ have "([(x, x2)] \ p \ (bn \)) \* ([(x, x2)] \ p \ (subject \))" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + then have "bn([(x, x2)] \ p \ \) \* subject([(x, x2)] \ p \ \)" + by(simp add: eqvts) + moreover from \distinct(bn \)\ have "distinct([(x, x2)] \ p \ (bn \))" by simp + then have "distinct(bn([(x, x2)] \ p \ \))" by(simp add: eqvts) + ultimately show "cP = \\x2\([(x, x2)] \ P) \ + cRs = ([(x, x2)] \ p \ \) \ \\x2\([(x, x2)] \ p \ P') \ + xvec5 = bn ([(x, x2)] \ p \ \) \ + \ \ [(x, x2)] \ P \ ([(x, x2)] \ p \ \) \ [(x, x2)] \ p \ P' \ + x2 \ \ \ + x2 \ [(x, x2)] \ p \ \ \ + bn ([(x, x2)] \ p \ \) \* subject ([(x, x2)] \ p \ \) \ + distinct (bn ([(x, x2)] \ p \ \))" by blast + qed +next + case(cBang P) + then show ?thesis by(auto intro: rBang) +qed + +lemma resResidEq: + fixes xvec :: "name list" + and P :: "('a, 'b, 'c) psi" + and Q :: "('a, 'b, 'c) psi" + and M :: 'a + and N :: 'a + and N' :: 'a + +assumes "\M\\*xvec\\N\ \ P = \M\\*xvec\\N'\ \ Q" + and "xvec \* M" + +shows "\M\\*xvec\\N\ = \M\\*xvec\\N'\" + using assms +proof(induct xvec) + case Nil + then show ?case by(simp add: residualInject) +next + case (Cons x xvec) + from \(x # xvec) \* M\ + have "x \ M" and "xvec \* M" by simp+ + from \\M\\*(x # xvec)\\N\ \ P = \M\\*(x # xvec)\\N'\ \ Q\ \x \ M\ + have "\M\\*(xvec)\\N\ \ P = \M\\*(xvec)\\N'\ \ Q" + by (metis action.inject(4) assms(1) bn.simps(4) residualInject'') + then have "\M\\*xvec\\N\ = \M\\*xvec\\N'\" using \xvec \* M\ + by(rule Cons(1)) + then show ?case + by(simp add: action.inject) +qed + +lemma parCases[consumes 5, case_names cPar1 cPar2 cComm1 cComm2 cBrMerge cBrComm1 cBrComm2]: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and Q :: "('a, 'b, 'c) psi" + and \ :: "'a action" + and T :: "('a, 'b, 'c) psi" + and C :: "'f::fs_name" + and M :: 'a + and N :: 'a + +assumes Trans: "\ \ P \ Q \\ \ T" + and "bn \ \* \" + and "bn \ \* P" + and "bn \ \* Q" + and "bn \ \* subject \" + and rPar1: "\P' A\<^sub>Q \\<^sub>Q. \\ \ \\<^sub>Q \ P \\ \ P'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + A\<^sub>Q \* \; A\<^sub>Q \* P; A\<^sub>Q \* Q; A\<^sub>Q \* \; A\<^sub>Q \* P'; A\<^sub>Q \* C\ \ Prop \ (P' \ Q)" + and rPar2: "\Q' A\<^sub>P \\<^sub>P. \\ \ \\<^sub>P \ Q \\ \ Q'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* Q; A\<^sub>P \* \; A\<^sub>P \* Q'; A\<^sub>P \* C\ \ Prop \ (P \ Q')" + and rComm1: "\\\<^sub>Q M N P' A\<^sub>P \\<^sub>P K xvec Q' A\<^sub>Q. + \\ \ \\<^sub>Q \ P \M\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \ \ \\<^sub>P \ Q \K\\*xvec\\N\ \ Q'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + \ \ \\<^sub>P \ \\<^sub>Q \ M \ K; distinct xvec; \ = \; + A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* P'; A\<^sub>P \* Q; A\<^sub>P \* xvec; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* C; + A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; A\<^sub>Q \* P; A\<^sub>Q \* K; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* Q; A\<^sub>Q \* xvec; A\<^sub>Q \* Q'; A\<^sub>Q \* C; + xvec \* \; xvec \* \\<^sub>P; xvec \* P; xvec \* M; xvec \* K; xvec \* Q; xvec \* \\<^sub>Q; xvec \* C\ \ + Prop (\) (\\*xvec\(P' \ Q'))" + and rComm2: "\\\<^sub>Q M xvec N P' A\<^sub>P \\<^sub>P K Q' A\<^sub>Q. + \\ \ \\<^sub>Q \ P \M\\*xvec\\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \ \ \\<^sub>P \ Q \K\N\ \ Q'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + \ \ \\<^sub>P \ \\<^sub>Q \ M \ K; distinct xvec; \ = \; + A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* P'; A\<^sub>P \* Q; A\<^sub>P \* xvec; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* C; + A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; A\<^sub>Q \* P; A\<^sub>Q \* K; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* Q; A\<^sub>Q \* xvec; A\<^sub>Q \* Q'; A\<^sub>Q \* C; + xvec \* \; xvec \* \\<^sub>P; xvec \* P; xvec \* M; xvec \* K; xvec \* Q; xvec \* \\<^sub>Q; xvec \* C\ \ + Prop (\) (\\*xvec\(P' \ Q'))" + and rBrMerge: "\\\<^sub>Q M N P' A\<^sub>P \\<^sub>P Q' A\<^sub>Q. + \\ \ \\<^sub>Q \ P \ \M\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \ \ \\<^sub>P \ Q \ \M\N\ \ Q'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* N; A\<^sub>P \* P'; + A\<^sub>P \* Q; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* M; A\<^sub>Q \* M; + A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; A\<^sub>Q \* P; A\<^sub>Q \* N; A\<^sub>Q \* P'; + A\<^sub>Q \* Q; A\<^sub>Q \* Q'; A\<^sub>P \* C; A\<^sub>Q \* C; \ = \M\N\\ \ + Prop (\M\N\) (P' \ Q')" + and rBrComm1: "\\\<^sub>Q M N P' A\<^sub>P \\<^sub>P xvec Q' A\<^sub>Q. + \\ \ \\<^sub>Q \ P \\M\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \ \ \\<^sub>P \ Q \\M\\*xvec\\N\ \ Q'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + distinct xvec; + A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* N; A\<^sub>P \* P'; A\<^sub>P \* Q; A\<^sub>P \* xvec; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* C; + A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; A\<^sub>Q \* P; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* Q; A\<^sub>Q \* xvec; A\<^sub>Q \* Q'; A\<^sub>Q \* C; + A\<^sub>P \* M; A\<^sub>Q \* M; xvec \* M; \M\\*xvec\\N\ = \; + xvec \* \; xvec \* \\<^sub>P; xvec \* P; xvec \* Q; xvec \* \\<^sub>Q\ \ + Prop (\M\\*xvec\\N\) (P' \ Q')" + and rBrComm2: "\\\<^sub>Q M xvec N P' A\<^sub>P \\<^sub>P Q' A\<^sub>Q. + \\ \ \\<^sub>Q \ P \\M\\*xvec\\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \ \ \\<^sub>P \ Q \\M\N\ \ Q'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + distinct xvec; + A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* N; A\<^sub>P \* P'; A\<^sub>P \* Q; A\<^sub>P \* xvec; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* C; + A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; A\<^sub>Q \* P; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* Q; A\<^sub>Q \* xvec; A\<^sub>Q \* Q'; A\<^sub>Q \* C; + A\<^sub>P \* M; A\<^sub>Q \* M; xvec \* M; \M\\*xvec\\N\ = \; + xvec \* \; xvec \* \\<^sub>P; xvec \* P; xvec \* Q; xvec \* \\<^sub>Q\ \ + Prop (\M\\*xvec\\N\) (P' \ Q')" + +shows "Prop \ T" +proof - + from Trans have "distinct(bn \)" by(auto dest: boundOutputDistinct) + have "length(bn \) = residualLength(\ \ T)" by simp + note Trans + moreover have "length [] = inputLength(P \ Q)" and "distinct []" + by(auto simp add: inputLength_inputLength'_inputLength''.simps) + moreover have "length [] = inputLength(P \ Q)" and "distinct []" + by(auto simp add: inputLength_inputLength'_inputLength''.simps) + moreover note \length(bn \) = residualLength(\ \ T)\ \distinct(bn \)\ + moreover note \length(bn \) = residualLength(\ \ T)\ \distinct(bn \)\ + moreover note \length(bn \) = residualLength(\ \ T)\ \distinct(bn \)\ + moreover note \length(bn \) = residualLength(\ \ T)\ \distinct(bn \)\ + moreover note \length(bn \) = residualLength(\ \ T)\ \distinct(bn \)\ + moreover note \length(bn \) = residualLength(\ \ T)\ \distinct(bn \)\ + moreover note \length(bn \) = residualLength(\ \ T)\ \distinct(bn \)\ + moreover obtain x::name where "x \ \" and "x \ P" and "x \ Q" and "x \ \" and "x \ T" + by(generate_fresh "name") auto + ultimately show ?thesis using \bn \ \* \\ \bn \ \* P\ \bn \ \* Q\ \bn \ \* subject \\ + proof(cases rule: semanticsCases[of _ _ _ _ _ _ _ _ _ _ _ _ _ C x x x x]) + case cInput + then show ?thesis + by(simp add: residualInject) + next + case cBrInput + then show ?thesis + by(simp add: residualInject) + next + case cOutput + then show ?thesis + by(simp add: residualInject) + next + case cBrOutput + then show ?thesis + by(simp add: residualInject) + next + case cCase + then show ?thesis + by(simp add: residualInject) + next + case (cPar1 \\<^sub>Q P \' P' Q A\<^sub>Q) + then show ?thesis using assms + by(force simp add: psi.inject residualInject residualInject' intro: rPar1) + next + case (cPar2 \\<^sub>P Q \' Q' P A\<^sub>P) + then show ?thesis using assms + by(force simp add: psi.inject residualInject residualInject' intro: rPar1) + next + case cComm1 + then show ?thesis using assms + by(force simp add: psi.inject residualInject residualInject' intro: rComm1) + next + case cComm2 + then show ?thesis using assms + by(force simp add: psi.inject residualInject residualInject' intro: rComm2) + next + case cBrMerge + then show ?thesis using assms + by(force simp add: psi.inject residualInject residualInject' intro: rBrMerge) + next + case (cBrComm1 \\<^sub>Q P1 M N P' A\<^sub>P \\<^sub>P Q1 Q' A\<^sub>Q) + note \bn \ \* \\ + moreover from \bn \ \* P\ \bn \ \* Q\ have "bn \ \* (P \ Q)" by simp + moreover from \bn \ \* subject \\ have "bn \ \* (\ \ T)" by simp + ultimately have all: + "P \ Q = P1 \ Q1 \ + \ \ T = \M\\*bn \\\N\ \ P' \ Q' \ + \ \ \\<^sub>Q \ P1 \ \M\N\ \ P' \ + extractFrame P1 = \A\<^sub>P, \\<^sub>P\ \ + distinct A\<^sub>P \ + \ \ \\<^sub>P \ Q1 \ \M\\*bn \\\N\ \ Q' \ + extractFrame Q1 = \A\<^sub>Q, \\<^sub>Q\ \ + distinct A\<^sub>Q \ + A\<^sub>P \* \ \ + A\<^sub>P \* \\<^sub>Q \ + A\<^sub>P \* P1 \ + A\<^sub>P \* N \ + A\<^sub>P \* P' \ + A\<^sub>P \* Q1 \ + A\<^sub>P \* Q' \ + A\<^sub>P \* A\<^sub>Q \ + A\<^sub>P \* bn \ \ + A\<^sub>Q \* \ \ + A\<^sub>Q \* \\<^sub>P \ + A\<^sub>Q \* P1 \ + A\<^sub>Q \* N \ + A\<^sub>Q \* P' \ + A\<^sub>Q \* Q1 \ + A\<^sub>Q \* Q' \ + A\<^sub>Q \* bn \ \ + bn \ \* \ \ + bn \ \* \\<^sub>P \ + bn \ \* \\<^sub>Q \ + bn \ \* P1 \ + bn \ \* Q1 \ + A\<^sub>P \* M \ A\<^sub>Q \* M \ bn \ \* M \ A\<^sub>P \* C \ A\<^sub>Q \* C \ distinct (bn \)" + by(rule cBrComm1(1)) + + from all have "bn \ \* M" and "\ \ T = \M\\*bn \\\N\ \ P' \ Q'" + by simp+ + + from \\ \ T = \M\\*bn \\\N\ \ P' \ Q'\ have "\ \ T = RBrOut M (\\*bn \\N \' (P' \ Q'))" + by(simp add: residualInject) + + then obtain xvec N' where "\ = \M\\*xvec\\N'\" + by(auto simp add: residualInject) + then have "bn \ = xvec" by simp + from \\ = \M\\*xvec\\N'\\ \\ \ T = \M\\*bn \\\N\ \ P' \ Q'\ have resEq: "\M\\*xvec\\N'\ \ T = \M\\*bn \\\N\ \ P' \ Q'" + by simp + then have "\M\\*bn \\\N'\ \ T = \M\\*bn \\\N\ \ P' \ Q'" using \bn \ = xvec\ + by simp + then have "\M\\*bn \\\N'\ = \M\\*bn \\\N\" using \bn \ \* M\ + by(rule resResidEq) + with \\ = \M\\*xvec\\N'\\ have "\ = \M\\*xvec\\N\" by simp + + moreover from all have "P \ Q = P1 \ Q1" + and "\ \ T = \M\\*bn \\\N\ \ P' \ Q'" + and "\ \ \\<^sub>Q \ P1 \ \M\N\ \ P'" + and "extractFrame P1 = \A\<^sub>P, \\<^sub>P\" + and "distinct A\<^sub>P" + and "\ \ \\<^sub>P \ Q1 \ \M\\*bn \\\N\ \ Q'" + and "extractFrame Q1 = \A\<^sub>Q, \\<^sub>Q\" + and "distinct A\<^sub>Q" + and "A\<^sub>P \* \" + and "A\<^sub>P \* \\<^sub>Q" + and "A\<^sub>P \* P1" + and "A\<^sub>P \* N" + and "A\<^sub>P \* P'" + and "A\<^sub>P \* Q1" + and "A\<^sub>P \* Q'" + and "A\<^sub>P \* A\<^sub>Q" + and "A\<^sub>P \* bn \" + and "A\<^sub>Q \* \" + and "A\<^sub>Q \* \\<^sub>P" + and "A\<^sub>Q \* P1" + and "A\<^sub>Q \* N" + and "A\<^sub>Q \* P'" + and "A\<^sub>Q \* Q1" + and "A\<^sub>Q \* Q'" + and "A\<^sub>Q \* bn \" + and "bn \ \* \" + and "bn \ \* \\<^sub>P" + and "bn \ \* \\<^sub>Q" + and "bn \ \* P1" + and "bn \ \* Q1" + and "A\<^sub>P \* M" + and "A\<^sub>Q \* M" + and "bn \ \* M" + and "A\<^sub>P \* C" + and "A\<^sub>Q \* C" + and "distinct (bn \)" + by auto + + moreover then have "P = P1" and "Q = Q1" + by(auto simp add: psi.inject) + + ultimately have "Prop (\M\\*(bn \)\\N\) (P' \ Q')" + by(force intro: rBrComm1) + then show ?thesis using \\ \ T = \M\\*bn \\\N\ \ P' \ Q'\[symmetric] + by(force simp add: residualInject) + next + case (cBrComm2 \\<^sub>Q P1 M N P' A\<^sub>P \\<^sub>P Q1 Q' A\<^sub>Q) + note \bn \ \* \\ + moreover from \bn \ \* P\ \bn \ \* Q\ have "bn \ \* (P \ Q)" by simp + moreover from \bn \ \* subject \\ have "bn \ \* (\ \ T)" by simp + ultimately have all: + "P \ Q = P1 \ Q1 \ + \ \ T = \M\\*bn \\\N\ \ P' \ Q' \ + \ \ \\<^sub>Q \ P1 \ \M\\*bn \\\N\ \ P' \ + extractFrame P1 = \A\<^sub>P, \\<^sub>P\ \ + distinct A\<^sub>P \ + \ \ \\<^sub>P \ Q1 \ \M\N\ \ Q' \ + extractFrame Q1 = \A\<^sub>Q, \\<^sub>Q\ \ + distinct A\<^sub>Q \ + A\<^sub>P \* \ \ + A\<^sub>P \* \\<^sub>Q \ + A\<^sub>P \* P1 \ + A\<^sub>P \* N \ + A\<^sub>P \* P' \ + A\<^sub>P \* Q1 \ + A\<^sub>P \* Q' \ + A\<^sub>P \* A\<^sub>Q \ + A\<^sub>P \* bn \ \ + A\<^sub>Q \* \ \ + A\<^sub>Q \* \\<^sub>P \ + A\<^sub>Q \* P1 \ + A\<^sub>Q \* N \ + A\<^sub>Q \* P' \ + A\<^sub>Q \* Q1 \ + A\<^sub>Q \* Q' \ + A\<^sub>Q \* bn \ \ + bn \ \* \ \ + bn \ \* \\<^sub>P \ + bn \ \* \\<^sub>Q \ + bn \ \* P1 \ + bn \ \* Q1 \ + A\<^sub>P \* M \ A\<^sub>Q \* M \ bn \ \* M \ A\<^sub>P \* C \ A\<^sub>Q \* C \ distinct (bn \)" + by(rule cBrComm2(1)) + + from all have "bn \ \* M" and "\ \ T = \M\\*bn \\\N\ \ P' \ Q'" + by simp+ + + from \\ \ T = \M\\*bn \\\N\ \ P' \ Q'\ have "\ \ T = RBrOut M (\\*bn \\N \' (P' \ Q'))" + by(simp add: residualInject) + + then obtain xvec N' where "\ = \M\\*xvec\\N'\" + by(auto simp add: residualInject) + then have "bn \ = xvec" by simp + from \\ = \M\\*xvec\\N'\\ \\ \ T = \M\\*bn \\\N\ \ P' \ Q'\ have resEq: "\M\\*xvec\\N'\ \ T = \M\\*bn \\\N\ \ P' \ Q'" + by simp + then have "\M\\*bn \\\N'\ \ T = \M\\*bn \\\N\ \ P' \ Q'" using \bn \ = xvec\ + by simp + then have "\M\\*bn \\\N'\ = \M\\*bn \\\N\" using \bn \ \* M\ + by(rule resResidEq) + with \\ = \M\\*xvec\\N'\\ have "\ = \M\\*xvec\\N\" by simp + moreover from all have "P \ Q = P1 \ Q1" + and "\ \ T = \M\\*bn \\\N\ \ P' \ Q'" + and "\ \ \\<^sub>Q \ P1 \ \M\\*bn \\\N\ \ P'" + and "extractFrame P1 = \A\<^sub>P, \\<^sub>P\" + and "distinct A\<^sub>P" + and "\ \ \\<^sub>P \ Q1 \ \M\N\ \ Q'" + and "extractFrame Q1 = \A\<^sub>Q, \\<^sub>Q\" + and "distinct A\<^sub>Q" + and "A\<^sub>P \* \" + and "A\<^sub>P \* \\<^sub>Q" + and "A\<^sub>P \* P1" + and "A\<^sub>P \* N" + and "A\<^sub>P \* P'" + and "A\<^sub>P \* Q1" + and "A\<^sub>P \* Q'" + and "A\<^sub>P \* A\<^sub>Q" + and "A\<^sub>P \* bn \" + and "A\<^sub>Q \* \" + and "A\<^sub>Q \* \\<^sub>P" + and "A\<^sub>Q \* P1" + and "A\<^sub>Q \* N" + and "A\<^sub>Q \* P'" + and "A\<^sub>Q \* Q1" + and "A\<^sub>Q \* Q'" + and "A\<^sub>Q \* bn \" + and "bn \ \* \" + and "bn \ \* \\<^sub>P" + and "bn \ \* \\<^sub>Q" + and "bn \ \* P1" + and "bn \ \* Q1" + and "A\<^sub>P \* M" + and "A\<^sub>Q \* M" + and "bn \ \* M" + and "A\<^sub>P \* C" + and "A\<^sub>Q \* C" + and "distinct (bn \)" + by auto + + moreover then have "P = P1" and "Q = Q1" + by(auto simp add: psi.inject) + + ultimately have "Prop (\M\\*(bn \)\\N\) (P' \ Q')" + by(force intro: rBrComm2) + then show ?thesis using \\ \ T = \M\\*bn \\\N\ \ P' \ Q'\[symmetric] + by(force simp add: residualInject) + next + case cBrClose + then show ?thesis using \x \ \\ \x \ P\ \x \ Q\ \x \ \\ \x \ T\ + by(simp add: residualInject) + next + case cOpen + then show ?thesis using assms \x \ \\ \x \ P\ \x \ Q\ \x \ \\ \x \ T\ + by(simp add: residualInject) + next + case cBrOpen + then show ?thesis using assms \x \ \\ \x \ P\ \x \ Q\ \x \ \\ \x \ T\ + by(simp add: residualInject) + next + case cScope + then show ?thesis using assms \x \ \\ \x \ P\ \x \ Q\ \x \ \\ \x \ T\ + by(simp add: residualInject) + next + case cBang + then show ?thesis + by(simp add: residualInject) + qed +qed + +lemma parInputCases[consumes 1, case_names cPar1 cPar2]: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and Q :: "('a, 'b, 'c) psi" + and M :: 'a + and N :: 'a + and R :: "('a, 'b, 'c) psi" + and C :: "'f::fs_name" + +assumes Trans: "\ \ P \ Q \M\N\ \ R" + and rPar1: "\P' A\<^sub>Q \\<^sub>Q. \\ \ \\<^sub>Q \ P \M\N\ \ P'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + A\<^sub>Q \* \; A\<^sub>Q \* P; A\<^sub>Q \* Q; A\<^sub>Q \* M; A\<^sub>Q \* N; A\<^sub>Q \* C\ \ Prop (P' \ Q)" + and rPar2: "\Q' A\<^sub>P \\<^sub>P. \\ \ \\<^sub>P \ Q \M\N\ \ Q'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* Q; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* C\ \ Prop (P \ Q')" +shows "Prop R" +proof - + from Trans obtain \ where "\ \ P \ Q \\ \ R" and "bn \ \* \" and "bn \ \* P" and "bn \ \* Q" and "bn \ \* subject \" and "\ = M\N\" by auto + then show ?thesis using rPar1 rPar2 + by(induct rule: parCases) (auto simp add: residualInject) +qed + +lemma parBrInputCases[consumes 1, case_names cPar1 cPar2 cBrMerge]: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and Q :: "('a, 'b, 'c) psi" + and M :: 'a + and N :: 'a + and R :: "('a, 'b, 'c) psi" + and C :: "'f::fs_name" + +assumes Trans: "\ \ P \ Q \\M\N\ \ R" + and rPar1: "\P' A\<^sub>Q \\<^sub>Q. \\ \ \\<^sub>Q \ P \\M\N\ \ P'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + A\<^sub>Q \* \; A\<^sub>Q \* P; A\<^sub>Q \* Q; A\<^sub>Q \* M; A\<^sub>Q \* N; A\<^sub>Q \* C\ \ Prop (P' \ Q)" + and rPar2: "\Q' A\<^sub>P \\<^sub>P. \\ \ \\<^sub>P \ Q \\M\N\ \ Q'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* Q; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* C\ \ Prop (P \ Q')" + and rBrMerge: "\\\<^sub>Q P' A\<^sub>P \\<^sub>P Q' A\<^sub>Q. + \\ \ \\<^sub>Q \ P \ \M\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \ \ \\<^sub>P \ Q \ \M\N\ \ Q'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* N; A\<^sub>P \* P'; + A\<^sub>P \* Q; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* M; A\<^sub>Q \* M; + A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; A\<^sub>Q \* P; A\<^sub>Q \* N; A\<^sub>Q \* P'; + A\<^sub>Q \* Q; A\<^sub>Q \* Q'; A\<^sub>P \* C; A\<^sub>Q \* C\ \ + Prop (P' \ Q')" +shows "Prop R" +proof - + from Trans obtain \ where "\ \ P \ Q \\ \ R" and "bn \ \* \" and "bn \ \* P" and "bn \ \* Q" and "bn \ \* subject \" and "\ = \M\N\" by auto + then show ?thesis using rPar1 rPar2 rBrMerge + by(induct rule: parCases) (auto simp add: residualInject action.inject) +qed + +lemma parOutputCases[consumes 5, case_names cPar1 cPar2]: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and Q :: "('a, 'b, 'c) psi" + and M :: 'a + and xvec :: "name list" + and N :: 'a + and R :: "('a, 'b, 'c) psi" + and C :: "'f::fs_name" + +assumes Trans: "\ \ P \ Q \M\\*xvec\\N\ \ R" + and "xvec \* \" + and "xvec \* P" + and "xvec \* Q" + and "xvec \* M" + and rPar1: "\P' A\<^sub>Q \\<^sub>Q. \\ \ \\<^sub>Q \ P \M\\*xvec\\N\ \ P'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + A\<^sub>Q \* \; A\<^sub>Q \* P; A\<^sub>Q \* Q; A\<^sub>Q \* M; A\<^sub>Q \* xvec; A\<^sub>Q \* N; A\<^sub>Q \* C; A\<^sub>Q \* xvec; distinct xvec\ \ Prop (P' \ Q)" + and rPar2: "\Q' A\<^sub>P \\<^sub>P. \\ \ \\<^sub>P \ Q \M\\*xvec\\N\ \ Q'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* Q; A\<^sub>P \* M; A\<^sub>P \* xvec; A\<^sub>P \* N; A\<^sub>P \* C; A\<^sub>P \* xvec; distinct xvec\ \ Prop (P \ Q')" +shows "Prop R" +proof - + from Trans have "distinct xvec" by(auto dest: boundOutputDistinct) + obtain \ where "\=M\\*xvec\\N\" by simp + with Trans \xvec \* \\ \xvec \* P\ \xvec \* Q\ \xvec \* M\ + have "\ \ P \ Q \\ \ R" and "bn \ \* \" and "bn \ \* P" and "bn \ \* Q" "bn \ \* subject \" + by simp+ + then show ?thesis using \\=M\\*xvec\\N\\ rPar1 rPar2 \distinct xvec\ + by(induct rule: parCases[where C="(xvec, C)"]) (auto simp add: residualInject) +qed + +lemma parBrOutputCases[consumes 5, case_names cPar1 cPar2 cBrComm1 cBrComm2]: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and Q :: "('a, 'b, 'c) psi" + and M :: 'a + and xvec :: "name list" + and N :: 'a + and R :: "('a, 'b, 'c) psi" + and C :: "'f::fs_name" + +assumes Trans: "\ \ P \ Q \\M\\*xvec\\N\ \ R" + and "xvec \* \" + and "xvec \* P" + and "xvec \* Q" + and "xvec \* M" + and rPar1: "\P' A\<^sub>Q \\<^sub>Q. \\ \ \\<^sub>Q \ P \\M\\*xvec\\N\ \ P'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + A\<^sub>Q \* \; A\<^sub>Q \* P; A\<^sub>Q \* Q; A\<^sub>Q \* M; A\<^sub>Q \* xvec; A\<^sub>Q \* N; A\<^sub>Q \* C; A\<^sub>Q \* xvec; distinct xvec\ \ Prop (P' \ Q)" + and rPar2: "\Q' A\<^sub>P \\<^sub>P. \\ \ \\<^sub>P \ Q \\M\\*xvec\\N\ \ Q'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* Q; A\<^sub>P \* M; A\<^sub>P \* xvec; A\<^sub>P \* N; A\<^sub>P \* C; A\<^sub>P \* xvec; distinct xvec\ \ Prop (P \ Q')" + and rBrComm1: "\\\<^sub>Q P' A\<^sub>P \\<^sub>P Q' A\<^sub>Q. + \\ \ \\<^sub>Q \ P \\M\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \ \ \\<^sub>P \ Q \\M\\*xvec\\N\ \ Q'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + distinct xvec; + A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* N; A\<^sub>P \* P'; A\<^sub>P \* Q; A\<^sub>P \* xvec; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* C; + A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; A\<^sub>Q \* P; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* Q; A\<^sub>Q \* xvec; A\<^sub>Q \* Q'; A\<^sub>Q \* C; + A\<^sub>P \* M; A\<^sub>Q \* M; xvec \* M; + xvec \* \; xvec \* \\<^sub>P; xvec \* P; xvec \* Q; xvec \* \\<^sub>Q\ \ + Prop (P' \ Q')" + and rBrComm2: "\\\<^sub>Q P' A\<^sub>P \\<^sub>P Q' A\<^sub>Q. + \\ \ \\<^sub>Q \ P \\M\\*xvec\\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \ \ \\<^sub>P \ Q \\M\N\ \ Q'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + distinct xvec; + A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* N; A\<^sub>P \* P'; A\<^sub>P \* Q; A\<^sub>P \* xvec; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* C; + A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; A\<^sub>Q \* P; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* Q; A\<^sub>Q \* xvec; A\<^sub>Q \* Q'; A\<^sub>Q \* C; + A\<^sub>P \* M; A\<^sub>Q \* M; xvec \* M; + xvec \* \; xvec \* \\<^sub>P; xvec \* P; xvec \* Q; xvec \* \\<^sub>Q\ \ + Prop (P' \ Q')" +shows "Prop R" +proof - + from Trans have "distinct xvec" by(auto dest: boundOutputDistinct) + obtain \ where "\=\M\\*xvec\\N\" by simp + with Trans \xvec \* \\ \xvec \* P\ \xvec \* Q\ \xvec \* M\ + have "\ \ P \ Q \\ \ R" and "bn \ \* \" and "bn \ \* P" and "bn \ \* Q" "bn \ \* subject \" + by simp+ + then show ?thesis using \\=\M\\*xvec\\N\\ rPar1 rPar2 rBrComm1 rBrComm2 \distinct xvec\ + by(induct rule: parCases[where C="(xvec, C)"]) (auto simp add: residualInject action.inject) +qed + +lemma theEqvt[eqvt_force]: + fixes p :: "name prm" + and \ :: "'a action" + +assumes "\ \ \" + +shows "(p \ the(subject \)) = the(p \ (subject \))" + using assms + by(induct rule: actionCases[where \=\]) auto + +lemma theSubjectFresh[simp]: + fixes \ :: "'a action" + and x :: name + +assumes "\ \ \" + +shows "x \ the(subject \) = x \ subject \" + using assms + by(cases rule: actionCases) auto + +lemma theSubjectFreshChain[simp]: + fixes \ :: "'a action" + and xvec :: "name list" + +assumes "\ \ \" + +shows "xvec \* the(subject \) = xvec \* subject \" + using assms + by(cases rule: actionCases) auto + +lemma inputObtainPrefix: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and P' :: "('a, 'b, 'c) psi" + and A\<^sub>P :: "name list" + and \\<^sub>P :: 'b + and N :: 'a + and K :: 'a + and B :: "name list" + +assumes "\ \ P \K\N\ \ P'" + and "extractFrame P = \A\<^sub>P, \\<^sub>P\" + and "distinct A\<^sub>P" + and "B \* P" + and "A\<^sub>P \* \" + and "A\<^sub>P \* B" + and "A\<^sub>P \* P" + and "A\<^sub>P \* K" + +obtains M where "\ \ \\<^sub>P \ K \ M" and "B \* M" + using assms +proof(nominal_induct avoiding: B arbitrary: thesis rule: inputFrameInduct) + case(cAlpha \ P K N P' A\<^sub>P \\<^sub>P p B) + then obtain M where subjEq: "\ \ \\<^sub>P \ K \ M" and "B \* M" + by(auto intro: cAlpha) + from \\ \ \\<^sub>P \ K \ M\ + have "p \ (\ \ \\<^sub>P \ K \ M)" by simp + with \set p \ set A\<^sub>P \ set (p \ A\<^sub>P)\ + \A\<^sub>P \* \\ \(p \ A\<^sub>P) \* \\ \A\<^sub>P \* K\ \(p \ A\<^sub>P) \* K\ + have permEq: "\ \ (p \ \\<^sub>P) \ K \ (p \ M)" by(simp add: eqvts) + from \B \* M\ have "p \ (B \* M)" by simp + with \set p \ set A\<^sub>P \ set (p \ A\<^sub>P)\ + \A\<^sub>P \* B\ \(p \ A\<^sub>P) \* B\ + have permFresh: "B \* (p \ M)" by(simp add: eqvts) + + show ?case using cAlpha permEq permFresh + by auto +next + case(cInput \ M K xvec N Tvec P B) + from \\ \ M \ K\ have "\ \ \ \ M \ K" + by(blast intro: statEqEnt AssertionStatEqSym[OF Identity]) + then have "\ \ \ \ K \ M" by(rule chanEqSym) + moreover from \B \* (M\\*xvec N\.P)\ have "B \* M" by simp + ultimately show ?case by(auto intro: cInput) +next + case(cCase \ P K N P' \ Cs A\<^sub>P \\<^sub>P B) + then obtain M where "\ \ \\<^sub>P \ K \ M" and "B \* M" + by - (rule cCase, auto dest: memFreshChain) + with \\\<^sub>P \ \\ show ?case by(blast intro: cCase statEqEnt compositionSym Identity) +next + case(cPar1 \ \\<^sub>Q P K N P' A\<^sub>Q Q A\<^sub>P \\<^sub>P B) + then obtain M where "(\ \ \\<^sub>Q) \ \\<^sub>P \ K \ M" and "B \* M" + by (metis freshCompChain(1) psiFreshVec(4)) + then show ?case + by(metis cPar1 statEqEnt Associativity Commutativity AssertionStatEqTrans Composition) +next + case(cPar2 \ \\<^sub>P Q K N Q' A\<^sub>P P A\<^sub>Q \\<^sub>Q B) + then obtain M where "(\ \ \\<^sub>P) \ \\<^sub>Q \ K \ M" and "B \* M" + by - (rule cPar2, auto) + then show ?case by(metis cPar2 statEqEnt Associativity) +next + case(cScope \ P K N P' x A\<^sub>P \\<^sub>P B) + then obtain M where "\ \ \\<^sub>P \ K \ M" and "B \* M" + by - (rule cScope, auto) + then show ?case by(auto intro: cScope) +next + case(cBang \ P K N P' A\<^sub>P \\<^sub>P B) + then obtain M where "\ \ \\<^sub>P \ \ \ K \ M" and "B \* M" + by - (rule cBang, auto) + with \\\<^sub>P \ \\ show ?case by(metis cBang statEqEnt compositionSym Identity) +qed + +lemma outputObtainPrefix: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and P' :: "('a, 'b, 'c) psi" + and A\<^sub>P :: "name list" + and \\<^sub>P :: 'b + and N :: 'a + and K :: 'a + and xvec :: "name list" + and B :: "name list" + +assumes "\ \ P \ ROut K (\\*xvec\N \' P')" + and "extractFrame P = \A\<^sub>P, \\<^sub>P\" + and "distinct A\<^sub>P" + and "xvec \* K" + and "distinct xvec" + and "B \* P" + and "A\<^sub>P \* \" + and "A\<^sub>P \* B" + and "A\<^sub>P \* P" + and "A\<^sub>P \* K" + +obtains M where "\ \ \\<^sub>P \ K \ M" and "B \* M" + using assms +proof(nominal_induct avoiding: B xvec arbitrary: thesis rule: outputFrameInduct) + case(cAlpha \ P K A\<^sub>P \\<^sub>P p \ B xvec) + then obtain M where subjEq: "\ \ \\<^sub>P \ K \ M" and "B \* M" + by(auto intro: cAlpha) + + from \\ \ \\<^sub>P \ K \ M\ + have "p \ (\ \ \\<^sub>P \ K \ M)" by simp + + with \set p \ set A\<^sub>P \ set (p \ A\<^sub>P)\ + \A\<^sub>P \* \\ \(p \ A\<^sub>P) \* \\ \A\<^sub>P \* K\ \(p \ A\<^sub>P) \* K\ + have permEq: "\ \ (p \ \\<^sub>P) \ K \ (p \ M)" by(simp add: eqvts) + + from \B \* M\ have "p \ (B \* M)" by simp + with \set p \ set A\<^sub>P \ set (p \ A\<^sub>P)\ + \A\<^sub>P \* B\ \(p \ A\<^sub>P) \* B\ + have permFresh: "B \* (p \ M)" by(simp add: eqvts) + + show ?case using cAlpha permEq permFresh + by auto +next + case(cOutput \ M K N P B xvec) + from \\ \ M \ K\ have "\ \ \ \ M \ K" + by(blast intro: statEqEnt AssertionStatEqSym[OF Identity]) + then have "\ \ \ \ K \ M" + by(rule chanEqSym) + moreover from \B \* (M\N\.P)\ have "B \* M" by simp + ultimately show ?case by(auto intro: cOutput) +next + case(cCase \ P K P' \ Cs A\<^sub>P \\<^sub>P B xvec) + then obtain M where "\ \ \\<^sub>P \ K \ M" and "B \* M" + by - (rule cCase, auto dest: memFreshChain) + with \\\<^sub>P \ \\ show ?case by(blast intro: cCase statEqEnt compositionSym Identity) +next + case(cPar1 \ \\<^sub>Q P K yvec N P' A\<^sub>Q Q A\<^sub>P \\<^sub>P B xvec) + then obtain M where "(\ \ \\<^sub>Q) \ \\<^sub>P \ K \ M" and "B \* M" + by (metis freshCompChain(1) psiFreshVec(4)) + then show ?case by(metis cPar1 statEqEnt Associativity Commutativity AssertionStatEqTrans Composition) +next + case(cPar2 \ \\<^sub>P Q K yvec N Q' A\<^sub>P P A\<^sub>Q \\<^sub>Q B xvec) + then obtain M where "(\ \ \\<^sub>P) \ \\<^sub>Q \ K \ M" and "B \* M" + by (metis freshCompChain(1) psiFreshVec(4)) + then show ?case by(metis cPar2 statEqEnt Associativity) +next + case(cOpen \ P M zvec yvec N P' x A\<^sub>P \\<^sub>P B xvec) + then obtain K where "\ \ \\<^sub>P \ M \ K" and "B \* K" + by (metis abs_fresh_list_star' psiFreshVec(5)) + then show ?case by(auto intro: cOpen) +next + case(cScope \ P K yvec N P' x A\<^sub>P \\<^sub>P B xvec) + then obtain M where "\ \ \\<^sub>P \ K \ M" and "B \* M" + by (metis abs_fresh_list_star' psiFreshVec(5)) + then show ?case by(auto intro: cScope) +next + case(cBang \ P K P' A\<^sub>P \\<^sub>P B xvec) + then obtain M where "\ \ \\<^sub>P \ \ \ K \ M" and "B \* M" + by - (rule cBang, auto) + with \\\<^sub>P \ \\ show ?case by(metis cBang statEqEnt compositionSym Identity) +qed + +lemma inputRenameSubject: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and M :: 'a + and N :: 'a + and P' :: "('a, 'b, 'c) psi" + and A\<^sub>P :: "name list" + and \\<^sub>P :: 'b + +assumes "\ \ P \M\N\ \ P'" + and "extractFrame P = \A\<^sub>P, \\<^sub>P\" + and "distinct A\<^sub>P" + and "\ \ \\<^sub>P \ M \ K" + and "A\<^sub>P \* \" + and "A\<^sub>P \* P" + and "A\<^sub>P \* M" + and "A\<^sub>P \* K" + +shows "\ \ P \K\N\ \ P'" + using assms +proof(nominal_induct avoiding: K rule: inputFrameInduct) + case(cAlpha \ P M N P' A\<^sub>P \\<^sub>P p K) + have S: "set p \ set A\<^sub>P \ set (p \ A\<^sub>P)" by fact + from \\ \ (p \ \\<^sub>P) \ M \ K\ have "(p \ (\ \ (p \ \\<^sub>P))) \ (p \ M) \ (p \ K)" + by(rule chanEqClosed) + with S \distinctPerm p\ \A\<^sub>P \* \\ \A\<^sub>P \* M\ \A\<^sub>P \* K\ \(p \ A\<^sub>P) \* \\ \(p \ A\<^sub>P) \* M\ \(p \ A\<^sub>P) \* K\ + have "\ \ \\<^sub>P \ M \ K" by(simp add: eqvts) + with \A\<^sub>P \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* M\ \A\<^sub>P \* K\ + \\\ \ \\<^sub>P \ M \ K; A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* K\ \ \ \ P \K\N\ \ P'\ + show ?case by blast +next + case(cInput \ M K xvec N Tvec P K') + from \\ \ \ \ K \ K'\ have "\ \ K \ K'" + by(blast intro: statEqEnt Identity) + with \\ \ M \ K\ have "\ \ M \ K'" + by(rule chanEqTrans) + then show ?case using \distinct xvec\ \set xvec \ supp N\ \length xvec = length Tvec\ + by(rule Input) +next + case(cCase \ P M N P' \ Cs A\<^sub>P \\<^sub>P K) + from \\ \ \ \ M \ K\ \\\<^sub>P \ \\ have "\ \ \\<^sub>P \ M \ K" + by(blast intro: statEqEnt Identity compositionSym AssertionStatEqSym) + with \A\<^sub>P \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* M\ \A\<^sub>P \* K\ + \\K. \\ \ \\<^sub>P \ M \ K; A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* K\ \ \ \ P \K\N\ \ P'\ + have "\ \ P \K\N\ \ P'" by force + then show ?case using \(\, P) \ set Cs\ \\ \ \\ \guarded P\ by(rule Case) +next + case(cPar1 \ \\<^sub>Q P M N P' A\<^sub>Q Q A\<^sub>P \\<^sub>P K) + from \\ \ \\<^sub>P \ \\<^sub>Q \ M \ K\ have "(\ \ \\<^sub>Q) \ \\<^sub>P \ M \ K" + by(metis statEqEnt Associativity Composition AssertionStatEqTrans Commutativity) + with \A\<^sub>P \* \\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>P \* P\ \A\<^sub>P \* M\ \A\<^sub>P \* K\ + \\K. \(\ \ \\<^sub>Q) \ \\<^sub>P \ M \ K; A\<^sub>P \* (\ \ \\<^sub>Q); A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* K\ \ \ \ \\<^sub>Q \ P \K\N\ \ P'\ + have "\ \ \\<^sub>Q \ P \K\N\ \ P'" by force + then show ?case using \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ \A\<^sub>Q \* \\ \A\<^sub>Q \* P\ \A\<^sub>Q \* K\ \A\<^sub>Q \* N\ + by(auto intro: Par1) +next + case(cPar2 \ \\<^sub>P Q M N Q' A\<^sub>P P A\<^sub>Q \\<^sub>Q K) + from \\ \ \\<^sub>P \ \\<^sub>Q \ M \ K\ have "(\ \ \\<^sub>P) \ \\<^sub>Q \ M \ K" + by(rule statEqEnt[OF AssertionStatEqSym[OF Associativity]]) + with \A\<^sub>Q \* \\ \A\<^sub>Q \* \\<^sub>P\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* M\ \A\<^sub>Q \* K\ + \\K. \(\ \ \\<^sub>P) \ \\<^sub>Q \ M \ K; A\<^sub>Q \* (\ \ \\<^sub>P); A\<^sub>Q \* Q; A\<^sub>Q \* M; A\<^sub>Q \* K\ \ \ \ \\<^sub>P \ Q \K\N\ \ Q'\ + have "\ \ \\<^sub>P \ Q \K\N\ \ Q'" by force + then show ?case using \extractFrame P = \A\<^sub>P, \\<^sub>P\\ \A\<^sub>P \* \\ \A\<^sub>P \* Q\ \A\<^sub>P \* K\ \A\<^sub>P \* N\ + by(auto intro: Par2) +next + case(cScope \ P M N P' x A\<^sub>P \\<^sub>P) + then have "\ \ P \K\N\ \ P'" by force + with \x \ \\ \x \ K\ \x \ N\ show ?case + by(auto intro: Scope) +next + case(cBang \ P M N P' A\<^sub>P \\<^sub>P K) + from \\ \ \ \ M \ K\ \\\<^sub>P \ \\ have "\ \ \\<^sub>P \ \ \ M \ K" + by(blast intro: statEqEnt Identity compositionSym AssertionStatEqSym) + with \A\<^sub>P \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* M\ \A\<^sub>P \* K\ + \\K. \\ \ \\<^sub>P \ \ \ M \ K; A\<^sub>P \* \; A\<^sub>P \* (P \ !P); A\<^sub>P \* M; A\<^sub>P \* K\ \ \ \ P \ !P \K\N\ \ P'\ + have "\ \ P \ !P \K\N\ \ P'" by force + then show ?case using \guarded P\ by(rule Bang) +qed + +lemma outputRenameSubject: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and M :: 'a + and xvec :: "name list" + and N :: 'a + and P' :: "('a, 'b, 'c) psi" + and A\<^sub>P :: "name list" + and \\<^sub>P :: 'b + +assumes "\ \ P \M\\*xvec\\N\ \ P'" + and "extractFrame P = \A\<^sub>P, \\<^sub>P\" + and "distinct A\<^sub>P" + and "\ \ \\<^sub>P \ M \ K" + and "A\<^sub>P \* \" + and "A\<^sub>P \* P" + and "A\<^sub>P \* M" + and "A\<^sub>P \* K" + +shows "\ \ P \K\\*xvec\\N\ \ P'" + using assms unfolding residualInject +proof(nominal_induct avoiding: K rule: outputFrameInduct) + case(cAlpha \ P M A\<^sub>P \\<^sub>P p B K) + have S: "set p \ set A\<^sub>P \ set(p \ A\<^sub>P)" by fact + from \\ \ (p \ \\<^sub>P) \ M \ K\ have "(p \ (\ \ (p \ \\<^sub>P))) \ (p \ M) \ (p \ K)" + by(rule chanEqClosed) + with S \distinctPerm p\ \A\<^sub>P \* \\ \A\<^sub>P \* M\ \A\<^sub>P \* K\ \(p \ A\<^sub>P) \* \\ \(p \ A\<^sub>P) \* M\ \(p \ A\<^sub>P) \* K\ + have "\ \ \\<^sub>P \ M \ K" by(simp add: eqvts) + with \A\<^sub>P \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* M\ \A\<^sub>P \* K\ + show ?case by(blast intro: cAlpha) +next + case(cOutput \ M K N P K') + from \\ \ \ \ K \ K'\ have "\ \ K \ K'" + by(blast intro: statEqEnt Identity) + with \\ \ M \ K\ have "\ \ M \ K'" + by(rule chanEqTrans) + then show ?case using Output by(force simp add: residualInject) +next + case(cCase \ P M B \ Cs A\<^sub>P \\<^sub>P K) + from \\ \ \ \ M \ K\ \\\<^sub>P \ \\ have "\ \ \\<^sub>P \ M \ K" + by(blast intro: statEqEnt Identity compositionSym AssertionStatEqSym) + with \A\<^sub>P \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* M\ \A\<^sub>P \* K\ + \\K. \\ \ \\<^sub>P \ M \ K; A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* K\ \ \ \ P \(ROut K B)\ + have "\ \ P \ROut K B" by force + then show ?case using \(\, P) \ set Cs\ \\ \ \\ \guarded P\ by(rule Case) +next + case(cPar1 \ \\<^sub>Q P M xvec N P' A\<^sub>Q Q A\<^sub>P \\<^sub>P K) + from \\ \ \\<^sub>P \ \\<^sub>Q \ M \ K\ have "(\ \ \\<^sub>Q) \ \\<^sub>P \ M \ K" + by(metis statEqEnt Associativity Composition AssertionStatEqTrans Commutativity) + with \A\<^sub>P \* \\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>P \* P\ \A\<^sub>P \* M\ \A\<^sub>P \* K\ + \\K. \(\ \ \\<^sub>Q) \ \\<^sub>P \ M \ K; A\<^sub>P \* (\ \ \\<^sub>Q); A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* K\ \ \ \ \\<^sub>Q \ P \(ROut K (\\*xvec\N \' P'))\ + have "\ \ \\<^sub>Q \ P \K\\*xvec\\N\ \ P'" by(force simp add: residualInject) + then show ?case using \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ \xvec \* Q\ \A\<^sub>Q \* \\ \A\<^sub>Q \* P\ \A\<^sub>Q \* K\ \A\<^sub>Q \* xvec\ \A\<^sub>Q \* N\ Par1[where \="K\\*xvec\\N\"] + by(auto simp add: residualInject) +next + case(cPar2 \ \\<^sub>P Q M xvec N Q' A\<^sub>P P A\<^sub>Q \\<^sub>Q K) + from \\ \ \\<^sub>P \ \\<^sub>Q \ M \ K\ have "(\ \ \\<^sub>P) \ \\<^sub>Q \ M \ K" + by(rule statEqEnt[OF AssertionStatEqSym[OF Associativity]]) + with \A\<^sub>Q \* \\ \A\<^sub>Q \* \\<^sub>P\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* M\ \A\<^sub>Q \* K\ + \\K. \(\ \ \\<^sub>P) \ \\<^sub>Q \ M \ K; A\<^sub>Q \* (\ \ \\<^sub>P); A\<^sub>Q \* Q; A\<^sub>Q \* M; A\<^sub>Q \* K\ \ \ \ \\<^sub>P \ Q \ROut K (\\*xvec\N \' Q')\ + have "\ \ \\<^sub>P \ Q \ROut K (\\*xvec\N \' Q')" by force + then show ?case using \extractFrame P = \A\<^sub>P, \\<^sub>P\\ \xvec \* P\ \A\<^sub>P \* \\ \A\<^sub>P \* Q\ \A\<^sub>P \* K\ \A\<^sub>P \* xvec\ \A\<^sub>P \* N\ Par2[where \="K\\*xvec\\N\"] + by(auto simp add: residualInject) +next + case(cOpen \ P M xvec yvec N P' x A\<^sub>P \\<^sub>P) + then have "\ \ P \K\\*(xvec@yvec)\\N\ \ P'" by(force simp add: residualInject) + with \x \ supp N\ \x \ \\ \x \ K\ \x \ xvec\ \x \ yvec\ Open show ?case + by(auto simp add: residualInject) +next + case(cScope \ P M xvec N P' x A\<^sub>P \\<^sub>P) + then have "\ \ P \K\\*xvec\\N\ \ P'" by(force simp add: residualInject) + with \x \ \\ \x \ K\ \x \ xvec\ \x \ N\ Scope[where \="K\\*xvec\\N\"] show ?case + by(auto simp add: residualInject) +next + case(cBang \ P M B A\<^sub>P \\<^sub>P K) + from \\ \ \ \ M \ K\ \\\<^sub>P \ \\ have "\ \ \\<^sub>P \ \ \ M \ K" + by(blast intro: statEqEnt Identity compositionSym AssertionStatEqSym) + with \A\<^sub>P \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* M\ \A\<^sub>P \* K\ + \\K. \\ \ \\<^sub>P \ \ \ M \ K; A\<^sub>P \* \; A\<^sub>P \* (P \ !P); A\<^sub>P \* M; A\<^sub>P \* K\ \ \ \ P \ !P \ROut K B\ + have "\ \ P \ !P \ROut K B" by force + then show ?case using \guarded P\ by(rule Bang) +qed + +lemma parCasesSubject[consumes 7, case_names cPar1 cPar2 cComm1 cComm2 cBrMerge cBrComm1 cBrComm2]: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and Q :: "('a, 'b, 'c) psi" + and \ :: "'a action" + and R :: "('a, 'b, 'c) psi" + and C :: "'f::fs_name" + and yvec :: "name list" + +assumes Trans: "\ \ P \ Q \\ \ R" + and "bn \ \* \" + and "bn \ \* P" + and "bn \ \* Q" + and "bn \ \* subject \" + and "yvec \* P" + and "yvec \* Q" + and rPar1: "\P' A\<^sub>Q \\<^sub>Q. \\ \ \\<^sub>Q \ P \\ \ P'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + A\<^sub>Q \* \; A\<^sub>Q \* P; A\<^sub>Q \* \; A\<^sub>Q \* C\ \ Prop \ (P' \ Q)" + and rPar2: "\Q' A\<^sub>P \\<^sub>P. \\ \ \\<^sub>P \ Q \\ \ Q'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + A\<^sub>P \* \; A\<^sub>P \* Q; A\<^sub>P \* \; A\<^sub>P \* C\ \ Prop \ (P \ Q')" + and rComm1: "\\\<^sub>Q M N P' A\<^sub>P \\<^sub>P K xvec Q' A\<^sub>Q. + \\ \ \\<^sub>Q \ P \M\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \ \ \\<^sub>P \ Q \K\\*xvec\\N\ \ Q'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + \ \ \\<^sub>P \ \\<^sub>Q \ M \ K; yvec \* M; yvec \* K; distinct xvec; \ = \; + A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* P'; A\<^sub>P \* Q; A\<^sub>P \* xvec; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* C; + A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; A\<^sub>Q \* P; A\<^sub>Q \* K; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* Q; A\<^sub>Q \* xvec; A\<^sub>Q \* Q'; A\<^sub>Q \* C; + xvec \* \; xvec \* \\<^sub>P; xvec \* P; xvec \* M; xvec \* K; xvec \* Q; xvec \* \\<^sub>Q; xvec \* C\ \ + Prop (\) (\\*xvec\(P' \ Q'))" + and rComm2: "\\\<^sub>Q M xvec N P' A\<^sub>P \\<^sub>P K Q' A\<^sub>Q. + \\ \ \\<^sub>Q \ P \M\\*xvec\\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \ \ \\<^sub>P \ Q \K\N\ \ Q'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + \ \ \\<^sub>P \ \\<^sub>Q \ M \ K; yvec \* M; yvec \* K; distinct xvec; \ = \; + A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* P'; A\<^sub>P \* Q; A\<^sub>P \* xvec; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* C; + A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; A\<^sub>Q \* P; A\<^sub>Q \* K; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* Q; A\<^sub>Q \* xvec; A\<^sub>Q \* Q'; A\<^sub>Q \* C; + xvec \* \; xvec \* \\<^sub>P; xvec \* P; xvec \* M; xvec \* K; xvec \* Q; xvec \* \\<^sub>Q; xvec \* C\ \ + Prop (\) (\\*xvec\(P' \ Q'))" + and rBrMerge: "\\\<^sub>Q M N P' A\<^sub>P \\<^sub>P Q' A\<^sub>Q. + \\ \ \\<^sub>Q \ P \ \M\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \ \ \\<^sub>P \ Q \ \M\N\ \ Q'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* N; A\<^sub>P \* P'; + A\<^sub>P \* Q; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* M; A\<^sub>Q \* M; + A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; A\<^sub>Q \* P; A\<^sub>Q \* N; A\<^sub>Q \* P'; + A\<^sub>Q \* Q; A\<^sub>Q \* Q'; A\<^sub>P \* C; A\<^sub>Q \* C; \ = \M\N\\ \ + Prop (\M\N\) (P' \ Q')" + and rBrComm1: "\\\<^sub>Q M N P' A\<^sub>P \\<^sub>P xvec Q' A\<^sub>Q. + \\ \ \\<^sub>Q \ P \\M\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \ \ \\<^sub>P \ Q \\M\\*xvec\\N\ \ Q'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + distinct xvec; + A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* N; A\<^sub>P \* P'; A\<^sub>P \* Q; A\<^sub>P \* xvec; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* C; + A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; A\<^sub>Q \* P; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* Q; A\<^sub>Q \* xvec; A\<^sub>Q \* Q'; A\<^sub>Q \* C; + A\<^sub>P \* M; A\<^sub>Q \* M; xvec \* M; \M\\*xvec\\N\ = \; + xvec \* \; xvec \* \\<^sub>P; xvec \* P; xvec \* Q; xvec \* \\<^sub>Q\ \ + Prop (\M\\*xvec\\N\) (P' \ Q')" + and rBrComm2: "\\\<^sub>Q M xvec N P' A\<^sub>P \\<^sub>P Q' A\<^sub>Q. + \\ \ \\<^sub>Q \ P \\M\\*xvec\\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \ \ \\<^sub>P \ Q \\M\N\ \ Q'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + distinct xvec; + A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* N; A\<^sub>P \* P'; A\<^sub>P \* Q; A\<^sub>P \* xvec; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* C; + A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; A\<^sub>Q \* P; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* Q; A\<^sub>Q \* xvec; A\<^sub>Q \* Q'; A\<^sub>Q \* C; + A\<^sub>P \* M; A\<^sub>Q \* M; xvec \* M; \M\\*xvec\\N\ = \; + xvec \* \; xvec \* \\<^sub>P; xvec \* P; xvec \* Q; xvec \* \\<^sub>Q\ \ + Prop (\M\\*xvec\\N\) (P' \ Q')" + +shows "Prop \ R" + using Trans \bn \ \* \\ \bn \ \* P\ \bn \ \* Q\ \bn \ \* subject \\ +proof(induct rule: parCases[where C="(C, yvec)"]) + case(cPar1 P' A\<^sub>Q \\<^sub>Q) + then show ?case by(auto intro: rPar1) +next + case(cPar2 Q' A\<^sub>P \\<^sub>P) + then show ?case by(auto intro: rPar2) +next + case(cComm1 \\<^sub>Q M N P' A\<^sub>P \\<^sub>P K xvec Q' A\<^sub>Q) + from \A\<^sub>P \* (C, yvec)\ \A\<^sub>Q \* (C, yvec)\ \xvec \* (C, yvec)\ + have "A\<^sub>P \* C" and "A\<^sub>Q \* C" and "xvec \* C" and "A\<^sub>P \* yvec" and "A\<^sub>Q \* yvec" and "xvec \* yvec" + by simp+ + + have FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" and FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" + and MeqK: "\ \ \\<^sub>P \ \\<^sub>Q \ M \ K" by fact+ + + from \\ \ \\<^sub>Q \ P \M\N\ \ P'\ FrP \distinct A\<^sub>P\ \A\<^sub>P \* P\ \A\<^sub>Q \* P\ \yvec \* P\ \A\<^sub>P \* \\ + \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* yvec\ \A\<^sub>P \* xvec\ \A\<^sub>P \* P\ \A\<^sub>P \* M\ \xvec \* P\ \A\<^sub>P \* \\<^sub>Q\ + obtain M' where MeqM': "(\ \ \\<^sub>Q) \ \\<^sub>P \ M \ M'" and "xvec \* M'" and "yvec \* M'" and "A\<^sub>Q \* M'" + by - (rule inputObtainPrefix[where B="xvec@yvec@A\<^sub>Q"], (assumption | force)+) + from \\ \ \\<^sub>P \ Q \K\\*xvec\\N\ \ Q'\ have "\ \ \\<^sub>P \ Q \ ROut K (\\*xvec\N \' Q')" + by(simp add: residualInject) + with FrQ \distinct A\<^sub>Q\ \A\<^sub>P \* Q\ \A\<^sub>Q \* Q\ \yvec \* Q\ \A\<^sub>Q \* \\ + \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>Q \* yvec\ \A\<^sub>Q \* xvec\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* K\ \xvec \* Q\ \A\<^sub>Q \* \\<^sub>P\ \xvec \* K\ \distinct xvec\ + obtain K' where KeqK': "(\ \ \\<^sub>P) \ \\<^sub>Q \ K \ K'" and "xvec \* K'" and "yvec \* K'" and "A\<^sub>P \* K'" + by - (rule outputObtainPrefix[where B="xvec@yvec@A\<^sub>P"], (assumption | force | metis freshChainSym)+) + + from MeqK KeqK' have "(\ \ \\<^sub>Q) \ \\<^sub>P \ M \ K'" + by(metis statEqEnt Associativity Commutativity Composition chanEqTrans) + with \\ \ \\<^sub>Q \ P \M\N\ \ P'\ FrP \distinct A\<^sub>P\ + have "\ \ \\<^sub>Q \ P \K'\N\ \ P'" using \A\<^sub>P \* \\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>P \* P\ \A\<^sub>P \* M\ \A\<^sub>P \* K'\ + by - (rule inputRenameSubject, (assumption | force)+) + moreover note FrP \distinct A\<^sub>P\ + moreover from MeqK MeqM' have "(\ \ \\<^sub>P) \ \\<^sub>Q \ K \ M'" + by(metis statEqEnt Associativity Commutativity Composition chanEqTrans chanEqSym) + with \\ \ \\<^sub>P \ Q \K\\*xvec\\N\ \ Q'\ FrQ \distinct A\<^sub>Q\ + have "\ \ \\<^sub>P \ Q \M'\\*xvec\\N\ \ Q'" using \A\<^sub>Q \* \\ \A\<^sub>Q \* \\<^sub>P\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* K\ \A\<^sub>Q \* M'\ + by - (rule outputRenameSubject, (assumption | force)+) + moreover note FrQ \distinct A\<^sub>Q\ + moreover from MeqM' KeqK' MeqK have "\ \ \\<^sub>P \ \\<^sub>Q \ K' \ M'" + by(metis statEqEnt Associativity Commutativity Composition chanEqTrans chanEqSym) + moreover note \A\<^sub>P \* \\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>P \* P\ \A\<^sub>P \* K'\ \A\<^sub>P \* N\ \A\<^sub>P \* P'\ \A\<^sub>P \* Q\ \A\<^sub>P \* xvec\ \A\<^sub>P \* Q'\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* C\ + \A\<^sub>Q \* \\ \A\<^sub>Q \* \\<^sub>P\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* M'\ \A\<^sub>Q \* N\ \A\<^sub>Q \* Q'\ \A\<^sub>Q \* P\ \A\<^sub>Q \* xvec\ \A\<^sub>Q \* P'\ \A\<^sub>Q \* C\ \\ = \\ + \xvec \* \\ \xvec \* \\<^sub>P\ \xvec \* P\ \xvec \* M'\ \xvec \* K'\ \xvec \* Q\ \xvec \* \\<^sub>Q\ \xvec \* C\ \yvec \* M'\ \yvec \* K'\ \distinct xvec\ + ultimately show ?case + by(metis rComm1) +next + case(cComm2 \\<^sub>Q M xvec N P' A\<^sub>P \\<^sub>P K Q' A\<^sub>Q) + from \A\<^sub>P \* (C, yvec)\ \A\<^sub>Q \* (C, yvec)\ \xvec \* (C, yvec)\ + have "A\<^sub>P \* C" and "A\<^sub>Q \* C" and "xvec \* C" and "A\<^sub>P \* yvec" and "A\<^sub>Q \* yvec" and "xvec \* yvec" + by simp+ + + have FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" and FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" + and MeqK: "\ \ \\<^sub>P \ \\<^sub>Q \ M \ K" by fact+ + + from \\ \ \\<^sub>Q \ P \M\\*xvec\\N\ \ P'\ have "\ \ \\<^sub>Q \ P \ ROut M (\\*xvec\N \' P')" + by(simp add: residualInject) + with FrP \distinct A\<^sub>P\ \A\<^sub>P \* P\ \A\<^sub>Q \* P\ \yvec \* P\ \A\<^sub>P \* \\ + \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* yvec\ \A\<^sub>P \* xvec\ \A\<^sub>P \* P\ \A\<^sub>P \* M\ \xvec \* P\ \A\<^sub>P \* \\<^sub>Q\ \xvec \* M\ \distinct xvec\ + obtain M' where MeqM': "(\ \ \\<^sub>Q) \ \\<^sub>P \ M \ M'" and "xvec \* M'" and "yvec \* M'" and "A\<^sub>Q \* M'" + by - (rule outputObtainPrefix[where B="xvec@yvec@A\<^sub>Q"], (assumption | force)+) + from \\ \ \\<^sub>P \ Q \K\N\ \ Q'\ FrQ \distinct A\<^sub>Q\ \A\<^sub>P \* Q\ \A\<^sub>Q \* Q\ \yvec \* Q\ \A\<^sub>Q \* \\ + \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>Q \* yvec\ \A\<^sub>Q \* xvec\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* K\ \xvec \* Q\ \A\<^sub>Q \* \\<^sub>P\ + obtain K' where KeqK': "(\ \ \\<^sub>P) \ \\<^sub>Q \ K \ K'" and "xvec \* K'" and "yvec \* K'" and "A\<^sub>P \* K'" + by - (rule inputObtainPrefix[where B="xvec@yvec@A\<^sub>P"], (assumption | force | metis freshChainSym)+) + + from MeqK KeqK' have "(\ \ \\<^sub>Q) \ \\<^sub>P \ M \ K'" + by(metis statEqEnt Associativity Commutativity Composition chanEqTrans) + with \\ \ \\<^sub>Q \ P \M\\*xvec\\N\ \ P'\ FrP \distinct A\<^sub>P\ + have "\ \ \\<^sub>Q \ P \K'\\*xvec\\N\ \ P'" using \A\<^sub>P \* \\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>P \* P\ \A\<^sub>P \* M\ \A\<^sub>P \* K'\ + by - (rule outputRenameSubject, (assumption | force)+) + moreover note FrP \distinct A\<^sub>P\ + moreover from MeqK MeqM' have "(\ \ \\<^sub>P) \ \\<^sub>Q \ K \ M'" + by(metis statEqEnt Associativity Commutativity Composition chanEqTrans chanEqSym) + with \\ \ \\<^sub>P \ Q \K\N\ \ Q'\ FrQ \distinct A\<^sub>Q\ + have "\ \ \\<^sub>P \ Q \M'\N\ \ Q'" using \A\<^sub>Q \* \\ \A\<^sub>Q \* \\<^sub>P\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* K\ \A\<^sub>Q \* M'\ + by - (rule inputRenameSubject, (assumption | force)+) + moreover note FrQ \distinct A\<^sub>Q\ + moreover from MeqM' KeqK' MeqK have "\ \ \\<^sub>P \ \\<^sub>Q \ K' \ M'" + by(metis statEqEnt Associativity Commutativity Composition chanEqTrans chanEqSym) + moreover note \A\<^sub>P \* \\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>P \* P\ \A\<^sub>P \* K'\ \A\<^sub>P \* N\ \A\<^sub>P \* P'\ \A\<^sub>P \* Q\ \A\<^sub>P \* xvec\ \A\<^sub>P \* Q'\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* C\ + \A\<^sub>Q \* \\ \A\<^sub>Q \* \\<^sub>P\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* M'\ \A\<^sub>Q \* N\ \A\<^sub>Q \* Q'\ \A\<^sub>Q \* P\ \A\<^sub>Q \* xvec\ \A\<^sub>Q \* P'\ \A\<^sub>Q \* C\ \\ = \\ + \xvec \* \\ \xvec \* \\<^sub>P\ \xvec \* P\ \xvec \* M'\ \xvec \* K'\ \xvec \* Q\ \xvec \* \\<^sub>Q\ \xvec \* C\ \yvec \* M'\ \yvec \* K'\ \distinct xvec\ + ultimately show ?case + by(metis rComm2) +next + case cBrMerge + then show ?case by (simp add: rBrMerge) +next + case cBrComm1 + then show ?case by (auto intro: rBrComm1) +next + case cBrComm2 + then show ?case by (auto intro: rBrComm2) +qed + +lemma inputCases[consumes 1, case_names cInput cBrInput]: + fixes \ :: 'b + and M :: 'a + and xvec :: "name list" + and N :: 'a + and P :: "('a, 'b, 'c) psi" + and \ :: "'a action" + and P' :: "('a, 'b, 'c) psi" + +assumes Trans: "\ \ M\\*xvec N\.P \\ \ P'" + and rInput: "\K Tvec. \\ \ M \ K; set xvec \ supp N; length xvec = length Tvec; distinct xvec\ \ Prop (K\N[xvec::=Tvec]\) (P[xvec::=Tvec])" + and rBrInput: "\K Tvec. \\ \ K \ M; set xvec \ supp N; length xvec = length Tvec; distinct xvec\ \ Prop (\K\N[xvec::=Tvec]\) (P[xvec::=Tvec])" + +shows "Prop \ P'" +proof - + { + fix xvec N P + assume Trans: "\ \ M\\*xvec N\.P \\ \ P'" + and "xvec \* \" and "xvec \* M" and "xvec \* \" and "xvec \* P'" and "distinct xvec" + and rInput: "\K Tvec. \\ \ M \ K; set xvec \ supp N; length xvec = length Tvec; distinct xvec\ \ Prop (K\N[xvec::=Tvec]\) (P[xvec::=Tvec])" + and rBrInput: "\K Tvec. \\ \ K \ M; set xvec \ supp N; length xvec = length Tvec; distinct xvec\ \ Prop (\K\N[xvec::=Tvec]\) (P[xvec::=Tvec])" + + from Trans have "bn \ = []" + apply - + by(ind_cases "\ \ M\\*xvec N\.P \\ \ P'") (auto simp add: residualInject) + from Trans have "distinct(bn \)" by(auto dest: boundOutputDistinct) + have "length(bn \) = residualLength(\ \ P')" by simp + note Trans + moreover have "length xvec = inputLength(M\\*xvec N\.P)" by auto + moreover note \distinct xvec\ + moreover have "length xvec = inputLength(M\\*xvec N\.P)" by auto + moreover note \distinct xvec\ + moreover note \length(bn \) = residualLength(\ \ P')\ \distinct(bn \)\ + moreover note \length(bn \) = residualLength(\ \ P')\ \distinct(bn \)\ + moreover note \length(bn \) = residualLength(\ \ P')\ \distinct(bn \)\ + moreover note \length(bn \) = residualLength(\ \ P')\ \distinct(bn \)\ + moreover note \length(bn \) = residualLength(\ \ P')\ \distinct(bn \)\ + moreover note \length(bn \) = residualLength(\ \ P')\ \distinct(bn \)\ + moreover note \length(bn \) = residualLength(\ \ P')\ \distinct(bn \)\ + moreover obtain x::name where "x \ \" and "x \ P" and "x \ M" and "x \ xvec" and "x \ \" and "x \ P'" and "x \ N" + by(generate_fresh "name") auto + ultimately have "Prop \ P'" using \bn \ = []\ \xvec \* \\\xvec \* M\ \xvec \* \\ \xvec \* P'\ + apply(cases rule: semanticsCases[of _ _ _ _ _ _ _ _ _ _ _ _ _ C x x x x]) + apply(force simp add: residualInject psi.inject rInput) + apply(force simp add: residualInject psi.inject rBrInput) + by(auto simp add: residualInject psi.inject inputChainFresh)+ + } + note Goal = this + moreover obtain p :: "name prm" where "(p \ xvec) \* \" and "(p \ xvec) \* M" and "(p \ xvec) \* N" and "(p \ xvec) \* P" + and "(p \ xvec) \* \" and "(p \ xvec) \* P'" and S: "set p \ set xvec \ set(p \ xvec)" + and "distinctPerm p" + by(rule name_list_avoiding[where xvec=xvec and c="(\, M, N, P, \, P')"]) auto + from Trans \(p \ xvec) \* N\ \(p \ xvec) \* P\ S have "\ \ M\\*(p \ xvec) (p \ N)\.(p \ P) \\ \ P'" + by(simp add: inputChainAlpha') + moreover { + fix K Tvec + assume "\ \ M \ K" + moreover assume "set(p \ xvec) \ supp(p \ N)" + then have "(p \ set(p \ xvec)) \ (p \ supp(p \ N))" by simp + with \distinctPerm p\ have "set xvec \ supp N" by(simp add: eqvts) + moreover assume "length(p \ xvec) = length(Tvec::'a list)" + then have "length xvec = length Tvec" by simp + moreover assume "distinct xvec" + ultimately have "Prop (K\N[xvec::=Tvec]\) (P[xvec::=Tvec])" + by(rule rInput) + with \length xvec = length Tvec\ S \distinctPerm p\ \(p \ xvec) \* N\ \(p \ xvec) \* P\ + have "Prop (K\(p \ N)[(p \ xvec)::=Tvec]\) ((p \ P)[(p \ xvec)::=Tvec])" + by(simp add: renaming substTerm.renaming) + } + moreover { + fix K Tvec + assume "\ \ K \ M" + moreover assume "set(p \ xvec) \ supp(p \ N)" + then have "(p \ set(p \ xvec)) \ (p \ supp(p \ N))" by simp + with \distinctPerm p\ have "set xvec \ supp N" by(simp add: eqvts) + moreover assume "length(p \ xvec) = length(Tvec::'a list)" + then have "length xvec = length Tvec" by simp + moreover assume "distinct xvec" + ultimately have "Prop (\K\N[xvec::=Tvec]\) (P[xvec::=Tvec])" + by(rule rBrInput) + with \length xvec = length Tvec\ S \distinctPerm p\ \(p \ xvec) \* N\ \(p \ xvec) \* P\ + have "Prop (\K\(p \ N)[(p \ xvec)::=Tvec]\) ((p \ P)[(p \ xvec)::=Tvec])" + by(simp add: renaming substTerm.renaming) + } + moreover from Trans have "distinct xvec" by(rule inputDistinct) + then have "distinct(p \ xvec)" by simp + ultimately show ?thesis using \(p \ xvec) \* \\ \(p \ xvec) \* M\ \(p \ xvec) \* \\ \(p \ xvec) \* P'\ \distinct xvec\ + by(metis Goal) +qed + +lemma outputCases[consumes 1, case_names cOutput 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 "\ \ M\N\.P \\ \ P'" + and "\K. \ \ M \ K \ Prop (K\N\) P" + and "\K. \ \ M \ K \ Prop (\K\N\) P" + +shows "Prop \ P'" + using assms + by(cases rule: semantics.cases) (auto simp add: residualInject psi.inject) + +lemma caseCases[consumes 1, case_names cCase]: + fixes \ :: 'b + and Cs :: "('c \ ('a, 'b, 'c) psi) list" + and \ :: "'a action" + and P' :: "('a, 'b, 'c) psi" + +assumes Trans: "\ \ (Cases Cs) \ Rs" + and rCase: "\\ P. \\ \ P \ Rs; (\, P) \ set Cs; \ \ \; guarded P\ \ Prop" + +shows "Prop" + using assms + by(cases rule: semantics.cases) (auto simp add: residualInject psi.inject) + +lemma resCases[consumes 7, case_names cOpen cBrOpen cRes cBrClose]: + fixes \ :: 'b + and x :: name + 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 "x \ \" + and "x \ \" + and "x \ P'" + and "bn \ \* \" + and "bn \ \* P" + and "bn \ \* subject \" + and rOpen: "\M xvec yvec y N P'. \\ \ P \M\\*(xvec@yvec)\\([(x, y)] \ N)\ \ ([(x, y)] \ 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 rBrOpen: "\M xvec yvec y N P'. \\ \ P \\M\\*(xvec@yvec)\\([(x, y)] \ N)\ \ ([(x, y)] \ 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')" + and rBrClose: "\M xvec N P'. + \\ \ P \ \M\\*xvec\\N\ \ P'; + x \ supp M; + distinct xvec; xvec \* \; xvec \* P; + xvec \* M; + x \ \; x \ xvec; + xvec \* C\ \ Prop (\) (\\x\(\\*xvec\P'))" +shows "Prop \ P'" +proof - + from Trans have "distinct(bn \)" + by(auto dest: boundOutputDistinct) + note facts = \bn \ \* \\ \bn \ \* P\ \bn \ \* subject \\ \x \ \\ \x \ \\ \x \ P'\ \distinct(bn \)\ + have "length(bn \) = residualLength(\ \ P')" by simp + note Trans + moreover have "length [] = inputLength(\\x\P)" and "distinct []" + by(auto simp add: inputLength_inputLength'_inputLength''.simps) + moreover have "length [] = inputLength(\\x\P)" and "distinct []" + by(auto simp add: inputLength_inputLength'_inputLength''.simps) + moreover note \length(bn \) = residualLength(\ \ P')\ \distinct(bn \)\ + moreover note \length(bn \) = residualLength(\ \ P')\ \distinct(bn \)\ + moreover note \length(bn \) = residualLength(\ \ P')\ \distinct(bn \)\ + moreover note \length(bn \) = residualLength(\ \ P')\ \distinct(bn \)\ + moreover note \length(bn \) = residualLength(\ \ P')\ \distinct(bn \)\ + moreover note \length(bn \) = residualLength(\ \ P')\ \distinct(bn \)\ + moreover note \length(bn \) = residualLength(\ \ P')\ \distinct(bn \)\ + ultimately show ?thesis using facts + proof(cases rule: semanticsCases[of _ _ _ _ _ _ _ _ _ _ _ _ _ C x x x x]) + case (cOpen P M xvec y yvec N P') + moreover then have "y \ supp ([(x, y)] \ N)" using facts + apply(clarsimp simp add: psi.inject alpha abs_fresh residualInject boundOutputApp boundOutput.inject eqvts) + apply(drule pt_set_bij2[where pi="[(x, y)]", where x=x, OF pt_name_inst, OF at_name_inst]) + by(auto simp add: calc_atm eqvts fresh_def) + ultimately show ?thesis using facts + apply(clarsimp simp add: psi.inject alpha abs_fresh residualInject boundOutputApp boundOutput.inject eqvts) + by(rule rOpen) (auto simp add: residualInject boundOutputApp) + next + case (cBrOpen P M xvec y yvec N P') + moreover then have "y \ supp ([(x, y)] \ N)" using facts + apply(clarsimp simp add: psi.inject alpha abs_fresh residualInject boundOutputApp boundOutput.inject eqvts) + apply(drule pt_set_bij2[where pi="[(x, y)]", where x=x, OF pt_name_inst, OF at_name_inst]) + by(auto simp add: calc_atm eqvts fresh_def) + ultimately show ?thesis using facts + apply(clarsimp simp add: psi.inject alpha abs_fresh residualInject boundOutputApp boundOutput.inject eqvts) + by(rule rBrOpen) (auto simp add: residualInject boundOutputApp) + next + qed (auto simp add: psi.inject alpha abs_fresh residualInject boundOutputApp boundOutput.inject eqvts + intro: rScope rBrClose) +qed + +lemma resCases'[consumes 7, case_names cOpen cBrOpen cRes cBrClose]: + fixes \ :: 'b + and x :: name + 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 "x \ \" + and "x \ \" + and "x \ P'" + and "bn \ \* \" + and "bn \ \* P" + and "bn \ \* subject \" + 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 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\ \ + Prop (\M\\*(xvec@y#yvec)\\N\) P'" + and rScope: "\P'. \\ \ P \\ \ P'\ \ Prop \ (\\x\P')" + and rBrClose: "\M xvec N P'. + \\ \ P \ \M\\*xvec\\N\ \ P'; + x \ supp M; + distinct xvec; xvec \* \; xvec \* P; + xvec \* M; + x \ \; x \ xvec; + xvec \* C\ \ Prop (\) (\\x\(\\*xvec\P'))" + +shows "Prop \ P'" +proof - + from Trans have "distinct(bn \)" + by(auto dest: boundOutputDistinct) + note facts = \bn \ \* \\ \bn \ \* P\ \bn \ \* subject \\ \x \ \\ \x \ \\ \x \ P'\ \distinct(bn \)\ + have "length(bn \) = residualLength(\ \ P')" + by simp + note Trans + moreover have "length [] = inputLength(\\x\P)" and "distinct []" + by(auto simp add: inputLength_inputLength'_inputLength''.simps) + moreover have "length [] = inputLength(\\x\P)" and "distinct []" + by(auto simp add: inputLength_inputLength'_inputLength''.simps) + moreover note \length(bn \) = residualLength(\ \ P')\ \distinct(bn \)\ + moreover note \length(bn \) = residualLength(\ \ P')\ \distinct(bn \)\ + moreover note \length(bn \) = residualLength(\ \ P')\ \distinct(bn \)\ + moreover note \length(bn \) = residualLength(\ \ P')\ \distinct(bn \)\ + moreover note \length(bn \) = residualLength(\ \ P')\ \distinct(bn \)\ + moreover note \length(bn \) = residualLength(\ \ P')\ \distinct(bn \)\ + moreover note \length(bn \) = residualLength(\ \ P')\ \distinct(bn \)\ + ultimately show ?thesis using facts + proof(cases rule: semanticsCases[of _ _ _ _ _ _ _ _ _ _ _ _ _ C x x x x]) + case (cOpen P M xvec y yvec N P') + moreover then have "y \ supp ([(x, y)] \ N)" using facts + apply(clarsimp simp add: psi.inject alpha abs_fresh residualInject boundOutputApp boundOutput.inject eqvts) + apply(drule pt_set_bij2[where pi="[(x, y)]", OF pt_name_inst, OF at_name_inst]) + by(auto simp add: calc_atm eqvts fresh_def) + ultimately show ?thesis using facts + apply(clarsimp simp add: psi.inject alpha abs_fresh residualInject boundOutputApp boundOutput.inject eqvts) + apply(rule rOpen) \ \20 subgoals\ + apply(drule semantics.eqvt[where pi="[(x, y)]"]) + apply(auto simp add: eqvts residualInject boundOutputApp) + done + next + case (cBrOpen P M xvec y yvec N P') + moreover then have "y \ supp ([(x, y)] \ N)" using facts + apply(clarsimp simp add: psi.inject alpha abs_fresh residualInject boundOutputApp boundOutput.inject eqvts) + apply(drule pt_set_bij2[where pi="[(x, y)]", OF pt_name_inst, OF at_name_inst]) + by(auto simp add: calc_atm eqvts fresh_def) + ultimately show ?thesis using facts + apply(clarsimp simp add: psi.inject alpha abs_fresh residualInject boundOutputApp boundOutput.inject eqvts) + apply(rule rBrOpen) \ \20 subgoals\ + apply(drule semantics.eqvt[where pi="[(x, y)]"]) + apply(auto simp add: eqvts residualInject boundOutputApp) + done + qed (auto simp add: psi.inject alpha abs_fresh residualInject boundOutputApp boundOutput.inject eqvts + intro: rScope rBrClose) +qed + +lemma resInputCases'[consumes 4, case_names cRes]: + fixes \ :: 'b + and x :: name + and M :: 'a + and N :: 'a + and P :: "('a, 'b, 'c) psi" + and R :: "('a, 'b, 'c) psi" + and C :: "'f::fs_name" + +assumes Trans: "\ \ \\x\P \M\N\ \ R" + and 2: "x \ \" + and "x \ (M\N\)" + and 4: "x \ R" + and rScope: "\P'. \\ \ P \M\N\ \ P'\ \ Prop (\\x\P')" + +shows "Prop R" +proof - + from Trans obtain \ where 1: "\ \ \\x\P \\ \ R" and 5: "bn \ \* \" and 6: "bn \ \* P" and 7: "bn \ \* subject \" and "\ = M\N\" by auto + from \x \ (M\N\)\ \\ = (M\N\)\ have 3: "x \ \" by simp + show ?thesis using 1 2 3 4 5 6 7 \\ = M\N\\ rScope + by(induct rule: resCases') (auto simp add: residualInject) +qed + +lemma resBrInputCases'[consumes 4, case_names cRes]: + fixes \ :: 'b + and x :: name + and M :: 'a + and N :: 'a + and P :: "('a, 'b, 'c) psi" + and R :: "('a, 'b, 'c) psi" + and C :: "'f::fs_name" + +assumes Trans: "\ \ \\x\P \\M\N\ \ R" + and 2: "x \ \" + and "x \ (\M\N\)" + and 4: "x \ R" + and rScope: "\P'. \\ \ P \\M\N\ \ P'\ \ Prop (\\x\P')" + +shows "Prop R" +proof - + from Trans obtain \ where 1: "\ \ \\x\P \\ \ R" and 5: "bn \ \* \" and 6: "bn \ \* P" and 7: "bn \ \* subject \" and "\ = \M\N\" by auto + from \x \ (\M\N\)\ \\ = (\M\N\)\ have 3: "x \ \" by simp + show ?thesis using 1 2 3 4 5 6 7 \\ = \M\N\\ rScope + by(induct rule: resCases') (auto simp add: residualInject) +qed + +lemma resOutputCases'''[consumes 7, case_names cOpen cRes]: + fixes \ :: 'b + and x :: name + and M :: 'a + and N :: 'a + and P :: "('a, 'b, 'c) psi" + and R :: "('a, 'b, 'c) psi" + and C :: "'f::fs_name" + +assumes Trans: "\ \ \\x\P \M\\*(zvec1 @ zvec2)\\N\ \ R" + and 1: "x \ \" + and "x \ (M\\*(zvec1 @ zvec2)\\N\)" + and 3: "x \ R" + and "(zvec1 @ zvec2) \* \" + and "(zvec1 @ zvec2) \* P" + and "(zvec1 @ zvec2) \* 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\ \ + Prop P'" + and rScope: "\P'. \\ \ P \M\\*(zvec1 @ zvec2)\\N\ \ P'\ \ Prop (\\x\P')" + +shows "Prop R" +proof - + from Trans have "distinct (zvec1 @ zvec2)" by(auto dest: boundOutputDistinct) + obtain \ where "\=M\\*(zvec1 @ zvec2)\\N\" by simp + with Trans \(zvec1 @ zvec2) \* \\ \(zvec1 @ zvec2) \* P\ \(zvec1 @ zvec2) \* M\ + have \Trans: "\ \ \\x\P \\ \ R" and 4: "bn \ \* \" and 5: "bn \ \* P" and 6: "bn \ \* subject \" + by simp+ + from \x \ (M\\*(zvec1 @ zvec2)\\N\)\ \\=M\\*(zvec1 @ zvec2)\\N\\ have 2: "x \ \" by simp + show ?thesis using \Trans 1 2 3 4 5 6 \\=M\\*(zvec1 @ zvec2)\\N\\ rOpen rScope + proof(induct rule: resCases'[where C="(zvec1, zvec2, 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') + then show ?case + by(auto simp add: residualInject boundOutputApp) + qed +qed + +lemma resOutputCases''[consumes 7, case_names cOpen cRes]: + fixes \ :: 'b + and x :: name + and z :: name + and M :: 'a + and N :: 'a + and P :: "('a, 'b, 'c) psi" + and R :: "('a, 'b, 'c) psi" + and C :: "'f::fs_name" + +assumes Trans: "\ \ \\x\P \M\\*(zvec1@y#zvec2)\\N\ \ R" + and 1: "x \ \" + and "x \ (M\\*(zvec1@y#zvec2)\\N\)" + and 3: "x \ R" + and "(zvec1@y#zvec2) \* \" + and "(zvec1@y#zvec2) \* P" + and "(zvec1@y#zvec2) \* 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\ \ + Prop P'" + and rScope: "\P'. \\ \ P \M\\*(zvec1@y#zvec2)\\N\ \ P'\ \ Prop (\\x\P')" + +shows "Prop R" +proof - + from Trans have "distinct (zvec1@y#zvec2)" by(auto dest: boundOutputDistinct) + obtain \ where "\=M\\*(zvec1@y#zvec2)\\N\" by simp + with Trans \(zvec1@y#zvec2) \* \\ \(zvec1@y#zvec2) \* P\ \(zvec1@y#zvec2) \* M\ + have \Trans: "\ \ \\x\P \\ \ R" and 4: "bn \ \* \" and 5: "bn \ \* P" and 6: "bn \ \* subject \" + by simp+ + from \x \ (M\\*(zvec1@y#zvec2)\\N\)\ \\=M\\*(zvec1@y#zvec2)\\N\\ have 2: "x \ \" by simp + show ?thesis using \Trans 1 2 3 4 5 6 \\=M\\*(zvec1@y#zvec2)\\N\\ rOpen rScope + proof(induct rule: resCases'[where C="(zvec1, zvec2, z, 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') + then show ?case + by(auto simp add: residualInject boundOutputApp) + qed +qed + +abbreviation + statImpJudge ("_ \ _" [80, 80] 80) + where "\ \ \' \ AssertionStatImp \ \'" + +lemma statEqTransition: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and Rs :: "('a, 'b, 'c) residual" + and \' :: 'b + +assumes "\ \ P \ Rs" + and "\ \ \'" + +shows "\' \ P \ Rs" + using assms +proof(nominal_induct avoiding: \' rule: semantics.strong_induct) + case(cInput \ M K xvec N Tvec P \') + from \\ \ \'\ \\ \ M \ K\ have "\' \ M \ K" + by(simp add: AssertionStatImp_def AssertionStatEq_def) + then show ?case using \distinct xvec\ \set xvec \ supp N\ \length xvec = length Tvec\ + by(rule Input) +next + case(cBrInput \ K M xvec N Tvec P \') + from \\ \ \'\ \\ \ K \ M\ have "\' \ K \ M" + by(simp add: AssertionStatImp_def AssertionStatEq_def) + then show ?case using \distinct xvec\ \set xvec \ supp N\ \length xvec = length Tvec\ + by(rule BrInput) +next + case(Output \ M K N P \') + from \\ \ \'\ \\ \ M \ K\ have "\' \ M \ K" + by(simp add: AssertionStatImp_def AssertionStatEq_def) + then show ?case by(rule semantics.Output) +next + case(BrOutput \ M K N P \') + from \\ \ \'\ \\ \ M \ K\ have "\' \ M \ K" + by(simp add: AssertionStatImp_def AssertionStatEq_def) + then show ?case by(rule semantics.BrOutput) +next + case(Case \ P Rs \ Cs \') + then have "\' \ P \ Rs" by(intro Case) + moreover note \(\, P) \ set Cs\ + moreover from \\ \ \'\ \\ \ \\ have "\' \ \" + by(simp add: AssertionStatImp_def AssertionStatEq_def) + ultimately show ?case using \guarded P\ by(rule semantics.Case) +next + case(cPar1 \ \Q P \ P' xvec Q \') + then show ?case + by(intro Par1) (auto intro: Composition) +next + case(cPar2 \ \P Q \ Q' xvec P \') + then show ?case + by(intro Par2) (auto intro: Composition) +next + case(cComm1 \ \Q P M N P' xvec \P Q K zvec Q' yvec \') + then show ?case + by(clarsimp, intro Comm1) (blast intro: Composition statEqEnt)+ +next + case(cComm2 \ \Q P M zvec N P' xvec \P Q K Q' yvec \') + then show ?case + by(clarsimp, intro Comm2) (blast intro: Composition statEqEnt)+ +next + case(cBrMerge \ \\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q Q' A\<^sub>Q \') + then show ?case + by(clarsimp, intro BrMerge) (blast intro: Composition)+ +next + case(cBrComm1 \ \\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q xvec Q' A\<^sub>Q \') + then show ?case + by(clarsimp, intro BrComm1) (blast intro: Composition)+ +next + case(cBrComm2 \ \\<^sub>Q P M xvec N P' A\<^sub>P \\<^sub>P Q Q' A\<^sub>Q \') + then show ?case + by(clarsimp, intro BrComm2) (blast intro: Composition)+ +next + case(cBrClose \ P M xvec N P' x \') + then show ?case by(force intro: BrClose) +next + case(cOpen \ P M xvec N P' x yvec \') + then show ?case by(force intro: Open) +next + case(cBrOpen \ P M xvec N P' x yvec \') + then show ?case by(force intro: BrOpen) +next + case(cScope \ P \ P' x \') + then show ?case by(force intro: Scope) +next + case(Bang \ P Rs \') + then show ?case by(force intro: semantics.Bang) +qed + +lemma brInputTermSupp: + fixes \ :: "'b::fs_name" + and P :: "('a, 'b, ('c::fs_name)) psi" + and P' :: "('a, 'b, 'c) psi" + and N :: "'a::fs_name" + and K :: "'a::fs_name" + +assumes "\ \ P \ \K\N\ \ P'" + +shows "(supp K) \ ((supp P)::name set)" + using assms +proof(nominal_induct rule: brInputInduct) + case(cBrInput \ K M xvec N Tvec P) + from \\ \ K \ M\ have "(supp K) \ ((supp M)::name set)" + by(simp add: chanInConSupp) + then show ?case + by (metis Un_commute Un_upper2 psi.supp(3) subset_trans) +next + case(cCase \ P M N P' \ Cs) + then have "supp M \ ((supp P)::name set)" + by simp + from \(\, P) \ set Cs\ + have "{(\, P)} \ set Cs" + by simp + moreover have "finite {(\, P)}" by simp + moreover have "finite (set Cs)" by simp + ultimately have "((supp {(\, P)})::name set) \ ((supp (set Cs))::name set)" + by(simp add: supp_subset) + + moreover have "supp {(\, P)} = ((supp (\, P))::name set)" + by (meson supp_singleton) + moreover have "supp P \ supp (\, P)" + by (metis Un_upper2 supp_prod) + ultimately have "((supp P)::name set) \ ((supp Cs)::name set)" + by (auto simp add: supp_list_set) + + moreover have "((supp Cs)::name set) = ((supp (Cases Cs))::name set)" + unfolding psi.supp + apply(induct rule: psiCases.induct) + apply(metis psiCase.supp(1) psiCases.simps(1) set_empty2 supp_list_nil) + by simp (metis Un_assoc psiCase.supp(2) supp_list_cons supp_prod) + + ultimately have "((supp P)::name set) \ ((supp (Cases Cs))::name set)" + by simp + + with \supp M \ supp P\ + show ?case by simp +next + case(cPar1 \ \\<^sub>Q P M N P' A\<^sub>Q Q) + then have "((supp M)::name set) \ ((supp P)::name set)" + by auto + then show ?case + by(auto simp add: psi.supp) +next + case(cPar2 \ \\<^sub>P Q M N Q' A\<^sub>P P) + then have "((supp M)::name set) \ ((supp Q)::name set)" + by auto + then show ?case + by(auto simp add: psi.supp) +next + case(cBrMerge \ \\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q Q' A\<^sub>Q) + then show ?case + by(auto simp add: psi.supp) +next + case(cScope \ P M N P' x) + then have "((supp M)::name set) \ ((supp P)::name set)" + by simp + then show ?case + by(simp add: psi.supp) (metis abs_supp cScope.hyps fresh_def insert_Diff_single subset_insert_iff) +next + case(cBang \ P M N P') + then show ?case + by(simp add: psi.supp) +qed + +lemma brOutputTermSupp: + fixes \ :: "'b::fs_name" + and P :: "('a, 'b, ('c::fs_name)) psi" + and P' :: "('a, 'b, 'c) psi" + and N :: "'a::fs_name" + and K :: "'a::fs_name" + and xvec :: "name list" + +assumes "\ \ P \ RBrOut K (\\*xvec\N \' P')" + +shows "(supp K) \ ((supp P)::name set)" + using assms +proof(nominal_induct rule: brOutputInduct) + case(cBrOutput \ M K N P) + from \\ \ M \ K\ have "(supp K) \ ((supp M)::name set)" + by(simp add: chanOutConSupp) + then show ?case + by (metis Un_commute Un_upper2 psi.supp(2) subset_iff_psubset_eq subset_trans) +next + case(cCase \ P M B \ Cs) + then have "supp M \ ((supp P)::name set)" + by simp + from \(\, P) \ set Cs\ + have "{(\, P)} \ set Cs" + by simp + moreover have "finite {(\, P)}" by simp + moreover have "finite (set Cs)" by simp + ultimately have "((supp {(\, P)})::name set) \ ((supp (set Cs))::name set)" + by(simp add: supp_subset) + + moreover have "supp {(\, P)} = ((supp (\, P))::name set)" + by(meson supp_singleton) + moreover have "supp P \ supp (\, P)" + by (metis Un_upper2 supp_prod) + ultimately have "((supp P)::name set) \ ((supp Cs)::name set)" + by (auto simp add: supp_list_set) + + moreover have "((supp Cs)::name set) = ((supp (Cases Cs))::name set)" + unfolding psi.supp + apply(induct rule: psiCases.induct) + apply(metis psiCase.supp(1) psiCases.simps(1) set_empty2 supp_list_nil) + by simp (metis Un_assoc psiCase.supp(2) supp_list_cons supp_prod) + + ultimately have "((supp P)::name set) \ ((supp (Cases Cs))::name set)" + by simp + + with \supp M \ supp P\ + show ?case by simp +next + case(cPar1 \ \\<^sub>Q P M xvec N P' A\<^sub>Q Q) + then have "((supp M)::name set) \ ((supp P)::name set)" + by auto + then show ?case + by(auto simp add: psi.supp) +next + case(cPar2 \ \\<^sub>P Q M xvec N Q' A\<^sub>P P) + then have "((supp M)::name set) \ ((supp Q)::name set)" + by auto + then show ?case + by(auto simp add: psi.supp) +next + case(cBrComm1 \ \\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q xvec Q' A\<^sub>Q) + then have "((supp M)::name set) \ ((supp Q)::name set)" + by auto + then show ?case + by(auto simp add: psi.supp) +next + case(cBrComm2 \ \\<^sub>Q P M xvec N P' A\<^sub>P \\<^sub>P Q Q' A\<^sub>Q) + then have "((supp M)::name set) \ ((supp P)::name set)" + by auto + then show ?case + by(auto simp add: psi.supp) +next + case(cBrOpen \ P M xvec yvec N P' x) + then have "((supp M)::name set) \ ((supp P)::name set)" + by simp + with \x \ M\ + show ?case + by(simp add: psi.supp) (metis abs_supp cBrOpen.hyps fresh_def insert_Diff_single subset_insert_iff) +next + case(cScope \ P M xvec N P' x) + then have "((supp M)::name set) \ ((supp P)::name set)" + by simp + then show ?case + by(simp add: psi.supp) (metis abs_supp cScope.hyps fresh_def insert_Diff_single subset_insert_iff) +next + case cBang + then show ?case + by(simp add: psi.supp) +qed + +lemma actionPar1Dest': + fixes \ :: "('a::fs_name) action" + and P :: "('a, 'b::fs_name, 'c::fs_name) psi" + and \ :: "('a::fs_name) action" + and Q :: "('a, 'b, 'c) psi" + and R :: "('a, 'b, 'c) psi" + +assumes "\ \ P = \ \ (Q \ R)" + and "bn \ \* R" + and "bn \ \* R" + +obtains T where "P = T \ R" and "\ \ T = \ \ Q" + using assms + apply(cases rule: actionCases[where \=\]) + apply (metis residualInject'(1)) + apply (metis residualInject'(7)) + apply (smt (z3) bn.simps(3) boundOutputPar1Dest create_residual.simps(3) residualInject'(8)) + apply (smt (z3) bn.simps(4) boundOutputPar1Dest create_residual.simps(4) residualInject'(9)) + by (metis residualInject'(10)) + +lemma actionPar2Dest': + fixes \ :: "('a::fs_name) action" + and P :: "('a, 'b::fs_name, 'c::fs_name) psi" + and \ :: "('a::fs_name) action" + and Q :: "('a, 'b, 'c) psi" + and R :: "('a, 'b, 'c) psi" + +assumes "\ \ P = \ \ (Q \ R)" + and "bn \ \* Q" + and "bn \ \* Q" + +obtains T where "P = Q \ T" and "\ \ T = \ \ R" + using assms + apply(cases rule: actionCases[where \=\]) + apply (metis residualInject'(1)) + apply (metis residualInject'(7)) + apply (smt (z3) bn.simps(3) boundOutputPar2Dest create_residual.simps(3) residualInject'(8)) + apply (smt (z3) bn.simps(4) boundOutputPar2Dest create_residual.simps(4) residualInject'(9)) + by (metis residualInject'(10)) + +lemma expandNonTauFrame: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and \ :: "'a action" + and P' :: "('a, 'b, 'c) psi" + and A\<^sub>P :: "name list" + and \\<^sub>P :: 'b + and C :: "'f::fs_name" + and C' :: "'g::fs_name" + +assumes Trans: "\ \ P \\ \ P'" + and "extractFrame P = \A\<^sub>P, \\<^sub>P\" + and "distinct A\<^sub>P" + and "bn \ \* subject \" + and "A\<^sub>P \* P" + and "A\<^sub>P \* \" + and "A\<^sub>P \* C" + and "A\<^sub>P \* C'" + and "bn \ \* P" + and "bn \ \* C'" + and "\ \ \" + +obtains p \' A\<^sub>P' \\<^sub>P' where "set p \ set(bn \) \ set(bn(p \ \))" and "(p \ \\<^sub>P) \ \' \ \\<^sub>P'" and "distinctPerm p" and + "extractFrame P' = \A\<^sub>P', \\<^sub>P'\" and "A\<^sub>P' \* P'" and "A\<^sub>P' \* \" and "A\<^sub>P' \* (p \ \)" and + "A\<^sub>P' \* C" and "(bn(p \ \)) \* C'" and "(bn(p \ \)) \* \" and "(bn(p \ \)) \* P'" and "distinct A\<^sub>P'" +proof - + assume A: "\p \' \\<^sub>P' A\<^sub>P'. + \set p \ set(bn \) \ set(bn(p \ \)); (p \ \\<^sub>P) \ \' \ \\<^sub>P'; distinctPerm p; + extractFrame P' = \A\<^sub>P', \\<^sub>P'\; A\<^sub>P' \* P'; A\<^sub>P' \* \; A\<^sub>P' \* (p \ \); + A\<^sub>P' \* C; (bn(p \ \)) \* C'; (bn(p \ \)) \* \; (bn(p \ \)) \* P'; distinct A\<^sub>P'\ + \ thesis" + + from Trans have "distinct(bn \)" by(auto dest: boundOutputDistinct) + + with Trans \bn \ \* subject \\ \A\<^sub>P \* P\ \A\<^sub>P \* \\ have "A\<^sub>P \* P'" + by(drule_tac freeFreshChainDerivative) auto + + { + fix V :: "'a list" + and W :: "('a action) list" + and X :: "name list" + and Y :: "'b list" + and Z :: "('a, 'b, 'c) psi list" + + assume "bn \ \* V" and "bn \ \* W" and "bn \ \* X" and "bn \ \* Y" and "bn \ \* Z" and "A\<^sub>P \* V" and "A\<^sub>P \* W" and "A\<^sub>P \* X" and "A\<^sub>P \* Y" and "A\<^sub>P \* Z" + + with assms obtain p \' A\<^sub>P' \\<^sub>P' where "set p \ set(bn \) \ set(bn(p \ \))" and "(p \ \\<^sub>P) \ \' \ \\<^sub>P'" and "distinctPerm p" + and "extractFrame P' = \A\<^sub>P', \\<^sub>P'\" and "A\<^sub>P' \* P'" and "A\<^sub>P' \* \" and "A\<^sub>P' \* (p \ \)" + and "A\<^sub>P' \* C" and "(bn(p \ \)) \* C'" and "(bn(p \ \)) \* \" and "(bn(p \ \)) \* P'" + and "A\<^sub>P' \* V" and "A\<^sub>P' \* W" and "A\<^sub>P' \* X" and "A\<^sub>P' \* Y" and "A\<^sub>P' \* Z" and "distinct A\<^sub>P'" + and "(bn(p \ \)) \* V" and "(bn(p \ \)) \* W" and "(bn(p \ \)) \* X" and "(bn(p \ \)) \* Y" and "(bn(p \ \)) \* Z" + using \A\<^sub>P \* P'\ \distinct(bn \)\ + proof(nominal_induct \ P Rs=="\ \ P'" A\<^sub>P \\<^sub>P avoiding: C C' \ P' V W X Y Z arbitrary: thesis rule: semanticsFrameInduct) + case(cAlpha \ P A\<^sub>P \\<^sub>P p C C' \ P' V W X Y Z) + then obtain q \' A\<^sub>P' \\<^sub>P' where Sq: "set q \ set(bn \) \ set(bn(q \ \))" and PeqP': "(q \ \\<^sub>P) \ \' \ \\<^sub>P'" and "distinctPerm q" + and FrP': "extractFrame P' = \A\<^sub>P', \\<^sub>P'\" and "A\<^sub>P' \* P'" and "A\<^sub>P' \* \" and "A\<^sub>P' \* (q \ \)" + and "A\<^sub>P' \* C" and "(bn(q \ \)) \* C'" and "(bn(q \ \)) \* \" and "(bn(q \ \)) \* P'" + and "A\<^sub>P' \* V" and "A\<^sub>P' \* W" and "A\<^sub>P' \* X" and "A\<^sub>P' \* Y" and "A\<^sub>P' \* Z" and "distinct A\<^sub>P'" + and "(bn(q \ \)) \* V" and "(bn(q \ \)) \* W" and "(bn(q \ \)) \* X" and "(bn(q \ \)) \* Y" and "(bn(q \ \)) \* Z" + by metis + + have Sp: "set p \ set A\<^sub>P \ set (p \ A\<^sub>P)" by fact + + from Sq have "(p \ set q) \ p \ (set(bn \) \ set(bn(q \ \)))" + by(simp add: subsetClosed) + then have "set(p \ q) \ set(bn(p \ \)) \ set(p \ bn(q \ \))" + by(simp add: eqvts) + with \A\<^sub>P \* \\ \(p \ A\<^sub>P) \* \\ Sp have "set(p \ q) \ set(bn \) \ set(bn((p \ q) \ \))" + by(simp add: perm_compose bnEqvt[symmetric]) + moreover from PeqP' have "(p \ (q \ \\<^sub>P) \ \') \ (p \ \\<^sub>P')" + by(simp add: AssertionStatEqClosed) + then have "((p \ q) \ p \ \\<^sub>P) \ (p \ \') \ (p \ \\<^sub>P')" + apply(subst perm_compose[symmetric]) + by(simp add: eqvts) + moreover from \distinctPerm q\ have "distinctPerm (p \ q)" + by simp + moreover from \(bn(q \ \)) \* C'\ have "(p \ bn(q \ \)) \* (p \ C')" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>P \* \\ \(p \ A\<^sub>P) \* \\ \A\<^sub>P \* C'\ \(p \ A\<^sub>P) \* C'\ Sp have "bn((p \ q) \ \) \* C'" + by(simp add: perm_compose bnEqvt[symmetric]) + moreover from FrP' have "(p \ extractFrame P') = p \ \A\<^sub>P', \\<^sub>P'\" by simp + with \A\<^sub>P \* P'\ \(p \ A\<^sub>P) \* P'\ Sp have "extractFrame P' = \p \ A\<^sub>P', p \ \\<^sub>P'\" + by(simp add: eqvts) + moreover from \A\<^sub>P' \* P'\ have "(p \ A\<^sub>P') \* (p \ P')" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>P \* P'\ \(p \ A\<^sub>P) \* P'\ Sp have "(p \ A\<^sub>P') \* P'" by simp + moreover from \A\<^sub>P' \* \\ have "(p \ A\<^sub>P') \* (p \ \)" + by(simp only: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>P \* \\ \(p \ A\<^sub>P) \* \\ Sp have "(p \ A\<^sub>P') \* \" by simp + moreover from \A\<^sub>P' \* C\ have "(p \ A\<^sub>P') \* (p \ C)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>P \* C\ \(p \ A\<^sub>P) \* C\ Sp have "(p \ A\<^sub>P') \* C" by simp + moreover from \(bn(q \ \)) \* \\ have "(p \ bn(q \ \)) \* (p \ \)" + by(simp only: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>P \* \\ \(p \ A\<^sub>P) \* \\ \A\<^sub>P \* \\ \(p \ A\<^sub>P) \* \\ Sp have "bn((p \ q) \ \) \* \" + by(simp add: perm_compose eqvts) + moreover from \(bn(q \ \)) \* P'\ have "(p \ bn(q \ \)) \* (p \ P')" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>P \* \\ \(p \ A\<^sub>P) \* \\ \A\<^sub>P \* P'\ \(p \ A\<^sub>P) \* P'\ Sp have "bn((p \ q) \ \) \* P'" + by(simp add: perm_compose eqvts) + moreover from \A\<^sub>P' \* (q \ \)\ have "(p \ A\<^sub>P') \* (p \ q \ \)" + by(simp only: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with Sp \A\<^sub>P \* \\ \(p \ A\<^sub>P) \* \\ have "(p \ A\<^sub>P') \* ((p \ q) \ \)" + by(simp add: perm_compose) + moreover from \A\<^sub>P' \* V\ have "(p \ A\<^sub>P') \* (p \ V)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>P \* V\ \(p \ A\<^sub>P) \* V\ Sp have "(p \ A\<^sub>P') \* V" by simp + moreover from \A\<^sub>P' \* W\ have "(p \ A\<^sub>P') \* (p \ W)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>P \* W\ \(p \ A\<^sub>P) \* W\ Sp have "(p \ A\<^sub>P') \* W" by simp + moreover from \A\<^sub>P' \* X\ have "(p \ A\<^sub>P') \* (p \ X)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>P \* X\ \(p \ A\<^sub>P) \* X\ Sp have "(p \ A\<^sub>P') \* X" by simp + moreover from \A\<^sub>P' \* Y\ have "(p \ A\<^sub>P') \* (p \ Y)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>P \* Y\ \(p \ A\<^sub>P) \* Y\ Sp have "(p \ A\<^sub>P') \* Y" by simp + moreover from \A\<^sub>P' \* Z\ have "(p \ A\<^sub>P') \* (p \ Z)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>P \* Z\ \(p \ A\<^sub>P) \* Z\ Sp have "(p \ A\<^sub>P') \* Z" by simp + moreover from \(bn(q \ \)) \* V\ have "(p \ bn(q \ \)) \* (p \ V)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>P \* \\ \(p \ A\<^sub>P) \* \\ \A\<^sub>P \* V\ \(p \ A\<^sub>P) \* V\ Sp have "bn((p \ q) \ \) \* V" + by(simp add: perm_compose eqvts) + moreover from \(bn(q \ \)) \* W\ have "(p \ bn(q \ \)) \* (p \ W)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>P \* \\ \(p \ A\<^sub>P) \* \\ \A\<^sub>P \* W\ \(p \ A\<^sub>P) \* W\ Sp have "bn((p \ q) \ \) \* W" + by(simp add: perm_compose eqvts) + moreover from \(bn(q \ \)) \* X\ have "(p \ bn(q \ \)) \* (p \ X)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>P \* \\ \(p \ A\<^sub>P) \* \\ \A\<^sub>P \* X\ \(p \ A\<^sub>P) \* X\ Sp have "bn((p \ q) \ \) \* X" + by(simp add: perm_compose eqvts) + moreover from \(bn(q \ \)) \* Y\ have "(p \ bn(q \ \)) \* (p \ Y)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>P \* \\ \(p \ A\<^sub>P) \* \\ \A\<^sub>P \* Y\ \(p \ A\<^sub>P) \* Y\ Sp have "bn((p \ q) \ \) \* Y" + by(simp add: perm_compose eqvts) + moreover from \(bn(q \ \)) \* Z\ have "(p \ bn(q \ \)) \* (p \ Z)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>P \* \\ \(p \ A\<^sub>P) \* \\ \A\<^sub>P \* Z\ \(p \ A\<^sub>P) \* Z\ Sp have "bn((p \ q) \ \) \* Z" + by(simp add: perm_compose eqvts) + moreover from \distinct A\<^sub>P'\ have "distinct(p \ A\<^sub>P')" by simp + ultimately show ?case + by(elim cAlpha) + next + case(cInput \ M K xvec N Tvec P C C' \ P' V W X Y Z) + moreover obtain A\<^sub>P \\<^sub>P where "extractFrame(P[xvec::=Tvec]) = \A\<^sub>P, \\<^sub>P\" and "distinct A\<^sub>P" + and "A\<^sub>P \* (C, P[xvec::=Tvec], \, P', V, W, X, Y, Z, N)" + by(rule freshFrame) + moreover have "\ \ \\<^sub>P \ \\<^sub>P" + by(blast intro: Identity Commutativity AssertionStatEqTrans) + ultimately show ?case + by(intro cInput) (assumption | simp add: residualInject)+ + next + case(cBrInput \ M K xvec N Tvec P C C' \ P' V W X Y Z) + moreover obtain A\<^sub>P \\<^sub>P where "extractFrame(P[xvec::=Tvec]) = \A\<^sub>P, \\<^sub>P\" and "distinct A\<^sub>P" + and "A\<^sub>P \* (C, P[xvec::=Tvec], \, P', V, W, X, Y, Z, N)" + by(rule freshFrame) + moreover have "\ \ \\<^sub>P \ \\<^sub>P" + by(blast intro: Identity Commutativity AssertionStatEqTrans) + ultimately show ?case + by(intro cBrInput) (assumption | simp add: residualInject)+ + next + case(cOutput \ M K N P C C' \ P' V W X Y Z) + moreover obtain A\<^sub>P \\<^sub>P where "extractFrame P = \A\<^sub>P, \\<^sub>P\" and "distinct A\<^sub>P" + and "A\<^sub>P \* (C, C', P, \, N, P', V, W, X, Y, Z)" + by(rule freshFrame) + moreover have "\ \ \\<^sub>P \ \\<^sub>P" + by(blast intro: Identity Commutativity AssertionStatEqTrans) + ultimately show ?case by(simp add: residualInject) + next + case(cBrOutput \ M K N P C C' \ P' V W X Y Z) + moreover obtain A\<^sub>P \\<^sub>P where "extractFrame P = \A\<^sub>P, \\<^sub>P\" and "distinct A\<^sub>P" + and "A\<^sub>P \* (C, C', P, \, N, P', V, W, X, Y, Z)" + by(rule freshFrame) + moreover have "\ \ \\<^sub>P \ \\<^sub>P" + by(blast intro: Identity Commutativity AssertionStatEqTrans) + ultimately show ?case by(simp add: residualInject) + next + case(cCase \ P \ Cs A\<^sub>P \\<^sub>P C C' \ P' V W X Y Z) + moreover from \bn \ \* (Cases Cs)\ \(\, P) \ set Cs\ have "bn \ \* P" by(auto dest: memFreshChain) + ultimately obtain p \' A\<^sub>P' \\<^sub>P' where S: "set p \ set(bn \) \ set(bn(p \ \))" + and FrP': "extractFrame P' = \A\<^sub>P', \\<^sub>P'\" + and PeqP': "(p \ \\<^sub>P) \ \' \ \\<^sub>P'" and "distinct A\<^sub>P'" + and "A\<^sub>P' \* C" and "A\<^sub>P' \* P'" and "A\<^sub>P' \* \" and "A\<^sub>P' \* (p \ \)" + and "A\<^sub>P' \* V" and "A\<^sub>P' \* W" and "A\<^sub>P' \* X" and "A\<^sub>P' \* Y" and "A\<^sub>P' \* Z" + and "distinctPerm p" and "(bn(p \ \)) \* \" and "(bn(p \ \)) \* P'" + and "(bn(p \ \)) \* C'" and "(bn(p \ \)) \* V" and "(bn(p \ \)) \* W" and "(bn(p \ \)) \* X" and "(bn(p \ \)) \* Y" and "(bn(p \ \)) \* Z" + apply - + apply (rule cCase) + apply (assumption | simp (no_asm_use))+ + done + moreover from \\\<^sub>P \ \\ have "(p \ \\<^sub>P) \ (p \ \)" + by(simp add: AssertionStatEqClosed) + then have "(p \ \\<^sub>P) \ \" by(simp add: permBottom) + with PeqP' have "(\ \ \') \ \\<^sub>P'" + by(metis Identity AssertionStatEqTrans composition' Commutativity Associativity AssertionStatEqSym) + ultimately show ?case using cCase \bn \ \* P\ + by(intro cCase(22)) (assumption | simp)+ + next + case(cPar1 \ \\<^sub>Q P \ P' A\<^sub>Q Q A\<^sub>P \\<^sub>P C C' \' PQ' V W X Y Z) + have FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" and FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" + by fact+ + + note \bn \' \* subject \'\ + moreover from \bn \' \* (P \ Q)\ have "bn \' \* P" and "bn \' \* Q" by simp+ + moreover with FrP FrQ \A\<^sub>P \* \'\ \A\<^sub>Q \* \'\ have "bn \' \* \\<^sub>P" and "bn \' \* \\<^sub>Q" + by(force dest: extractFrameFreshChain)+ + + moreover note \bn \' \* V\ \bn \' \* W\ + moreover from \bn \' \* X\ \A\<^sub>Q \* \'\ have "bn \' \* (X@A\<^sub>Q)" by simp + moreover from \bn \' \* Y\ \bn \' \* \\<^sub>Q\ have "bn \' \* (\\<^sub>Q#Y)" by simp + moreover from \bn \' \* Z\ \bn \' \* Q\ have "bn \' \* (Q#Z)" by simp + moreover note \A\<^sub>P \* V\ \A\<^sub>P \* W\ + moreover from \A\<^sub>P \* X\ \A\<^sub>P \* A\<^sub>Q\ have "A\<^sub>P \* (X@A\<^sub>Q)" by simp + moreover from \A\<^sub>P \* Y\ \A\<^sub>P \* \\<^sub>Q\ have "A\<^sub>P \* (\\<^sub>Q#Y)" by force + moreover from \A\<^sub>P \* Z\ \A\<^sub>P \* Q\ have "A\<^sub>P \* (Q#Z)" by simp + moreover from \\ \ (P' \ Q) = \' \ PQ'\ \bn \ \* Q\ \bn \' \* Q\ \bn \ \* \'\ + obtain P'' where A: "\ \ P' = \' \ P''" and "PQ' = P'' \ Q" + by(metis actionPar1Dest') + moreover from \A\<^sub>P \* PQ'\ \PQ' = P'' \ Q\ have "A\<^sub>P \* P''" by simp + ultimately obtain p \' \\<^sub>P' A\<^sub>P' where S: "set p \ set(bn \') \ set (bn(p \ \'))" and PeqP': "((p \ \\<^sub>P) \ \') \ \\<^sub>P'" + and "distinctPerm p" and "(bn(p \ \')) \* C'" and FrP': "extractFrame P'' = \A\<^sub>P', \\<^sub>P'\" + and "A\<^sub>P' \* P''" and "A\<^sub>P' \* \'" "A\<^sub>P' \* (p \ \')" and "A\<^sub>P' \* C" + and "(bn(p \ \')) \* \'" and "(bn(p \ \')) \* P''" and "distinct A\<^sub>P'" + and "A\<^sub>P' \* V" and "A\<^sub>P' \* W" and "A\<^sub>P' \* (X @ A\<^sub>Q)" and "A\<^sub>P' \* (\\<^sub>Q#Y)" + and "A\<^sub>P' \* (Q#Z)" and "(bn(p \ \')) \* V" and "(bn(p \ \')) \* W" and "(bn(p \ \')) \* (X @ A\<^sub>Q)" and "(bn(p \ \')) \* (\\<^sub>Q#Y)" + and "(bn(p \ \')) \* (Q#Z)" using cPar1 + by(elim cPar1) + + then have "A\<^sub>P' \* Q" and "A\<^sub>P' \* Z" and "A\<^sub>P' \* A\<^sub>Q" and "A\<^sub>P' \* X" and "A\<^sub>P' \* \\<^sub>Q" and "A\<^sub>P' \* Y" + and "(bn(p \ \')) \* A\<^sub>Q" and "(bn(p \ \')) \* X" and "(bn(p \ \')) \* Y" and "(bn(p \ \')) \* Z" and "(bn(p \ \')) \* \\<^sub>Q" + and "(bn(p \ \')) \* Q" + by(simp del: freshChainSimps)+ + + from \A\<^sub>Q \* PQ'\ \PQ' = P'' \ Q\ \A\<^sub>P' \* A\<^sub>Q\ FrP' have "A\<^sub>Q \* \\<^sub>P'" + by(force dest: extractFrameFreshChain) + note S + moreover from PeqP' have "((p \ (\\<^sub>P \ \\<^sub>Q)) \ \') \ \\<^sub>P' \ (p \ \\<^sub>Q)" + by(simp add: eqvts) (metis Composition Associativity AssertionStatEqTrans AssertionStatEqSym Commutativity) + with \(bn(p \ \')) \* \\<^sub>Q\ \bn \' \* \\<^sub>Q\ S have "((p \ (\\<^sub>P \ \\<^sub>Q)) \ \') \ \\<^sub>P' \ \\<^sub>Q" + by simp + moreover from \PQ' = P'' \ Q\ \A\<^sub>P' \* A\<^sub>Q\ \A\<^sub>P' \* \\<^sub>Q\ \A\<^sub>Q \* \\<^sub>P'\ \A\<^sub>Q \* PQ'\ FrP' FrQ have "extractFrame PQ' = \(A\<^sub>P'@A\<^sub>Q), \\<^sub>P' \ \\<^sub>Q\" + by simp + moreover note \distinctPerm p\ \(bn(p \ \')) \* C'\ + moreover from \A\<^sub>P' \* P''\ \A\<^sub>P' \* Q\ \PQ' = P'' \ Q\ have "A\<^sub>P' \* PQ'" by simp + moreover note \A\<^sub>Q \* PQ'\ \A\<^sub>P' \* \'\ \A\<^sub>Q \* \'\ \A\<^sub>P' \* C\ \A\<^sub>Q \* C\ \(bn(p \ \')) \* \'\ + moreover from \bn \' \* Q\ have "(bn(p \ \')) \* (p \ Q)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst] bnEqvt[symmetric]) + with \bn \ \* Q\ \(bn(p \ \')) \* Q\ S have "(bn(p \ \')) \* Q" by simp + with \(bn(p \ \')) \* P''\ \PQ' = P'' \ Q\ have "(bn(p \ \')) \* PQ'" by simp + moreover from \A\<^sub>P' \* \'\ \A\<^sub>Q \* \'\ have "(A\<^sub>P'@A\<^sub>Q) \* \'" by simp + moreover from \A\<^sub>P' \* C\ \A\<^sub>Q \* C\ have "(A\<^sub>P'@A\<^sub>Q) \* C" by simp + moreover from \A\<^sub>P' \* V\ \A\<^sub>Q \* V\ have "(A\<^sub>P'@A\<^sub>Q) \* V" by simp + moreover from \A\<^sub>P' \* W\ \A\<^sub>Q \* W\ have "(A\<^sub>P'@A\<^sub>Q) \* W" by simp + moreover from \A\<^sub>P' \* X\ \A\<^sub>Q \* X\ have "(A\<^sub>P'@A\<^sub>Q) \* X" by simp + moreover from \A\<^sub>P' \* Y\ \A\<^sub>Q \* Y\ have "(A\<^sub>P'@A\<^sub>Q) \* Y" by simp + moreover from \A\<^sub>P' \* Z\ \A\<^sub>Q \* Z\ have "(A\<^sub>P'@A\<^sub>Q) \* Z" by simp + moreover from \A\<^sub>P' \* PQ'\ \A\<^sub>Q \* PQ'\ have "(A\<^sub>P'@A\<^sub>Q) \* PQ'" by simp + moreover from \A\<^sub>P' \* \'\ \A\<^sub>Q \* \'\ have "(A\<^sub>P'@A\<^sub>Q) \* \'" by simp + moreover from \A\<^sub>Q \* \'\ have "(p \ A\<^sub>Q) \* (p \ \')" + by(simp only: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst] bnEqvt[symmetric]) + with S \A\<^sub>Q \* \'\ \bn(p \ \') \* A\<^sub>Q\ have "A\<^sub>Q \* (p \ \')" + by simp + with \A\<^sub>P' \* (p \ \')\ \A\<^sub>Q \* \'\ \bn(p \ \') \* A\<^sub>Q\ S have "(A\<^sub>P'@A\<^sub>Q) \* (p \ \')" + by simp + moreover from \A\<^sub>P' \* A\<^sub>Q\ \distinct A\<^sub>P'\ \distinct A\<^sub>Q\ have "distinct(A\<^sub>P'@A\<^sub>Q)" by auto + moreover note \(bn(p \ \')) \* V\ \(bn(p \ \')) \* W\ \(bn(p \ \')) \* X\ \(bn(p \ \')) \* Y\ \(bn(p \ \')) \* Z\ + ultimately show ?case using cPar1 + by metis + next + case(cPar2 \ \\<^sub>P Q \ Q' A\<^sub>P P A\<^sub>Q \\<^sub>Q C C' \' PQ' V W X Y Z) + have FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" and FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" + by fact+ + + note \bn \' \* subject \'\ + moreover from \bn \' \* (P \ Q)\ have "bn \' \* Q" and "bn \' \* P" by simp+ + moreover with FrP FrQ \A\<^sub>P \* \'\ \A\<^sub>Q \* \'\ have "bn \' \* \\<^sub>P" and "bn \' \* \\<^sub>Q" + by(force dest: extractFrameFreshChain)+ + + moreover note \bn \' \* V\ \bn \' \* W\ + moreover from \bn \' \* X\ \A\<^sub>P \* \'\ have "bn \' \* (X@A\<^sub>P)" by simp + moreover from \bn \' \* Y\ \bn \' \* \\<^sub>P\ have "bn \' \* (\\<^sub>P#Y)" by simp + moreover from \bn \' \* Z\ \bn \' \* P\ have "bn \' \* (P#Z)" by simp + moreover note \A\<^sub>Q \* V\ \A\<^sub>Q \* W\ + moreover from \A\<^sub>Q \* X\ \A\<^sub>P \* A\<^sub>Q\ have "A\<^sub>Q \* (X@A\<^sub>P)" by simp + moreover from \A\<^sub>Q \* Y\ \A\<^sub>Q \* \\<^sub>P\ have "A\<^sub>Q \* (\\<^sub>P#Y)" by force + moreover from \A\<^sub>Q \* Z\ \A\<^sub>Q \* P\ have "A\<^sub>Q \* (P#Z)" by simp + moreover from \\ \ (P \ Q') = \' \ PQ'\ \bn \ \* P\ \bn \' \* P\ \bn \ \* \'\ + obtain Q'' where A: "\ \ Q' = \' \ Q''" and "PQ' = P \ Q''" + by(metis actionPar2Dest') + moreover from \A\<^sub>Q \* PQ'\ \PQ' = P \ Q''\ have "A\<^sub>Q \* Q''" by simp + ultimately obtain p \' A\<^sub>Q' \\<^sub>Q' where S: "set p \ set(bn \') \ set (bn(p \ \'))" and QeqQ': "((p \ \\<^sub>Q) \ \') \ \\<^sub>Q'" + and "distinctPerm p" and "(bn(p \ \')) \* C'" and FrQ': "extractFrame Q'' = \A\<^sub>Q', \\<^sub>Q'\" + and "A\<^sub>Q' \* Q''" and "A\<^sub>Q' \* \'" "A\<^sub>Q' \* (p \ \')" and "A\<^sub>Q' \* C" + and "(bn(p \ \')) \* \'" and "(bn(p \ \')) \* Q''" and "distinct A\<^sub>Q'" + and "A\<^sub>Q' \* V" and "A\<^sub>Q' \* W" and "A\<^sub>Q' \* (X @ A\<^sub>P)" and "A\<^sub>Q' \* (\\<^sub>P#Y)" + and "A\<^sub>Q' \* (P#Z)" and "(bn(p \ \')) \* V" and "(bn(p \ \')) \* W" and "(bn(p \ \')) \* (X @ A\<^sub>P)" and "(bn(p \ \')) \* (\\<^sub>P#Y)" + and "(bn(p \ \')) \* (P#Z)" using cPar2 + by(elim cPar2) + + then have "A\<^sub>Q' \* P" and "A\<^sub>Q' \* Z" and "A\<^sub>Q' \* A\<^sub>P" and "A\<^sub>Q' \* X" and "A\<^sub>Q' \* \\<^sub>P" and "A\<^sub>Q' \* Y" + and "(bn(p \ \')) \* A\<^sub>P" and "(bn(p \ \')) \* X" and "(bn(p \ \')) \* Y" and "(bn(p \ \')) \* Z" and "(bn(p \ \')) \* \\<^sub>P" + and "(bn(p \ \')) \* P" + by(simp del: freshChainSimps)+ + + from \A\<^sub>P \* PQ'\ \PQ' = P \ Q''\ \A\<^sub>Q' \* A\<^sub>P\ FrQ' have "A\<^sub>P \* \\<^sub>Q'" + by(force dest: extractFrameFreshChain) + note S + moreover from QeqQ' have "((p \ (\\<^sub>P \ \\<^sub>Q)) \ \') \ (p \ \\<^sub>P) \ \\<^sub>Q'" + by(simp add: eqvts) (metis Composition Associativity AssertionStatEqTrans AssertionStatEqSym Commutativity) + with \(bn(p \ \')) \* \\<^sub>P\ \bn \' \* \\<^sub>P\ S have "((p \ (\\<^sub>P \ \\<^sub>Q)) \ \') \ \\<^sub>P \ \\<^sub>Q'" + by simp + moreover from \PQ' = P \ Q''\ \A\<^sub>Q' \* A\<^sub>P\ \A\<^sub>Q' \* \\<^sub>P\ \A\<^sub>P \* \\<^sub>Q'\ \A\<^sub>P \* PQ'\ FrQ' FrP have "extractFrame PQ' = \(A\<^sub>P@A\<^sub>Q'), \\<^sub>P \ \\<^sub>Q'\" + by simp + moreover note \distinctPerm p\ \(bn(p \ \')) \* C'\ + moreover from \A\<^sub>Q' \* Q''\ \A\<^sub>Q' \* P\ \PQ' = P \ Q''\ have "A\<^sub>Q' \* PQ'" by simp + moreover note \A\<^sub>Q \* PQ'\ \A\<^sub>Q' \* \'\ \A\<^sub>Q \* \'\ \A\<^sub>Q' \* C\ \A\<^sub>Q \* C\ \(bn(p \ \')) \* \'\ + moreover from \bn \' \* Q\ have "(bn(p \ \')) \* (p \ Q)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst] bnEqvt[symmetric]) + with \bn \ \* P\ \(bn(p \ \')) \* P\ S have "(bn(p \ \')) \* P" by simp + with \(bn(p \ \')) \* Q''\ \PQ' = P \ Q''\ have "(bn(p \ \')) \* PQ'" by simp + moreover from \A\<^sub>Q' \* \'\ \A\<^sub>P \* \'\ have "(A\<^sub>P@A\<^sub>Q') \* \'" by simp + moreover from \A\<^sub>Q' \* C\ \A\<^sub>P \* C\ have "(A\<^sub>P@A\<^sub>Q') \* C" by simp + moreover from \A\<^sub>Q' \* V\ \A\<^sub>P \* V\ have "(A\<^sub>P@A\<^sub>Q') \* V" by simp + moreover from \A\<^sub>Q' \* W\ \A\<^sub>P \* W\ have "(A\<^sub>P@A\<^sub>Q') \* W" by simp + moreover from \A\<^sub>Q' \* X\ \A\<^sub>P \* X\ have "(A\<^sub>P@A\<^sub>Q') \* X" by simp + moreover from \A\<^sub>Q' \* Y\ \A\<^sub>P \* Y\ have "(A\<^sub>P@A\<^sub>Q') \* Y" by simp + moreover from \A\<^sub>Q' \* Z\ \A\<^sub>P \* Z\ have "(A\<^sub>P@A\<^sub>Q') \* Z" by simp + moreover from \A\<^sub>Q' \* PQ'\ \A\<^sub>P \* PQ'\ have "(A\<^sub>P@A\<^sub>Q') \* PQ'" by simp + moreover from \A\<^sub>Q' \* \'\ \A\<^sub>P \* \'\ have "(A\<^sub>P@A\<^sub>Q') \* \'" by simp + moreover from \A\<^sub>P \* \'\ have "(p \ A\<^sub>P) \* (p \ \')" + by(simp only: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst] bnEqvt[symmetric]) + with S \A\<^sub>P \* \'\ \bn(p \ \') \* A\<^sub>P\ have "A\<^sub>P \* (p \ \')" + by simp + with \A\<^sub>Q' \* (p \ \')\ \A\<^sub>P \* \'\ \bn(p \ \') \* A\<^sub>P\ S have "(A\<^sub>P@A\<^sub>Q') \* (p \ \')" + by simp + moreover note \(bn(p \ \')) \* V\ \(bn(p \ \')) \* W\ \(bn(p \ \')) \* X\ \(bn(p \ \')) \* Y\ \(bn(p \ \')) \* Z\ + moreover from \A\<^sub>Q' \* A\<^sub>P\ \distinct A\<^sub>P\ \distinct A\<^sub>Q'\ have "distinct(A\<^sub>P@A\<^sub>Q')" by auto + ultimately show ?case using cPar2 + by metis + next + case cComm1 + then show ?case by(simp add: residualInject) + next + case cComm2 + then show ?case by(simp add: residualInject) + next + case(cBrMerge \ \\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q Q' A\<^sub>Q C C' \ PQ' V W X Y Z) + have FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" and FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" + by fact+ + + have "bn (\M\N\) \* Q" by simp + have "bn (\M\N\) \* Q'" by simp + have "bn (\M\N\) \* \" by simp + have "bn (\M\N\) \* P" by simp + have "bn (\M\N\) \* P'" by simp + have "bn (\M\N\) \* \" by simp + + from \\M\N\ \ P' \ Q' = \ \ PQ'\ + have empty\: "bn \ = ([]::name list)" and + "\ = (\M\N\)" + by(simp add: residualInject)+ + then have "bn \ \* Q" and "bn \ \* Q'" + and "bn \ \* P" and "bn \ \* P'" by simp+ + + from \bn \ = []\ + have "bn \ \* subject \" and "bn \ \* P" and "bn \ \* Q" + and "bn \ \* \\<^sub>P" and "bn \ \* \\<^sub>Q" and "bn \ \* C'" + and "bn \ \* (X@A\<^sub>Q)" + and "bn \ \* (\\<^sub>Q#Y)" and "bn \ \* (Q#Z)" by simp+ + + moreover note \A\<^sub>P \* P\ \A\<^sub>P \* \\ \A\<^sub>P \* C\ \A\<^sub>P \* C'\ + \\ \ \\ + + moreover from \A\<^sub>P \* X\ \A\<^sub>P \* A\<^sub>Q\ have "A\<^sub>P \* (X@A\<^sub>Q)" by simp + moreover from \A\<^sub>P \* Y\ \A\<^sub>P \* \\<^sub>Q\ have "A\<^sub>P \* (\\<^sub>Q#Y)" by force + moreover from \A\<^sub>P \* Z\ \A\<^sub>P \* Q\ have "A\<^sub>P \* (Q#Z)" by simp + moreover from \\M\N\ \ (P' \ Q') = \ \ PQ'\ \bn (\M\N\) \* Q'\ \bn \ \* Q'\ \bn (\M\N\) \* \\ + obtain P'' where A: "\M\N\ \ P' = \ \ P''" and "PQ' = P'' \ Q'" + by(metis actionPar1Dest') + moreover from \A\<^sub>P \* PQ'\ \PQ' = P'' \ Q'\ have "A\<^sub>P \* P''" by simp + ultimately obtain p P\' A\<^sub>P' \\<^sub>P' where Sp: "set p \ set(bn \) \ set (bn(p \ \))" and PeqP': "((p \ \\<^sub>P) \ P\') \ \\<^sub>P'" + and "distinctPerm p" and FrP': "extractFrame P'' = \A\<^sub>P', \\<^sub>P'\" + and "A\<^sub>P' \* P''" and "A\<^sub>P' \* \" "A\<^sub>P' \* (p \ \)" and "A\<^sub>P' \* C" and "(bn(p \ \)) \* C'" + and "(bn(p \ \)) \* \" and "(bn(p \ \)) \* P''" and "distinct A\<^sub>P'" + and "A\<^sub>P' \* V" and "A\<^sub>P' \* W" and "A\<^sub>P' \* (X @ A\<^sub>Q)" and "A\<^sub>P' \* (\\<^sub>Q#Y)" + and "A\<^sub>P' \* (Q#Z)" and "(bn(p \ \)) \* V" and "(bn(p \ \)) \* W" and "(bn(p \ \)) \* (X @ A\<^sub>Q)" and "(bn(p \ \)) \* (\\<^sub>Q#Y)" + and "(bn(p \ \)) \* (Q#Z)" using cBrMerge + by(elim cBrMerge) + + then have "A\<^sub>P' \* Q" and "A\<^sub>P' \* Z" and "A\<^sub>P' \* A\<^sub>Q" and "A\<^sub>P' \* X" and "A\<^sub>P' \* \\<^sub>Q" and "A\<^sub>P' \* Y" + and "(bn(p \ \)) \* A\<^sub>Q" and "(bn(p \ \)) \* X" and "(bn(p \ \)) \* Y" and "(bn(p \ \)) \* Z" and "(bn(p \ \)) \* \\<^sub>Q" + and "(bn(p \ \)) \* Q" + by(simp del: freshChainSimps)+ + + from \A\<^sub>Q \* PQ'\ \PQ' = (P'' \ Q')\ have "A\<^sub>Q \* P''" by simp + with \extractFrame P'' = \A\<^sub>P', \\<^sub>P'\\ \A\<^sub>P' \* A\<^sub>Q\ have "A\<^sub>Q \* \\<^sub>P'" + by (metis extractFrameFreshChain freshFrameDest) + + from \bn \ = []\ + have "bn \ \* subject \" and "bn \ \* Q" and "bn \ \* P''" + and "bn \ \* \\<^sub>Q" and "bn \ \* \\<^sub>P'" and "bn \ \* C'" + and "bn \ \* (X@A\<^sub>P')" + and "bn \ \* (\\<^sub>P'#Y)" and "bn \ \* (P''#Z)" by simp+ + + moreover note \A\<^sub>P' \* Q\ \A\<^sub>Q \* \\ \A\<^sub>Q \* C\ \A\<^sub>Q \* C'\ + \\ \ \\ + + moreover from \A\<^sub>Q \* X\ \A\<^sub>P' \* A\<^sub>Q\ have "A\<^sub>Q \* (X@A\<^sub>P')" by simp + moreover from \A\<^sub>Q \* Y\ \A\<^sub>Q \* \\<^sub>P'\ have "A\<^sub>Q \* (\\<^sub>P'#Y)" by force + moreover from \A\<^sub>Q \* Z\ \A\<^sub>Q \* P''\ have "A\<^sub>Q \* (P''#Z)" by simp + moreover from \\M\N\ \ (P' \ Q') = \ \ PQ'\ \bn (\M\N\) \* P'\ \bn \ \* P'\ \bn (\M\N\) \* \\ + obtain Q'' where A: "\M\N\ \ Q' = \ \ Q''" and "PQ' = P' \ Q''" + by(metis actionPar2Dest') + + moreover from \PQ' = P'' \ Q'\ \PQ' = P' \ Q''\ + have "PQ' = P'' \ Q''" + by (simp add: psi.inject) + + moreover from \A\<^sub>Q \* PQ'\ \PQ' = P'' \ Q''\ have "A\<^sub>Q \* Q''" by simp + ultimately obtain q \' A\<^sub>Q' \\<^sub>Q' where Sq: "set q \ set(bn \) \ set (bn(q \ \))" and QeqQ': "((q \ \\<^sub>Q) \ \') \ \\<^sub>Q'" + and "distinctPerm q" and FrQ': "extractFrame Q'' = \A\<^sub>Q', \\<^sub>Q'\" + and "A\<^sub>Q' \* Q''" and "A\<^sub>Q' \* \" "A\<^sub>Q' \* (q \ \)" and "A\<^sub>Q' \* C" and "(bn(q \ \)) \* C'" + and "(bn(q \ \)) \* \" and "(bn(q \ \)) \* Q''" and "distinct A\<^sub>Q'" + and "A\<^sub>Q' \* V" and "A\<^sub>Q' \* W" and "A\<^sub>Q' \* (X @ A\<^sub>P')" and "A\<^sub>Q' \* (\\<^sub>P'#Y)" + and "A\<^sub>Q' \* (P''#Z)" and "(bn(q \ \)) \* V" and "(bn(q \ \)) \* W" and "(bn(q \ \)) \* (X @ A\<^sub>P')" and "(bn(q \ \)) \* (\\<^sub>P'#Y)" + and "(bn(q \ \)) \* (P''#Z)" using cBrMerge + by(elim cBrMerge(6)[where bb=\]) (rule refl | assumption)+ + + then have "A\<^sub>Q' \* P''" and "A\<^sub>Q' \* Z" and "A\<^sub>Q' \* A\<^sub>P'" and "A\<^sub>Q' \* X" and "A\<^sub>Q' \* \\<^sub>P'" and "A\<^sub>Q' \* Y" + and "(bn(q \ \)) \* A\<^sub>P'" and "(bn(q \ \)) \* X" and "(bn(q \ \)) \* Y" and "(bn(q \ \)) \* Z" and "(bn(q \ \)) \* \\<^sub>P'" + and "(bn(q \ \)) \* P''" + by(simp del: freshChainSimps)+ + + from Sp Sq \bn \ = []\ have "p = ([]::name prm)" and "q = ([]::name prm)" and "p = q" + by simp+ + + from \A\<^sub>Q' \* P''\ \A\<^sub>Q' \* A\<^sub>P'\ \extractFrame P'' = \A\<^sub>P', \\<^sub>P'\\ have "A\<^sub>Q' \* \\<^sub>P'" + by (metis extractFrameFreshChain freshFrameDest) + from \A\<^sub>P' \* \\ \\ = \M\N\\ have "A\<^sub>P' \* M" and "A\<^sub>P' \* N" + by simp+ + + from \\ \ \\<^sub>P \ Q \ \M\N\ \ Q'\ \A\<^sub>P' \* Q\ \A\<^sub>P' \* N\ + have "A\<^sub>P' \* Q'" + by(simp add: brinputFreshChainDerivative) + with \PQ' = (P'' \ Q')\ \PQ' = (P'' \ Q'')\ have "A\<^sub>P' \* Q''" + by (simp add: psi.inject) + with FrQ' \A\<^sub>Q' \* A\<^sub>P'\ have "A\<^sub>P' \* \\<^sub>Q'" + by(metis extractFrameFreshChain freshFrameDest) + from \A\<^sub>P' \* P''\ \A\<^sub>P' \* Q''\ \PQ' = (P'' \ Q'')\ + have "A\<^sub>P' \* PQ'" by simp + from \A\<^sub>Q' \* P''\ \A\<^sub>Q' \* Q''\ \PQ' = (P'' \ Q'')\ + have "A\<^sub>Q' \* PQ'" by simp + + + from PeqP' have "((p \ (\\<^sub>P \ \\<^sub>Q)) \ P\') \ \\<^sub>P' \ (p \ \\<^sub>Q)" + by(simp add: eqvts) (metis Composition Associativity AssertionStatEqTrans AssertionStatEqSym Commutativity) + + with \p = []\ have "((\\<^sub>P \ \\<^sub>Q) \ P\') \ (\\<^sub>P' \ \\<^sub>Q)" by simp + with QeqQ' have "((q \ (\\<^sub>Q \ \\<^sub>P')) \ \') \ \\<^sub>Q' \ (q \ \\<^sub>P')" + by(simp add: eqvts) (metis Composition Associativity AssertionStatEqTrans AssertionStatEqSym Commutativity) + with PeqP' \p = []\ \q = []\ have "((q \ (\\<^sub>Q \ ((p \ \\<^sub>P) \ P\'))) \ \') \ \\<^sub>Q' \ (q \ \\<^sub>P')" + by(simp add: eqvts) (metis AssertionStatEqTrans Composition compositionSym) + with \p = []\ have "((q \ (\\<^sub>Q \ (\\<^sub>P \ P\'))) \ \') \ \\<^sub>Q' \ (q \ \\<^sub>P')" by simp + with \q = []\ have "(q \ (\\<^sub>P \ \\<^sub>Q)) \ (P\' \ \') \ \\<^sub>P' \ \\<^sub>Q'" + by(simp add: eqvts) (metis AssertionStatEqTrans Commutativity Composition associativitySym) + + moreover note \set q \ set(bn \) \ set (bn (q \ \))\ + + moreover from \PQ' = P'' \ Q''\ \A\<^sub>Q' \* A\<^sub>P'\ \A\<^sub>P' \* \\<^sub>Q'\ \A\<^sub>Q' \* \\<^sub>P'\ \A\<^sub>Q' \* PQ'\ FrP' FrQ' have "extractFrame PQ' = \(A\<^sub>P'@A\<^sub>Q'), \\<^sub>P' \ \\<^sub>Q'\" + by simp + + moreover note \distinctPerm q\ + + moreover from \A\<^sub>P' \* PQ'\ \A\<^sub>Q' \* PQ'\ + have "(A\<^sub>P'@A\<^sub>Q') \* PQ'" by simp + moreover from \A\<^sub>P' \* \\ \A\<^sub>Q' \* \\ + have "(A\<^sub>P'@A\<^sub>Q') \* \" by simp + moreover with \q = []\ + have "(A\<^sub>P'@A\<^sub>Q') \* (q \ \)" by simp + moreover from \A\<^sub>P' \* C\ \A\<^sub>Q' \* C\ + have "(A\<^sub>P'@A\<^sub>Q') \* C" by simp + moreover note \bn (q \ \) \* C'\ \bn (q \ \) \* \\ + moreover from \bn \ = []\ + have "bn (q \ \) \* PQ'" + by (metis Nominal.nil_eqvt bnEqvt freshSets) + moreover from \A\<^sub>P' \* V\ \A\<^sub>Q' \* V\ + have "(A\<^sub>P'@A\<^sub>Q') \* V" by simp + moreover from \A\<^sub>P' \* W\ \A\<^sub>Q' \* W\ + have "(A\<^sub>P'@A\<^sub>Q') \* W" by simp + moreover from \A\<^sub>P' \* X\ \A\<^sub>Q' \* X\ + have "(A\<^sub>P'@A\<^sub>Q') \* X" by simp + moreover from \A\<^sub>P' \* Y\ \A\<^sub>Q' \* Y\ + have "(A\<^sub>P'@A\<^sub>Q') \* Y" by simp + moreover from \A\<^sub>P' \* Z\ \A\<^sub>Q' \* Z\ + have "(A\<^sub>P'@A\<^sub>Q') \* Z" by simp + moreover from \A\<^sub>Q' \* A\<^sub>P'\ \distinct A\<^sub>P'\ \distinct A\<^sub>Q'\ have "distinct(A\<^sub>P'@A\<^sub>Q')" by simp + moreover note \(bn(q \ \)) \* V\ \(bn(q \ \)) \* W\ \(bn(q \ \)) \* X\ \(bn(q \ \)) \* Y\ \(bn(q \ \)) \* Z\ + ultimately show ?case + by(intro cBrMerge(47)) + next + case(cBrComm1 \ \\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q xvec Q' A\<^sub>Q C C' \ PQ' V W X Y Z) + have FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" and FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" + by fact+ + + from \xvec \* \\ have "xvec \* bn \" by simp + + from \\M\\*xvec\\N\ \ P' \ Q' = \ \ PQ'\ have "\ \ PQ' = \M\\*xvec\\N\ \ P' \ Q'" by simp + with \xvec \* (bn \)\ + obtain Q'' r where rPerm: "set r \ set (bn \) \ set xvec" + and "PQ' = (r \ P') \ Q''" and "\ \ Q'' = \M\\*xvec\\N\ \ Q'" + by(elim actionPar2Dest) (assumption | simp)+ + then have "\M\\*xvec\\N\ \ Q' = \ \ Q''" by simp + + from \A\<^sub>Q \* PQ'\ \PQ' = (r \ P') \ Q''\ have "A\<^sub>Q \* Q''" + by simp + + from \A\<^sub>P \* PQ'\ \PQ' = (r \ P') \ Q''\ have "A\<^sub>P \* Q''" + by simp + + from \bn \ \* (P \ Q)\ \A\<^sub>P \* \\ + have "bn \ \* P" and "A\<^sub>P \* bn \" by simp+ + with \extractFrame P = \A\<^sub>P, \\<^sub>P\\ have "bn \ \* \\<^sub>P" + by (metis extractFrameFreshChain freshFrameDest) + + from \bn \ \* (P \ Q)\ \A\<^sub>Q \* \\ + have "bn \ \* Q" and "A\<^sub>Q \* bn \" by simp+ + with \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ have "bn \ \* \\<^sub>Q" + by (metis extractFrameFreshChain freshFrameDest) + + from rPerm \A\<^sub>P \* xvec\ \xvec \* \\<^sub>P\ \A\<^sub>P \* bn \\ \bn \ \* \\<^sub>P\ + have "r \ A\<^sub>P = A\<^sub>P" and "r \ \\<^sub>P = \\<^sub>P" by simp+ + + have "\M\N\ \ P' = \M\N\ \ P'" by simp + moreover note \A\<^sub>P \* P\ \A\<^sub>P \* C\ \A\<^sub>P \* C'\ + moreover from \A\<^sub>P \* M\ \A\<^sub>P \* N\ have "A\<^sub>P \* (\M\N\)" by simp + moreover note \A\<^sub>P \* V\ + moreover from \A\<^sub>P \* W\ \A\<^sub>P \* \\ have "A\<^sub>P \* (\#W)" by simp + moreover from \A\<^sub>P \* X\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* xvec\ have "A\<^sub>P \* (X@A\<^sub>Q@xvec)" by simp + moreover from \A\<^sub>P \* Y\ \A\<^sub>P \* \\<^sub>Q\ have "A\<^sub>P \* (\\<^sub>Q#Y)" by force + moreover from \A\<^sub>P \* Z\ \A\<^sub>P \* Q\ \A\<^sub>P \* Q'\ \A\<^sub>P \* Q''\ have "A\<^sub>P \* (Q#(Q'#(Q''#Z)))" by simp + moreover note \A\<^sub>P \* P'\ \A\<^sub>P \* Q\ + + ultimately obtain p P\' A\<^sub>P' \\<^sub>P' where Sp: "set p \ set(bn (\M\N\)) \ set (bn(p \ (\M\N\)))" and "((p \ \\<^sub>P) \ P\') \ \\<^sub>P'" + and "distinctPerm p" and FrP': "extractFrame P' = \A\<^sub>P', \\<^sub>P'\" + and "A\<^sub>P' \* P'" and "A\<^sub>P' \* (\M\N\)" "A\<^sub>P' \* (p \ (\M\N\))" and "A\<^sub>P' \* C" and "(bn(p \ (\M\N\))) \* C'" + and "(bn(p \ (\M\N\))) \* (\M\N\)" and "(bn(p \ (\M\N\))) \* P'" and "distinct A\<^sub>P'" + and "A\<^sub>P' \* V" and "A\<^sub>P' \* (\#W)" and "A\<^sub>P' \* (X @ A\<^sub>Q @ xvec)" and "A\<^sub>P' \* (\\<^sub>Q#Y)" + and "A\<^sub>P' \* (Q#(Q'#(Q''#Z)))" and "(bn(p \ (\M\N\))) \* (\#W)" and "(bn(p \ (\M\N\))) \* (X @ A\<^sub>Q @ xvec)" and "(bn(p \ (\M\N\))) \* (\\<^sub>Q#Y)" + and "(bn(p \ (\M\N\))) \* (Q#(Q'#(Q''#Z)))" + by(elim cBrComm1(4)) (assumption | simp)+ + + then have PeqP': "\\<^sub>P \ P\' \ \\<^sub>P'" + and "A\<^sub>P' \* P'" and "A\<^sub>P' \* (\M\N\)" and "A\<^sub>P' \* C" and "distinct A\<^sub>P'" + and "A\<^sub>P' \* Q" and "A\<^sub>P' \* Q'" and "A\<^sub>P' \* Q''" and "A\<^sub>P' \* Z" and "A\<^sub>P' \* A\<^sub>Q" and "A\<^sub>P' \* X" and "A\<^sub>P' \* \\<^sub>Q" and "A\<^sub>P' \* Y" + and "A\<^sub>P' \* xvec" and "A\<^sub>P' \* \" and "A\<^sub>P' \* (subject \)" and "A\<^sub>P' \* (bn \)" and "A\<^sub>P' \* (object \)" + and "A\<^sub>P' \* V" and "A\<^sub>P' \* W" + by(simp del: freshChainSimps)+ + + from rPerm \bn \ \* \\<^sub>P\ \xvec \* \\<^sub>P\ + have "(r \ \\<^sub>P) = \\<^sub>P" by simp + + from PeqP' have "r \ (\\<^sub>P \ P\' \ \\<^sub>P')" + by simp + with \(r \ \\<^sub>P) = \\<^sub>P\ + have rPeqP': "\\<^sub>P \ (r \ P\') \ (r \ \\<^sub>P')" by(simp add: eqvts) + + from rPerm \A\<^sub>P' \* (bn \)\ \A\<^sub>P' \* xvec\ + have "r \ A\<^sub>P' = A\<^sub>P'" by simp + + from FrP' have "r \ (extractFrame P' = \A\<^sub>P', \\<^sub>P'\)" + by simp + + with \r \ A\<^sub>P' = A\<^sub>P'\ have rFrP': "extractFrame (r \ P') = \A\<^sub>P', (r \ \\<^sub>P')\" + by(simp add: eqvts) + + from \xvec \* \\ have "(bn \) \* xvec" by simp + + from \(A\<^sub>P @ A\<^sub>Q) \* \\ have "A\<^sub>Q \* bn \" by simp + + from \A\<^sub>Q \* M\ \A\<^sub>Q \* xvec\ \A\<^sub>Q \* N\ have "A\<^sub>Q \* (\M\\*xvec\\N\)" by simp + from \\ \ \\<^sub>P \ Q \ \M\\*xvec\\N\ \ Q'\ \xvec \* M\ \distinct xvec\ \bn \ \* Q\ \xvec \* bn \\ + have "bn \ \* Q'" by(simp add: broutputFreshChainDerivative) + + from \\ \ \\<^sub>P \ Q \ \M\\*xvec\\N\ \ Q'\ \xvec \* M\ \distinct xvec\ \bn \ \* Q\ \bn \ \* xvec\ + have "bn \ \* N" by(simp add: broutputFreshChainDerivative) + + note \bn \ \* subject \\ + + moreover from \bn \ \* (P \ Q)\ have "bn \ \* Q" and "bn \ \* P" by simp+ + moreover from \bn \ \* V\ \bn \ \* N\ have "bn \ \* (N#V)" by simp + moreover note \bn \ \* \\<^sub>P\ \bn \ \* \\<^sub>Q\ \bn \ \* W\ + moreover from \bn \ \* X\ \A\<^sub>P \* bn \\ \A\<^sub>P' \* bn \\ \bn \ \* xvec\ have "bn \ \* (X@xvec@A\<^sub>P@A\<^sub>P')" by simp + moreover from \bn \ \* Y\ \bn \ \* \\<^sub>P\ have "bn \ \* (\\<^sub>P#Y)" by simp + moreover from \bn \ \* Z\ \bn \ \* P\ \bn \ \* Q\ have "bn \ \* (P#Q#Z)" by simp + moreover from \A\<^sub>Q \* V\ \A\<^sub>Q \* N\ have "A\<^sub>Q \* (N#V)" by simp + moreover note \A\<^sub>Q \* W\ + moreover from \A\<^sub>Q \* X\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P' \* A\<^sub>Q\ \A\<^sub>Q \* xvec\ have "A\<^sub>Q \* (X@xvec@A\<^sub>P@A\<^sub>P')" by simp + moreover from \A\<^sub>Q \* Y\ \A\<^sub>Q \* \\<^sub>P\ have "A\<^sub>Q \* (\\<^sub>P#Y)" by force + moreover from \A\<^sub>Q \* Z\ \A\<^sub>Q \* P\ \A\<^sub>Q \* Q\ have "A\<^sub>Q \* (P#Q#Z)" by simp + + moreover note \\M\\*xvec\\N\ \ Q' = \ \ Q''\ + ultimately obtain q Q\' A\<^sub>Q' \\<^sub>Q' where Sq: "set q \ set (bn \) \ set (bn (q \ \))" and QeqQ': "((q \ \\<^sub>Q) \ Q\') \ \\<^sub>Q'" + and "distinctPerm q" and "bn(q \ \) \* C'" and FrQ': "extractFrame Q'' = \A\<^sub>Q', \\<^sub>Q'\" + and "A\<^sub>Q' \* Q''" and "A\<^sub>Q' \* \" and "A\<^sub>Q' \* (q \ \)" and "A\<^sub>Q' \* C" + and "bn(q \ \) \* \" and "bn(q \ \) \* Q''" and "distinct A\<^sub>Q'" + and "A\<^sub>Q' \* (N#V)" and "A\<^sub>Q' \* W" and "A\<^sub>Q' \* (X @ xvec @ A\<^sub>P @ A\<^sub>P')" and "A\<^sub>Q' \* (\\<^sub>P#Y)" + and "A\<^sub>Q' \* (P#Q#Z)" and "bn(q \ \) \* (N#V)" and "bn(q \ \) \* W" and "bn(q \ \) \* (X @ xvec @ A\<^sub>P @ A\<^sub>P')" and "bn(q \ \) \* (\\<^sub>P#Y)" + and "bn(q \ \) \* (P#Q#Z)" + using \A\<^sub>Q \* Q\ \A\<^sub>Q \* \\ \A\<^sub>Q \* C\ \A\<^sub>Q \* C'\ + \bn \ \* C'\ \\ \ \\ \A\<^sub>Q \* Q''\ \distinct (bn \)\ + by(elim cBrComm1(8)[where b=C and ba=C' and bb=\ and bc=Q'' and bf="(X @ xvec @ A\<^sub>P @ A\<^sub>P')" and bg="(\\<^sub>P#Y)" and bh="(P#Q#Z)"]) (assumption | simp)+ + then have "A\<^sub>Q' \* P" and "A\<^sub>Q' \* Z" and "A\<^sub>Q' \* A\<^sub>P" and "A\<^sub>Q' \* A\<^sub>P'" and "A\<^sub>Q' \* X" and "A\<^sub>Q' \* \\<^sub>P" and "A\<^sub>Q' \* Y" and "A\<^sub>Q' \* Q" and "A\<^sub>Q' \* N" + and "A\<^sub>Q' \* xvec" and "A\<^sub>Q' \* V" and "A\<^sub>Q' \* W" + and "bn(q \ \) \* A\<^sub>P" and "bn(q \ \) \* A\<^sub>P'" and "bn(q \ \) \* X" and "bn(q \ \) \* Y" and "bn(q \ \) \* Z" and "bn(q \ \) \* \\<^sub>P" + and "bn(q \ \) \* P" and "bn(q \ \) \* W" and "bn(q \ \) \* V" and "bn(q \ \) \* N" and "bn(q \ \) \* xvec" + by(simp del: freshChainSimps)+ + + from \A\<^sub>P' \* Q''\ \A\<^sub>Q' \* A\<^sub>P'\ FrQ' + have "A\<^sub>P' \* \\<^sub>Q'" + by(metis extractFrameFreshChain freshFrameDest) + + from \\ \ \\<^sub>Q \ P \ \M\N\ \ P'\ \A\<^sub>Q' \* P\ \A\<^sub>Q' \* N\ + have "A\<^sub>Q' \* P'" + by(simp add: brinputFreshChainDerivative) + with FrP' \A\<^sub>Q' \* A\<^sub>P'\ have "A\<^sub>Q' \* \\<^sub>P'" + by(metis extractFrameFreshChain freshFrameDest) + + have "(\\<^sub>P \ (r \ P\')) \ ((q \ \\<^sub>Q) \ Q\') \ (r \ \\<^sub>P') \ \\<^sub>Q'" + by(metis Composition' rPeqP' QeqQ') + then have "(\\<^sub>P \ ((r \ P\') \ ((q \ \\<^sub>Q) \ Q\'))) \ (r \ \\<^sub>P') \ \\<^sub>Q'" + by (metis AssertionStatEqSym AssertionStatEqTrans Associativity) + then have "(\\<^sub>P \ (((q \ \\<^sub>Q) \ Q\') \ (r \ P\'))) \ (r \ \\<^sub>P') \ \\<^sub>Q'" + by (metis AssertionStatEqSym AssertionStatEqTrans Associativity associativitySym) + then have "(\\<^sub>P \ ((q \ \\<^sub>Q) \ (Q\' \ (r \ P\')))) \ (r \ \\<^sub>P') \ \\<^sub>Q'" + by (metis AssertionStatEqSym AssertionStatEqTrans Associativity compositionSym) + then have "((\\<^sub>P \ (q \ \\<^sub>Q)) \ (Q\' \ (r \ P\'))) \ (r \ \\<^sub>P') \ \\<^sub>Q'" + by (metis AssertionStatEqTrans Associativity) + then have "((\\<^sub>P \ (q \ \\<^sub>Q)) \ ((r \ P\') \ Q\')) \ (r \ \\<^sub>P') \ \\<^sub>Q'" + by (metis AssertionStatEqSym AssertionStatEqTrans Associativity associativitySym) + with Sq \bn \ \* \\<^sub>P\ \bn(q \ \) \* \\<^sub>P\ have "((q \ (\\<^sub>P \ \\<^sub>Q)) \ ((r \ P\') \ Q\')) \ (r \ \\<^sub>P') \ \\<^sub>Q'" + by(simp add: eqvts) + + from Sq \A\<^sub>P' \* \\ \bn(q \ \) \* A\<^sub>P'\ have "A\<^sub>P' \* (q \ \)" + by (metis actionFreshChain freshChainSym freshStarChainSimps fresh_star_set_eq) + + from \A\<^sub>Q' \* \\ have "A\<^sub>Q' \* bn \" by simp + from rPerm \A\<^sub>P' \* P'\ \A\<^sub>P' \* xvec\ \A\<^sub>P' \* bn \\ + have "A\<^sub>P' \* (r \ P')" + by (metis freshAlphaPerm freshChainSym name_list_set_fresh permStarFresh) + from rPerm \A\<^sub>Q' \* P'\ \A\<^sub>Q' \* xvec\ \A\<^sub>Q' \* bn \\ + have "A\<^sub>Q' \* (r \ P')" + by (metis freshAlphaPerm freshChainSym name_list_set_fresh permStarFresh) + from rPerm \A\<^sub>Q' \* \\<^sub>P'\ \A\<^sub>Q' \* xvec\ \A\<^sub>Q' \* bn \\ + have "A\<^sub>Q' \* (r \ \\<^sub>P')" + by (metis freshAlphaPerm freshChainSym name_list_set_fresh permStarFresh) + + with \A\<^sub>P' \* \\<^sub>Q'\ \A\<^sub>Q' \* A\<^sub>P'\ rFrP' FrQ' + have "extractFrame ((r \ P') \ Q'') = \(A\<^sub>P'@A\<^sub>Q'), ((r \ \\<^sub>P') \ \\<^sub>Q')\" + by simp + with \PQ' = ((r \ P') \ Q'')\ + have "extractFrame PQ' = \(A\<^sub>P'@A\<^sub>Q'), ((r \ \\<^sub>P') \ \\<^sub>Q')\" + by simp + + from \\ \ \\<^sub>Q \ P \ \M\N\ \ P'\ \bn(q \ \) \* P\ \bn(q \ \) \* N\ + have "bn(q \ \) \* P'" by(simp add: brinputFreshChainDerivative) + + with rPerm \bn(q \ \) \* \\ \bn(q \ \) \* xvec\ + have "bn(q \ \) \* (r \ P')" + by (metis actionFreshChain freshAlphaPerm freshChainSym name_list_set_fresh permStarFresh) + + from \A\<^sub>P' \* Q''\ \A\<^sub>P' \* (r \ P')\ \PQ' = (r \ P') \ Q''\ have "A\<^sub>P' \* PQ'" by simp + from \A\<^sub>Q' \* Q''\ \A\<^sub>Q' \* (r \ P')\ \PQ' = (r \ P') \ Q''\ have "A\<^sub>Q' \* PQ'" by simp + + note \set q \ (set (bn \)) \ set (bn(q \ \))\ + \((q \ (\\<^sub>P \ \\<^sub>Q)) \ ((r \ P\') \ Q\')) \ (r \ \\<^sub>P') \ \\<^sub>Q'\ \distinctPerm q\ + \extractFrame PQ' = \(A\<^sub>P'@A\<^sub>Q'), ((r \ \\<^sub>P') \ \\<^sub>Q')\\ + + moreover from \A\<^sub>P' \* PQ'\ \A\<^sub>Q' \* PQ'\ have "(A\<^sub>P'@A\<^sub>Q') \* PQ'" by simp + moreover from \A\<^sub>P' \* \\ \A\<^sub>Q' \* \\ have "(A\<^sub>P'@A\<^sub>Q') \* \" by simp + moreover from \A\<^sub>P' \* (q \ \)\ \A\<^sub>Q' \* (q \ \)\ have "(A\<^sub>P'@A\<^sub>Q') \* (q \ \)" by simp + moreover from \A\<^sub>P' \* C\ \A\<^sub>Q' \* C\ have "(A\<^sub>P'@A\<^sub>Q') \* C" by simp + moreover note \bn(q \ \) \* C'\ \bn(q \ \) \* \\ + moreover from \bn(q \ \) \* (r \ P')\ \bn(q \ \) \* Q''\ \PQ' = (r \ P') \ Q''\ + have "bn(q \ \) \* PQ'" by simp + moreover from \A\<^sub>P' \* V\ \A\<^sub>Q' \* V\ have "(A\<^sub>P'@A\<^sub>Q') \* V" by simp + moreover from \A\<^sub>P' \* W\ \A\<^sub>Q' \* W\ have "(A\<^sub>P'@A\<^sub>Q') \* W" by simp + moreover from \A\<^sub>P' \* X\ \A\<^sub>Q' \* X\ have "(A\<^sub>P'@A\<^sub>Q') \* X" by simp + moreover from \A\<^sub>P' \* Y\ \A\<^sub>Q' \* Y\ have "(A\<^sub>P'@A\<^sub>Q') \* Y" by simp + moreover from \A\<^sub>P' \* Z\ \A\<^sub>Q' \* Z\ have "(A\<^sub>P'@A\<^sub>Q') \* Z" by simp + moreover from \distinct A\<^sub>P'\ \distinct A\<^sub>Q'\ \A\<^sub>Q' \* A\<^sub>P'\ + have "distinct(A\<^sub>P'@A\<^sub>Q')" by simp + moreover note \bn(q \ \) \* V\ \bn(q \ \) \* W\ + \bn(q \ \) \* X\ \bn(q \ \) \* Y\ \bn(q \ \) \* Z\ + ultimately show ?case + by(rule cBrComm1(63)) + next + case(cBrComm2 \ \\<^sub>Q P M xvec N P' A\<^sub>P \\<^sub>P Q Q' A\<^sub>Q C C' \ PQ' V W X Y Z) + have FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" and FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" + by fact+ + + from \xvec \* \\ have "xvec \* bn \" by simp + + from \\M\\*xvec\\N\ \ P' \ Q' = \ \ PQ'\ have "\ \ PQ' = \M\\*xvec\\N\ \ P' \ Q'" by simp + with \xvec \* (bn \)\ + obtain P'' r where rPerm: "set r \ set (bn \) \ set xvec" + and "PQ' = P'' \ (r \ Q')" and "\ \ P'' = \M\\*xvec\\N\ \ P'" + by(elim actionPar1Dest) (assumption | simp)+ + then have "\M\\*xvec\\N\ \ P' = \ \ P''" by simp + + from \A\<^sub>Q \* PQ'\ \PQ' = P'' \ (r \ Q')\ have "A\<^sub>Q \* P''" + by simp + + from \A\<^sub>P \* PQ'\ \PQ' = P'' \ (r \ Q')\ have "A\<^sub>P \* P''" + by simp + + from \bn \ \* (P \ Q)\ \A\<^sub>P \* \\ + have "bn \ \* P" and "A\<^sub>P \* bn \" by simp+ + with \extractFrame P = \A\<^sub>P, \\<^sub>P\\ have "bn \ \* \\<^sub>P" + by (metis extractFrameFreshChain freshFrameDest) + + from \bn \ \* (P \ Q)\ \A\<^sub>Q \* \\ + have "bn \ \* Q" and "A\<^sub>Q \* bn \" by simp+ + with \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ have "bn \ \* \\<^sub>Q" + by (metis extractFrameFreshChain freshFrameDest) + + from rPerm \A\<^sub>Q \* xvec\ \xvec \* \\<^sub>Q\ \A\<^sub>Q \* bn \\ \bn \ \* \\<^sub>Q\ + have "r \ A\<^sub>Q = A\<^sub>Q" and "r \ \\<^sub>Q = \\<^sub>Q" by simp+ + + have "\M\N\ \ Q' = \M\N\ \ Q'" by simp + moreover note \A\<^sub>Q \* P\ \A\<^sub>Q \* C\ \A\<^sub>Q \* C'\ + moreover from \A\<^sub>Q \* M\ \A\<^sub>Q \* N\ have "A\<^sub>Q \* (\M\N\)" by simp + moreover note \A\<^sub>Q \* V\ + moreover from \A\<^sub>Q \* W\ \A\<^sub>Q \* \\ have "A\<^sub>Q \* (\#W)" by simp + moreover from \A\<^sub>Q \* X\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>Q \* xvec\ have "A\<^sub>Q \* (X@A\<^sub>P@xvec)" by simp + moreover from \A\<^sub>Q \* Y\ \A\<^sub>Q \* \\<^sub>P\ have "A\<^sub>Q \* (\\<^sub>P#Y)" by force + moreover from \A\<^sub>Q \* Z\ \A\<^sub>Q \* P\ \A\<^sub>Q \* P'\ \A\<^sub>Q \* P''\ have "A\<^sub>Q \* (P#(P'#(P''#Z)))" by simp + moreover note \A\<^sub>Q \* Q'\ \A\<^sub>Q \* Q\ + + ultimately obtain q Q\' A\<^sub>Q' \\<^sub>Q' where Sq: "set q \ set(bn (\M\N\)) \ set (bn(q \ (\M\N\)))" and "((q \ \\<^sub>Q) \ Q\') \ \\<^sub>Q'" + and "distinctPerm q" and FrQ': "extractFrame Q' = \A\<^sub>Q', \\<^sub>Q'\" + and "A\<^sub>Q' \* Q'" and "A\<^sub>Q' \* (\M\N\)" "A\<^sub>Q' \* (q \ (\M\N\))" and "A\<^sub>Q' \* C" and "(bn(q \ (\M\N\))) \* C'" + and "(bn(q \ (\M\N\))) \* (\M\N\)" and "(bn(q \ (\M\N\))) \* Q'" and "distinct A\<^sub>Q'" + and "A\<^sub>Q' \* V" and "A\<^sub>Q' \* (\#W)" and "A\<^sub>Q' \* (X @ A\<^sub>P @ xvec)" and "A\<^sub>Q' \* (\\<^sub>P#Y)" + and "A\<^sub>Q' \* (P#(P'#(P''#Z)))" and "(bn(q \ (\M\N\))) \* (\#W)" and "(bn(q \ (\M\N\))) \* (X @ A\<^sub>P @ xvec)" and "(bn(q \ (\M\N\))) \* (\\<^sub>P#Y)" + and "(bn(q \ (\M\N\))) \* (P#(P'#(P''#Z)))" + by(elim cBrComm2(8)) (assumption | simp)+ + + then have QeqQ': "\\<^sub>Q \ Q\' \ \\<^sub>Q'" + and "A\<^sub>Q' \* Q'" and "A\<^sub>Q' \* (\M\N\)" and "A\<^sub>Q' \* C" and "distinct A\<^sub>Q'" + and "A\<^sub>Q' \* P" and "A\<^sub>Q' \* P'" and "A\<^sub>Q' \* P''" and "A\<^sub>Q' \* Z" and "A\<^sub>Q' \* A\<^sub>P" and "A\<^sub>Q' \* X" and "A\<^sub>Q' \* \\<^sub>P" and "A\<^sub>Q' \* Y" + and "A\<^sub>Q' \* xvec" and "A\<^sub>Q' \* \" and "A\<^sub>Q' \* (subject \)" and "A\<^sub>Q' \* (bn \)" and "A\<^sub>Q' \* (object \)" + and "A\<^sub>Q' \* V" and "A\<^sub>Q' \* W" + by(simp del: freshChainSimps)+ + + from rPerm \bn \ \* \\<^sub>Q\ \xvec \* \\<^sub>Q\ + have "(r \ \\<^sub>Q) = \\<^sub>Q" by simp + + from QeqQ' have "r \ (\\<^sub>Q \ Q\' \ \\<^sub>Q')" + by simp + with \(r \ \\<^sub>Q) = \\<^sub>Q\ + have rQeqQ': "\\<^sub>Q \ (r \ Q\') \ (r \ \\<^sub>Q')" by(simp add: eqvts) + + from rPerm \A\<^sub>Q' \* (bn \)\ \A\<^sub>Q' \* xvec\ + have "r \ A\<^sub>Q' = A\<^sub>Q'" by simp + + from FrQ' have "r \ (extractFrame Q' = \A\<^sub>Q', \\<^sub>Q'\)" + by simp + + with \r \ A\<^sub>Q' = A\<^sub>Q'\ have rFrQ': "extractFrame (r \ Q') = \A\<^sub>Q', (r \ \\<^sub>Q')\" + by(simp add: eqvts) + + from \xvec \* \\ have "(bn \) \* xvec" by simp + + from \(A\<^sub>P @ A\<^sub>Q) \* \\ have "A\<^sub>P \* bn \" by simp + + from \A\<^sub>P \* M\ \A\<^sub>P \* xvec\ \A\<^sub>P \* N\ have "A\<^sub>P \* (\M\\*xvec\\N\)" by simp + from \\ \ \\<^sub>Q \ P \ \M\\*xvec\\N\ \ P'\ \xvec \* M\ \distinct xvec\ \bn \ \* P\ \xvec \* bn \\ + have "bn \ \* P'" by(simp add: broutputFreshChainDerivative) + + from \\ \ \\<^sub>Q \ P \ \M\\*xvec\\N\ \ P'\ \xvec \* M\ \distinct xvec\ \bn \ \* P\ \bn \ \* xvec\ + have "bn \ \* N" by(simp add: broutputFreshChainDerivative) + + note \bn \ \* subject \\ + + moreover from \bn \ \* (P \ Q)\ have "bn \ \* Q" and "bn \ \* P" by simp+ + moreover from \bn \ \* V\ \bn \ \* N\ have "bn \ \* (N#V)" by simp + moreover note \bn \ \* \\<^sub>P\ \bn \ \* \\<^sub>Q\ \bn \ \* W\ + moreover from \bn \ \* X\ \A\<^sub>Q \* bn \\ \A\<^sub>Q' \* bn \\ \bn \ \* xvec\ have "bn \ \* (X@xvec@A\<^sub>Q@A\<^sub>Q')" by simp + moreover from \bn \ \* Y\ \bn \ \* \\<^sub>Q\ have "bn \ \* (\\<^sub>Q#Y)" by simp + moreover from \bn \ \* Z\ \bn \ \* P\ \bn \ \* Q\ have "bn \ \* (P#Q#Z)" by simp + moreover from \A\<^sub>P \* V\ \A\<^sub>P \* N\ have "A\<^sub>P \* (N#V)" by simp + moreover note \A\<^sub>P \* W\ + moreover from \A\<^sub>P \* X\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>Q' \* A\<^sub>P\ \A\<^sub>P \* xvec\ have "A\<^sub>P \* (X@xvec@A\<^sub>Q@A\<^sub>Q')" by simp + moreover from \A\<^sub>P \* Y\ \A\<^sub>P \* \\<^sub>Q\ have "A\<^sub>P \* (\\<^sub>Q#Y)" by force + moreover from \A\<^sub>P \* Z\ \A\<^sub>P \* Q\ \A\<^sub>P \* P\ have "A\<^sub>P \* (P#Q#Z)" by simp + + moreover note \\M\\*xvec\\N\ \ P' = \ \ P''\ + ultimately obtain p P\' A\<^sub>P' \\<^sub>P' where Sp: "set p \ set (bn \) \ set (bn (p \ \))" and PeqP': "((p \ \\<^sub>P) \ P\') \ \\<^sub>P'" + and "distinctPerm p" and "bn(p \ \) \* C'" and FrP': "extractFrame P'' = \A\<^sub>P', \\<^sub>P'\" + and "A\<^sub>P' \* P''" and "A\<^sub>P' \* \" and "A\<^sub>P' \* (p \ \)" and "A\<^sub>P' \* C" + and "bn(p \ \) \* \" and "bn(p \ \) \* P''" and "distinct A\<^sub>P'" + and "A\<^sub>P' \* (N#V)" and "A\<^sub>P' \* W" and "A\<^sub>P' \* (X @ xvec @ A\<^sub>Q @ A\<^sub>Q')" and "A\<^sub>P' \* (\\<^sub>Q#Y)" + and "A\<^sub>P' \* (P#Q#Z)" and "bn(p \ \) \* (N#V)" and "bn(p \ \) \* W" and "bn(p \ \) \* (X @ xvec @ A\<^sub>Q @ A\<^sub>Q')" and "bn(p \ \) \* (\\<^sub>Q#Y)" + and "bn(p \ \) \* (P#Q#Z)" + using \A\<^sub>P \* P\ \A\<^sub>P \* \\ \A\<^sub>P \* C\ \A\<^sub>P \* C'\ + \bn \ \* C'\ \\ \ \\ \A\<^sub>P \* P''\ \distinct (bn \)\ + by(elim cBrComm2(4)[where b=C and ba=C' and bb=\ and bc=P'' and bf="(X @ xvec @ A\<^sub>Q @ A\<^sub>Q')" and bg="(\\<^sub>Q#Y)" and bh="(P#Q#Z)"]) (assumption | simp)+ + then have "A\<^sub>P' \* Q" and "A\<^sub>P' \* Z" and "A\<^sub>P' \* A\<^sub>Q" and "A\<^sub>Q' \* A\<^sub>P'" and "A\<^sub>P' \* X" and "A\<^sub>P' \* \\<^sub>Q" and "A\<^sub>P' \* Y" and "A\<^sub>P' \* P" and "A\<^sub>P' \* N" + and "A\<^sub>P' \* xvec" and "A\<^sub>P' \* V" and "A\<^sub>P' \* W" + and "bn(p \ \) \* A\<^sub>Q" and "bn(p \ \) \* A\<^sub>Q'" and "bn(p \ \) \* X" and "bn(p \ \) \* Y" and "bn(p \ \) \* Z" and "bn(p \ \) \* \\<^sub>Q" + and "bn(p \ \) \* Q" and "bn(p \ \) \* W" and "bn(p \ \) \* V" and "bn(p \ \) \* N" and "bn(p \ \) \* xvec" + by(simp del: freshChainSimps)+ + + from \A\<^sub>Q' \* P''\ \A\<^sub>Q' \* A\<^sub>P'\ FrP' + have "A\<^sub>Q' \* \\<^sub>P'" + by(metis extractFrameFreshChain freshFrameDest) + + from \\ \ \\<^sub>P \ Q \ \M\N\ \ Q'\ \A\<^sub>P' \* Q\ \A\<^sub>P' \* N\ + have "A\<^sub>P' \* Q'" + by(simp add: brinputFreshChainDerivative) + with FrQ' \A\<^sub>Q' \* A\<^sub>P'\ have "A\<^sub>P' \* \\<^sub>Q'" + by(metis extractFrameFreshChain freshFrameDest) + + have "((p \ \\<^sub>P) \ P\') \ (\\<^sub>Q \ (r \ Q\')) \ \\<^sub>P' \ (r \ \\<^sub>Q')" + by (metis Composition' PeqP' rQeqQ') + then have "((p \ \\<^sub>P) \ (P\' \ (\\<^sub>Q \ (r \ Q\')))) \ \\<^sub>P' \ (r \ \\<^sub>Q')" + by (metis AssertionStatEqSym AssertionStatEqTrans Associativity) + then have "((p \ \\<^sub>P) \ ((\\<^sub>Q \ (r \ Q\')) \ P\')) \ \\<^sub>P' \ (r \ \\<^sub>Q')" + by (metis AssertionStatEqSym AssertionStatEqTrans Associativity associativitySym) + then have "(p \ \\<^sub>P) \ (\\<^sub>Q \ ((r \ Q\') \ P\')) \ \\<^sub>P' \ (r \ \\<^sub>Q')" + by (metis AssertionStatEqSym AssertionStatEqTrans Associativity compositionSym) + then have "((p \ \\<^sub>P) \ \\<^sub>Q) \ ((r \ Q\') \ P\') \ \\<^sub>P' \ (r \ \\<^sub>Q')" + by (metis AssertionStatEqTrans Associativity) + then have "((p \ \\<^sub>P) \ \\<^sub>Q) \ (P\' \ (r \ Q\')) \ \\<^sub>P' \ (r \ \\<^sub>Q')" + by (metis AssertionStatEqSym AssertionStatEqTrans Associativity associativitySym) + with Sp \bn \ \* \\<^sub>Q\ \bn (p \ \) \* \\<^sub>Q\ have "(p \ (\\<^sub>P \ \\<^sub>Q)) \ (P\' \ (r \ Q\')) \ \\<^sub>P' \ (r \ \\<^sub>Q')" + by (simp add: eqvts) + + from Sp \A\<^sub>Q' \* \\ \bn(p \ \) \* A\<^sub>Q'\ have "A\<^sub>Q' \* (p \ \)" + by (metis actionFreshChain freshChainSym freshStarChainSimps fresh_star_set_eq) + + from \A\<^sub>P' \* \\ have "A\<^sub>P' \* bn \" by simp + from rPerm \A\<^sub>Q' \* Q'\ \A\<^sub>Q' \* xvec\ \A\<^sub>Q' \* bn \\ + have "A\<^sub>Q' \* (r \ Q')" + by (metis freshAlphaPerm freshChainSym name_list_set_fresh permStarFresh) + from rPerm \A\<^sub>P' \* Q'\ \A\<^sub>P' \* xvec\ \A\<^sub>P' \* bn \\ + have "A\<^sub>P' \* (r \ Q')" + by (metis freshAlphaPerm freshChainSym name_list_set_fresh permStarFresh) + from rPerm \A\<^sub>P' \* \\<^sub>Q'\ \A\<^sub>P' \* xvec\ \A\<^sub>P' \* bn \\ + have "A\<^sub>P' \* (r \ \\<^sub>Q')" + by (metis freshAlphaPerm freshChainSym name_list_set_fresh permStarFresh) + + with \A\<^sub>Q' \* \\<^sub>P'\ \A\<^sub>Q' \* A\<^sub>P'\ FrP' rFrQ' + have "extractFrame (P'' \ (r \ Q')) = \(A\<^sub>P'@A\<^sub>Q'), (\\<^sub>P' \ (r \ \\<^sub>Q'))\" + by simp + with \PQ' = (P'' \ (r \ Q'))\ + have "extractFrame PQ' = \(A\<^sub>P'@A\<^sub>Q'), (\\<^sub>P' \ (r \ \\<^sub>Q'))\" + by simp + + from \\ \ \\<^sub>P \ Q \ \M\N\ \ Q'\ \bn(p \ \) \* Q\ \bn(p \ \) \* N\ + have "bn(p \ \) \* Q'" by(simp add: brinputFreshChainDerivative) + + with rPerm \bn(p \ \) \* \\ \bn(p \ \) \* xvec\ + have "bn(p \ \) \* (r \ Q')" + by (metis actionFreshChain freshAlphaPerm freshChainSym name_list_set_fresh permStarFresh) + + from \A\<^sub>P' \* P''\ \A\<^sub>P' \* (r \ Q')\ \PQ' = P'' \ (r \ Q')\ have "A\<^sub>P' \* PQ'" by simp + from \A\<^sub>Q' \* P''\ \A\<^sub>Q' \* (r \ Q')\ \PQ' = P'' \ (r \ Q')\ have "A\<^sub>Q' \* PQ'" by simp + + note \set p \ (set (bn \)) \ set (bn(p \ \))\ + \((p \ (\\<^sub>P \ \\<^sub>Q)) \ (P\' \ (r \ Q\'))) \ \\<^sub>P' \ (r \ \\<^sub>Q')\ \distinctPerm p\ + \extractFrame PQ' = \(A\<^sub>P'@A\<^sub>Q'), (\\<^sub>P' \ (r \ \\<^sub>Q'))\\ + + moreover from \A\<^sub>P' \* PQ'\ \A\<^sub>Q' \* PQ'\ have "(A\<^sub>P'@A\<^sub>Q') \* PQ'" by simp + moreover from \A\<^sub>P' \* \\ \A\<^sub>Q' \* \\ have "(A\<^sub>P'@A\<^sub>Q') \* \" by simp + moreover from \A\<^sub>P' \* (p \ \)\ \A\<^sub>Q' \* (p \ \)\ have "(A\<^sub>P'@A\<^sub>Q') \* (p \ \)" by simp + moreover from \A\<^sub>P' \* C\ \A\<^sub>Q' \* C\ have "(A\<^sub>P'@A\<^sub>Q') \* C" by simp + moreover note \bn(p \ \) \* C'\ \bn(p \ \) \* \\ + moreover from \bn(p \ \) \* (r \ Q')\ \bn(p \ \) \* P''\ \PQ' = P'' \ (r \ Q')\ + have "bn(p \ \) \* PQ'" by simp + moreover from \A\<^sub>P' \* V\ \A\<^sub>Q' \* V\ have "(A\<^sub>P'@A\<^sub>Q') \* V" by simp + moreover from \A\<^sub>P' \* W\ \A\<^sub>Q' \* W\ have "(A\<^sub>P'@A\<^sub>Q') \* W" by simp + moreover from \A\<^sub>P' \* X\ \A\<^sub>Q' \* X\ have "(A\<^sub>P'@A\<^sub>Q') \* X" by simp + moreover from \A\<^sub>P' \* Y\ \A\<^sub>Q' \* Y\ have "(A\<^sub>P'@A\<^sub>Q') \* Y" by simp + moreover from \A\<^sub>P' \* Z\ \A\<^sub>Q' \* Z\ have "(A\<^sub>P'@A\<^sub>Q') \* Z" by simp + moreover from \distinct A\<^sub>P'\ \distinct A\<^sub>Q'\ \A\<^sub>Q' \* A\<^sub>P'\ + have "distinct(A\<^sub>P'@A\<^sub>Q')" by simp + moreover note \bn(p \ \) \* V\ \bn(p \ \) \* W\ + \bn(p \ \) \* X\ \bn(p \ \) \* Y\ \bn(p \ \) \* Z\ + ultimately show ?case + by(rule cBrComm2(63)) + next + case cBrClose + then show ?case + by(simp add: residualInject) + next + case(cOpen \ P M xvec1 xvec2 N P' x A\<^sub>P \\<^sub>P C C' \ P'' V W X Y Z) + from \M\\*(xvec1@x#xvec2)\\N\ \ P' = \ \ P''\ \x \ xvec1\ \x \ xvec2\ \x \ \\ \x \ P''\ \distinct(bn \)\ \A\<^sub>P \* \\ \x \ \\ + obtain yvec1 y yvec2 N' where yvecEq: "bn \ = yvec1@y#yvec2" and P'eqP'': "\\*(xvec1@xvec2)\N \' P' = \\*(yvec1@yvec2)\([(x, y)] \ N') \' ([(x, y)] \ P'')" and "A\<^sub>P \* N'" and Subj: "subject \ = Some M" and "x \ N'" and \eq: "\ = M\\*(yvec1@y#yvec2)\\N'\" + apply(cases rule: actionCases[where \=\]) + apply(simp add: residualInject) + apply(simp add: residualInject) + apply(simp add: residualInject) + apply(metis boundOutputOpenDest) + apply(simp add: residualInject) + by(simp add: residualInject) + + note \A\<^sub>P \* P\ \A\<^sub>P \* M\ + moreover from Subj yvecEq \bn \ \* subject \\ have "yvec1 \* M" "yvec2 \* M" by simp+ + moreover from yvecEq \A\<^sub>P \* \\ have "A\<^sub>P \* (yvec1@yvec2)" by simp + moreover note \A\<^sub>P \* C\ + moreover from yvecEq \bn \ \* \\x\P\ \x \ \\ have "(yvec1@yvec2) \* P" by simp + moreover from yvecEq \bn \ \* C'\ \bn \ \* V\ \bn \ \* W\ \bn \ \* X\ \bn \ \* Y\ \bn \ \* Z\ \distinct(bn \)\ \x \ \\ + have "(yvec1@yvec2) \* C'" and "(yvec1@yvec2) \* V" and "(yvec1@yvec2) \* W" and "(yvec1@yvec2) \* (x#y#X)" and "(yvec1@yvec2) \* Y" and "(yvec1@yvec2) \* Z" + by simp+ + moreover note \A\<^sub>P \* V\ \A\<^sub>P \* W\ + moreover from \A\<^sub>P \* X\ \x \ A\<^sub>P\ \A\<^sub>P \* \\ yvecEq have "A\<^sub>P \* (x#y#X)" by simp + moreover note \A\<^sub>P \* Y\ \A\<^sub>P \* Z\ + moreover from \A\<^sub>P \* N'\ \A\<^sub>P \* P''\ \x \ A\<^sub>P\ \A\<^sub>P \* \\ yvecEq have "A\<^sub>P \* ([(x, y)] \ N')" and "A\<^sub>P \* ([(x, y)] \ P'')" + by simp+ + moreover from yvecEq \distinct(bn \)\ have "distinct(yvec1@yvec2)" by simp + moreover from P'eqP'' have "M\\*(xvec1@xvec2)\\N\ \ P' = M\\*(yvec1@yvec2)\\([(x, y)] \ N')\ \ ([(x, y)] \ P'')" + by(simp add: residualInject) + ultimately obtain p \' A\<^sub>P' \\<^sub>P' where S: "set p \ set (yvec1@yvec2) \ set (p \ (yvec1@yvec2))" and PeqP': "((p \ \\<^sub>P) \ \') \ \\<^sub>P'" + and "distinctPerm p" and "(p \ (yvec1@yvec2)) \* C'" and FrP': "extractFrame([(x, y)] \ P'') = \A\<^sub>P', \\<^sub>P'\" + and "A\<^sub>P' \* ([(x, y)] \ P'')" and "A\<^sub>P' \* ([(x, y)] \ N')" and "A\<^sub>P' \* C" and "(p \ (yvec1@yvec2)) \* ([(x, y)] \ N')" and "A\<^sub>P' \* M" and "(p \ (yvec1@yvec2)) \* (yvec1@yvec2)" and "(p \ (yvec1@yvec2)) \* M" and "distinct A\<^sub>P'" + and "(p \ (yvec1@yvec2)) \* ([(x, y)] \ P'')" and "(yvec1@yvec2) \* A\<^sub>P'" and "(p \ (yvec1@yvec2)) \* A\<^sub>P'" + and "A\<^sub>P' \* V" and "A\<^sub>P' \* W" and "A\<^sub>P' \* (x#y#X)" and "A\<^sub>P' \* Y" and "A\<^sub>P' \* Z" and "(p \ (yvec1@yvec2)) \* (x#y#X)" + and "(p \ (yvec1@yvec2)) \* V" and "(p \ (yvec1@yvec2)) \* W" and "(p \ (yvec1@yvec2)) \* Y" and "(p \ (yvec1@yvec2)) \* Z" using \A\<^sub>P \* C'\ + by(elim cOpen(4)[where b=C and ba=C' and bd=V and be=W and bf="x#y#X" and bg=Y and bh=Z]) (assumption | simp)+ + + from \A\<^sub>P' \* (x#y#X)\ have "x \ A\<^sub>P'" and "y \ A\<^sub>P'" and "A\<^sub>P' \* X" by simp+ + from \(p \ (yvec1@yvec2)) \* (x#y#X)\ have "x \ (p \ (yvec1@yvec2))" and "y \ (p \ (yvec1@yvec2))" and "(p \ (yvec1@yvec2)) \* X" by simp+ + + from \x \ \\ yvecEq have "x \ yvec1" and "x \ y" and "x \ yvec2" by simp+ + from \distinct(bn \)\ yvecEq have "yvec1 \* yvec2" and "y \ yvec1" and "y \ yvec2" by simp+ + from \bn \ \* C'\ yvecEq have "yvec1 \* C'" and "y \ C'" and "yvec2 \* C'" by simp+ + + from S \x \ \\ \x \ p \ (yvec1@yvec2)\ yvecEq have "x \ p" by(intro freshAlphaSwap) (assumption | simp)+ + from S \distinct(bn \)\ \y \ p \ (yvec1@yvec2)\ yvecEq have "y \ p" by(intro freshAlphaSwap) (assumption | simp)+ + + from yvecEq S \x \ yvec1\ \x \ yvec2\ \y \ yvec1\ \y \ yvec2\ \x \ p \ (yvec1@yvec2)\ \y \ p \ (yvec1@yvec2)\ + have "set ((y, x)#p) \ set(bn \) \ set(bn(((y, x)#p) \ \))" + apply(simp add: bnEqvt[symmetric]) + by(auto simp add: eqvts calc_atm) + + moreover from PeqP' have "([(y, x)] \ ((p \ \\<^sub>P) \ \')) \ [(y, x)] \ \\<^sub>P'" + by(simp add: AssertionStatEqClosed) + then have "(((y, x)#p) \ \\<^sub>P) \ ([(y, x)] \ \') \ ([(y, x)] \ \\<^sub>P')" + by(simp add: eqvts) + moreover from \distinctPerm p\ S \x \ y\ \x \ p\ \y \ p\ have "distinctPerm((y, x)#p)" + by simp + moreover from FrP' have "([(x, y)] \ (extractFrame([(x, y)] \ P''))) = ([(x, y)] \ \A\<^sub>P', \\<^sub>P'\)" + by simp + with \x \ A\<^sub>P'\ \y \ A\<^sub>P'\ have "extractFrame P'' = \A\<^sub>P', ([(y, x)] \ \\<^sub>P')\" + by(simp add: eqvts name_swap) + moreover from \x \ p\ \y \ p\ \x \ N'\ \(p \ (yvec1@yvec2)) \* ([(x, y)] \ N')\ have "([(y, x)] \ p \ (yvec1@yvec2)) \* ([(y, x)] \ [(x, y)] \ N')" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + then have "(((y, x)#p) \ (yvec1@yvec2)) \* N'" by(simp add: name_swap) + with \x \ yvec1\ \x \ yvec2\ \x \ y\ \x \ C\ \y \ yvec1\ \y \ yvec2\ \x \ p\ \y \ p\ \x \ N'\ yvecEq + have "bn(((y, x)#p) \ \) \* N'" by(simp add: bnEqvt[symmetric]) (simp add: eqvts perm_compose calc_atm freshChainSimps) + moreover from \x \ p\ \y \ p\ \x \ N'\ \(p \ (yvec1@yvec2)) \* ([(x, y)] \ P'')\ have "([(y, x)] \ p \ (yvec1@yvec2)) \* ([(y, x)] \ [(x, y)] \ P'')" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + then have "(((y, x)#p) \ (yvec1@yvec2)) \* P''" by(simp add: name_swap) + with \x \ yvec1\ \x \ yvec2\ \x \ y\ \y \ yvec1\ \y \ yvec2\ \x \ p\ \y \ p\ \x \ P''\ yvecEq + have "bn(((y, x)#p) \ \) \* P''" by(simp add: bnEqvt[symmetric]) (simp add: perm_compose calc_atm eqvts freshChainSimps) + moreover from \x \ p\ \y \ p\ \x \ A\<^sub>P'\ \(p \ (yvec1@yvec2)) \* A\<^sub>P'\ have "(p \ (yvec1@x#yvec2)) \* A\<^sub>P'" + by(simp add: eqvts freshChainSimps) + then have "([(y, x)] \ p \ (yvec1@x#yvec2)) \* ([(y, x)] \ A\<^sub>P')" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \x \ yvec1\ \x \ yvec2\ \x \ y\ \x \ A\<^sub>P'\ \y \ yvec1\ \y \ yvec2\ \x \ p\ \y \ p\ \y \ A\<^sub>P'\ yvecEq + have "bn(((y, x)#p) \ \) \* A\<^sub>P'" by(simp add: bnEqvt[symmetric]) (simp add: perm_compose calc_atm eqvts freshChainSimps) + moreover from \x \ p\ \y \ p\ \x \ C'\ \(p \ (yvec1@yvec2)) \* C'\ have "(p \ (yvec1@x#yvec2)) \* C'" + by(simp add: eqvts freshChainSimps) + then have "([(y, x)] \ p \ (yvec1@x#yvec2)) \* ([(y, x)] \ C')" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \x \ yvec1\ \x \ yvec2\ \x \ y\ \x \ C'\ \y \ yvec1\ \y \ yvec2\ \x \ p\ \y \ p\ \y \ C'\ yvecEq + have "bn(((y, x)#p) \ \) \* C'" by(simp add: bnEqvt[symmetric]) (simp add: perm_compose calc_atm eqvts freshChainSimps) + moreover from \x \ p\ \y \ p\ \x \ X\ \(p \ (yvec1@yvec2)) \* X\ have "(p \ (yvec1@x#yvec2)) \* X" + by(simp add: eqvts freshChainSimps) + then have "([(y, x)] \ p \ (yvec1@x#yvec2)) \* ([(y, x)] \ X)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \x \ yvec1\ \x \ yvec2\ \x \ y\ \x \ X\ \y \ yvec1\ \y \ yvec2\ \x \ p\ \y \ p\ \bn \ \* X\ yvecEq + have "bn(((y, x)#p) \ \) \* X" by(simp add: bnEqvt[symmetric]) (simp add: perm_compose calc_atm eqvts freshChainSimps) + moreover from \x \ p\ \y \ p\ \x \ V\ \(p \ (yvec1@yvec2)) \* V\ have "(p \ (yvec1@x#yvec2)) \* V" + by(simp add: eqvts freshChainSimps) + then have "([(y, x)] \ p \ (yvec1@x#yvec2)) \* ([(y, x)] \ V)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \x \ yvec1\ \x \ yvec2\ \x \ y\ \x \ V\ \y \ yvec1\ \y \ yvec2\ \x \ p\ \y \ p\ \bn \ \* V\ yvecEq + have "bn(((y, x)#p) \ \) \* V" by(simp add: bnEqvt[symmetric]) (simp add: perm_compose calc_atm eqvts freshChainSimps) + moreover from \x \ p\ \y \ p\ \x \ W\ \(p \ (yvec1@yvec2)) \* W\ have "(p \ (yvec1@x#yvec2)) \* W" + by(simp add: eqvts freshChainSimps) + then have "([(y, x)] \ p \ (yvec1@x#yvec2)) \* ([(y, x)] \ W)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \x \ yvec1\ \x \ yvec2\ \x \ y\ \x \ W\ \y \ yvec1\ \y \ yvec2\ \x \ p\ \y \ p\ \bn \ \* W\ yvecEq + have "bn(((y, x)#p) \ \) \* W" by(simp add: bnEqvt[symmetric]) (simp add: perm_compose calc_atm eqvts freshChainSimps) + moreover from \x \ p\ \y \ p\ \x \ Y\ \(p \ (yvec1@yvec2)) \* Y\ have "(p \ (yvec1@x#yvec2)) \* Y" + by(simp add: eqvts freshChainSimps) + then have "([(y, x)] \ p \ (yvec1@x#yvec2)) \* ([(y, x)] \ Y)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \x \ yvec1\ \x \ yvec2\ \x \ y\ \x \ Y\ \y \ yvec1\ \y \ yvec2\ \x \ p\ \y \ p\ \bn \ \* Y\ yvecEq + have "bn(((y, x)#p) \ \) \* Y" by(simp add: bnEqvt[symmetric]) (simp add: perm_compose calc_atm eqvts freshChainSimps) + moreover from \x \ p\ \y \ p\ \x \ Z\ \(p \ (yvec1@yvec2)) \* Z\ have "(p \ (yvec1@x#yvec2)) \* Z" + by(simp add: eqvts freshChainSimps) + then have "([(y, x)] \ p \ (yvec1@x#yvec2)) \* ([(y, x)] \ Z)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \x \ yvec1\ \x \ yvec2\ \x \ y\ \x \ Z\ \y \ yvec1\ \y \ yvec2\ \x \ p\ \y \ p\ \bn \ \* Z\ yvecEq + have "bn(((y, x)#p) \ \) \* Z" by(simp add: bnEqvt[symmetric]) (simp add: perm_compose calc_atm eqvts freshChainSimps) + moreover from \(yvec1@yvec2) \* A\<^sub>P'\ \y \ A\<^sub>P'\ yvecEq have "bn \ \* A\<^sub>P'" + by simp + moreover from \A\<^sub>P' \* ([(x, y)] \ N')\ have "([(x, y)] \ A\<^sub>P') \* ([(x, y)] \ [(x, y)] \ N')" + by(simp only: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \x \ A\<^sub>P'\ \y \ A\<^sub>P'\ have "A\<^sub>P' \* N'" by simp + with \A\<^sub>P' \* M\ \(yvec1@yvec2) \* A\<^sub>P'\ \y \ A\<^sub>P'\ \eq have "A\<^sub>P' \* \" by simp + moreover then have "(((y, x)#p) \ A\<^sub>P') \* (((y, x)#p) \ \)" + by(simp only: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \x \ A\<^sub>P'\ \y \ A\<^sub>P'\ S \(yvec1@yvec2) \* A\<^sub>P'\ \(p \ (yvec1@yvec2)) \* A\<^sub>P'\ + have "A\<^sub>P' \* (((y, x)#p) \ \)" by(simp add: eqvts) + moreover from \A\<^sub>P' \* ([(x, y)] \ P'')\ have "([(x, y)] \ A\<^sub>P') \* ([(x, y)] \ [(x, y)] \ P'')" + by(simp only: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \x \ A\<^sub>P'\ \y \ A\<^sub>P'\ have "A\<^sub>P' \* P''" by simp + moreover from yvecEq \eq \(p \ (yvec1@yvec2)) \* (yvec1@yvec2)\ \y \ p\ \x \ \\ S \(p \ (yvec1@yvec2)) \* M\\(p \ (yvec1@yvec2)) \* ([(x, y)] \ N')\ \y \ yvec1\\y \ yvec2\ \x \ p\ + have "bn(((y, x)#p) \ \) \* \" + apply(simp add: eqvts del: set_append) + apply(intro conjI) + apply(simp add: perm_compose eqvts del: set_append) + apply(simp add: perm_compose freshChainSimps(6) calc_atm eqvts del: set_append) + apply(simp add: perm_compose eqvts del: set_append) + apply(simp add: perm_compose eqvts swapStarFresh del: set_append) + apply(simp add: perm_compose freshChainSimps(6) calc_atm eqvts del: set_append) + apply(simp add: perm_compose freshChainSimps(6) calc_atm eqvts del: set_append) + apply(simp add: perm_compose freshChainSimps(6) calc_atm eqvts del: set_append) + apply(simp add: perm_compose freshChainSimps(6) swapStarFresh calc_atm eqvts del: set_append) + apply(simp add: perm_compose freshChainSimps(6) calc_atm eqvts del: set_append) + apply(subst pt_fresh_star_bij[symmetric, OF pt_name_inst, OF at_name_inst, where pi="[(x, y)]"]) + apply(simp add: perm_compose freshChainSimps(6) calc_atm eqvts del: set_append) + apply(simp add: perm_compose freshChainSimps(6) calc_atm eqvts del: set_append) + apply(subst pt_fresh_star_bij[symmetric, OF pt_name_inst, OF at_name_inst, where pi="[(x, y)]"]) + by(simp add: perm_compose freshChainSimps(6) calc_atm eqvts del: set_append) + moreover note \A\<^sub>P' \* C\ \A\<^sub>P' \* V\ \A\<^sub>P' \* W\ \A\<^sub>P' \* X\ \A\<^sub>P' \* Y\ \A\<^sub>P' \* Z\ \distinct A\<^sub>P'\ + + ultimately show ?case + by(elim cOpen) + next + case(cBrOpen \ P M xvec1 xvec2 N P' x A\<^sub>P \\<^sub>P C C' \ P'' V W X Y Z) + from \\M\\*(xvec1@x#xvec2)\\N\ \ P' = \ \ P''\ \x \ xvec1\ \x \ xvec2\ \x \ \\ \x \ P''\ \distinct(bn \)\ \A\<^sub>P \* \\ \x \ \\ + obtain yvec1 y yvec2 N' where yvecEq: "bn \ = yvec1@y#yvec2" and P'eqP'': "\\*(xvec1@xvec2)\N \' P' = \\*(yvec1@yvec2)\([(x, y)] \ N') \' ([(x, y)] \ P'')" and "A\<^sub>P \* N'" and Subj: "subject \ = Some M" and "x \ N'" and \eq: "\ = \M\\*(yvec1@y#yvec2)\\N'\" + apply(cases rule: actionCases[where \=\]) + apply(simp_all add: residualInject) + by (metis boundOutputOpenDest) + + note \A\<^sub>P \* P\ \A\<^sub>P \* M\ + moreover from Subj yvecEq \bn \ \* subject \\ have "yvec1 \* M" "yvec2 \* M" by simp+ + moreover from yvecEq \A\<^sub>P \* \\ have "A\<^sub>P \* (yvec1@yvec2)" by simp + moreover note \A\<^sub>P \* C\ + moreover from yvecEq \bn \ \* \\x\P\ \x \ \\ have "(yvec1@yvec2) \* P" by simp + moreover from yvecEq \bn \ \* C'\ \bn \ \* V\ \bn \ \* W\ \bn \ \* X\ \bn \ \* Y\ \bn \ \* Z\ \distinct(bn \)\ \x \ \\ + have "(yvec1@yvec2) \* C'" and "(yvec1@yvec2) \* V" and "(yvec1@yvec2) \* W" and "(yvec1@yvec2) \* (x#y#X)" and "(yvec1@yvec2) \* Y" and "(yvec1@yvec2) \* Z" + by simp+ + moreover note \A\<^sub>P \* V\ \A\<^sub>P \* W\ + moreover from \A\<^sub>P \* X\ \x \ A\<^sub>P\ \A\<^sub>P \* \\ yvecEq have "A\<^sub>P \* (x#y#X)" by simp + moreover note \A\<^sub>P \* Y\ \A\<^sub>P \* Z\ + moreover from \A\<^sub>P \* N'\ \A\<^sub>P \* P''\ \x \ A\<^sub>P\ \A\<^sub>P \* \\ yvecEq have "A\<^sub>P \* ([(x, y)] \ N')" and "A\<^sub>P \* ([(x, y)] \ P'')" + by simp+ + moreover from yvecEq \distinct(bn \)\ have "distinct(yvec1@yvec2)" by simp + moreover from P'eqP'' have "\M\\*(xvec1@xvec2)\\N\ \ P' = \M\\*(yvec1@yvec2)\\([(x, y)] \ N')\ \ ([(x, y)] \ P'')" + by(simp add: residualInject) + ultimately obtain p \' A\<^sub>P' \\<^sub>P' where S: "set p \ set (yvec1@yvec2) \ set (p \ (yvec1@yvec2))" and PeqP': "((p \ \\<^sub>P) \ \') \ \\<^sub>P'" + and "distinctPerm p" and "(p \ (yvec1@yvec2)) \* C'" and FrP': "extractFrame([(x, y)] \ P'') = \A\<^sub>P', \\<^sub>P'\" + and "A\<^sub>P' \* ([(x, y)] \ P'')" and "A\<^sub>P' \* ([(x, y)] \ N')" and "A\<^sub>P' \* C" and "(p \ (yvec1@yvec2)) \* ([(x, y)] \ N')" and "A\<^sub>P' \* M" and "(p \ (yvec1@yvec2)) \* (yvec1@yvec2)" and "(p \ (yvec1@yvec2)) \* M" and "distinct A\<^sub>P'" + and "(p \ (yvec1@yvec2)) \* ([(x, y)] \ P'')" and "(yvec1@yvec2) \* A\<^sub>P'" and "(p \ (yvec1@yvec2)) \* A\<^sub>P'" + and "A\<^sub>P' \* V" and "A\<^sub>P' \* W" and "A\<^sub>P' \* (x#y#X)" and "A\<^sub>P' \* Y" and "A\<^sub>P' \* Z" and "(p \ (yvec1@yvec2)) \* (x#y#X)" + and "(p \ (yvec1@yvec2)) \* V" and "(p \ (yvec1@yvec2)) \* W" and "(p \ (yvec1@yvec2)) \* Y" and "(p \ (yvec1@yvec2)) \* Z" using \A\<^sub>P \* C'\ + by(elim cBrOpen(4)[where b=C and ba=C' and bd=V and be=W and bf="x#y#X" and bg=Y and bh=Z]) (assumption | simp)+ + + from \A\<^sub>P' \* (x#y#X)\ have "x \ A\<^sub>P'" and "y \ A\<^sub>P'" and "A\<^sub>P' \* X" by simp+ + from \(p \ (yvec1@yvec2)) \* (x#y#X)\ have "x \ (p \ (yvec1@yvec2))" and "y \ (p \ (yvec1@yvec2))" and "(p \ (yvec1@yvec2)) \* X" by simp+ + + from \x \ \\ yvecEq have "x \ yvec1" and "x \ y" and "x \ yvec2" by simp+ + from \distinct(bn \)\ yvecEq have "yvec1 \* yvec2" and "y \ yvec1" and "y \ yvec2" by simp+ + from \bn \ \* C'\ yvecEq have "yvec1 \* C'" and "y \ C'" and "yvec2 \* C'" by simp+ + + from S \x \ \\ \x \ p \ (yvec1@yvec2)\ yvecEq have "x \ p" by(intro freshAlphaSwap) (assumption | simp)+ + from S \distinct(bn \)\ \y \ p \ (yvec1@yvec2)\ yvecEq have "y \ p" by(intro freshAlphaSwap) (assumption | simp)+ + + from yvecEq S \x \ yvec1\ \x \ yvec2\ \y \ yvec1\ \y \ yvec2\ \x \ p \ (yvec1@yvec2)\ \y \ p \ (yvec1@yvec2)\ + have "set ((y, x)#p) \ set(bn \) \ set(bn(((y, x)#p) \ \))" + apply(simp add: bnEqvt[symmetric]) + by(auto simp add: eqvts calc_atm) + + moreover from PeqP' have "([(y, x)] \ ((p \ \\<^sub>P) \ \')) \ [(y, x)] \ \\<^sub>P'" + by(simp add: AssertionStatEqClosed) + then have "(((y, x)#p) \ \\<^sub>P) \ ([(y, x)] \ \') \ ([(y, x)] \ \\<^sub>P')" + by(simp add: eqvts) + moreover from \distinctPerm p\ S \x \ y\ \x \ p\ \y \ p\ have "distinctPerm((y, x)#p)" + by simp + moreover from FrP' have "([(x, y)] \ (extractFrame([(x, y)] \ P''))) = ([(x, y)] \ \A\<^sub>P', \\<^sub>P'\)" + by simp + with \x \ A\<^sub>P'\ \y \ A\<^sub>P'\ have "extractFrame P'' = \A\<^sub>P', ([(y, x)] \ \\<^sub>P')\" + by(simp add: eqvts name_swap) + moreover from \x \ p\ \y \ p\ \x \ N'\ \(p \ (yvec1@yvec2)) \* ([(x, y)] \ N')\ have "([(y, x)] \ p \ (yvec1@yvec2)) \* ([(y, x)] \ [(x, y)] \ N')" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + then have "(((y, x)#p) \ (yvec1@yvec2)) \* N'" by(simp add: name_swap) + with \x \ yvec1\ \x \ yvec2\ \x \ y\ \x \ C\ \y \ yvec1\ \y \ yvec2\ \x \ p\ \y \ p\ \x \ N'\ yvecEq + have "bn(((y, x)#p) \ \) \* N'" by(simp add: bnEqvt[symmetric]) (simp add: eqvts perm_compose calc_atm freshChainSimps) + moreover from \x \ p\ \y \ p\ \x \ N'\ \(p \ (yvec1@yvec2)) \* ([(x, y)] \ P'')\ have "([(y, x)] \ p \ (yvec1@yvec2)) \* ([(y, x)] \ [(x, y)] \ P'')" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + then have "(((y, x)#p) \ (yvec1@yvec2)) \* P''" by(simp add: name_swap) + with \x \ yvec1\ \x \ yvec2\ \x \ y\ \y \ yvec1\ \y \ yvec2\ \x \ p\ \y \ p\ \x \ P''\ yvecEq + have "bn(((y, x)#p) \ \) \* P''" by(simp add: bnEqvt[symmetric]) (simp add: perm_compose calc_atm eqvts freshChainSimps) + moreover from \x \ p\ \y \ p\ \x \ A\<^sub>P'\ \(p \ (yvec1@yvec2)) \* A\<^sub>P'\ have "(p \ (yvec1@x#yvec2)) \* A\<^sub>P'" + by(simp add: eqvts freshChainSimps) + then have "([(y, x)] \ p \ (yvec1@x#yvec2)) \* ([(y, x)] \ A\<^sub>P')" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \x \ yvec1\ \x \ yvec2\ \x \ y\ \x \ A\<^sub>P'\ \y \ yvec1\ \y \ yvec2\ \x \ p\ \y \ p\ \y \ A\<^sub>P'\ yvecEq + have "bn(((y, x)#p) \ \) \* A\<^sub>P'" by(simp add: bnEqvt[symmetric]) (simp add: perm_compose calc_atm eqvts freshChainSimps) + moreover from \x \ p\ \y \ p\ \x \ C'\ \(p \ (yvec1@yvec2)) \* C'\ have "(p \ (yvec1@x#yvec2)) \* C'" + by(simp add: eqvts freshChainSimps) + then have "([(y, x)] \ p \ (yvec1@x#yvec2)) \* ([(y, x)] \ C')" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \x \ yvec1\ \x \ yvec2\ \x \ y\ \x \ C'\ \y \ yvec1\ \y \ yvec2\ \x \ p\ \y \ p\ \y \ C'\ yvecEq + have "bn(((y, x)#p) \ \) \* C'" by(simp add: bnEqvt[symmetric]) (simp add: perm_compose calc_atm eqvts freshChainSimps) + moreover from \x \ p\ \y \ p\ \x \ X\ \(p \ (yvec1@yvec2)) \* X\ have "(p \ (yvec1@x#yvec2)) \* X" + by(simp add: eqvts freshChainSimps) + then have "([(y, x)] \ p \ (yvec1@x#yvec2)) \* ([(y, x)] \ X)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \x \ yvec1\ \x \ yvec2\ \x \ y\ \x \ X\ \y \ yvec1\ \y \ yvec2\ \x \ p\ \y \ p\ \bn \ \* X\ yvecEq + have "bn(((y, x)#p) \ \) \* X" by(simp add: bnEqvt[symmetric]) (simp add: perm_compose calc_atm eqvts freshChainSimps) + moreover from \x \ p\ \y \ p\ \x \ V\ \(p \ (yvec1@yvec2)) \* V\ have "(p \ (yvec1@x#yvec2)) \* V" + by(simp add: eqvts freshChainSimps) + then have "([(y, x)] \ p \ (yvec1@x#yvec2)) \* ([(y, x)] \ V)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \x \ yvec1\ \x \ yvec2\ \x \ y\ \x \ V\ \y \ yvec1\ \y \ yvec2\ \x \ p\ \y \ p\ \bn \ \* V\ yvecEq + have "bn(((y, x)#p) \ \) \* V" by(simp add: bnEqvt[symmetric]) (simp add: perm_compose calc_atm eqvts freshChainSimps) + moreover from \x \ p\ \y \ p\ \x \ W\ \(p \ (yvec1@yvec2)) \* W\ have "(p \ (yvec1@x#yvec2)) \* W" + by(simp add: eqvts freshChainSimps) + then have "([(y, x)] \ p \ (yvec1@x#yvec2)) \* ([(y, x)] \ W)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \x \ yvec1\ \x \ yvec2\ \x \ y\ \x \ W\ \y \ yvec1\ \y \ yvec2\ \x \ p\ \y \ p\ \bn \ \* W\ yvecEq + have "bn(((y, x)#p) \ \) \* W" by(simp add: bnEqvt[symmetric]) (simp add: perm_compose calc_atm eqvts freshChainSimps) + moreover from \x \ p\ \y \ p\ \x \ Y\ \(p \ (yvec1@yvec2)) \* Y\ have "(p \ (yvec1@x#yvec2)) \* Y" + by(simp add: eqvts freshChainSimps) + then have "([(y, x)] \ p \ (yvec1@x#yvec2)) \* ([(y, x)] \ Y)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \x \ yvec1\ \x \ yvec2\ \x \ y\ \x \ Y\ \y \ yvec1\ \y \ yvec2\ \x \ p\ \y \ p\ \bn \ \* Y\ yvecEq + have "bn(((y, x)#p) \ \) \* Y" by(simp add: bnEqvt[symmetric]) (simp add: perm_compose calc_atm eqvts freshChainSimps) + moreover from \x \ p\ \y \ p\ \x \ Z\ \(p \ (yvec1@yvec2)) \* Z\ have "(p \ (yvec1@x#yvec2)) \* Z" + by(simp add: eqvts freshChainSimps) + then have "([(y, x)] \ p \ (yvec1@x#yvec2)) \* ([(y, x)] \ Z)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \x \ yvec1\ \x \ yvec2\ \x \ y\ \x \ Z\ \y \ yvec1\ \y \ yvec2\ \x \ p\ \y \ p\ \bn \ \* Z\ yvecEq + have "bn(((y, x)#p) \ \) \* Z" by(simp add: bnEqvt[symmetric]) (simp add: perm_compose calc_atm eqvts freshChainSimps) + moreover from \(yvec1@yvec2) \* A\<^sub>P'\ \y \ A\<^sub>P'\ yvecEq have "bn \ \* A\<^sub>P'" + by simp + moreover from \A\<^sub>P' \* ([(x, y)] \ N')\ have "([(x, y)] \ A\<^sub>P') \* ([(x, y)] \ [(x, y)] \ N')" + by(simp only: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \x \ A\<^sub>P'\ \y \ A\<^sub>P'\ have "A\<^sub>P' \* N'" by simp + with \A\<^sub>P' \* M\ \(yvec1@yvec2) \* A\<^sub>P'\ \y \ A\<^sub>P'\ \eq have "A\<^sub>P' \* \" by simp + moreover then have "(((y, x)#p) \ A\<^sub>P') \* (((y, x)#p) \ \)" + by(simp only: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \x \ A\<^sub>P'\ \y \ A\<^sub>P'\ S \(yvec1@yvec2) \* A\<^sub>P'\ \(p \ (yvec1@yvec2)) \* A\<^sub>P'\ + have "A\<^sub>P' \* (((y, x)#p) \ \)" by(simp add: eqvts) + moreover from \A\<^sub>P' \* ([(x, y)] \ P'')\ have "([(x, y)] \ A\<^sub>P') \* ([(x, y)] \ [(x, y)] \ P'')" + by(simp only: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \x \ A\<^sub>P'\ \y \ A\<^sub>P'\ have "A\<^sub>P' \* P''" by simp + moreover from yvecEq \eq \(p \ (yvec1@yvec2)) \* (yvec1@yvec2)\ \y \ p\ \x \ \\ S \(p \ (yvec1@yvec2)) \* M\\(p \ (yvec1@yvec2)) \* ([(x, y)] \ N')\ \y \ yvec1\\y \ yvec2\ \x \ p\ + have "bn(((y, x)#p) \ \) \* \" + apply(simp add: eqvts del: set_append) + apply(intro conjI) + apply(simp add: perm_compose eqvts del: set_append) + apply(simp add: perm_compose freshChainSimps(6) calc_atm eqvts del: set_append) + apply(simp add: perm_compose eqvts del: set_append) + apply(simp add: perm_compose eqvts swapStarFresh del: set_append) + apply(simp add: perm_compose freshChainSimps(6) calc_atm eqvts del: set_append) + apply(simp add: perm_compose freshChainSimps(6) calc_atm eqvts del: set_append) + apply(simp add: perm_compose freshChainSimps(6) calc_atm eqvts del: set_append) + apply(simp add: perm_compose freshChainSimps(6) swapStarFresh calc_atm eqvts del: set_append) + apply(simp add: perm_compose freshChainSimps(6) calc_atm eqvts del: set_append) + apply(subst pt_fresh_star_bij[symmetric, OF pt_name_inst, OF at_name_inst, where pi="[(x, y)]"]) + apply(simp add: perm_compose freshChainSimps(6) calc_atm eqvts del: set_append) + apply(simp add: perm_compose freshChainSimps(6) calc_atm eqvts del: set_append) + apply(subst pt_fresh_star_bij[symmetric, OF pt_name_inst, OF at_name_inst, where pi="[(x, y)]"]) + by(simp add: perm_compose freshChainSimps(6) calc_atm eqvts del: set_append) + moreover note \A\<^sub>P' \* C\ \A\<^sub>P' \* V\ \A\<^sub>P' \* W\ \A\<^sub>P' \* X\ \A\<^sub>P' \* Y\ \A\<^sub>P' \* Z\ \distinct A\<^sub>P'\ + + ultimately show ?case + by(elim cBrOpen) + next + case(cScope \ P \ P' x A\<^sub>P \\<^sub>P C C' \' P'' V W X Y Z) + from \\ \ \\x\P' = \' \ P''\ \x \ \\ \x \ \'\ + obtain P''' where "\ \ P' = \' \ P'''" and "P'' = \\x\P'''" + apply(cases rule: actionCases[where \=\]) + apply(simp_all add: residualInject) + apply (metis bn.simps(3) boundOutputScopeDest) + by (metis bn.simps(4) boundOutputScopeDest) + then obtain p \' A\<^sub>P' \\<^sub>P' where S: "set p \ set(bn \') \ set(bn(p \ \'))" and PeqP': "((p \ \\<^sub>P) \ \') \ \\<^sub>P'" + and "distinctPerm p" and "bn(p \ \') \* C'" and FrP': "extractFrame P''' = \A\<^sub>P', \\<^sub>P'\" + and "A\<^sub>P' \* P'''" and "A\<^sub>P' \* \'" and "A\<^sub>P' \* (p \ \')" and "A\<^sub>P' \* C" and "distinct A\<^sub>P'" + and "bn(p \ \') \* P'''" and "A\<^sub>P' \* V" and "A\<^sub>P' \* W" and "A\<^sub>P' \* (x#X)" and "A\<^sub>P' \* Y" and "bn(p \ \') \* \'" + and "A\<^sub>P' \* Z" and "bn(p \ \') \* V" and "bn(p \ \') \* W" and "bn(p \ \') \* (x#X)" and "bn(p \ \') \* Y" + and "bn(p \ \') \* Z" using cScope + by(elim cScope) (assumption | simp)+ + from \A\<^sub>P' \* (x#X)\ have "x \ A\<^sub>P'" and "A\<^sub>P' \* X" by simp+ + from \bn(p \ \') \* (x#X)\ have "x \ bn(p \ \')" and "bn(p \ \') \* X" by simp+ + + note S PeqP' \distinctPerm p\ \bn(p \ \') \* C'\ + moreover from FrP' \P'' = \\x\P'''\ have "extractFrame P'' = \(x#A\<^sub>P'), \\<^sub>P'\" by simp + moreover from \A\<^sub>P' \* P'''\ \P'' = \\x\P'''\ \x \ A\<^sub>P'\ have "(x#A\<^sub>P') \* P''" by(simp add: abs_fresh) + moreover from \A\<^sub>P' \* \'\ \A\<^sub>P' \* C\ \x \ \'\ \x \ C\ have "(x#A\<^sub>P') \* \'" and "(x#A\<^sub>P') \* C" by simp+ + moreover note \bn(p \ \') \* \'\ + moreover from \bn(p \ \') \* P'''\ \P'' = \\x\P'''\ \x \ bn(p \ \')\ have "bn(p \ \') \* P''" by simp + moreover from \A\<^sub>P' \* \'\ \x \ \'\ have "(x#A\<^sub>P') \* \'" by simp + moreover from \A\<^sub>P' \* (p \ \')\ \x \ \'\ S \x \ bn(p \ \')\ have "(x#A\<^sub>P') \* (p \ \')" + by(simp add: subjectEqvt[symmetric] bnEqvt[symmetric] okjectEqvt[symmetric] freshChainSimps) + moreover from \A\<^sub>P' \* V\ \x \ V\ have "(x#A\<^sub>P') \* V" by simp+ + moreover from \A\<^sub>P' \* W\ \x \ W\ have "(x#A\<^sub>P') \* W" by simp+ + moreover from \A\<^sub>P' \* X\ \x \ X\ have "(x#A\<^sub>P') \* X" by simp+ + moreover from \A\<^sub>P' \* Y\ \x \ Y\ have "(x#A\<^sub>P') \* Y" by simp+ + moreover from \A\<^sub>P' \* Z\ \x \ Z\ have "(x#A\<^sub>P') \* Z" by simp+ + moreover note \bn(p \ \') \* V\ \bn(p \ \') \* W\ \bn(p \ \') \* X\ \bn(p \ \') \* Y\ \bn(p \ \') \* Z\ + moreover from \distinct A\<^sub>P'\ \x \ A\<^sub>P'\ have "distinct(x#A\<^sub>P')" by simp + ultimately show ?case by(elim cScope) + next + case(cBang \ P A\<^sub>P \\<^sub>P C C' \ P' V W X Y Z) + then obtain p \' A\<^sub>P' \\<^sub>P' where S: "set p \ set(bn \) \ set(bn(p \ \))" + and FrP': "extractFrame P' = \A\<^sub>P', \\<^sub>P'\" + and PeqP': "(p \ (\\<^sub>P \ \)) \ \' \ \\<^sub>P'" + and "A\<^sub>P' \* C" and "A\<^sub>P' \* P'" and "A\<^sub>P' \* \" and "A\<^sub>P' \* (p \ \)" + and "A\<^sub>P' \* V" and "A\<^sub>P' \* W" and "A\<^sub>P' \* X" and "A\<^sub>P' \* Y" and "A\<^sub>P' \* Z" and "distinct A\<^sub>P'" + and "distinctPerm p" and "(bn(p \ \)) \* \" and "(bn(p \ \)) \* P'" + and "(bn(p \ \)) \* C'" and "(bn(p \ \)) \* V" and "(bn(p \ \)) \* W" and "(bn(p \ \)) \* X" and "(bn(p \ \)) \* Y" and "(bn(p \ \)) \* Z" + apply - + by(rule cBang)(assumption | simp (no_asm_use))+ + moreover from \\\<^sub>P \ \\ have "(p \ \\<^sub>P) \ (p \ \)" + by(simp add: AssertionStatEqClosed) + then have "(p \ \\<^sub>P) \ \" by(simp add: permBottom) + with PeqP' have "(\ \ \') \ \\<^sub>P'" + by(simp add: eqvts permBottom) (metis Identity AssertionStatEqTrans composition' Commutativity Associativity AssertionStatEqSym) + ultimately show ?case using cBang + by (metis permBottom) + qed + + with A have ?thesis by blast + } + moreover have "bn \ \* ([]::'a list)" and "bn \ \* ([]::('a action) list)" and "bn \ \* ([]::name list)" and "bn \ \* ([]::'b list)" and "bn \ \* ([]::('a, 'b, 'c) psi list)" + and "A\<^sub>P \* ([]::'a list)" and "A\<^sub>P \* ([]::('a action) list)" and "A\<^sub>P \* ([]::name list)" and "A\<^sub>P \* ([]::'b list)" and "A\<^sub>P \* ([]::('a, 'b, 'c) psi list)" + by simp+ + ultimately show ?thesis by blast +qed + +lemma expandTauFrame: + fixes \ :: 'b + and P :: "('a, 'b,'c) psi" + and P' :: "('a, 'b, 'c) psi" + and A\<^sub>P :: "name list" + and \\<^sub>P :: 'b + and C :: "'f::fs_name" + +assumes "\ \ P \\ \ P'" + and "extractFrame P = \A\<^sub>P, \\<^sub>P\" + and "distinct A\<^sub>P" + and "A\<^sub>P \* P" + and "A\<^sub>P \* C" + +obtains \' A\<^sub>P' \\<^sub>P' where "extractFrame P' = \A\<^sub>P', \\<^sub>P'\" and "\\<^sub>P \ \' \ \\<^sub>P'" and "A\<^sub>P' \* C" and "A\<^sub>P' \* P'" and "distinct A\<^sub>P'" +proof - + assume A: "\A\<^sub>P' \\<^sub>P' \'. + \extractFrame P' = \A\<^sub>P', \\<^sub>P'\; \\<^sub>P \ \' \ \\<^sub>P'; A\<^sub>P' \* C; A\<^sub>P' \* P'; distinct A\<^sub>P'\ + \ thesis" + + from \\ \ P \\ \P'\ \A\<^sub>P \* P\ have "A\<^sub>P \* P'" by(rule tauFreshChainDerivative) + + { + fix X :: "name list" + and Y :: "'b list" + and Z :: "('a, 'b, 'c) psi list" + + assume "A\<^sub>P \* X" + and "A\<^sub>P \* Y" + and "A\<^sub>P \* Z" + + with assms \A\<^sub>P \* P'\ obtain \' A\<^sub>P' \\<^sub>P' where "extractFrame P' = \A\<^sub>P', \\<^sub>P'\" and "\\<^sub>P \ \' \ \\<^sub>P'" and "A\<^sub>P' \* C" + and "A\<^sub>P' \* P'" and "A\<^sub>P' \* X" and "A\<^sub>P' \* Y" and "A\<^sub>P' \* Z" and "distinct A\<^sub>P'" + proof(nominal_induct avoiding: C X Y Z arbitrary: thesis rule: tauFrameInduct) + case(cAlpha \ P P' A\<^sub>P \\<^sub>P p C X Y Z) + then obtain \' A\<^sub>P' \\<^sub>P' where FrP': "extractFrame P' = \A\<^sub>P', \\<^sub>P'\" and "\\<^sub>P \ \' \ \\<^sub>P'" and "distinct A\<^sub>P'" + and "A\<^sub>P' \* C" and "A\<^sub>P' \* P'" and "A\<^sub>P' \* X" and "A\<^sub>P' \* Y" and "A\<^sub>P' \* Z" + by metis + + have S: "set p \ set A\<^sub>P \ set(p \ A\<^sub>P)" by fact + + from FrP' have "(p \ extractFrame P') = p \ \A\<^sub>P', \\<^sub>P'\" by simp + with \A\<^sub>P \* P'\ \(p \ A\<^sub>P) \* P'\ S have "extractFrame P' = \(p \ A\<^sub>P'), (p \ \\<^sub>P')\" by(simp add: eqvts) + moreover from \\\<^sub>P \ \' \ \\<^sub>P'\ have "(p \ (\\<^sub>P \ \')) \ (p \ \\<^sub>P')" by(rule AssertionStatEqClosed) + then have "(p \ \\<^sub>P) \ (p \ \') \ (p \ \\<^sub>P')" by(simp add: eqvts) + moreover from \A\<^sub>P' \* C\ have "(p \ A\<^sub>P') \* (p \ C)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>P \* C\ \(p \ A\<^sub>P) \* C\ S have "(p \ A\<^sub>P') \* C" by simp + moreover from \A\<^sub>P' \* P'\ have "(p \ A\<^sub>P') \* (p \ P')" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>P \* P'\ \(p \ A\<^sub>P) \* P'\ S have "(p \ A\<^sub>P') \* P'" by simp + moreover from \A\<^sub>P' \* X\ have "(p \ A\<^sub>P') \* (p \ X)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>P \* X\ \(p \ A\<^sub>P) \* X\ S have "(p \ A\<^sub>P') \* X" by simp + moreover from \A\<^sub>P' \* Y\ have "(p \ A\<^sub>P') \* (p \ Y)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>P \* Y\ \(p \ A\<^sub>P) \* Y\ S have "(p \ A\<^sub>P') \* Y" by simp + moreover from \A\<^sub>P' \* Z\ have "(p \ A\<^sub>P') \* (p \ Z)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>P \* Z\ \(p \ A\<^sub>P) \* Z\ S have "(p \ A\<^sub>P') \* Z" by simp + moreover from \distinct A\<^sub>P'\ have "distinct(p \ A\<^sub>P')" by simp + ultimately show ?case by(rule cAlpha) + next + case(cCase \ P P' \ Cs A\<^sub>P \\<^sub>P C B Y Z thesis) + then obtain \' A\<^sub>P' \\<^sub>P' where FrP': "extractFrame P' = \A\<^sub>P', \\<^sub>P'\" + and "\\<^sub>P \ \' \ \\<^sub>P'" and "distinct A\<^sub>P'" + and "A\<^sub>P' \* C" and "A\<^sub>P' \* P'" + and "A\<^sub>P' \* B" and "A\<^sub>P' \* Y" and "A\<^sub>P' \* Z" + apply - + by(rule cCase) (assumption | simp (no_asm_use))+ + with \\\<^sub>P \ \\ \\\<^sub>P \ \' \ \\<^sub>P'\ have "\ \ \' \ \\<^sub>P'" + by(metis Identity AssertionStatEqTrans composition' Commutativity Associativity AssertionStatEqSym) + then show ?case using FrP' \A\<^sub>P' \* P'\ \A\<^sub>P' \* C\ \A\<^sub>P' \* B\ \A\<^sub>P' \* Y\ \A\<^sub>P' \* Z\ \distinct A\<^sub>P'\ using cCase + by force + next + case(cPar1 \ \\<^sub>Q P P' A\<^sub>Q Q A\<^sub>P \\<^sub>P C X Y Z) + moreover from \A\<^sub>P \* X\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* Y\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>P \* Q\ \A\<^sub>P \* Z\ + have "A\<^sub>P \* (X@A\<^sub>Q)" and "A\<^sub>P \* (\\<^sub>Q#Y)" and "A\<^sub>P \* (Q#Z)" + by simp+ + ultimately obtain \' A\<^sub>P' \\<^sub>P' where FrP': "extractFrame P' = \A\<^sub>P', \\<^sub>P'\" and "distinct A\<^sub>P'" + and "\\<^sub>P \ \' \ \\<^sub>P'" and "A\<^sub>P \* P" and "A\<^sub>P \* C" and "A\<^sub>P' \* C" and "A\<^sub>P' \* P'" + and "A\<^sub>P' \* (X@A\<^sub>Q)" and "A\<^sub>P' \* (\\<^sub>Q#Y)" and "A\<^sub>P' \* (Q#Z)" + by metis + + then have "A\<^sub>P' \* X" and "A\<^sub>P' \* A\<^sub>Q" and "A\<^sub>P' \* Y" and "A\<^sub>P' \* \\<^sub>Q" and "A\<^sub>P' \* Q" and "A\<^sub>P' \* Z" + by simp+ + + from \A\<^sub>P' \* A\<^sub>Q\ \A\<^sub>Q \* P'\ FrP' have "A\<^sub>Q \* \\<^sub>P'" by(force dest: extractFrameFreshChain) + with \A\<^sub>P' \* \\<^sub>Q\ \A\<^sub>P' \* A\<^sub>Q\ \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ FrP' + have "extractFrame(P' \ Q) = \(A\<^sub>P'@A\<^sub>Q), \\<^sub>P' \ \\<^sub>Q\" by simp + + moreover from \\\<^sub>P \ \' \ \\<^sub>P'\have "(\\<^sub>P \ \\<^sub>Q) \ \' \ \\<^sub>P' \ \\<^sub>Q" + by(metis Associativity Commutativity Composition AssertionStatEqTrans AssertionStatEqSym) + + moreover from \A\<^sub>P' \* C\ \A\<^sub>Q \* C\ have "(A\<^sub>P'@A\<^sub>Q) \* C" by simp + moreover from \A\<^sub>P' \* P'\ \A\<^sub>Q \* P'\ \A\<^sub>P' \* Q\ \A\<^sub>Q \* Q\ have "(A\<^sub>P'@A\<^sub>Q) \* (P' \ Q)" by simp + moreover from \A\<^sub>P' \* X\ \A\<^sub>Q \* X\ have "(A\<^sub>P'@A\<^sub>Q) \* X" by simp + moreover from \A\<^sub>P' \* Y\ \A\<^sub>Q \* Y\ have "(A\<^sub>P'@A\<^sub>Q) \* Y" by simp + moreover from \A\<^sub>P' \* Z\ \A\<^sub>Q \* Z\ have "(A\<^sub>P'@A\<^sub>Q) \* Z" by simp + moreover from \A\<^sub>P' \* A\<^sub>Q\ \distinct A\<^sub>P'\ \distinct A\<^sub>Q\ have "distinct(A\<^sub>P'@A\<^sub>Q)" by simp + ultimately show ?case by(rule cPar1) + next + case(cPar2 \ \\<^sub>P Q Q' A\<^sub>P P A\<^sub>Q \\<^sub>Q C X Y Z) + moreover from \A\<^sub>Q \* X\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>Q \* Y\ \A\<^sub>Q \* \\<^sub>P\ \A\<^sub>Q \* P\ \A\<^sub>Q \* Z\ + have "A\<^sub>Q \* (X@A\<^sub>P)" and "A\<^sub>Q \* (\\<^sub>P#Y)" and "A\<^sub>Q \* (P#Z)" + by(simp add: freshChainSimps)+ + ultimately obtain \' A\<^sub>Q' \\<^sub>Q' where FrQ': "extractFrame Q' = \A\<^sub>Q', \\<^sub>Q'\" and "distinct A\<^sub>Q'" + and "\\<^sub>Q \ \' \ \\<^sub>Q'"and "A\<^sub>Q' \* C" and "A\<^sub>Q' \* Q'" + and "A\<^sub>Q' \* (X@A\<^sub>P)" and "A\<^sub>Q' \* (\\<^sub>P#Y)" and "A\<^sub>Q' \* (P#Z)" + by metis + + then have "A\<^sub>Q' \* X" and "A\<^sub>Q' \* A\<^sub>P" and "A\<^sub>Q' \* Y" and "A\<^sub>Q' \* \\<^sub>P" and "A\<^sub>Q' \* P" and "A\<^sub>Q' \* Z" + by simp+ + + from \A\<^sub>Q' \* A\<^sub>P\ \A\<^sub>P \* Q'\ FrQ' have "A\<^sub>P \* \\<^sub>Q'" by(force dest: extractFrameFreshChain) + with \A\<^sub>Q' \* \\<^sub>P\ \A\<^sub>Q' \* A\<^sub>P\ \extractFrame P = \A\<^sub>P, \\<^sub>P\\ FrQ' + have "extractFrame(P \ Q') = \(A\<^sub>P@A\<^sub>Q'), \\<^sub>P \ \\<^sub>Q'\" by simp + + moreover from \\\<^sub>Q \ \' \ \\<^sub>Q'\have "(\\<^sub>P \ \\<^sub>Q) \ \' \ \\<^sub>P \ \\<^sub>Q'" + by(metis Associativity Commutativity Composition AssertionStatEqTrans AssertionStatEqSym) + moreover from \A\<^sub>P \* C\ \A\<^sub>Q' \* C\ have "(A\<^sub>P@A\<^sub>Q') \* C" by simp + moreover from \A\<^sub>P \* P\ \A\<^sub>Q' \* P\ \A\<^sub>P \* Q'\ \A\<^sub>Q' \* Q'\ have "(A\<^sub>P@A\<^sub>Q') \* (P \ Q')" by simp + moreover from \A\<^sub>P \* X\ \A\<^sub>Q' \* X\ have "(A\<^sub>P@A\<^sub>Q') \* X" by simp + moreover from \A\<^sub>P \* Y\ \A\<^sub>Q' \* Y\ have "(A\<^sub>P@A\<^sub>Q') \* Y" by simp + moreover from \A\<^sub>P \* Z\ \A\<^sub>Q' \* Z\ have "(A\<^sub>P@A\<^sub>Q') \* Z" by simp + moreover from \A\<^sub>Q' \* A\<^sub>P\ \distinct A\<^sub>P\ \distinct A\<^sub>Q'\ have "distinct(A\<^sub>P@A\<^sub>Q')" by simp + ultimately show ?case by(rule cPar2) + next + case(cComm1 \ \\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q K xvec Q' A\<^sub>Q C X Y Z) + have PTrans: "\ \ \\<^sub>Q \ P \M\N\ \ P'" and QTrans: "\ \ \\<^sub>P \ Q \K\\*xvec\\N\ \ Q'" by fact+ + from PTrans \extractFrame P = \A\<^sub>P, \\<^sub>P\\ \distinct A\<^sub>P\ \A\<^sub>P \* P\ \A\<^sub>P \* N\ \A\<^sub>P \* M\ + \A\<^sub>P \* Q'\ \A\<^sub>P \* C\ \A\<^sub>P \* X\ \A\<^sub>P \* Y\ \A\<^sub>P \* Z\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* xvec\ + obtain \' A\<^sub>P' \\<^sub>P' where FrP': "extractFrame P' = \A\<^sub>P', \\<^sub>P'\" and PeqP': "\\<^sub>P \ \' \ \\<^sub>P'" + and "A\<^sub>P' \* Q'" and "A\<^sub>P' \* C" and "A\<^sub>P' \* X" and "A\<^sub>P' \* Y" and "distinct A\<^sub>P'" + and "A\<^sub>P' \* Z" and "A\<^sub>P' \* A\<^sub>Q" and "A\<^sub>P' \* xvec" and "A\<^sub>P' \* P'" + by(elim expandNonTauFrame[where C="(Q', C, X, Y, Z, A\<^sub>Q, xvec)" and C'="(Q', C, X, Y, Z, A\<^sub>Q, xvec)"]) auto + moreover from QTrans have "distinct xvec" by(auto dest: boundOutputDistinct) + from QTrans \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ \distinct A\<^sub>Q\ \A\<^sub>Q \* P\ \A\<^sub>Q \* xvec\ \A\<^sub>Q \* K\ \xvec \* K\ \distinct xvec\ + \A\<^sub>P' \* A\<^sub>Q\ \A\<^sub>P' \* xvec\ \A\<^sub>Q \* Q\ \xvec \* Q\ \A\<^sub>Q \* \\<^sub>P\ \xvec \* \\<^sub>P\ \A\<^sub>Q \* N\ + \A\<^sub>Q \* C\ \A\<^sub>Q \* X\ \A\<^sub>Q \* Y\ \A\<^sub>Q \* Z\ \xvec \* P\ \xvec \* C\ \xvec \* X\ \xvec \* Y\ \xvec \* Z\ + obtain p \'' A\<^sub>Q' \\<^sub>Q' where S: "set p \ set xvec \ set(p \ xvec)" and QeqQ': "(p \ \\<^sub>Q) \ \'' \ \\<^sub>Q'" and FrQ': "extractFrame Q' = \A\<^sub>Q', \\<^sub>Q'\" + and "A\<^sub>Q' \* A\<^sub>P'" and "A\<^sub>Q' \* C" and "A\<^sub>Q' \* X" and "A\<^sub>Q' \* Y" and "A\<^sub>Q' \* Z" and "A\<^sub>Q' \* P" and "A\<^sub>Q' \* N" and "distinct A\<^sub>Q'" + and "(p \ xvec) \* A\<^sub>P'" and "(p \ xvec) \* C" and "(p \ xvec) \* X" and "(p \ xvec) \* Y" and "(p \ xvec) \* P" + and "(p \ xvec) \* Z" and "(p \ xvec) \* N" and "(p \ xvec) \* \\<^sub>P" and "(p \ xvec) \* A\<^sub>Q'"and "(p \ xvec) \* Q'" + and "distinctPerm p" and "A\<^sub>Q' \* xvec" and "A\<^sub>Q' \* Q'" + by(elim expandNonTauFrame[where C="(P, C, X, Y, Z, A\<^sub>P', \\<^sub>P)" and C'="(P, C, X, Y, Z, A\<^sub>P', \\<^sub>P)"]) (assumption | simp)+ + + from PTrans \A\<^sub>Q' \* P\ \A\<^sub>Q' \* N\ \(p \ xvec) \* P\ \(p \ xvec) \* N\ + have "A\<^sub>Q' \* P'" and "(p \ xvec) \* P'" by(force dest: inputFreshChainDerivative)+ + with FrP' \A\<^sub>Q' \* A\<^sub>P'\ \(p \ xvec) \* A\<^sub>P'\ have "A\<^sub>Q' \* \\<^sub>P'" and "(p \ xvec) \* \\<^sub>P'" by(force dest: extractFrameFreshChain)+ + from FrQ' \A\<^sub>Q' \* A\<^sub>P'\ \A\<^sub>P' \* Q'\ \(p \ xvec) \* A\<^sub>Q'\ \(p \ xvec) \* Q'\ have "A\<^sub>P' \* \\<^sub>Q'" and "(p \ xvec) \* \\<^sub>Q'" + by(force dest: extractFrameFreshChain)+ + + have "extractFrame(\\*xvec\(P' \ Q')) = \((p \ xvec)@A\<^sub>P'@A\<^sub>Q'), (p \ \\<^sub>P') \ (p \ \\<^sub>Q')\" + proof - + from FrP' FrQ' \A\<^sub>P' \* \\<^sub>Q'\ \A\<^sub>Q' \* A\<^sub>P'\ \A\<^sub>Q' \* \\<^sub>P'\ have "extractFrame(P' \ Q') = \(A\<^sub>P'@A\<^sub>Q'), \\<^sub>P' \ \\<^sub>Q'\" + by simp + then have "extractFrame(\\*xvec\(P' \ Q')) = \(xvec@A\<^sub>P'@A\<^sub>Q'), \\<^sub>P' \ \\<^sub>Q'\" + by(induct xvec) auto + moreover from \(p \ xvec) \* \\<^sub>P'\ \(p \ xvec) \* \\<^sub>Q'\ S + have "\\*xvec\(\\*(A\<^sub>P'@A\<^sub>Q')\(FAssert (\\<^sub>P' \ \\<^sub>Q'))) = \\*(p \ xvec)\(p \ \\*(A\<^sub>P'@A\<^sub>Q')\(FAssert(\\<^sub>P' \ \\<^sub>Q')))" + by(intro frameChainAlpha) (auto simp add: fresh_star_def frameResChainFresh) + then have "\\*xvec\(\\*(A\<^sub>P'@A\<^sub>Q')\(FAssert (\\<^sub>P' \ \\<^sub>Q'))) = \\*(p \ xvec)\(\\*(A\<^sub>P'@A\<^sub>Q')\(FAssert((p \ \\<^sub>P') \ (p \ \\<^sub>Q'))))" + using \A\<^sub>P' \* xvec\ \(p \ xvec) \* A\<^sub>P'\ \A\<^sub>Q' \* xvec\ \(p \ xvec) \* A\<^sub>Q'\ S + by(auto simp add: eqvts) + ultimately show ?thesis + by(simp add: frameChainAppend) + qed + + moreover have "(\\<^sub>P \ \\<^sub>Q) \ ((p \ \') \ (p \ \'')) \ (p \ \\<^sub>P') \ (p \ \\<^sub>Q')" + proof - + have "(\\<^sub>P \ (p \ \\<^sub>Q)) \ (\' \ \'') \ (\\<^sub>P \ \') \ ((p \ \\<^sub>Q) \ \'')" + by(metis Associativity Commutativity Composition AssertionStatEqTrans) + moreover from PeqP' QeqQ' have "(\\<^sub>P \ \') \ ((p \ \\<^sub>Q) \ \'') \ \\<^sub>P' \ \\<^sub>Q'" + by(metis Associativity Commutativity Composition AssertionStatEqTrans) + ultimately have "(\\<^sub>P \ (p \ \\<^sub>Q)) \ (\' \ \'') \ \\<^sub>P' \ \\<^sub>Q'" + by(metis AssertionStatEqTrans) + then have "(p \ ((\\<^sub>P \ (p \ \\<^sub>Q)) \ (\' \ \''))) \ (p \ (\\<^sub>P' \ \\<^sub>Q'))" + by(rule AssertionStatEqClosed) + with \xvec \* \\<^sub>P\ \(p \ xvec) \* \\<^sub>P\ S \distinctPerm p\ show ?thesis + by(simp add: eqvts) + qed + + moreover from \(p \ xvec) \* C\ \A\<^sub>P' \* C\ \A\<^sub>Q' \* C\ have "((p \ xvec)@A\<^sub>P'@A\<^sub>Q') \* C" by simp + moreover from \(p \ xvec) \* X\ \A\<^sub>P' \* X\ \A\<^sub>Q' \* X\ have "((p \ xvec)@A\<^sub>P'@A\<^sub>Q') \* X" by simp + moreover from \(p \ xvec) \* Y\ \A\<^sub>P' \* Y\ \A\<^sub>Q' \* Y\ have "((p \ xvec)@A\<^sub>P'@A\<^sub>Q') \* Y" by simp + moreover from \(p \ xvec) \* Z\ \A\<^sub>P' \* Z\ \A\<^sub>Q' \* Z\ have "((p \ xvec)@A\<^sub>P'@A\<^sub>Q') \* Z" by simp + moreover from \(p \ xvec) \* P'\ \(p \ xvec) \* Q'\ \A\<^sub>P' \* P'\ \A\<^sub>P' \* Q' \\A\<^sub>Q' \* P'\ \A\<^sub>Q' \* Q'\ + have "((p \ xvec)@A\<^sub>P'@A\<^sub>Q') \* (\\*xvec\(P' \ Q'))" by(auto simp add: resChainFresh fresh_star_def) + moreover from \(p \ xvec) \* A\<^sub>P'\ \(p \ xvec) \* A\<^sub>Q'\ \A\<^sub>Q' \* A\<^sub>P'\ \distinct xvec\ \distinct A\<^sub>P'\ \distinct A\<^sub>Q'\ + have "distinct((p \ xvec)@A\<^sub>P'@A\<^sub>Q')" + by auto (simp add: name_list_supp fresh_star_def fresh_def)+ + + ultimately show ?case using cComm1 + by metis + next + case(cComm2 \ \\<^sub>Q P M xvec N P' A\<^sub>P \\<^sub>P Q K Q' A\<^sub>Q C X Y Z) + have PTrans: "\ \ \\<^sub>Q \ P \M\\*xvec\\N\ \ P'" and QTrans: "\ \ \\<^sub>P \ Q \K\N\ \ Q'" by fact+ + from PTrans have "distinct xvec" by(auto dest: boundOutputDistinct) + from PTrans \extractFrame P = \A\<^sub>P, \\<^sub>P\\ \distinct A\<^sub>P\ \A\<^sub>P \* P\ \A\<^sub>P \* Q\ \xvec \* Q\ \distinct xvec\ \xvec \* M\ + \A\<^sub>P \* C\ \A\<^sub>P \* X\ \A\<^sub>P \* Y\ \A\<^sub>P \* Z\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* xvec\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>P \* M\ \A\<^sub>P \* N\ + \xvec \* P\ \xvec \* C\ \xvec \* X\ \xvec \* Y\ \xvec \* Z\ \A\<^sub>Q \* xvec\ \xvec \* \\<^sub>Q\ + obtain p \' A\<^sub>P' \\<^sub>P' where S: "set p \ set xvec \ set(p \ xvec)" + and FrP': "extractFrame P' = \A\<^sub>P', \\<^sub>P'\" and PeqP': "(p \ \\<^sub>P) \ \' \ \\<^sub>P'" and "distinct A\<^sub>P'" + and "A\<^sub>P' \* C" and "A\<^sub>P' \* X" and "A\<^sub>P' \* Y" and "A\<^sub>P' \* N" and "A\<^sub>P' \* Q" and "(p \ xvec) \* Q" + and "A\<^sub>P' \* Z" and "A\<^sub>P' \* A\<^sub>Q" and "A\<^sub>P' \* xvec" and "A\<^sub>P' \* P'" and "(p \ xvec) \* N" and "(p \ xvec) \* \\<^sub>Q" + and "(p \ xvec) \* A\<^sub>P'" and "(p \ xvec) \* C" and "(p \ xvec) \* X" and "(p \ xvec) \* A\<^sub>Q" + and "(p \ xvec) \* Y" and "(p \ xvec) \* Z" and "(p \ xvec) \* P'" and "distinctPerm p" + by(elim expandNonTauFrame[where C="(C, X, Y, Z, A\<^sub>Q, Q, \\<^sub>Q)" and C'="(C, X, Y, Z, A\<^sub>Q, Q, \\<^sub>Q)"]) (assumption | simp)+ + + from QTrans \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ \distinct A\<^sub>Q\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* xvec\ \A\<^sub>Q \* P'\ \(p \ xvec) \* A\<^sub>Q\ + \A\<^sub>P' \* A\<^sub>Q\ \A\<^sub>Q \* P\ \A\<^sub>Q \* \\<^sub>P\ \A\<^sub>Q \* C\ \A\<^sub>Q \* X\ \A\<^sub>Q \* Y\ \A\<^sub>Q \* Z\ \A\<^sub>Q \* N\ \A\<^sub>Q \* K\ + obtain \'' A\<^sub>Q' \\<^sub>Q' where QeqQ': "\\<^sub>Q \ \'' \ \\<^sub>Q'" and FrQ': "extractFrame Q' = \A\<^sub>Q', \\<^sub>Q'\" and "distinct A\<^sub>Q'" + and "A\<^sub>Q' \* xvec" and "A\<^sub>Q' \* Q'" and "A\<^sub>Q' \* xvec" and "A\<^sub>Q' \* P'" and "A\<^sub>Q' \* (p \ xvec)" + and "A\<^sub>Q' \* A\<^sub>P'" and "A\<^sub>Q' \* C" and "A\<^sub>Q' \* X" and "A\<^sub>Q' \* Y" and "A\<^sub>Q' \* Z" and "A\<^sub>Q' \* P" + by(elim expandNonTauFrame[where C="(P, C, P', X, Y, Z, A\<^sub>P', xvec, (p \ xvec), \\<^sub>P)" and C'="(P, C, P', X, Y, Z, A\<^sub>P', xvec, (p \ xvec), \\<^sub>P)"]) (assumption | simp)+ + + from QTrans \A\<^sub>P' \* Q\ \A\<^sub>P' \* N\ \(p \ xvec) \* Q\ \(p \ xvec) \* N\ + have "A\<^sub>P' \* Q'" and "(p \ xvec) \* Q'" by(force dest: inputFreshChainDerivative)+ + with FrQ' \A\<^sub>Q' \* A\<^sub>P'\ \A\<^sub>Q' \* (p \ xvec)\ have "A\<^sub>P' \* \\<^sub>Q'" and "(p \ xvec) \* \\<^sub>Q'" by(force dest: extractFrameFreshChain)+ + from FrP' \A\<^sub>Q' \* A\<^sub>P'\ \A\<^sub>Q' \* P'\ \(p \ xvec) \* A\<^sub>P'\ \(p \ xvec) \* P'\ have "A\<^sub>Q' \* \\<^sub>P'" and "(p \ xvec) \* \\<^sub>P'" + by(force dest: extractFrameFreshChain)+ + + have "extractFrame(\\*xvec\(P' \ Q')) = \((p \ xvec)@A\<^sub>P'@A\<^sub>Q'), (p \ \\<^sub>P') \ (p \ \\<^sub>Q')\" + proof - + from FrP' FrQ' \A\<^sub>P' \* \\<^sub>Q'\ \A\<^sub>Q' \* A\<^sub>P'\ \A\<^sub>Q' \* \\<^sub>P'\ have "extractFrame(P' \ Q') = \(A\<^sub>P'@A\<^sub>Q'), \\<^sub>P' \ \\<^sub>Q'\" + by simp + then have "extractFrame(\\*xvec\(P' \ Q')) = \(xvec@A\<^sub>P'@A\<^sub>Q'), \\<^sub>P' \ \\<^sub>Q'\" + by(induct xvec) auto + moreover from \(p \ xvec) \* \\<^sub>P'\ \(p \ xvec) \* \\<^sub>Q'\ S + have "\\*xvec\(\\*(A\<^sub>P'@A\<^sub>Q')\(FAssert (\\<^sub>P' \ \\<^sub>Q'))) = \\*(p \ xvec)\(p \ \\*(A\<^sub>P'@A\<^sub>Q')\(FAssert(\\<^sub>P' \ \\<^sub>Q')))" + by(intro frameChainAlpha) (auto simp add: fresh_star_def frameResChainFresh) + then have "\\*xvec\(\\*(A\<^sub>P'@A\<^sub>Q')\(FAssert (\\<^sub>P' \ \\<^sub>Q'))) = \\*(p \ xvec)\(\\*(A\<^sub>P'@A\<^sub>Q')\(FAssert((p \ \\<^sub>P') \ (p \ \\<^sub>Q'))))" + using \A\<^sub>P' \* xvec\ \(p \ xvec) \* A\<^sub>P'\ \A\<^sub>Q' \* xvec\ \A\<^sub>Q' \* (p \ xvec)\ S + by(auto simp add: eqvts) + ultimately show ?thesis + by(simp add: frameChainAppend) + qed + + moreover have "(\\<^sub>P \ \\<^sub>Q) \ ((p \ \') \ (p \ \'')) \ (p \ \\<^sub>P') \ (p \ \\<^sub>Q')" + proof - + have "((p \ \\<^sub>P) \ \\<^sub>Q) \ (\' \ \'') \ ((p \ \\<^sub>P) \ \') \ (\\<^sub>Q \ \'')" + by(metis Associativity Commutativity Composition AssertionStatEqTrans) + moreover from PeqP' QeqQ' have "((p \ \\<^sub>P) \ \') \ (\\<^sub>Q \ \'') \ \\<^sub>P' \ \\<^sub>Q'" + by(metis Associativity Commutativity Composition AssertionStatEqTrans) + ultimately have "((p \ \\<^sub>P) \ \\<^sub>Q) \ (\' \ \'') \ \\<^sub>P' \ \\<^sub>Q'" + by(metis AssertionStatEqTrans) + then have "(p \ ((p \ \\<^sub>P) \ \\<^sub>Q) \ (\' \ \'')) \ (p \ (\\<^sub>P' \ \\<^sub>Q'))" + by(rule AssertionStatEqClosed) + with \xvec \* \\<^sub>Q\ \(p \ xvec) \* \\<^sub>Q\ S \distinctPerm p\ show ?thesis + by(simp add: eqvts) + qed + + moreover from \(p \ xvec) \* C\ \A\<^sub>P' \* C\ \A\<^sub>Q' \* C\ have "((p \ xvec)@A\<^sub>P'@A\<^sub>Q') \* C" by simp + moreover from \(p \ xvec) \* X\ \A\<^sub>P' \* X\ \A\<^sub>Q' \* X\ have "((p \ xvec)@A\<^sub>P'@A\<^sub>Q') \* X" by simp + moreover from \(p \ xvec) \* Y\ \A\<^sub>P' \* Y\ \A\<^sub>Q' \* Y\ have "((p \ xvec)@A\<^sub>P'@A\<^sub>Q') \* Y" by simp + moreover from \(p \ xvec) \* Z\ \A\<^sub>P' \* Z\ \A\<^sub>Q' \* Z\ have "((p \ xvec)@A\<^sub>P'@A\<^sub>Q') \* Z" by simp + moreover from \(p \ xvec) \* P'\ \(p \ xvec) \* Q'\ \A\<^sub>P' \* P'\ \A\<^sub>P' \* Q' \\A\<^sub>Q' \* P'\ \A\<^sub>Q' \* Q'\ + have "((p \ xvec)@A\<^sub>P'@A\<^sub>Q') \* (\\*xvec\(P' \ Q'))" by(auto simp add: resChainFresh fresh_star_def) + moreover from \(p \ xvec) \* A\<^sub>P'\ \ A\<^sub>Q' \* (p \ xvec)\ \A\<^sub>Q' \* A\<^sub>P'\ \distinct xvec\ \distinct A\<^sub>P'\ \distinct A\<^sub>Q'\ + have "distinct((p \ xvec)@A\<^sub>P'@A\<^sub>Q')" + by auto (simp add: name_list_supp fresh_star_def fresh_def)+ + + ultimately show ?case using cComm2 by metis + next + case(cBrClose \ P M xvec N P' A\<^sub>P \\<^sub>P x C X Y Z) + from \(x # A\<^sub>P) \* C\ have "A\<^sub>P \* C" by simp + from \(x # A\<^sub>P) \* X\ \x \ A\<^sub>P\ have "A\<^sub>P \* (x#X)" by simp + from \(x # A\<^sub>P) \* Y\ have "A\<^sub>P \* Y" by simp + from \(x # A\<^sub>P) \* Z\ have "A\<^sub>P \* Z" by simp + from \x \ xvec\ \xvec \* X\ have "xvec \* (x#X)" by simp + + have PTrans: "\ \ P \\M\\*xvec\\N\ \ P'" by fact+ + from PTrans have "distinct xvec" by(auto dest: boundOutputDistinct) + from PTrans \extractFrame P = \A\<^sub>P, \\<^sub>P\\ \distinct A\<^sub>P\ \xvec \* M\ \A\<^sub>P \* P\ \A\<^sub>P \* M\ \A\<^sub>P \* xvec\ \A\<^sub>P \* N\ \A\<^sub>P \* C\ + \A\<^sub>P \* (x#X)\ \A\<^sub>P \* Y\ \A\<^sub>P \* Z\ \xvec \* P\ \xvec \* C\ \xvec \* (x#X)\ \xvec \* Y\ \xvec \* Z\ + + obtain p \' A\<^sub>P' \\<^sub>P' where S: "set p \ set xvec \ set(p \ xvec)" + and FrP': "extractFrame P' = \A\<^sub>P', \\<^sub>P'\" and PeqP': "(p \ \\<^sub>P) \ \' \ \\<^sub>P'" and "distinct A\<^sub>P'" + and "A\<^sub>P' \* C" and "A\<^sub>P' \* (x#X)" and "A\<^sub>P' \* Y" and "A\<^sub>P' \* N" + and "A\<^sub>P' \* Z" and "A\<^sub>P' \* xvec" and "A\<^sub>P' \* P'" and "(p \ xvec) \* xvec" and "(p \ xvec) \* N" + and "(p \ xvec) \* A\<^sub>P'" and "(p \ xvec) \* C" and "(p \ xvec) \* (x#X)" + and "(p \ xvec) \* Y" and "(p \ xvec) \* Z" and "(p \ xvec) \* P'" and "distinctPerm p" + by(elim expandNonTauFrame[where C="(C, (x#X), Y, Z)" and C'="(C, (x#X), Y, Z)"]) (assumption | simp)+ + then have "A\<^sub>P' \* X" and "x \ A\<^sub>P'" and "(p \ xvec) \* X" and "x \ (p \ xvec)" by simp+ + from FrP' \(p \ xvec) \* A\<^sub>P'\ \(p \ xvec) \* P'\ have "(p \ xvec) \* \\<^sub>P'" + by (metis extractFrameFreshChain freshFrameDest) + from S \A\<^sub>P' \* xvec\ \(p \ xvec) \* A\<^sub>P'\ have "p \ A\<^sub>P' = A\<^sub>P'" by simp + + from \(p \ xvec) \* \\<^sub>P'\ S + have "\\*xvec\(\A\<^sub>P', \\<^sub>P'\) = \\*(p \ xvec)\(p \ \A\<^sub>P', \\<^sub>P'\)" + by(intro frameChainAlpha) (auto simp add: fresh_star_def frameResChainFresh) + then have "\(xvec@A\<^sub>P'), \\<^sub>P'\ = (p \ \(xvec@A\<^sub>P'), \\<^sub>P'\)" + by (metis frameChainAppend frameResChainEqvt) + with \p \ A\<^sub>P' = A\<^sub>P'\ have "\(xvec@A\<^sub>P'), \\<^sub>P'\ = \((p \ xvec)@A\<^sub>P'), (p \ \\<^sub>P')\" + by(simp add: eqvts) + then have "\(x#(xvec@A\<^sub>P')), \\<^sub>P'\ = \(x#((p \ xvec)@A\<^sub>P')), (p \ \\<^sub>P')\" + by simp + + moreover from FrP' have "extractFrame (\\x\(\\*xvec\P')) = \(x#(xvec@A\<^sub>P')), \\<^sub>P'\" + by (metis extractFrameResChain extractFrame_extractFrame'_extractFrame''.simps frameChainAppend frameResChain.step) + + ultimately have FrP'2: "extractFrame (\\x\(\\*xvec\P')) = \(x#((p \ xvec)@A\<^sub>P')), (p \ \\<^sub>P')\" + by simp + + from PeqP' have "(p \ ((p \ \\<^sub>P) \ \')) \ (p \ \\<^sub>P')" + by (metis AssertionStatEqClosed) + with \distinctPerm p\ have PeqP'2: "\\<^sub>P \ (p \ \') \ (p \ \\<^sub>P')" + by(simp add: eqvts) + + note FrP'2 PeqP'2 + + moreover from \x \ C\ \(p \ xvec) \* C\ \A\<^sub>P' \* C\ + have "(x#((p \ xvec)@A\<^sub>P')) \* C" by simp + moreover from \(x # A\<^sub>P) \* \\x\(\\*xvec\P')\ \A\<^sub>P' \* (x#X)\ \A\<^sub>P' \* P'\ \A\<^sub>P' \* xvec\ + \x \ (p \ xvec)\ \(p \ xvec) \* xvec\ \(p \ xvec) \* P'\ + have "(x#((p \ xvec)@A\<^sub>P')) \* (\\x\(\\*xvec\P'))" + by simp + moreover from \x \ X\ \(p \ xvec) \* X\ \A\<^sub>P' \* X\ + have "(x#((p \ xvec)@A\<^sub>P')) \* X" by simp + moreover from \x \ Y\ \(p \ xvec) \* Y\ \A\<^sub>P' \* Y\ + have "(x#((p \ xvec)@A\<^sub>P')) \* Y" by simp + moreover from \x \ Z\ \(p \ xvec) \* Z\ \A\<^sub>P' \* Z\ + have "(x#((p \ xvec)@A\<^sub>P')) \* Z" by simp + moreover from \x \ (p \ xvec)\ \x \ A\<^sub>P'\ \(p \ xvec) \* A\<^sub>P'\ + \distinct A\<^sub>P'\ \distinct xvec\ \distinctPerm p\ + have "distinct (x#((p \ xvec)@A\<^sub>P'))" by simp + ultimately show ?case + by(rule cBrClose) + next + case(cScope \ P P' x A\<^sub>P \\<^sub>P C X Y Z) + then obtain \' A\<^sub>P' \\<^sub>P' where FrP': "extractFrame P' = \A\<^sub>P', \\<^sub>P'\" and "distinct A\<^sub>P'" + and "\\<^sub>P \ \' \ \\<^sub>P'" and "A\<^sub>P' \* C" and "A\<^sub>P' \* P'" + and "A\<^sub>P' \* (x#X)" and "A\<^sub>P' \* Y" and "A\<^sub>P' \* Z" + using cScope(4)[where ba="x#X"] by (metis freshStarAtom fresh_star_list_cons(1)) + from \A\<^sub>P' \* (x#X)\ have "x \ A\<^sub>P'" and "A\<^sub>P' \* X" by simp+ + moreover from FrP' have "extractFrame(\\x\P') = \(x#A\<^sub>P'), \\<^sub>P'\" by simp + moreover note \\\<^sub>P \ \' \ \\<^sub>P'\ + moreover from \x \ C\ \A\<^sub>P' \* C\ have "(x#A\<^sub>P') \* C" by simp + moreover from \A\<^sub>P' \* P'\ have "(x#A\<^sub>P') \* (\\x\P')" by(simp add: abs_fresh fresh_star_def) + moreover from \x \ X\ \A\<^sub>P' \* X\ have "(x#A\<^sub>P') \* X" by simp + moreover from \x \ Y\ \A\<^sub>P' \* Y\ have "(x#A\<^sub>P') \* Y" by simp + moreover from \x \ Z\ \A\<^sub>P' \* Z\ have "(x#A\<^sub>P') \* Z" by simp + moreover from \x \ A\<^sub>P'\ \distinct A\<^sub>P'\ have "distinct(x#A\<^sub>P')" by simp + ultimately show ?case by(elim cScope) + next + case(cBang \ P P' A\<^sub>P \\<^sub>P C B Y Z) + then obtain \' A\<^sub>P' \\<^sub>P' where FrP': "extractFrame P' = \A\<^sub>P', \\<^sub>P'\" + and "(\\<^sub>P \ \) \ \' \ \\<^sub>P'" + and "A\<^sub>P' \* C" and "A\<^sub>P' \* P'" and "distinct A\<^sub>P'" + and "A\<^sub>P' \* B" and "A\<^sub>P' \* Y" and "A\<^sub>P' \* Z" + using cBang by (metis psiFreshVec(4) psiFreshVec(7)) + with \\\<^sub>P \ \\ \(\\<^sub>P \ \) \ \' \ \\<^sub>P'\ have "\ \ \' \ \\<^sub>P'" + by(metis Identity AssertionStatEqTrans composition' Commutativity Associativity AssertionStatEqSym) + then show ?case using FrP' \A\<^sub>P' \* P'\ \A\<^sub>P' \* C\ \A\<^sub>P' \* B\ \A\<^sub>P' \* Y\ \A\<^sub>P' \* Z\ \distinct A\<^sub>P'\ + by(elim cBang) + qed + with A have ?thesis by simp + } + moreover have "A\<^sub>P \* ([]::name list)" and "A\<^sub>P \* ([]::'b list)" and "A\<^sub>P \* ([]::('a, 'b, 'c) psi list)" by simp+ + ultimately show ?thesis by blast +qed + +lemma expandFrame: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and \ :: "'a action" + and P' :: "('a, 'b, 'c) psi" + and A\<^sub>P :: "name list" + and \\<^sub>P :: 'b + and C :: "'f::fs_name" + and C' :: "'g::fs_name" + +assumes Trans: "\ \ P \\ \ P'" + and "extractFrame P = \A\<^sub>P, \\<^sub>P\" + and "distinct A\<^sub>P" + and "bn \ \* subject \" + and "distinct(bn \)" + and "A\<^sub>P \* \" + and "A\<^sub>P \* P" + and "A\<^sub>P \* C" + and "A\<^sub>P \* C'" + and "bn \ \* P" + and "bn \ \* C'" + +obtains p \' A\<^sub>P' \\<^sub>P' where "set p \ set(bn \) \ set(bn(p \ \))" and "(p \ \\<^sub>P) \ \' \ \\<^sub>P'" and "distinctPerm p" and + "extractFrame P' = \A\<^sub>P', \\<^sub>P'\" and "A\<^sub>P' \* P'" and "A\<^sub>P' \* \" and "A\<^sub>P' \* (p \ \)" and + "A\<^sub>P' \* C" and "(bn(p \ \)) \* C'" and "(bn(p \ \)) \* \" and "(bn(p \ \)) \* P'" and "distinct A\<^sub>P'" + using assms + apply(cases "\=\") + by(auto intro: expandTauFrame[where C=C] expandNonTauFrame[where C=C and C'=C']) + +abbreviation + frameImpJudge ("_ \\<^sub>F _" [80, 80] 80) + where "F \\<^sub>F G \ FrameStatImp F G" + +lemma FrameStatEqImpCompose: + fixes F :: "'b frame" + and G :: "'b frame" + and H :: "'b frame" + and I :: "'b frame" + +assumes "F \\<^sub>F G" + and "G \\<^sub>F H" + and "H \\<^sub>F I" + +shows "F \\<^sub>F I" + using assms + by(auto simp add: FrameStatEq_def) (blast intro: FrameStatImpTrans) + +lemma transferNonTauFrame: + fixes \\<^sub>F :: 'b + and P :: "('a, 'b, 'c) psi" + and \ :: "'a action" + and P' :: "('a, 'b, 'c) psi" + and A\<^sub>F :: "name list" + and A\<^sub>G :: "name list" + and \\<^sub>G :: 'b + +assumes "\\<^sub>F \ P \\ \ P'" + and "extractFrame P = \A\<^sub>P, \\<^sub>P\" + and "distinct A\<^sub>P" + and "distinct(bn \)" + and "\A\<^sub>F, \\<^sub>F \ \\<^sub>P\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ \\<^sub>P\" + and "A\<^sub>F \* P" + and "A\<^sub>G \* P" + and "A\<^sub>F \* subject \" + and "A\<^sub>G \* subject \" + and "A\<^sub>P \* A\<^sub>F" + and "A\<^sub>P \* A\<^sub>G" + and "A\<^sub>P \* \\<^sub>F" + and "A\<^sub>P \* \\<^sub>G" + and "\ \ \" + +shows "\\<^sub>G \ P \\ \ P'" + using assms +proof(nominal_induct \\<^sub>F P Rs=="\ \ P'" A\<^sub>P \\<^sub>P avoiding: \ P' \\<^sub>G A\<^sub>F A\<^sub>G rule: semanticsFrameInduct) + case(cAlpha \\<^sub>F P A\<^sub>P \\<^sub>P p \ P' \\<^sub>G A\<^sub>F A\<^sub>G) + from \\A\<^sub>F, \\<^sub>F \ (p \ \\<^sub>P)\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ (p \ \\<^sub>P)\\ + have "(p \ (\A\<^sub>F, \\<^sub>F \ (p \ \\<^sub>P)\)) \\<^sub>F (p \ (\A\<^sub>G, \\<^sub>G \ (p \ \\<^sub>P)\))" + by(rule FrameStatImpClosed) + with \A\<^sub>P \* A\<^sub>F\ \(p \ A\<^sub>P) \* A\<^sub>F\ \A\<^sub>P \* \\<^sub>F\ \(p \ A\<^sub>P) \* \\<^sub>F\ \A\<^sub>P \* A\<^sub>G\ \(p \ A\<^sub>P) \* A\<^sub>G\ \A\<^sub>P \* \\<^sub>G\ \(p \ A\<^sub>P) \* \\<^sub>G\ + \distinctPerm p\ \set p \ set A\<^sub>P \ set (p \ A\<^sub>P)\ have "\A\<^sub>F, \\<^sub>F \ \\<^sub>P\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ \\<^sub>P\" + by(simp add: eqvts) + with cAlpha show ?case by force +next + case(cInput \\<^sub>F M K xvec N Tvec P \ P' \\<^sub>G A\<^sub>F A\<^sub>G) + from cInput have "A\<^sub>F \* K" and "A\<^sub>G \* K" by(auto simp add: residualInject) + + from \A\<^sub>F \* (M\\*xvec N\.P)\ \A\<^sub>G \* (M\\*xvec N\.P)\ have "A\<^sub>F \* M" and "A\<^sub>G \* M" by simp+ + from \\\<^sub>F \ M \ K\ + have "\\<^sub>F \ \ \ M \ K" + by(blast intro: statEqEnt Identity AssertionStatEqSym) + with \A\<^sub>F \* M\ \A\<^sub>F \* K\ + have "(\A\<^sub>F, \\<^sub>F \ \\) \\<^sub>F M \ K" + by(force intro: frameImpI) + with \\A\<^sub>F, \\<^sub>F \ \\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ \\\ + have "(\A\<^sub>G, \\<^sub>G \ \\) \\<^sub>F M \ K" + by(simp add: FrameStatEq_def FrameStatImp_def) + with \A\<^sub>G \* M\ \A\<^sub>G \* K\ + have "\\<^sub>G \ \ \ M \ K" by(force dest: frameImpE) + then have "\\<^sub>G \ M \ K" by(blast intro: statEqEnt Identity) + then show ?case using \distinct xvec\ \set xvec \ supp N\ \length xvec = length Tvec\ using cInput Input + by(force simp add: residualInject) +next + case(cBrInput \\<^sub>F M K xvec N Tvec P \ P' \\<^sub>G A\<^sub>F A\<^sub>G) + from cBrInput have "A\<^sub>F \* K" and "A\<^sub>G \* K" by(auto simp add: residualInject) + + from \A\<^sub>F \* (M\\*xvec N\.P)\ \A\<^sub>G \* (M\\*xvec N\.P)\ have "A\<^sub>F \* M" and "A\<^sub>G \* M" by simp+ + from \\\<^sub>F \ K \ M\ + have "\\<^sub>F \ \ \ K \ M" + by(blast intro: statEqEnt Identity AssertionStatEqSym) + with \A\<^sub>F \* M\ \A\<^sub>F \* K\ + have "(\A\<^sub>F, \\<^sub>F \ \\) \\<^sub>F K \ M" + by(force intro: frameImpI) + with \\A\<^sub>F, \\<^sub>F \ \\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ \\\ + have "(\A\<^sub>G, \\<^sub>G \ \\) \\<^sub>F K \ M" + by(simp add: FrameStatEq_def FrameStatImp_def) + with \A\<^sub>G \* M\ \A\<^sub>G \* K\ + have "\\<^sub>G \ \ \ K \ M" by(force dest: frameImpE) + then have "\\<^sub>G \ K \ M" by(blast intro: statEqEnt Identity) + then show ?case using \distinct xvec\ \set xvec \ supp N\ \length xvec = length Tvec\ using cBrInput BrInput + by(force simp add: residualInject) +next + case(cOutput \\<^sub>F M K N P \ P' \\<^sub>G A\<^sub>F A\<^sub>G) + from cOutput have "A\<^sub>F \* K" and "A\<^sub>G \* K" by(auto simp add: residualInject) + + from \A\<^sub>F \* (M\N\.P)\ \A\<^sub>G \* (M\N\.P)\ have "A\<^sub>F \* M" and "A\<^sub>G \* M" by simp+ + from \\\<^sub>F \ M \ K\ + have "\\<^sub>F \ \ \ M \ K" + by(blast intro: statEqEnt Identity AssertionStatEqSym) + with \A\<^sub>F \* M\ \A\<^sub>F \* K\ + have "(\A\<^sub>F, \\<^sub>F \ \\) \\<^sub>F M \ K" + by(force intro: frameImpI) + with \\A\<^sub>F, \\<^sub>F \ \\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ \\\ + have "(\A\<^sub>G, \\<^sub>G \ \\) \\<^sub>F M \ K" + by(simp add: FrameStatImp_def) + with \A\<^sub>G \* M\ \A\<^sub>G \* K\ + have "\\<^sub>G \ \ \ M \ K" by(force dest: frameImpE) + then have "\\<^sub>G \ M \ K" by(blast intro: statEqEnt Identity) + then show ?case using cOutput Output by(force simp add: residualInject) +next + case(cBrOutput \\<^sub>F M K N P \ P' \\<^sub>G A\<^sub>F A\<^sub>G) + from cBrOutput have "A\<^sub>F \* K" and "A\<^sub>G \* K" by(auto simp add: residualInject) + + from \A\<^sub>F \* (M\N\.P)\ \A\<^sub>G \* (M\N\.P)\ have "A\<^sub>F \* M" and "A\<^sub>G \* M" by simp+ + from \\\<^sub>F \ M \ K\ + have "\\<^sub>F \ \ \ M \ K" + by(blast intro: statEqEnt Identity AssertionStatEqSym) + with \A\<^sub>F \* M\ \A\<^sub>F \* K\ + have "(\A\<^sub>F, \\<^sub>F \ \\) \\<^sub>F M \ K" + by(force intro: frameImpI) + with \\A\<^sub>F, \\<^sub>F \ \\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ \\\ + have "(\A\<^sub>G, \\<^sub>G \ \\) \\<^sub>F M \ K" + by(simp add: FrameStatImp_def) + with \A\<^sub>G \* M\ \A\<^sub>G \* K\ + have "\\<^sub>G \ \ \ M \ K" by(force dest: frameImpE) + then have "\\<^sub>G \ M \ K" by(blast intro: statEqEnt Identity) + then show ?case using cBrOutput BrOutput by(force simp add: residualInject) +next + case(cCase \\<^sub>F P \ Cs A\<^sub>P \\<^sub>P \ P' \\<^sub>G A\<^sub>F A\<^sub>G) + from \\\<^sub>P \ \\ have "\A\<^sub>F, \\<^sub>F \ \\<^sub>P\ \\<^sub>F \A\<^sub>F, \\<^sub>F \ \\" + by(metis frameIntCompositionSym Identity AssertionStatEqTrans) + moreover note \\A\<^sub>F, \\<^sub>F \ \\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ \\\ + moreover from \\\<^sub>P \ \\ have "\A\<^sub>G, \\<^sub>G \ \\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ \\<^sub>P\" + by(metis frameIntCompositionSym Identity AssertionStatEqTrans AssertionStatEqSym) + ultimately have "\A\<^sub>F, \\<^sub>F \ \\<^sub>P\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ \\<^sub>P\" + by(rule FrameStatEqImpCompose) + with cCase have "\\<^sub>G \ P \\ \ P'" + by (metis freshStarPair(2) memFreshChain(1) psiCasesFreshChain(1) psiFreshVec(3)) + moreover note \(\, P) \ set Cs\ + moreover from \A\<^sub>F \* (Cases Cs)\\A\<^sub>G \* (Cases Cs)\ \(\, P) \ set Cs\ have "A\<^sub>F \* \" and "A\<^sub>G \* \" + by(auto dest: memFreshChain) + from \\\<^sub>F \ \\ have "\\<^sub>F \ \ \ \" by(blast intro: statEqEnt Identity AssertionStatEqSym) + with \A\<^sub>F \* \\ have "(\A\<^sub>F, \\<^sub>F \ \\) \\<^sub>F \" by(force intro: frameImpI) + with \\A\<^sub>F, \\<^sub>F \ \\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ \\\ have "(\A\<^sub>G, \\<^sub>G \ \\) \\<^sub>F \" + by(simp add: FrameStatImp_def) + with \A\<^sub>G \* \\ have "\\<^sub>G \ \ \ \" by(force dest: frameImpE) + then have "\\<^sub>G \ \" by(blast intro: statEqEnt Identity) + ultimately show ?case using \guarded P\ cCase Case by(force intro: residualInject) +next + case(cPar1 \\<^sub>F \\<^sub>Q P \ P' A\<^sub>Q Q A\<^sub>P \\<^sub>P \' PQ' \\<^sub>G A\<^sub>F A\<^sub>G) + from \A\<^sub>F \* (P \ Q)\ \A\<^sub>G \* (P \ Q)\ have "A\<^sub>F \* P" and "A\<^sub>G \* P" and "A\<^sub>F \* Q" and "A\<^sub>G \* Q" + by simp+ + have FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" by fact + have "\A\<^sub>F, (\\<^sub>F \ \\<^sub>Q) \ \\<^sub>P\ \\<^sub>F \A\<^sub>F, \\<^sub>F \ \\<^sub>P \ \\<^sub>Q\" + by(metis Associativity Composition AssertionStatEqSym AssertionStatEqTrans Commutativity frameResChainPres frameNilStatEq) + moreover note \\A\<^sub>F, \\<^sub>F \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ \\<^sub>P \ \\<^sub>Q\\ + moreover have "\A\<^sub>G, \\<^sub>G \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>G, (\\<^sub>G \ \\<^sub>Q) \ \\<^sub>P\ " + by(metis Associativity Composition AssertionStatEqSym AssertionStatEqTrans Commutativity frameResChainPres frameNilStatEq) + ultimately have "\A\<^sub>F, (\\<^sub>F \ \\<^sub>Q) \ \\<^sub>P\ \\<^sub>F \A\<^sub>G, (\\<^sub>G \ \\<^sub>Q) \ \\<^sub>P\" + by(rule FrameStatEqImpCompose) + moreover from \A\<^sub>F \* Q\ \A\<^sub>G \* Q\ FrQ \A\<^sub>Q \* A\<^sub>F\ \A\<^sub>Q \* A\<^sub>G\ have "A\<^sub>F \* \\<^sub>Q" and "A\<^sub>G \* \\<^sub>Q" + by(force dest: extractFrameFreshChain)+ + moreover note \A\<^sub>F \* P\ \A\<^sub>G \* P\ \A\<^sub>F \* subject \'\ \A\<^sub>G \* subject \'\ \A\<^sub>P \* A\<^sub>F\ \A\<^sub>P \* A\<^sub>G\ \A\<^sub>P \* \\<^sub>F\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>P \* A\<^sub>G\ \A\<^sub>P \* \\<^sub>G\ + moreover from \\ \ P' \ Q = \' \ PQ'\ \bn \ \* \'\ + obtain p P'' where "\ \ P' = \' \ P''" and "set p \ set(bn \') \ set(bn \)" and "PQ' = P'' \ (p \ Q)" + apply(drule_tac sym) + by(rule actionPar1Dest) (assumption | simp | blast dest: sym)+ + ultimately have "\\<^sub>G \ \\<^sub>Q \ P \\ \ P'" using \\' \ \\ by(force intro: cPar1) + then show ?case using FrQ \(bn \) \* Q\ \A\<^sub>Q \* \\<^sub>G\ \A\<^sub>Q \* P\ \A\<^sub>Q \* \\ using cPar1 + by(metis Par1) +next + case(cPar2 \\<^sub>F \\<^sub>P Q \ Q' A\<^sub>P P A\<^sub>Q \\<^sub>Q \' PQ' \\<^sub>G A\<^sub>F A\<^sub>G) + from \A\<^sub>F \* (P \ Q)\ \A\<^sub>G \* (P \ Q)\ have "A\<^sub>F \* P" and "A\<^sub>G \* P" and "A\<^sub>F \* Q" and "A\<^sub>G \* Q" + by simp+ + have FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" by fact + have "\A\<^sub>F, (\\<^sub>F \ \\<^sub>P) \ \\<^sub>Q\ \\<^sub>F \A\<^sub>F, \\<^sub>F \ \\<^sub>P \ \\<^sub>Q\" + by(metis Associativity frameResChainPres frameNilStatEq) + moreover note \\A\<^sub>F, \\<^sub>F \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ \\<^sub>P \ \\<^sub>Q\\ + moreover have "\A\<^sub>G, \\<^sub>G \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>G, (\\<^sub>G \ \\<^sub>P) \ \\<^sub>Q\ " + by(metis Associativity AssertionStatEqSym frameResChainPres frameNilStatEq) + ultimately have "\A\<^sub>F, (\\<^sub>F \ \\<^sub>P) \ \\<^sub>Q\ \\<^sub>F \A\<^sub>G, (\\<^sub>G \ \\<^sub>P) \ \\<^sub>Q\" + by(rule FrameStatEqImpCompose) + moreover from \A\<^sub>F \* P\ \A\<^sub>G \* P\ FrP \A\<^sub>P \* A\<^sub>F\ \A\<^sub>P \* A\<^sub>G\ have "A\<^sub>F \* \\<^sub>P" and "A\<^sub>G \* \\<^sub>P" + by(force dest: extractFrameFreshChain)+ + moreover note \A\<^sub>F \* Q\ \A\<^sub>G \* Q\ \A\<^sub>F \* subject \'\ \A\<^sub>G \* subject \'\ \A\<^sub>Q \* A\<^sub>F\ \A\<^sub>Q \* A\<^sub>G\ \A\<^sub>Q \* \\<^sub>F\ \A\<^sub>Q \* \\<^sub>P\ \A\<^sub>Q \* A\<^sub>G\ \A\<^sub>Q \* \\<^sub>G\ + moreover from \\ \ P \ Q' = \' \ PQ'\ \bn \ \* \'\ + obtain p Q'' where "\ \ Q' = \' \ Q''" and "set p \ set(bn \') \ set(bn \)" and "PQ' = (p \ P) \ Q''" + apply(drule_tac sym) + by(rule actionPar2Dest) (assumption | simp | blast dest: sym)+ + ultimately have "\\<^sub>G \ \\<^sub>P \ Q \\ \ Q'" using \\' \ \\ by(force intro: cPar2) + then show ?case using FrP \(bn \) \* P\ \A\<^sub>P \* \\<^sub>G\ \A\<^sub>P \* Q\ \A\<^sub>P \* \\ using cPar2 + by(metis Par2) +next + case cComm1 + then show ?case by(simp add: residualInject) +next + case cComm2 + then show ?case by(simp add: residualInject) +next + case(cBrMerge \\<^sub>F \\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q Q' A\<^sub>Q \ PQ' \\<^sub>G A\<^sub>F A\<^sub>G) + from \\A\<^sub>F, \\<^sub>F \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ \\<^sub>P \ \\<^sub>Q\\ + have "\A\<^sub>F, (\\<^sub>F \ \\<^sub>Q) \ \\<^sub>P\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ \\<^sub>P \ \\<^sub>Q\" + by (metis AssertionStatEqTrans AssertionStatEq_def Associativity FrameStatImpTrans associativitySym frameImpNilStatEq frameImpResChainPres) + moreover have "\A\<^sub>G, \\<^sub>G \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>G, (\\<^sub>G \ \\<^sub>Q) \ \\<^sub>P\" + by (metis AssertionStatEqTrans AssertionStatEq_def Associativity associativitySym frameImpNilStatEq frameImpResChainPres) + ultimately have FimpP: "\A\<^sub>F, (\\<^sub>F \ \\<^sub>Q) \ \\<^sub>P\ \\<^sub>F \A\<^sub>G, (\\<^sub>G \ \\<^sub>Q) \ \\<^sub>P\" + by (metis FrameStatImpTrans) + + from \\M\N\ \ P' \ Q' = \ \ PQ'\ have "\ = \M\N\" and "(P' \ Q') = PQ'" + by(simp add: residualInject)+ + + moreover note FimpP \A\<^sub>F \* (P \ Q)\ \A\<^sub>G \* (P \ Q)\ \A\<^sub>F \* subject \\ \A\<^sub>G \* subject \\ + \A\<^sub>P \* A\<^sub>F\ \A\<^sub>P \* A\<^sub>G\ \A\<^sub>P \* \\<^sub>F\ \A\<^sub>P \* \\<^sub>G\ \A\<^sub>P \* \\<^sub>Q\ + ultimately have TransP: "\\<^sub>G \ \\<^sub>Q \ P \ \M\N\ \ P'" + by(intro cBrMerge(2)[where bd="A\<^sub>G"]) auto + + from \\A\<^sub>F, \\<^sub>F \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ \\<^sub>P \ \\<^sub>Q\\ + have "\A\<^sub>F, (\\<^sub>F \ \\<^sub>P) \ \\<^sub>Q\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ \\<^sub>P \ \\<^sub>Q\" + by (metis FrameStatEq_def FrameStatImpTrans frameIntAssociativity) + moreover have "\A\<^sub>G, \\<^sub>G \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>G, (\\<^sub>G \ \\<^sub>P) \ \\<^sub>Q\" + by (metis FrameStatEq_def frameIntAssociativity) + ultimately have FimpQ: "\A\<^sub>F, (\\<^sub>F \ \\<^sub>P) \ \\<^sub>Q\ \\<^sub>F \A\<^sub>G, (\\<^sub>G \ \\<^sub>P) \ \\<^sub>Q\" + by (metis FrameStatImpTrans) + + note \\ = \M\N\\ \(P' \ Q') = PQ'\ FimpQ \A\<^sub>F \* (P \ Q)\ \A\<^sub>G \* (P \ Q)\ \A\<^sub>F \* subject \\ \A\<^sub>G \* subject \\ + \A\<^sub>Q \* A\<^sub>F\ \A\<^sub>Q \* A\<^sub>G\ \A\<^sub>Q \* \\<^sub>F\ \A\<^sub>Q \* \\<^sub>G\ \A\<^sub>Q \* \\<^sub>P\ + then have TransQ: "\\<^sub>G \ \\<^sub>P \ Q \ \M\N\ \ Q'" + by(intro cBrMerge(6)[where bd="A\<^sub>G"]) auto + + have "\\<^sub>G \ P \ Q \ \M\N\ \ P' \ Q'" using TransP TransQ cBrMerge + by(intro BrMerge) + with \\M\N\ \ P' \ Q' = \ \ PQ'\ + show ?case by simp +next + case(cBrComm1 \\<^sub>F \\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q xvec Q' A\<^sub>Q \ PQ' \\<^sub>G A\<^sub>F A\<^sub>G) + from \\M\\*xvec\\N\ \ P' \ Q' = \ \ PQ'\ \A\<^sub>F \* subject \\ \A\<^sub>G \* subject \\ + have "A\<^sub>F \* M" and "A\<^sub>G \* M" + by(auto simp add: residualInject) + + from \\A\<^sub>F, \\<^sub>F \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ \\<^sub>P \ \\<^sub>Q\\ + have "\A\<^sub>F, (\\<^sub>F \ \\<^sub>Q) \ \\<^sub>P\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ \\<^sub>P \ \\<^sub>Q\" + by (metis AssertionStatEqTrans AssertionStatEq_def Associativity FrameStatImpTrans associativitySym frameImpNilStatEq frameImpResChainPres) + moreover have "\A\<^sub>G, \\<^sub>G \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>G, (\\<^sub>G \ \\<^sub>Q) \ \\<^sub>P\" + by (metis AssertionStatEqTrans AssertionStatEq_def Associativity associativitySym frameImpNilStatEq frameImpResChainPres) + ultimately have FimpP: "\A\<^sub>F, (\\<^sub>F \ \\<^sub>Q) \ \\<^sub>P\ \\<^sub>F \A\<^sub>G, (\\<^sub>G \ \\<^sub>Q) \ \\<^sub>P\" + by (metis FrameStatImpTrans) + + note FimpP \A\<^sub>F \* (P \ Q)\ \A\<^sub>G \* (P \ Q)\ \A\<^sub>F \* M\ \A\<^sub>G \* M\ + \A\<^sub>P \* A\<^sub>F\ \A\<^sub>P \* A\<^sub>G\ \A\<^sub>P \* \\<^sub>F\ \A\<^sub>P \* \\<^sub>G\ \A\<^sub>P \* \\<^sub>Q\ + then have TransP: "\\<^sub>G \ \\<^sub>Q \ P \ \M\N\ \ P'" + by (intro cBrComm1(4)[where bc="A\<^sub>F" and bd="A\<^sub>G"]) auto + from \\A\<^sub>F, \\<^sub>F \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ \\<^sub>P \ \\<^sub>Q\\ + have "\A\<^sub>F, (\\<^sub>F \ \\<^sub>P) \ \\<^sub>Q\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ \\<^sub>P \ \\<^sub>Q\" + by (metis FrameStatEq_def FrameStatImpTrans frameIntAssociativity) + moreover have "\A\<^sub>G, \\<^sub>G \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>G, (\\<^sub>G \ \\<^sub>P) \ \\<^sub>Q\" + by (metis FrameStatEq_def frameIntAssociativity) + ultimately have FimpQ: "\A\<^sub>F, (\\<^sub>F \ \\<^sub>P) \ \\<^sub>Q\ \\<^sub>F \A\<^sub>G, (\\<^sub>G \ \\<^sub>P) \ \\<^sub>Q\" + by (metis FrameStatImpTrans) + note \distinct xvec\ FimpQ \A\<^sub>F \* (P \ Q)\ \A\<^sub>G \* (P \ Q)\ \A\<^sub>F \* M\ \A\<^sub>G \* M\ + \A\<^sub>Q \* A\<^sub>F\ \A\<^sub>Q \* A\<^sub>G\ \A\<^sub>Q \* \\<^sub>F\ \A\<^sub>Q \* \\<^sub>G\ \A\<^sub>Q \* \\<^sub>P\ + then have TransQ: "\\<^sub>G \ \\<^sub>P \ Q \ \M\\*xvec\\N\ \ Q'" + by (simp add: cBrComm1.hyps(8) freshCompChain(1)) + + from TransP TransQ cBrComm1 + have "\\<^sub>G \ P \ Q \ \M\\*xvec\\N\ \ P' \ Q'" + by(intro BrComm1) + with \\M\\*xvec\\N\ \ P' \ Q' = \ \ PQ'\ + show ?case + by simp +next + case(cBrComm2 \\<^sub>F \\<^sub>Q P M xvec N P' A\<^sub>P \\<^sub>P Q Q' A\<^sub>Q \ PQ' \\<^sub>G A\<^sub>F A\<^sub>G) + from \\M\\*xvec\\N\ \ P' \ Q' = \ \ PQ'\ \A\<^sub>F \* subject \\ \A\<^sub>G \* subject \\ + have "A\<^sub>F \* M" and "A\<^sub>G \* M" + by(auto simp add: residualInject) + + from \\A\<^sub>F, \\<^sub>F \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ \\<^sub>P \ \\<^sub>Q\\ + have "\A\<^sub>F, (\\<^sub>F \ \\<^sub>P) \ \\<^sub>Q\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ \\<^sub>P \ \\<^sub>Q\" + by (metis FrameStatEq_def FrameStatImpTrans frameIntAssociativity) + moreover have "\A\<^sub>G, \\<^sub>G \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>G, (\\<^sub>G \ \\<^sub>P) \ \\<^sub>Q\" + by (metis FrameStatEq_def frameIntAssociativity) + ultimately have FimpQ: "\A\<^sub>F, (\\<^sub>F \ \\<^sub>P) \ \\<^sub>Q\ \\<^sub>F \A\<^sub>G, (\\<^sub>G \ \\<^sub>P) \ \\<^sub>Q\" + by (metis FrameStatImpTrans) + + note FimpQ \A\<^sub>F \* (P \ Q)\ \A\<^sub>G \* (P \ Q)\ \A\<^sub>F \* M\ \A\<^sub>G \* M\ + \A\<^sub>Q \* A\<^sub>F\ \A\<^sub>Q \* A\<^sub>G\ \A\<^sub>Q \* \\<^sub>F\ \A\<^sub>Q \* \\<^sub>G\ \A\<^sub>Q \* \\<^sub>P\ + then have TransQ: "\\<^sub>G \ \\<^sub>P \ Q \ \M\N\ \ Q'" + by(intro cBrComm2(8)[where bc="A\<^sub>F" and bd="A\<^sub>G"]) auto + + from \\A\<^sub>F, \\<^sub>F \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ \\<^sub>P \ \\<^sub>Q\\ + have "\A\<^sub>F, (\\<^sub>F \ \\<^sub>Q) \ \\<^sub>P\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ \\<^sub>P \ \\<^sub>Q\" + by (metis AssertionStatEqTrans AssertionStatEq_def Associativity FrameStatImpTrans associativitySym frameImpNilStatEq frameImpResChainPres) + moreover have "\A\<^sub>G, \\<^sub>G \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>G, (\\<^sub>G \ \\<^sub>Q) \ \\<^sub>P\" + by (metis AssertionStatEqTrans AssertionStatEq_def Associativity associativitySym frameImpNilStatEq frameImpResChainPres) + ultimately have FimpP: "\A\<^sub>F, (\\<^sub>F \ \\<^sub>Q) \ \\<^sub>P\ \\<^sub>F \A\<^sub>G, (\\<^sub>G \ \\<^sub>Q) \ \\<^sub>P\" + by (metis FrameStatImpTrans) + + note \distinct xvec\ FimpP \A\<^sub>F \* (P \ Q)\ \A\<^sub>G \* (P \ Q)\ \A\<^sub>F \* M\ \A\<^sub>G \* M\ + \A\<^sub>P \* A\<^sub>F\ \A\<^sub>P \* A\<^sub>G\ \A\<^sub>P \* \\<^sub>F\ \A\<^sub>P \* \\<^sub>G\ \A\<^sub>P \* \\<^sub>Q\ + then have TransP: "\\<^sub>G \ \\<^sub>Q \ P \ \M\\*xvec\\N\ \ P'" + by(intro cBrComm2(4)[where bc="A\<^sub>F" and bd="A\<^sub>G"]) auto + + from TransP TransQ cBrComm2 + have "\\<^sub>G \ P \ Q \ \M\\*xvec\\N\ \ P' \ Q'" + by(intro BrComm2) + with \\M\\*xvec\\N\ \ P' \ Q' = \ \ PQ'\ + show ?case + by simp +next + case cBrClose + then show ?case by(simp add: residualInject) +next + case(cOpen \\<^sub>F P M xvec yvec N P' x A\<^sub>P \\<^sub>P \ P'' \\<^sub>G A\<^sub>F A\<^sub>G) + from \M\\*(xvec @ x # yvec)\\N\ \ P' = \ \ P''\ \x \ xvec\ \x \ yvec\ \x \ \\ \x \ P''\ \distinct(bn \)\ + obtain xvec' x' yvec' N' where "M\\*(xvec@yvec)\\N\ \ P' = M\\*(xvec'@yvec')\\([(x, x')] \ N')\ \ ([(x, x')] \ P'')" + and "\ = M\\*(xvec'@x'#yvec')\\N'\" + apply(cases rule: actionCases[where \=\]) + apply(simp add: residualInject) + apply(simp add: residualInject) + apply(simp add: residualInject) + apply(metis boundOutputOpenDest) + apply(simp add: residualInject) + by(simp add: residualInject) + + then have "\\<^sub>G \ P \M\\*(xvec@yvec)\\N\ \ P'" using cOpen + by(intro cOpen(4)[where bc="A\<^sub>F" and bd="A\<^sub>G"]) auto + with \x \ supp N\ \x \ \\<^sub>G\ \x \ M\ \x \ xvec\ \x \ yvec\ + have "\\<^sub>G \ \\x\P \M\\*(xvec@x#yvec)\\N\ \ P'" + by(intro Open) + then show ?case using \\ = M\\*(xvec'@x'#yvec')\\N'\\ \M\\*(xvec @ x # yvec)\\N\ \ P' = \ \ P''\ + by simp +next + case(cBrOpen \\<^sub>F P M xvec yvec N P' x A\<^sub>P \\<^sub>P \ P'' \\<^sub>G A\<^sub>F A\<^sub>G) + from \\M\\*(xvec @ x # yvec)\\N\ \ P' = \ \ P''\ \x \ xvec\ \x \ yvec\ \x \ \\ \x \ P''\ \distinct(bn \)\ + obtain xvec' x' yvec' N' where "\M\\*(xvec@yvec)\\N\ \ P' = \M\\*(xvec'@yvec')\\([(x, x')] \ N')\ \ ([(x, x')] \ P'')" + and "\ = \M\\*(xvec'@x'#yvec')\\N'\" + apply(cases rule: actionCases[where \=\]) + apply(simp add: residualInject) + apply(simp add: residualInject) + apply(simp add: residualInject) + apply(simp add: residualInject) + apply(metis boundOutputOpenDest) + by(simp add: residualInject) + + then have "\\<^sub>G \ P \\M\\*(xvec@yvec)\\N\ \ P'" using cBrOpen + by(intro cBrOpen) (assumption | simp)+ + with \x \ supp N\ \x \ \\<^sub>G\ \x \ M\ \x \ xvec\ \x \ yvec\ + have "\\<^sub>G \ \\x\P \\M\\*(xvec@x#yvec)\\N\ \ P'" + by(intro BrOpen) + then show ?case using \\ = \M\\*(xvec'@x'#yvec')\\N'\\ \\M\\*(xvec @ x # yvec)\\N\ \ P' = \ \ P''\ + by simp +next + case(cScope \\<^sub>F P \ P' x A\<^sub>P \\<^sub>P \' xP \\<^sub>G A\<^sub>F A\<^sub>G) + from \\ \ \\x\P' = \' \ xP\ \x \ \\ \x \ \'\ obtain P'' where "xP = \\x\P''" and "\ \ P' = \' \ P''" + by(drule_tac sym) (force intro: actionScopeDest) + then have "\\<^sub>G \ P \\ \ P'" using cScope by auto + with \x \ \\<^sub>G\ \x \ \'\ \\ \ P' = \' \ P''\ \xP = \\x\P''\ show ?case + by(metis Scope) +next + case(cBang \\<^sub>F P A\<^sub>P \\<^sub>P \ P' \\<^sub>G A\<^sub>F A\<^sub>G) + from \\\<^sub>P \ \\ have "\A\<^sub>F, \\<^sub>F \ \\<^sub>P \ \\ \\<^sub>F \A\<^sub>F, \\<^sub>F \ \\" + by(metis frameIntCompositionSym Identity AssertionStatEqTrans) + moreover note \\A\<^sub>F, \\<^sub>F \ \\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ \\\ + moreover from \\\<^sub>P \ \\ have "\A\<^sub>G, \\<^sub>G \ \\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ \\<^sub>P \ \\" + by(metis frameIntCompositionSym Identity AssertionStatEqTrans AssertionStatEqSym) + ultimately have "\A\<^sub>F, \\<^sub>F \ \\<^sub>P \ \\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ \\<^sub>P \ \\" + by(rule FrameStatEqImpCompose) + with cBang have "\\<^sub>G \ P \ !P \\ \ P'" by force + then show ?case using \guarded P\ using cBang by(metis Bang) +qed + +lemma transferTauFrame: + fixes \\<^sub>F :: 'b + and P :: "('a, 'b, 'c) psi" + and P' :: "('a, 'b, 'c) psi" + and A\<^sub>F :: "name list" + and A\<^sub>G :: "name list" + and \\<^sub>G :: 'b + +assumes "\\<^sub>F \ P \\ \ P'" + and "extractFrame P = \A\<^sub>P, \\<^sub>P\" + and "distinct A\<^sub>P" + and "\A\<^sub>F, \\<^sub>F \ \\<^sub>P\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ \\<^sub>P\" + and "A\<^sub>F \* P" + and "A\<^sub>G \* P" + and "A\<^sub>P \* A\<^sub>F" + and "A\<^sub>P \* A\<^sub>G" + and "A\<^sub>P \* \\<^sub>F" + and "A\<^sub>P \* \\<^sub>G" + +shows "\\<^sub>G \ P \\ \ P'" + using assms +proof(nominal_induct avoiding: \\<^sub>G A\<^sub>F A\<^sub>G rule: tauFrameInduct) + case(cAlpha \\<^sub>F P P' A\<^sub>P \\<^sub>P p \\<^sub>G A\<^sub>F A\<^sub>G) + from \\A\<^sub>F, \\<^sub>F \ (p \ \\<^sub>P)\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ (p \ \\<^sub>P)\\ + have "(p \ (\A\<^sub>F, \\<^sub>F \ (p \ \\<^sub>P)\)) \\<^sub>F (p \ (\A\<^sub>G, \\<^sub>G \ (p \ \\<^sub>P)\))" + by(rule FrameStatImpClosed) + with \A\<^sub>P \* A\<^sub>F\ \(p \ A\<^sub>P) \* A\<^sub>F\ \A\<^sub>P \* \\<^sub>F\ \(p \ A\<^sub>P) \* \\<^sub>F\ \A\<^sub>P \* A\<^sub>G\ \(p \ A\<^sub>P) \* A\<^sub>G\ \A\<^sub>P \* \\<^sub>G\ \(p \ A\<^sub>P) \* \\<^sub>G\ + \distinctPerm p\ \set p \ set A\<^sub>P \ set (p \ A\<^sub>P)\ have "\A\<^sub>F, \\<^sub>F \ \\<^sub>P\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ \\<^sub>P\" + by(simp add: eqvts) + with cAlpha show ?case by blast +next + case(cCase \\<^sub>F P P' \ Cs A\<^sub>P \\<^sub>P \\<^sub>G A\<^sub>F A\<^sub>G) + from \\\<^sub>P \ \\ have "\A\<^sub>F, \\<^sub>F \ \\<^sub>P\ \\<^sub>F \A\<^sub>F, \\<^sub>F \ \\" + by(metis frameIntCompositionSym Identity AssertionStatEqTrans) + moreover note \\A\<^sub>F, \\<^sub>F \ \\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ \\\ + moreover from \\\<^sub>P \ \\ have "\A\<^sub>G, \\<^sub>G \ \\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ \\<^sub>P\" + by(metis frameIntCompositionSym Identity AssertionStatEqTrans AssertionStatEqSym) + ultimately have "\A\<^sub>F, \\<^sub>F \ \\<^sub>P\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ \\<^sub>P\" + by(rule FrameStatEqImpCompose) + with cCase have "\\<^sub>G \ P \\ \ P'" + by (metis freshStarPair(2) memFreshChain(1) psiCasesFreshChain(1) psiFreshVec(3)) + moreover note \(\, P) \ set Cs\ + moreover from \A\<^sub>F \* (Cases Cs)\\A\<^sub>G \* (Cases Cs)\ \(\, P) \ set Cs\ have "A\<^sub>F \* \" and "A\<^sub>G \* \" + by(auto dest: memFreshChain) + from \\\<^sub>F \ \\ have "\\<^sub>F \ \ \ \" by(blast intro: statEqEnt Identity AssertionStatEqSym) + with \A\<^sub>F \* \\ have "(\A\<^sub>F, \\<^sub>F \ \\) \\<^sub>F \" by(force intro: frameImpI) + with \\A\<^sub>F, \\<^sub>F \ \\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ \\\ have "(\A\<^sub>G, \\<^sub>G \ \\) \\<^sub>F \" + by(simp add: FrameStatImp_def) + with \A\<^sub>G \* \\ have "\\<^sub>G \ \ \ \" by(force dest: frameImpE) + then have "\\<^sub>G \ \" by(blast intro: statEqEnt Identity) + ultimately show ?case using \guarded P\ by(rule Case) +next + case(cPar1 \\<^sub>F \\<^sub>Q P P' A\<^sub>Q Q A\<^sub>P \\<^sub>P \\<^sub>G A\<^sub>F A\<^sub>G) + from \A\<^sub>F \* (P \ Q)\ \A\<^sub>G \* (P \ Q)\ have "A\<^sub>F \* P" and "A\<^sub>G \* P" and "A\<^sub>F \* Q" and "A\<^sub>G \* Q" + by simp+ + have IH: "\\ A\<^sub>F A\<^sub>G. \\A\<^sub>F, (\\<^sub>F \ \\<^sub>Q) \ \\<^sub>P\ \\<^sub>F \A\<^sub>G, \ \ \\<^sub>P\; A\<^sub>F \* P; A\<^sub>G \* P; + A\<^sub>P \* A\<^sub>F; A\<^sub>P \* A\<^sub>G; A\<^sub>P \* (\\<^sub>F \ \\<^sub>Q); A\<^sub>P \* \\ \ \ \ P \\ \ P'" + by fact + have FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" by fact + have "\A\<^sub>F, (\\<^sub>F \ \\<^sub>Q) \ \\<^sub>P\ \\<^sub>F \A\<^sub>F, \\<^sub>F \ \\<^sub>P \ \\<^sub>Q\" + by(metis Associativity Composition AssertionStatEqSym AssertionStatEqTrans Commutativity frameResChainPres frameNilStatEq) + moreover note \\A\<^sub>F, \\<^sub>F \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ \\<^sub>P \ \\<^sub>Q\\ + moreover have "\A\<^sub>G, \\<^sub>G \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>G, (\\<^sub>G \ \\<^sub>Q) \ \\<^sub>P\ " + by(metis Associativity Composition AssertionStatEqSym AssertionStatEqTrans Commutativity frameResChainPres frameNilStatEq) + ultimately have "\A\<^sub>F, (\\<^sub>F \ \\<^sub>Q) \ \\<^sub>P\ \\<^sub>F \A\<^sub>G, (\\<^sub>G \ \\<^sub>Q) \ \\<^sub>P\" + by(rule FrameStatEqImpCompose) + moreover from \A\<^sub>F \* Q\ \A\<^sub>G \* Q\ FrQ \A\<^sub>Q \* A\<^sub>F\ \A\<^sub>Q \* A\<^sub>G\ have "A\<^sub>F \* \\<^sub>Q" and "A\<^sub>G \* \\<^sub>Q" + by(force dest: extractFrameFreshChain)+ + moreover note \A\<^sub>F \* P\ \A\<^sub>G \* P\ \A\<^sub>P \* A\<^sub>F\ \A\<^sub>P \* A\<^sub>G\ \A\<^sub>P \* \\<^sub>F\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>P \* A\<^sub>G\ \A\<^sub>P \* \\<^sub>G\ + ultimately have "\\<^sub>G \ \\<^sub>Q \ P \\ \ P'" + using IH by blast + then show ?case using FrQ \A\<^sub>Q \* \\<^sub>G\ \A\<^sub>Q \* P\ + by(intro Par1) auto +next + case(cPar2 \\<^sub>F \\<^sub>P Q Q' A\<^sub>P P A\<^sub>Q \\<^sub>Q \\<^sub>G A\<^sub>F A\<^sub>G) + from \A\<^sub>F \* (P \ Q)\ \A\<^sub>G \* (P \ Q)\ have "A\<^sub>F \* P" and "A\<^sub>G \* P" and "A\<^sub>F \* Q" and "A\<^sub>G \* Q" + by simp+ + have IH: "\\ A\<^sub>F A\<^sub>G. \\A\<^sub>F, (\\<^sub>F \ \\<^sub>P) \ \\<^sub>Q\ \\<^sub>F \A\<^sub>G, \ \ \\<^sub>Q\; A\<^sub>F \* Q; A\<^sub>G \* Q; + A\<^sub>Q \* A\<^sub>F; A\<^sub>Q \* A\<^sub>G; A\<^sub>Q \* (\\<^sub>F \ \\<^sub>P); A\<^sub>Q \* \\ \ \ \ Q \\ \ Q'" + by fact + have FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" by fact + have "\A\<^sub>F, (\\<^sub>F \ \\<^sub>P) \ \\<^sub>Q\ \\<^sub>F \A\<^sub>F, \\<^sub>F \ \\<^sub>P \ \\<^sub>Q\" + by(metis Associativity frameResChainPres frameNilStatEq) + moreover note \\A\<^sub>F, \\<^sub>F \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ \\<^sub>P \ \\<^sub>Q\\ + moreover have "\A\<^sub>G, \\<^sub>G \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>G, (\\<^sub>G \ \\<^sub>P) \ \\<^sub>Q\ " + by(metis Associativity AssertionStatEqSym frameResChainPres frameNilStatEq) + ultimately have "\A\<^sub>F, (\\<^sub>F \ \\<^sub>P) \ \\<^sub>Q\ \\<^sub>F \A\<^sub>G, (\\<^sub>G \ \\<^sub>P) \ \\<^sub>Q\" + by(rule FrameStatEqImpCompose) + moreover from \A\<^sub>F \* P\ \A\<^sub>G \* P\ FrP \A\<^sub>P \* A\<^sub>F\ \A\<^sub>P \* A\<^sub>G\ have "A\<^sub>F \* \\<^sub>P" and "A\<^sub>G \* \\<^sub>P" + by(force dest: extractFrameFreshChain)+ + moreover note \A\<^sub>F \* Q\ \A\<^sub>G \* Q\ \A\<^sub>Q \* A\<^sub>F\ \A\<^sub>Q \* A\<^sub>G\ \A\<^sub>Q \* \\<^sub>F\ \A\<^sub>Q \* \\<^sub>P\ \A\<^sub>Q \* A\<^sub>G\ \A\<^sub>Q \* \\<^sub>G\ + ultimately have "\\<^sub>G \ \\<^sub>P \ Q \\ \ Q'" + using IH by blast + then show ?case using FrP \A\<^sub>P \* \\<^sub>G\ \A\<^sub>P \* Q\ + by(intro Par2) auto +next + case(cComm1 \\<^sub>F \\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q K xvec Q' A\<^sub>Q \\<^sub>G A\<^sub>F A\<^sub>G) + have FimpG: "\A\<^sub>F, \\<^sub>F \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ \\<^sub>P \ \\<^sub>Q\" by fact + from \A\<^sub>F \* (P \ Q)\ \A\<^sub>G \* (P \ Q)\ have "A\<^sub>F \* P" and "A\<^sub>G \* P" and "A\<^sub>F \* Q" and "A\<^sub>G \* Q" + by simp+ + have FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" by fact + have FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" by fact + from \A\<^sub>F \* P\ \A\<^sub>G \* P\ FrP \A\<^sub>P \* A\<^sub>F\ \A\<^sub>P \* A\<^sub>G\ have "A\<^sub>F \* \\<^sub>P" and "A\<^sub>G \* \\<^sub>P" + by(force dest: extractFrameFreshChain)+ + from \A\<^sub>F \* Q\ \A\<^sub>G \* Q\ FrQ \A\<^sub>Q \* A\<^sub>F\ \A\<^sub>Q \* A\<^sub>G\ have "A\<^sub>F \* \\<^sub>Q" and "A\<^sub>G \* \\<^sub>Q" + by(force dest: extractFrameFreshChain)+ + + from \\\<^sub>F \ \\<^sub>P \ Q \K\\*xvec\\N\ \ Q'\ have "\\<^sub>F \ \\<^sub>P \ Q \ ROut K (\\*xvec\N \' Q')" + by(simp add: residualInject) + with FrQ \distinct A\<^sub>Q\ + obtain K' where KeqK': "(\\<^sub>F \ \\<^sub>P) \ \\<^sub>Q \ K \ K'" and "A\<^sub>P \* K'" and "A\<^sub>F \* K'" and "A\<^sub>G \* K'" + using \A\<^sub>P \* Q\ \A\<^sub>Q \* A\<^sub>F\ \A\<^sub>Q \* A\<^sub>G\ \A\<^sub>F \* Q\ \A\<^sub>G \* Q\ \A\<^sub>Q \* \\<^sub>F\ \A\<^sub>Q \* \\<^sub>P\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* K\ \xvec \* K\ \distinct xvec\ + by(elim outputObtainPrefix[where B="A\<^sub>P@A\<^sub>F@A\<^sub>G"]) force+ + have "\\<^sub>G \ \\<^sub>Q \ P \K'\N\ \ P'" + proof - + from KeqK' have "\\<^sub>F \ (\\<^sub>P \ \\<^sub>Q) \ K \ K'" by(rule statEqEnt[OF Associativity]) + with \\\<^sub>F \ (\\<^sub>P \ \\<^sub>Q) \ M \ K\ have "\\<^sub>F \ (\\<^sub>P \ \\<^sub>Q) \ M \ K'" + by(rule chanEqTrans) + then have "(\\<^sub>F \ \\<^sub>Q) \ \\<^sub>P \ M \ K'" + by(metis statEqEnt AssertionStatEqSym Associativity AssertionStatEqTrans compositionSym Commutativity) + with \\\<^sub>F \ \\<^sub>Q \ P \M\N\ \ P'\ FrP \distinct A\<^sub>P\ + have "\\<^sub>F \ \\<^sub>Q \ P \K'\N\ \ P'" using \A\<^sub>P \* \\<^sub>F\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>P \* P\ \A\<^sub>P \* M\ \A\<^sub>P \* K'\ + by(force intro: inputRenameSubject) + moreover have "\A\<^sub>F, (\\<^sub>F \ \\<^sub>Q) \ \\<^sub>P\ \\<^sub>F \A\<^sub>G, (\\<^sub>G \ \\<^sub>Q) \ \\<^sub>P\" + proof - + have "\A\<^sub>F, (\\<^sub>F \ \\<^sub>Q) \ \\<^sub>P\ \\<^sub>F \A\<^sub>F, \\<^sub>F \ \\<^sub>P \ \\<^sub>Q\" + by(metis Associativity Composition AssertionStatEqSym AssertionStatEqTrans Commutativity frameResChainPres frameNilStatEq) + moreover have "\A\<^sub>G, \\<^sub>G \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>G, (\\<^sub>G \ \\<^sub>Q) \ \\<^sub>P\ " + by(metis Associativity Composition AssertionStatEqSym AssertionStatEqTrans Commutativity frameResChainPres frameNilStatEq) + ultimately show ?thesis using FimpG + by(elim FrameStatEqImpCompose) + qed + ultimately show ?thesis using \A\<^sub>F \* P\ \A\<^sub>G \* P\ \A\<^sub>F \* K'\ + \A\<^sub>G \* K'\ \A\<^sub>P \* A\<^sub>F\ \A\<^sub>P \* A\<^sub>G\ \A\<^sub>P \* \\<^sub>F\ \A\<^sub>P \* \\<^sub>G\ \A\<^sub>P \* \\<^sub>Q\ FrP \distinct A\<^sub>P\ + by(auto intro: transferNonTauFrame) + qed + + moreover from FrP \\\<^sub>F \ \\<^sub>Q \ P \M\N\ \ P'\ \distinct A\<^sub>P\ + obtain M' where MeqM': "(\\<^sub>F \ \\<^sub>Q) \ \\<^sub>P \ M \ M'" and "A\<^sub>Q \* M'" and "A\<^sub>F \* M'" and "A\<^sub>G \* M'" + using \A\<^sub>Q \* P\ \A\<^sub>P \* A\<^sub>F\ \A\<^sub>P \* A\<^sub>G\ \A\<^sub>F \* P\ \A\<^sub>G \* P\ \A\<^sub>P \* \\<^sub>F\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* P\ \A\<^sub>P \* M\ + by(elim inputObtainPrefix[where B="A\<^sub>Q@A\<^sub>F@A\<^sub>G"]) force+ + + have "\\<^sub>G \ \\<^sub>P \ Q \M'\\*xvec\\N\ \ Q'" + proof - + from MeqM' have "\\<^sub>F \ (\\<^sub>Q \ \\<^sub>P) \ M \ M'" + by(rule statEqEnt[OF Associativity]) + with \\\<^sub>F \ (\\<^sub>P \ \\<^sub>Q) \ M \ K\ have "\\<^sub>F \ (\\<^sub>Q \ \\<^sub>P) \ K \ M'" + by(blast intro: chanEqTrans chanEqSym compositionSym Commutativity statEqEnt) + then have "(\\<^sub>F \ \\<^sub>P) \ \\<^sub>Q \ K \ M'" + by(blast intro: statEqEnt AssertionStatEqSym Associativity + AssertionStatEqTrans compositionSym Commutativity) + with \\\<^sub>F \ \\<^sub>P \ Q \K\\*xvec\\N\ \ Q'\ FrQ \distinct A\<^sub>Q\ + have "\\<^sub>F \ \\<^sub>P \ Q \M'\\*xvec\\N\ \ Q'" using \A\<^sub>Q \* \\<^sub>F\ \A\<^sub>Q \* \\<^sub>P\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* K\ \A\<^sub>Q \* M'\ + by(force intro: outputRenameSubject) + moreover have "\A\<^sub>F, (\\<^sub>F \ \\<^sub>P) \ \\<^sub>Q\ \\<^sub>F \A\<^sub>G, (\\<^sub>G \ \\<^sub>P) \ \\<^sub>Q\" + proof - + have "\A\<^sub>F, (\\<^sub>F \ \\<^sub>P) \ \\<^sub>Q\ \\<^sub>F \A\<^sub>F, \\<^sub>F \ \\<^sub>P \ \\<^sub>Q\" + by(metis Associativity frameResChainPres frameNilStatEq) + moreover have "\A\<^sub>G, \\<^sub>G \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>G, (\\<^sub>G \ \\<^sub>P) \ \\<^sub>Q\ " + by(metis Associativity AssertionStatEqSym frameResChainPres frameNilStatEq) + ultimately show ?thesis using FimpG + by(elim FrameStatEqImpCompose) + qed + + ultimately show ?thesis using \A\<^sub>F \* Q\ \A\<^sub>G \* Q\ \A\<^sub>F \* M'\ \A\<^sub>G \* M'\ + \A\<^sub>Q \* A\<^sub>F\ \A\<^sub>Q \* A\<^sub>G\ \A\<^sub>Q \* \\<^sub>F\ \A\<^sub>Q \* \\<^sub>G\ \A\<^sub>Q \* \\<^sub>P\ FrQ \distinct A\<^sub>Q\ \distinct xvec\ + by(auto intro: transferNonTauFrame) + qed + + moreover have "\\<^sub>G \ \\<^sub>P \ \\<^sub>Q \ K' \ M'" + proof - + from MeqM' have "\\<^sub>F \ \\<^sub>P \ \\<^sub>Q \ M' \ M" + by(blast intro: chanEqSym Associativity statEqEnt Commutativity compositionSym) + moreover from KeqK' have "\\<^sub>F \ \\<^sub>P \ \\<^sub>Q \ K \ K'" + by(blast intro: chanEqSym Associativity statEqEnt Commutativity compositionSym) + ultimately have "\\<^sub>F \ \\<^sub>P \ \\<^sub>Q \ K' \ M'" using \\\<^sub>F \ \\<^sub>P \ \\<^sub>Q \ M \ K\ + by(blast intro: chanEqSym chanEqTrans) + then show ?thesis using \A\<^sub>F \* M'\ \A\<^sub>F \* K'\ \A\<^sub>G \* M'\ \A\<^sub>G \* K'\ FimpG + apply(simp add: FrameStatImp_def) + apply(erule allE[where x="SChanEq' K' M'"]) + by(force intro: frameImpI dest: frameImpE) + qed + + ultimately show ?case using \A\<^sub>P \* \\<^sub>G\ \A\<^sub>P \* P\ \A\<^sub>P \* Q\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* K'\ \A\<^sub>Q \* \\<^sub>G\ \A\<^sub>Q \* P\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* M'\ \xvec \* P\ FrP FrQ + by(intro Comm1) (assumption | simp)+ +next + case(cComm2 \\<^sub>F \\<^sub>Q P M xvec N P' A\<^sub>P \\<^sub>P Q K Q' A\<^sub>Q \\<^sub>G A\<^sub>F A\<^sub>G) + have FimpG: "\A\<^sub>F, \\<^sub>F \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ \\<^sub>P \ \\<^sub>Q\" by fact + from \A\<^sub>F \* (P \ Q)\ \A\<^sub>G \* (P \ Q)\ have "A\<^sub>F \* P" and "A\<^sub>G \* P" and "A\<^sub>F \* Q" and "A\<^sub>G \* Q" + by simp+ + have FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" by fact + have FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" by fact + from \A\<^sub>F \* P\ \A\<^sub>G \* P\ FrP \A\<^sub>P \* A\<^sub>F\ \A\<^sub>P \* A\<^sub>G\ have "A\<^sub>F \* \\<^sub>P" and "A\<^sub>G \* \\<^sub>P" + by(force dest: extractFrameFreshChain)+ + from \A\<^sub>F \* Q\ \A\<^sub>G \* Q\ FrQ \A\<^sub>Q \* A\<^sub>F\ \A\<^sub>Q \* A\<^sub>G\ have "A\<^sub>F \* \\<^sub>Q" and "A\<^sub>G \* \\<^sub>Q" + by(force dest: extractFrameFreshChain)+ + from \\\<^sub>F \ \\<^sub>P \ Q \K\N\ \ Q'\ FrQ \distinct A\<^sub>Q\ + obtain K' where KeqK': "(\\<^sub>F \ \\<^sub>P) \ \\<^sub>Q \ K \ K'" and "A\<^sub>P \* K'" and "A\<^sub>F \* K'" and "A\<^sub>G \* K'" + using \A\<^sub>P \* Q\ \A\<^sub>Q \* A\<^sub>F\ \A\<^sub>Q \* A\<^sub>G\ \A\<^sub>F \* Q\ \A\<^sub>G \* Q\ \A\<^sub>Q \* \\<^sub>F\ \A\<^sub>Q \* \\<^sub>P\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* K\ + by(elim inputObtainPrefix[where B="A\<^sub>P@A\<^sub>F@A\<^sub>G"]) force+ + have "\\<^sub>G \ \\<^sub>Q \ P \K'\\*xvec\\N\ \ P'" + proof - + from KeqK' have "\\<^sub>F \ (\\<^sub>P \ \\<^sub>Q) \ K \ K'" + by(rule statEqEnt[OF Associativity]) + with \\\<^sub>F \ (\\<^sub>P \ \\<^sub>Q) \ M \ K\ have "\\<^sub>F \ (\\<^sub>P \ \\<^sub>Q) \ M \ K'" + by(rule chanEqTrans) + then have "(\\<^sub>F \ \\<^sub>Q) \ \\<^sub>P \ M \ K'" + by(metis statEqEnt AssertionStatEqSym Associativity AssertionStatEqTrans compositionSym Commutativity) + with \\\<^sub>F \ \\<^sub>Q \ P \M\\*xvec\\N\ \ P'\ FrP \distinct A\<^sub>P\ + have "\\<^sub>F \ \\<^sub>Q \ P \K'\\*xvec\\N\ \ P'" using \A\<^sub>P \* \\<^sub>F\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>P \* P\ \A\<^sub>P \* M\ \A\<^sub>P \* K'\ + by(force intro: outputRenameSubject) + moreover have "\A\<^sub>F, (\\<^sub>F \ \\<^sub>Q) \ \\<^sub>P\ \\<^sub>F \A\<^sub>G, (\\<^sub>G \ \\<^sub>Q) \ \\<^sub>P\" + proof - + have "\A\<^sub>F, (\\<^sub>F \ \\<^sub>Q) \ \\<^sub>P\ \\<^sub>F \A\<^sub>F, \\<^sub>F \ \\<^sub>P \ \\<^sub>Q\" + by(metis Associativity Composition AssertionStatEqSym AssertionStatEqTrans Commutativity frameResChainPres frameNilStatEq) + moreover have "\A\<^sub>G, \\<^sub>G \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>G, (\\<^sub>G \ \\<^sub>Q) \ \\<^sub>P\ " + by(metis Associativity Composition AssertionStatEqSym AssertionStatEqTrans Commutativity frameResChainPres frameNilStatEq) + ultimately show ?thesis using FimpG + by(elim FrameStatEqImpCompose) + qed + ultimately show ?thesis using \A\<^sub>F \* P\ \A\<^sub>G \* P\ \A\<^sub>F \* K'\ + \A\<^sub>G \* K'\ \A\<^sub>P \* A\<^sub>F\ \A\<^sub>P \* A\<^sub>G\ \A\<^sub>P \* \\<^sub>F\ \A\<^sub>P \* \\<^sub>G\ \A\<^sub>P \* \\<^sub>Q\ FrP \distinct A\<^sub>P\ + \distinct xvec\ + by(auto intro: transferNonTauFrame) + qed + + moreover from \\\<^sub>F \ \\<^sub>Q \ P \M\\*xvec\\N\ \ P'\ have "\\<^sub>F \ \\<^sub>Q \ P \ ROut M (\\*xvec\N \' P')" + by(simp add: residualInject) + moreover with FrP \distinct A\<^sub>P\ + obtain M' where MeqM': "(\\<^sub>F \ \\<^sub>Q) \ \\<^sub>P \ M \ M'" and "A\<^sub>Q \* M'" and "A\<^sub>F \* M'" and "A\<^sub>G \* M'" + using \A\<^sub>Q \* P\ \A\<^sub>P \* A\<^sub>F\ \A\<^sub>P \* A\<^sub>G\ \A\<^sub>F \* P\ \A\<^sub>G \* P\ \A\<^sub>P \* \\<^sub>F\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* P\ \A\<^sub>P \* M\ \xvec \* M\ \distinct xvec\ + by(elim outputObtainPrefix[where B="A\<^sub>Q@A\<^sub>F@A\<^sub>G"]) force+ + + have "\\<^sub>G \ \\<^sub>P \ Q \M'\N\ \ Q'" + proof - + from MeqM' have "\\<^sub>F \ (\\<^sub>Q \ \\<^sub>P) \ M \ M'" by(rule statEqEnt[OF Associativity]) + with \\\<^sub>F \ (\\<^sub>P \ \\<^sub>Q) \ M \ K\ have "\\<^sub>F \ (\\<^sub>Q \ \\<^sub>P) \ K \ M'" + by(blast intro: chanEqTrans chanEqSym compositionSym Commutativity statEqEnt) + then have "(\\<^sub>F \ \\<^sub>P) \ \\<^sub>Q \ K \ M'" + by(blast intro: statEqEnt AssertionStatEqSym Associativity + AssertionStatEqTrans compositionSym Commutativity) + with \\\<^sub>F \ \\<^sub>P \ Q \K\N\ \ Q'\ FrQ \distinct A\<^sub>Q\ + have "\\<^sub>F \ \\<^sub>P \ Q \M'\N\ \ Q'" using \A\<^sub>Q \* \\<^sub>F\ \A\<^sub>Q \* \\<^sub>P\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* K\ \A\<^sub>Q \* M'\ + by(force intro: inputRenameSubject) + moreover have "\A\<^sub>F, (\\<^sub>F \ \\<^sub>P) \ \\<^sub>Q\ \\<^sub>F \A\<^sub>G, (\\<^sub>G \ \\<^sub>P) \ \\<^sub>Q\" + proof - + have "\A\<^sub>F, (\\<^sub>F \ \\<^sub>P) \ \\<^sub>Q\ \\<^sub>F \A\<^sub>F, \\<^sub>F \ \\<^sub>P \ \\<^sub>Q\" + by(metis Associativity frameResChainPres frameNilStatEq) + moreover have "\A\<^sub>G, \\<^sub>G \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>G, (\\<^sub>G \ \\<^sub>P) \ \\<^sub>Q\ " + by(metis Associativity AssertionStatEqSym frameResChainPres frameNilStatEq) + ultimately show ?thesis using FimpG + by(elim FrameStatEqImpCompose) + qed + + ultimately show ?thesis using \A\<^sub>F \* Q\ \A\<^sub>G \* Q\ \A\<^sub>F \* M'\ \A\<^sub>G \* M'\ + \A\<^sub>Q \* A\<^sub>F\ \A\<^sub>Q \* A\<^sub>G\ \A\<^sub>Q \* \\<^sub>F\ \A\<^sub>Q \* \\<^sub>G\ \A\<^sub>Q \* \\<^sub>P\ FrQ \distinct A\<^sub>Q\ \distinct xvec\ + by(auto intro: transferNonTauFrame) + qed + + moreover have "\\<^sub>G \ \\<^sub>P \ \\<^sub>Q \ K' \ M'" + proof - + from MeqM' have "\\<^sub>F \ \\<^sub>P \ \\<^sub>Q \ M' \ M" + by(blast intro: chanEqSym Associativity statEqEnt Commutativity compositionSym) + moreover from KeqK' have "\\<^sub>F \ \\<^sub>P \ \\<^sub>Q \ K \ K'" + by(blast intro: chanEqSym Associativity statEqEnt Commutativity compositionSym) + ultimately have "\\<^sub>F \ \\<^sub>P \ \\<^sub>Q \ K' \ M'" using \\\<^sub>F \ \\<^sub>P \ \\<^sub>Q \ M \ K\ + by(blast intro: chanEqSym chanEqTrans) + then show ?thesis using \A\<^sub>F \* M'\ \A\<^sub>F \* K'\ \A\<^sub>G \* M'\ \A\<^sub>G \* K'\ FimpG + apply(simp add: FrameStatImp_def) + apply(erule allE[where x="SChanEq' K' M'"]) + by(force intro: frameImpI dest: frameImpE) + qed + + ultimately show ?case using \A\<^sub>P \* \\<^sub>G\ \A\<^sub>P \* P\ \A\<^sub>P \* Q\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* K'\ \A\<^sub>Q \* \\<^sub>G\ \A\<^sub>Q \* P\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* M'\ \xvec \* Q\ FrP FrQ + by(intro Comm2) (assumption | simp)+ +next + case(cBrClose \\<^sub>F P M xvec N P' A\<^sub>P \\<^sub>P x \\<^sub>G A\<^sub>F A\<^sub>G) + from \\\<^sub>F \ P \ \M\\*xvec\\N\ \ P'\ + have suppM: "((supp M)::name set) \ ((supp P)::name set)" + by(simp add: residualInject brOutputTermSupp) + + note \\\<^sub>F \ P \ \M\\*xvec\\N\ \ P'\ \extractFrame P = \A\<^sub>P, \\<^sub>P\\ + \distinct A\<^sub>P\ \distinct xvec\ \\A\<^sub>F, \\<^sub>F \ \\<^sub>P\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ \\<^sub>P\\ + + moreover from \x \ A\<^sub>F\ \A\<^sub>F \* \\x\P\ \x \ A\<^sub>G\ \A\<^sub>G \* \\x\P\ + have "A\<^sub>F \* P" and "A\<^sub>G \* P" by simp+ + moreover with suppM + have "A\<^sub>F \* M" and "A\<^sub>G \* M" + by(auto simp add: fresh_star_def fresh_def) + moreover note \A\<^sub>P \* A\<^sub>F\ \A\<^sub>P \* A\<^sub>G\ \A\<^sub>P \* \\<^sub>F\ \A\<^sub>P \* \\<^sub>G\ + ultimately have "\\<^sub>G \ P \ \M\\*xvec\\N\ \ P'" + by(simp add: transferNonTauFrame) + with \x \ supp M\ \x \ \\<^sub>G\ + show ?case + by(simp add: BrClose) +next + case(cScope \\<^sub>F P P' x A\<^sub>P \\<^sub>P \\<^sub>G A\<^sub>F A\<^sub>G) + then have "\\<^sub>G \ P \\ \ P'" by auto + with \x \ \\<^sub>G\ show ?case + by(intro Scope) auto +next + case(cBang \\<^sub>F P P' A\<^sub>P \\<^sub>P \\<^sub>G A\<^sub>F A\<^sub>G) + from \\\<^sub>P \ \\ have "\A\<^sub>F, \\<^sub>F \ \\<^sub>P \ \\ \\<^sub>F \A\<^sub>F, \\<^sub>F \ \\" + by(metis frameIntCompositionSym Identity AssertionStatEqTrans) + moreover note \\A\<^sub>F, \\<^sub>F \ \\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ \\\ + moreover from \\\<^sub>P \ \\ have "\A\<^sub>G, \\<^sub>G \ \\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ \\<^sub>P \ \\" + by(metis frameIntCompositionSym Identity AssertionStatEqTrans AssertionStatEqSym) + ultimately have "\A\<^sub>F, \\<^sub>F \ \\<^sub>P \ \\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ \\<^sub>P \ \\" + by(rule FrameStatEqImpCompose) + with cBang have "\\<^sub>G \ P \ !P \\ \ P'" by force + then show ?case using \guarded P\ by(rule Bang) +qed + +lemma transferFrame: + fixes \\<^sub>F :: 'b + and P :: "('a, 'b, 'c) psi" + and \ :: "'a action" + and P' :: "('a, 'b, 'c) psi" + and A\<^sub>F :: "name list" + and A\<^sub>G :: "name list" + and \\<^sub>G :: 'b + +assumes "\\<^sub>F \ P \\ \ P'" + and "extractFrame P = \A\<^sub>P, \\<^sub>P\" + and "distinct A\<^sub>P" + and "\A\<^sub>F, \\<^sub>F \ \\<^sub>P\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ \\<^sub>P\" + and "A\<^sub>F \* P" + and "A\<^sub>G \* P" + and "A\<^sub>F \* subject \" + and "A\<^sub>G \* subject \" + and "A\<^sub>P \* A\<^sub>F" + and "A\<^sub>P \* A\<^sub>G" + and "A\<^sub>P \* \\<^sub>F" + and "A\<^sub>P \* \\<^sub>G" + +shows "\\<^sub>G \ P \\ \ P'" + using assms +proof - + from \\\<^sub>F \ P \\ \ P'\ have "distinct(bn \)" by(auto dest: boundOutputDistinct) + then show ?thesis using assms + by(cases "\ = \") (auto intro: transferNonTauFrame transferTauFrame) +qed + +lemma parCasesInputFrame[consumes 7, case_names cPar1 cPar2]: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and Q :: "('a, 'b, 'c) psi" + and M :: 'a + and xvec :: "name list" + and N :: 'a + and T :: "('a, 'b, 'c) psi" + and C :: "'d::fs_name" + +assumes Trans: "\ \ P \ Q \M\N\ \ T" + and "extractFrame(P \ Q) = \A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q\" + and "distinct A\<^sub>P\<^sub>Q" + and "A\<^sub>P\<^sub>Q \* \" + and "A\<^sub>P\<^sub>Q \* P" + and "A\<^sub>P\<^sub>Q \* Q" + and "A\<^sub>P\<^sub>Q \* M" + and rPar1: "\P' A\<^sub>P \\<^sub>P A\<^sub>Q \\<^sub>Q. \\ \ \\<^sub>Q \ P \M\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; + distinct A\<^sub>P; distinct A\<^sub>Q; A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* Q; A\<^sub>P \* M; A\<^sub>Q \* \; A\<^sub>Q \* P; A\<^sub>Q \* Q; A\<^sub>Q \* M; + A\<^sub>P \* \\<^sub>Q; A\<^sub>Q \* \\<^sub>P; A\<^sub>P \* A\<^sub>Q; A\<^sub>P\<^sub>Q = A\<^sub>P@A\<^sub>Q; \\<^sub>P\<^sub>Q = \\<^sub>P \ \\<^sub>Q\ \ Prop (P' \ Q)" + and rPar2: "\Q' A\<^sub>P \\<^sub>P A\<^sub>Q \\<^sub>Q. \\ \ \\<^sub>P \ Q \M\N\ \ Q'; extractFrame P = \A\<^sub>P, \\<^sub>P\; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; + distinct A\<^sub>P; distinct A\<^sub>Q; A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* Q; A\<^sub>P \* M; A\<^sub>Q \* \; A\<^sub>Q \* P; A\<^sub>Q \* Q; A\<^sub>Q \* M; + A\<^sub>P \* \\<^sub>Q; A\<^sub>Q \* \\<^sub>P; A\<^sub>P \* A\<^sub>Q; A\<^sub>P\<^sub>Q = A\<^sub>P@A\<^sub>Q; \\<^sub>P\<^sub>Q = \\<^sub>P \ \\<^sub>Q\ \ Prop (P \ Q')" +shows "Prop T" + using Trans +proof(induct rule: parInputCases[of _ _ _ _ _ _ "(A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q)"]) + case(cPar1 P' A\<^sub>Q \\<^sub>Q) + from \A\<^sub>Q \* (A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q)\ have "A\<^sub>Q \* A\<^sub>P\<^sub>Q" and "A\<^sub>Q \* \\<^sub>P\<^sub>Q" by simp+ + obtain A\<^sub>P \\<^sub>P where FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" and "distinct A\<^sub>P" + "A\<^sub>P \* (P, Q, \, M, A\<^sub>Q, A\<^sub>P\<^sub>Q, \\<^sub>Q)" + by(rule freshFrame) + then have "A\<^sub>P \* P" and "A\<^sub>P \* Q" and "A\<^sub>P \* \" and "A\<^sub>P \* M" and "A\<^sub>P \* A\<^sub>Q" and "A\<^sub>P \* A\<^sub>P\<^sub>Q" and "A\<^sub>P \* \\<^sub>Q" + by simp+ + + have FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" by fact + + from \A\<^sub>Q \* P\ \A\<^sub>P \* A\<^sub>Q\ FrP have "A\<^sub>Q \* \\<^sub>P" + by(force dest: extractFrameFreshChain) + + from \extractFrame(P \ Q) = \A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q\\ FrP FrQ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>Q \* \\<^sub>P\ + have "\(A\<^sub>P@A\<^sub>Q), \\<^sub>P \ \\<^sub>Q\ = \A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q\" by simp + moreover from \distinct A\<^sub>P\ \distinct A\<^sub>Q\ \A\<^sub>P \* A\<^sub>Q\ have "distinct(A\<^sub>P@A\<^sub>Q)" + by(auto simp add: fresh_star_def fresh_def name_list_supp) + ultimately obtain p where S: "set p \ set(A\<^sub>P@A\<^sub>Q) \ set((p \ A\<^sub>P)@(p \ A\<^sub>Q))" and "distinctPerm p" + and \eq: "\\<^sub>P\<^sub>Q = (p \ \\<^sub>P) \ (p \ \\<^sub>Q)" and Aeq: "A\<^sub>P\<^sub>Q = (p \ A\<^sub>P)@(p \ A\<^sub>Q)" + using \A\<^sub>P \* A\<^sub>P\<^sub>Q\ \A\<^sub>Q \* A\<^sub>P\<^sub>Q\ \distinct A\<^sub>P\<^sub>Q\ + by(elim frameChainEq') (assumption | simp add: eqvts)+ + + from \\ \ \\<^sub>Q \ P \M\N\ \ P'\ S \A\<^sub>P\<^sub>Q \* P\ \A\<^sub>P \* P\ \A\<^sub>Q \* P\ \A\<^sub>P\<^sub>Q \* M\ \A\<^sub>P \* M\ \A\<^sub>Q \* M\ Aeq + have "(p \ (\ \ \\<^sub>Q)) \ P \M\N\ \ P'" + by(elim inputPermFrame) auto + with S \A\<^sub>P\<^sub>Q \* \\ \A\<^sub>P \* \\ \A\<^sub>Q \* \\ Aeq have "\ \ (p \ \\<^sub>Q) \ P \M\N\ \ P'" + by(simp add: eqvts) + moreover from FrP have "(p \ extractFrame P) = p \ \A\<^sub>P, \\<^sub>P\" by simp + with S \A\<^sub>P\<^sub>Q \* P\ \A\<^sub>P \* P\ \A\<^sub>Q \* P\ Aeq have "extractFrame P = \(p \ A\<^sub>P), p \ \\<^sub>P\" + by(simp add: eqvts) + moreover from FrQ have "(p \ extractFrame Q) = p \ \A\<^sub>Q, \\<^sub>Q\" by simp + with S \A\<^sub>P\<^sub>Q \* Q\ \A\<^sub>P \* Q\ \A\<^sub>Q \* Q\ Aeq have "extractFrame Q = \(p \ A\<^sub>Q), p \ \\<^sub>Q\" + by(simp add: eqvts) + moreover from \distinct A\<^sub>P\ \distinct A\<^sub>Q\ have "distinct(p \ A\<^sub>P)" and "distinct(p \ A\<^sub>Q)" + by simp+ + moreover from \A\<^sub>P \* A\<^sub>Q\ have "(p \ A\<^sub>P) \* (p \ A\<^sub>Q)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + moreover from \A\<^sub>P \* \\<^sub>Q\ have "(p \ A\<^sub>P) \* (p \ \\<^sub>Q)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + moreover from \A\<^sub>Q \* \\<^sub>P\ have "(p \ A\<^sub>Q) \* (p \ \\<^sub>P)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + ultimately show ?case using \A\<^sub>P\<^sub>Q \* \\ \A\<^sub>P\<^sub>Q \* P\ \A\<^sub>P\<^sub>Q \* Q\ \A\<^sub>P\<^sub>Q \* M\ Aeq \eq + by(intro rPar1) (assumption | simp)+ +next + case(cPar2 Q' A\<^sub>P \\<^sub>P) + from \A\<^sub>P \* (A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q)\ have "A\<^sub>P \* A\<^sub>P\<^sub>Q" and "A\<^sub>P \* \\<^sub>P\<^sub>Q" by simp+ + obtain A\<^sub>Q \\<^sub>Q where FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" and "distinct A\<^sub>Q" + "A\<^sub>Q \* (P, Q, \, M, A\<^sub>P, A\<^sub>P\<^sub>Q, \\<^sub>P)" + by(rule freshFrame) + then have "A\<^sub>Q \* P" and "A\<^sub>Q \* Q" and "A\<^sub>Q \* \" and "A\<^sub>Q \* M" and "A\<^sub>Q \* A\<^sub>P" and "A\<^sub>Q \* A\<^sub>P\<^sub>Q" and "A\<^sub>Q \* \\<^sub>P" + by simp+ + + have FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" by fact + + from \A\<^sub>P \* Q\ \A\<^sub>Q \* A\<^sub>P\ FrQ have "A\<^sub>P \* \\<^sub>Q" + by(force dest: extractFrameFreshChain) + + from \extractFrame(P \ Q) = \A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q\\ FrP FrQ \A\<^sub>Q \* A\<^sub>P\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>Q \* \\<^sub>P\ + have "\(A\<^sub>P@A\<^sub>Q), \\<^sub>P \ \\<^sub>Q\ = \A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q\" by simp + moreover from \distinct A\<^sub>P\ \distinct A\<^sub>Q\ \A\<^sub>Q \* A\<^sub>P\ have "distinct(A\<^sub>P@A\<^sub>Q)" + by(auto simp add: fresh_star_def fresh_def name_list_supp) + ultimately obtain p where S: "set p \ set(A\<^sub>P@A\<^sub>Q) \ set((p \ A\<^sub>P)@(p \ A\<^sub>Q))" and "distinctPerm p" + and \eq: "\\<^sub>P\<^sub>Q = (p \ \\<^sub>P) \ (p \ \\<^sub>Q)" and Aeq: "A\<^sub>P\<^sub>Q = (p \ A\<^sub>P)@(p \ A\<^sub>Q)" + using \A\<^sub>P \* A\<^sub>P\<^sub>Q\ \A\<^sub>Q \* A\<^sub>P\<^sub>Q\ \distinct A\<^sub>P\<^sub>Q\ + by(elim frameChainEq') (assumption | simp add: eqvts)+ + + from \\ \ \\<^sub>P \ Q \M\N\ \ Q'\ S \A\<^sub>P\<^sub>Q \* Q\ \A\<^sub>P \* Q\ \A\<^sub>Q \* Q\ \A\<^sub>P\<^sub>Q \* M\ \A\<^sub>P \* M\ \A\<^sub>Q \* M\ Aeq + have "(p \ (\ \ \\<^sub>P)) \ Q \M\N\ \ Q'" + by(elim inputPermFrame) auto + with S \A\<^sub>P\<^sub>Q \* \\ \A\<^sub>P \* \\ \A\<^sub>Q \* \\ Aeq have "\ \ (p \ \\<^sub>P) \ Q \M\N\ \ Q'" + by(simp add: eqvts) + moreover from FrP have "(p \ extractFrame P) = p \ \A\<^sub>P, \\<^sub>P\" by simp + with S \A\<^sub>P\<^sub>Q \* P\ \A\<^sub>P \* P\ \A\<^sub>Q \* P\ Aeq have "extractFrame P = \(p \ A\<^sub>P), p \ \\<^sub>P\" + by(simp add: eqvts) + moreover from FrQ have "(p \ extractFrame Q) = p \ \A\<^sub>Q, \\<^sub>Q\" by simp + with S \A\<^sub>P\<^sub>Q \* Q\ \A\<^sub>P \* Q\ \A\<^sub>Q \* Q\ Aeq have "extractFrame Q = \(p \ A\<^sub>Q), p \ \\<^sub>Q\" + by(simp add: eqvts) + moreover from \distinct A\<^sub>P\ \distinct A\<^sub>Q\ have "distinct(p \ A\<^sub>P)" and "distinct(p \ A\<^sub>Q)" + by simp+ + moreover from \A\<^sub>Q \* A\<^sub>P\ have "(p \ A\<^sub>P) \* (p \ A\<^sub>Q)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + moreover from \A\<^sub>P \* \\<^sub>Q\ have "(p \ A\<^sub>P) \* (p \ \\<^sub>Q)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + moreover from \A\<^sub>Q \* \\<^sub>P\ have "(p \ A\<^sub>Q) \* (p \ \\<^sub>P)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + ultimately show ?case using \A\<^sub>P\<^sub>Q \* \\ \A\<^sub>P\<^sub>Q \* P\ \A\<^sub>P\<^sub>Q \* Q\ \A\<^sub>P\<^sub>Q \* M\ Aeq \eq + by(intro rPar2) (assumption | simp)+ +qed + +lemma parCasesBrInputFrame[consumes 7, case_names cPar1 cPar2 cBrMerge]: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and Q :: "('a, 'b, 'c) psi" + and M :: 'a + and xvec :: "name list" + and N :: 'a + and T :: "('a, 'b, 'c) psi" + and C :: "'d::fs_name" + +assumes Trans: "\ \ P \ Q \\M\N\ \ T" + and "extractFrame(P \ Q) = \A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q\" + and "distinct A\<^sub>P\<^sub>Q" + and "A\<^sub>P\<^sub>Q \* \" + and "A\<^sub>P\<^sub>Q \* P" + and "A\<^sub>P\<^sub>Q \* Q" + and "A\<^sub>P\<^sub>Q \* M" + and rPar1: "\P' A\<^sub>P \\<^sub>P A\<^sub>Q \\<^sub>Q. \\ \ \\<^sub>Q \ P \\M\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; + distinct A\<^sub>P; distinct A\<^sub>Q; A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* Q; A\<^sub>P \* M; A\<^sub>Q \* \; A\<^sub>Q \* P; A\<^sub>Q \* Q; A\<^sub>Q \* M; + A\<^sub>P \* \\<^sub>Q; A\<^sub>Q \* \\<^sub>P; A\<^sub>P \* A\<^sub>Q; A\<^sub>P\<^sub>Q = A\<^sub>P@A\<^sub>Q; \\<^sub>P\<^sub>Q = \\<^sub>P \ \\<^sub>Q\ \ Prop (P' \ Q)" + and rPar2: "\Q' A\<^sub>P \\<^sub>P A\<^sub>Q \\<^sub>Q. \\ \ \\<^sub>P \ Q \\M\N\ \ Q'; extractFrame P = \A\<^sub>P, \\<^sub>P\; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; + distinct A\<^sub>P; distinct A\<^sub>Q; A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* Q; A\<^sub>P \* M; A\<^sub>Q \* \; A\<^sub>Q \* P; A\<^sub>Q \* Q; A\<^sub>Q \* M; + A\<^sub>P \* \\<^sub>Q; A\<^sub>Q \* \\<^sub>P; A\<^sub>P \* A\<^sub>Q; A\<^sub>P\<^sub>Q = A\<^sub>P@A\<^sub>Q; \\<^sub>P\<^sub>Q = \\<^sub>P \ \\<^sub>Q\ \ Prop (P \ Q')" + and rBrMerge: "\\\<^sub>Q P' A\<^sub>P \\<^sub>P Q' A\<^sub>Q. + \\ \ \\<^sub>Q \ P \ \M\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \ \ \\<^sub>P \ Q \ \M\N\ \ Q'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; + A\<^sub>P \* Q; A\<^sub>P \* A\<^sub>Q; + A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; A\<^sub>Q \* P; + A\<^sub>Q \* Q; + A\<^sub>P\<^sub>Q = A\<^sub>P@A\<^sub>Q; \\<^sub>P\<^sub>Q = \\<^sub>P \ \\<^sub>Q\ \ + Prop (P' \ Q')" +shows "Prop T" + using Trans +proof(induct rule: parBrInputCases[of _ _ _ _ _ _ "(A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q)"]) + case(cPar1 P' A\<^sub>Q \\<^sub>Q) + from \A\<^sub>Q \* (A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q)\ have "A\<^sub>Q \* A\<^sub>P\<^sub>Q" and "A\<^sub>Q \* \\<^sub>P\<^sub>Q" by simp+ + obtain A\<^sub>P \\<^sub>P where FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" and "distinct A\<^sub>P" + "A\<^sub>P \* (P, Q, \, M, A\<^sub>Q, A\<^sub>P\<^sub>Q, \\<^sub>Q)" + by(rule freshFrame) + then have "A\<^sub>P \* P" and "A\<^sub>P \* Q" and "A\<^sub>P \* \" and "A\<^sub>P \* M" and "A\<^sub>P \* A\<^sub>Q" and "A\<^sub>P \* A\<^sub>P\<^sub>Q" and "A\<^sub>P \* \\<^sub>Q" + by simp+ + + have FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" by fact + + from \A\<^sub>Q \* P\ \A\<^sub>P \* A\<^sub>Q\ FrP have "A\<^sub>Q \* \\<^sub>P" + by(force dest: extractFrameFreshChain) + + from \extractFrame(P \ Q) = \A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q\\ FrP FrQ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>Q \* \\<^sub>P\ + have "\(A\<^sub>P@A\<^sub>Q), \\<^sub>P \ \\<^sub>Q\ = \A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q\" by simp + moreover from \distinct A\<^sub>P\ \distinct A\<^sub>Q\ \A\<^sub>P \* A\<^sub>Q\ have "distinct(A\<^sub>P@A\<^sub>Q)" + by(auto simp add: fresh_star_def fresh_def name_list_supp) + ultimately obtain p where S: "set p \ set(A\<^sub>P@A\<^sub>Q) \ set((p \ A\<^sub>P)@(p \ A\<^sub>Q))" and "distinctPerm p" + and \eq: "\\<^sub>P\<^sub>Q = (p \ \\<^sub>P) \ (p \ \\<^sub>Q)" and Aeq: "A\<^sub>P\<^sub>Q = (p \ A\<^sub>P)@(p \ A\<^sub>Q)" + using \A\<^sub>P \* A\<^sub>P\<^sub>Q\ \A\<^sub>Q \* A\<^sub>P\<^sub>Q\ \distinct A\<^sub>P\<^sub>Q\ + by(elim frameChainEq') (assumption | simp add: eqvts)+ + + from \\ \ \\<^sub>Q \ P \\M\N\ \ P'\ S \A\<^sub>P\<^sub>Q \* P\ \A\<^sub>P \* P\ \A\<^sub>Q \* P\ \A\<^sub>P\<^sub>Q \* M\ \A\<^sub>P \* M\ \A\<^sub>Q \* M\ Aeq + have "(p \ (\ \ \\<^sub>Q)) \ P \\M\N\ \ P'" + by(elim brinputPermFrame) auto + with S \A\<^sub>P\<^sub>Q \* \\ \A\<^sub>P \* \\ \A\<^sub>Q \* \\ Aeq have "\ \ (p \ \\<^sub>Q) \ P \\M\N\ \ P'" + by(simp add: eqvts) + moreover from FrP have "(p \ extractFrame P) = p \ \A\<^sub>P, \\<^sub>P\" by simp + with S \A\<^sub>P\<^sub>Q \* P\ \A\<^sub>P \* P\ \A\<^sub>Q \* P\ Aeq have "extractFrame P = \(p \ A\<^sub>P), p \ \\<^sub>P\" + by(simp add: eqvts) + moreover from FrQ have "(p \ extractFrame Q) = p \ \A\<^sub>Q, \\<^sub>Q\" by simp + with S \A\<^sub>P\<^sub>Q \* Q\ \A\<^sub>P \* Q\ \A\<^sub>Q \* Q\ Aeq have "extractFrame Q = \(p \ A\<^sub>Q), p \ \\<^sub>Q\" + by(simp add: eqvts) + moreover from \distinct A\<^sub>P\ \distinct A\<^sub>Q\ have "distinct(p \ A\<^sub>P)" and "distinct(p \ A\<^sub>Q)" + by simp+ + moreover from \A\<^sub>P \* A\<^sub>Q\ have "(p \ A\<^sub>P) \* (p \ A\<^sub>Q)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + moreover from \A\<^sub>P \* \\<^sub>Q\ have "(p \ A\<^sub>P) \* (p \ \\<^sub>Q)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + moreover from \A\<^sub>Q \* \\<^sub>P\ have "(p \ A\<^sub>Q) \* (p \ \\<^sub>P)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + ultimately show ?case using \A\<^sub>P\<^sub>Q \* \\ \A\<^sub>P\<^sub>Q \* P\ \A\<^sub>P\<^sub>Q \* Q\ \A\<^sub>P\<^sub>Q \* M\ Aeq \eq + by(intro rPar1) (assumption | simp)+ +next + case(cPar2 Q' A\<^sub>P \\<^sub>P) + from \A\<^sub>P \* (A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q)\ have "A\<^sub>P \* A\<^sub>P\<^sub>Q" and "A\<^sub>P \* \\<^sub>P\<^sub>Q" by simp+ + obtain A\<^sub>Q \\<^sub>Q where FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" and "distinct A\<^sub>Q" + "A\<^sub>Q \* (P, Q, \, M, A\<^sub>P, A\<^sub>P\<^sub>Q, \\<^sub>P)" + by(rule freshFrame) + then have "A\<^sub>Q \* P" and "A\<^sub>Q \* Q" and "A\<^sub>Q \* \" and "A\<^sub>Q \* M" and "A\<^sub>Q \* A\<^sub>P" and "A\<^sub>Q \* A\<^sub>P\<^sub>Q" and "A\<^sub>Q \* \\<^sub>P" + by simp+ + + have FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" by fact + + from \A\<^sub>P \* Q\ \A\<^sub>Q \* A\<^sub>P\ FrQ have "A\<^sub>P \* \\<^sub>Q" + by(force dest: extractFrameFreshChain) + + from \extractFrame(P \ Q) = \A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q\\ FrP FrQ \A\<^sub>Q \* A\<^sub>P\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>Q \* \\<^sub>P\ + have "\(A\<^sub>P@A\<^sub>Q), \\<^sub>P \ \\<^sub>Q\ = \A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q\" by simp + moreover from \distinct A\<^sub>P\ \distinct A\<^sub>Q\ \A\<^sub>Q \* A\<^sub>P\ have "distinct(A\<^sub>P@A\<^sub>Q)" + by(auto simp add: fresh_star_def fresh_def name_list_supp) + ultimately obtain p where S: "set p \ set(A\<^sub>P@A\<^sub>Q) \ set((p \ A\<^sub>P)@(p \ A\<^sub>Q))" and "distinctPerm p" + and \eq: "\\<^sub>P\<^sub>Q = (p \ \\<^sub>P) \ (p \ \\<^sub>Q)" and Aeq: "A\<^sub>P\<^sub>Q = (p \ A\<^sub>P)@(p \ A\<^sub>Q)" + using \A\<^sub>P \* A\<^sub>P\<^sub>Q\ \A\<^sub>Q \* A\<^sub>P\<^sub>Q\ \distinct A\<^sub>P\<^sub>Q\ + by(elim frameChainEq') (assumption | simp add: eqvts)+ + + from \\ \ \\<^sub>P \ Q \\M\N\ \ Q'\ S \A\<^sub>P\<^sub>Q \* Q\ \A\<^sub>P \* Q\ \A\<^sub>Q \* Q\ \A\<^sub>P\<^sub>Q \* M\ \A\<^sub>P \* M\ \A\<^sub>Q \* M\ Aeq + have "(p \ (\ \ \\<^sub>P)) \ Q \\M\N\ \ Q'" + by(elim brinputPermFrame) auto + with S \A\<^sub>P\<^sub>Q \* \\ \A\<^sub>P \* \\ \A\<^sub>Q \* \\ Aeq have "\ \ (p \ \\<^sub>P) \ Q \\M\N\ \ Q'" + by(simp add: eqvts) + moreover from FrP have "(p \ extractFrame P) = p \ \A\<^sub>P, \\<^sub>P\" by simp + with S \A\<^sub>P\<^sub>Q \* P\ \A\<^sub>P \* P\ \A\<^sub>Q \* P\ Aeq have "extractFrame P = \(p \ A\<^sub>P), p \ \\<^sub>P\" + by(simp add: eqvts) + moreover from FrQ have "(p \ extractFrame Q) = p \ \A\<^sub>Q, \\<^sub>Q\" by simp + with S \A\<^sub>P\<^sub>Q \* Q\ \A\<^sub>P \* Q\ \A\<^sub>Q \* Q\ Aeq have "extractFrame Q = \(p \ A\<^sub>Q), p \ \\<^sub>Q\" + by(simp add: eqvts) + moreover from \distinct A\<^sub>P\ \distinct A\<^sub>Q\ have "distinct(p \ A\<^sub>P)" and "distinct(p \ A\<^sub>Q)" + by simp+ + moreover from \A\<^sub>Q \* A\<^sub>P\ have "(p \ A\<^sub>P) \* (p \ A\<^sub>Q)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + moreover from \A\<^sub>P \* \\<^sub>Q\ have "(p \ A\<^sub>P) \* (p \ \\<^sub>Q)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + moreover from \A\<^sub>Q \* \\<^sub>P\ have "(p \ A\<^sub>Q) \* (p \ \\<^sub>P)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + ultimately show ?case using \A\<^sub>P\<^sub>Q \* \\ \A\<^sub>P\<^sub>Q \* P\ \A\<^sub>P\<^sub>Q \* Q\ \A\<^sub>P\<^sub>Q \* M\ Aeq \eq + by(intro rPar2) (assumption | simp)+ +next + case(cBrMerge \\<^sub>Q P' A\<^sub>P \\<^sub>P Q' A\<^sub>Q) + then have FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" and FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" + and "A\<^sub>P \* A\<^sub>P\<^sub>Q" and "A\<^sub>Q \* A\<^sub>P\<^sub>Q" + by simp+ + + from \extractFrame(P \ Q) = \A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q\\ FrP FrQ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>Q \* \\<^sub>P\ + have "\(A\<^sub>P@A\<^sub>Q), \\<^sub>P \ \\<^sub>Q\ = \A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q\" by simp + moreover from \distinct A\<^sub>P\ \distinct A\<^sub>Q\ \A\<^sub>P \* A\<^sub>Q\ have "distinct(A\<^sub>P@A\<^sub>Q)" + by(auto simp add: fresh_star_def fresh_def name_list_supp) + ultimately obtain p where S: "set p \ set(A\<^sub>P@A\<^sub>Q) \ set((p \ A\<^sub>P)@(p \ A\<^sub>Q))" and "distinctPerm p" + and \eq: "\\<^sub>P\<^sub>Q = (p \ \\<^sub>P) \ (p \ \\<^sub>Q)" and Aeq: "A\<^sub>P\<^sub>Q = (p \ A\<^sub>P)@(p \ A\<^sub>Q)" + using \A\<^sub>P \* A\<^sub>P\<^sub>Q\ \A\<^sub>Q \* A\<^sub>P\<^sub>Q\ \distinct A\<^sub>P\<^sub>Q\ + by(elim frameChainEq') (assumption | simp add: eqvts)+ + + from \\ \ \\<^sub>Q \ P \\M\N\ \ P'\ S \A\<^sub>P\<^sub>Q \* P\ \A\<^sub>P \* P\ \A\<^sub>Q \* P\ \A\<^sub>P\<^sub>Q \* M\ \A\<^sub>P \* M\ \A\<^sub>Q \* M\ Aeq + have "(p \ (\ \ \\<^sub>Q)) \ P \\M\N\ \ P'" + by(elim brinputPermFrame) auto + with S \A\<^sub>P\<^sub>Q \* \\ \A\<^sub>P \* \\ \A\<^sub>Q \* \\ Aeq have "\ \ (p \ \\<^sub>Q) \ P \\M\N\ \ P'" + by(simp add: eqvts) + + from \\ \ \\<^sub>P \ Q \\M\N\ \ Q'\ S \A\<^sub>P\<^sub>Q \* Q\ \A\<^sub>Q \* Q\ \A\<^sub>P \* Q\ \A\<^sub>P\<^sub>Q \* M\ \A\<^sub>Q \* M\ \A\<^sub>P \* M\ Aeq + have "(p \ (\ \ \\<^sub>P)) \ Q \\M\N\ \ Q'" + by(elim brinputPermFrame) auto + with S \A\<^sub>P\<^sub>Q \* \\ \A\<^sub>P \* \\ \A\<^sub>Q \* \\ Aeq have "\ \ (p \ \\<^sub>P) \ Q \\M\N\ \ Q'" + by(simp add: eqvts) + + note \\ \ (p \ \\<^sub>Q) \ P \\M\N\ \ P'\ \\ \ (p \ \\<^sub>P) \ Q \\M\N\ \ Q'\ + moreover from FrP have "(p \ extractFrame P) = p \ \A\<^sub>P, \\<^sub>P\" by simp + with S \A\<^sub>P\<^sub>Q \* P\ \A\<^sub>P \* P\ \A\<^sub>Q \* P\ Aeq have "extractFrame P = \(p \ A\<^sub>P), p \ \\<^sub>P\" + by(simp add: eqvts) + moreover from FrQ have "(p \ extractFrame Q) = p \ \A\<^sub>Q, \\<^sub>Q\" by simp + with S \A\<^sub>P\<^sub>Q \* Q\ \A\<^sub>P \* Q\ \A\<^sub>Q \* Q\ Aeq have "extractFrame Q = \(p \ A\<^sub>Q), p \ \\<^sub>Q\" + by(simp add: eqvts) + moreover from \distinct A\<^sub>P\ \distinct A\<^sub>Q\ have "distinct(p \ A\<^sub>P)" and "distinct(p \ A\<^sub>Q)" + by simp+ + moreover from \A\<^sub>P \* A\<^sub>Q\ have "(p \ A\<^sub>P) \* (p \ A\<^sub>Q)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + moreover from \A\<^sub>P \* \\<^sub>Q\ have "(p \ A\<^sub>P) \* (p \ \\<^sub>Q)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + moreover from \A\<^sub>Q \* \\<^sub>P\ have "(p \ A\<^sub>Q) \* (p \ \\<^sub>P)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + ultimately show ?case using \A\<^sub>P\<^sub>Q \* \\ \A\<^sub>P\<^sub>Q \* P\ \A\<^sub>P\<^sub>Q \* Q\ \A\<^sub>P\<^sub>Q \* M\ Aeq \eq + by(intro rBrMerge) simp+ +qed + +lemma parCasesOutputFrame[consumes 11, case_names cPar1 cPar2]: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and Q :: "('a, 'b, 'c) psi" + and M :: 'a + and xvec :: "name list" + and N :: 'a + and T :: "('a, 'b, 'c) psi" + and C :: "'d::fs_name" + +assumes Trans: "\ \ P \ Q \M\\*xvec\\N\ \ T" + and "xvec \* \" + and "xvec \* P" + and "xvec \* Q" + and "xvec \* M" + and "extractFrame(P \ Q) = \A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q\" + and "distinct A\<^sub>P\<^sub>Q" + and "A\<^sub>P\<^sub>Q \* \" + and "A\<^sub>P\<^sub>Q \* P" + and "A\<^sub>P\<^sub>Q \* Q" + and "A\<^sub>P\<^sub>Q \* M" + and rPar1: "\P' A\<^sub>P \\<^sub>P A\<^sub>Q \\<^sub>Q. \\ \ \\<^sub>Q \ P \M\\*xvec\\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; + distinct A\<^sub>P; distinct A\<^sub>Q; A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* Q; A\<^sub>P \* M; A\<^sub>Q \* \; A\<^sub>Q \* P; A\<^sub>Q \* Q; A\<^sub>Q \* M; + A\<^sub>P \* \\<^sub>Q; A\<^sub>Q \* \\<^sub>P; A\<^sub>P \* A\<^sub>Q; A\<^sub>P\<^sub>Q = A\<^sub>P@A\<^sub>Q; \\<^sub>P\<^sub>Q = \\<^sub>P \ \\<^sub>Q\ \ Prop (P' \ Q)" + and rPar2: "\Q' A\<^sub>P \\<^sub>P A\<^sub>Q \\<^sub>Q. \\ \ \\<^sub>P \ Q \M\\*xvec\\N\ \ Q'; extractFrame P = \A\<^sub>P, \\<^sub>P\; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; + distinct A\<^sub>P; distinct A\<^sub>Q; A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* Q; A\<^sub>P \* M; A\<^sub>Q \* \; A\<^sub>Q \* P; A\<^sub>Q \* Q; A\<^sub>Q \* M; + A\<^sub>P \* \\<^sub>Q; A\<^sub>Q \* \\<^sub>P; A\<^sub>P \* A\<^sub>Q; A\<^sub>P\<^sub>Q = A\<^sub>P@A\<^sub>Q; \\<^sub>P\<^sub>Q = \\<^sub>P \ \\<^sub>Q\ \ Prop (P \ Q')" +shows "Prop T" + using Trans \xvec \* \\ \xvec \* P\ \xvec \* Q\ \xvec \* M\ +proof(induct rule: parOutputCases[of _ _ _ _ _ _ _ "(A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q)"]) + case(cPar1 P' A\<^sub>Q \\<^sub>Q) + from \A\<^sub>Q \* (A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q)\ have "A\<^sub>Q \* A\<^sub>P\<^sub>Q" and "A\<^sub>Q \* \\<^sub>P\<^sub>Q" by simp+ + obtain A\<^sub>P \\<^sub>P where FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" and "distinct A\<^sub>P" + "A\<^sub>P \* (P, Q, \, M, A\<^sub>Q, A\<^sub>P\<^sub>Q, \\<^sub>Q)" + by(rule freshFrame) + then have "A\<^sub>P \* P" and "A\<^sub>P \* Q" and "A\<^sub>P \* \" and "A\<^sub>P \* M" and "A\<^sub>P \* A\<^sub>Q" and "A\<^sub>P \* A\<^sub>P\<^sub>Q" and "A\<^sub>P \* \\<^sub>Q" + by simp+ + + have FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" by fact + + from \A\<^sub>Q \* P\ \A\<^sub>P \* A\<^sub>Q\ FrP have "A\<^sub>Q \* \\<^sub>P" + by(force dest: extractFrameFreshChain) + + from \extractFrame(P \ Q) = \A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q\\ FrP FrQ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>Q \* \\<^sub>P\ + have "\(A\<^sub>P@A\<^sub>Q), \\<^sub>P \ \\<^sub>Q\ = \A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q\" by simp + moreover from \distinct A\<^sub>P\ \distinct A\<^sub>Q\ \A\<^sub>P \* A\<^sub>Q\ have "distinct(A\<^sub>P@A\<^sub>Q)" + by(auto simp add: fresh_star_def fresh_def name_list_supp) + ultimately obtain p where S: "set p \ set(A\<^sub>P@A\<^sub>Q) \ set((p \ A\<^sub>P)@(p \ A\<^sub>Q))" and "distinctPerm p" + and \eq: "\\<^sub>P\<^sub>Q = (p \ \\<^sub>P) \ (p \ \\<^sub>Q)" and Aeq: "A\<^sub>P\<^sub>Q = (p \ A\<^sub>P)@(p \ A\<^sub>Q)" + using \A\<^sub>P \* A\<^sub>P\<^sub>Q\ \A\<^sub>Q \* A\<^sub>P\<^sub>Q\ \distinct A\<^sub>P\<^sub>Q\ + by(elim frameChainEq') (assumption | simp add: eqvts)+ + + from \\ \ \\<^sub>Q \ P \M\\*xvec\\N\ \ P'\ S \A\<^sub>P\<^sub>Q \* P\ \A\<^sub>P \* P\ \A\<^sub>Q \* P\ \A\<^sub>P\<^sub>Q \* M\ \A\<^sub>P \* M\ \A\<^sub>Q \* M\ Aeq + have "(p \ (\ \ \\<^sub>Q)) \ P \M\\*xvec\\N\ \ P'" + by(elim outputPermFrame) (assumption | simp)+ + + with S \A\<^sub>P\<^sub>Q \* \\ \A\<^sub>P \* \\ \A\<^sub>Q \* \\ Aeq have "\ \ (p \ \\<^sub>Q) \ P \M\\*xvec\\N\ \ P'" + by(simp add: eqvts) + moreover from FrP have "(p \ extractFrame P) = p \ \A\<^sub>P, \\<^sub>P\" by simp + with S \A\<^sub>P\<^sub>Q \* P\ \A\<^sub>P \* P\ \A\<^sub>Q \* P\ Aeq have "extractFrame P = \(p \ A\<^sub>P), p \ \\<^sub>P\" + by(simp add: eqvts) + moreover from FrQ have "(p \ extractFrame Q) = p \ \A\<^sub>Q, \\<^sub>Q\" by simp + with S \A\<^sub>P\<^sub>Q \* Q\ \A\<^sub>P \* Q\ \A\<^sub>Q \* Q\ Aeq have "extractFrame Q = \(p \ A\<^sub>Q), p \ \\<^sub>Q\" + by(simp add: eqvts) + moreover from \distinct A\<^sub>P\ \distinct A\<^sub>Q\ have "distinct(p \ A\<^sub>P)" and "distinct(p \ A\<^sub>Q)" + by simp+ + moreover from \A\<^sub>P \* A\<^sub>Q\ have "(p \ A\<^sub>P) \* (p \ A\<^sub>Q)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + moreover from \A\<^sub>P \* \\<^sub>Q\ have "(p \ A\<^sub>P) \* (p \ \\<^sub>Q)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + moreover from \A\<^sub>Q \* \\<^sub>P\ have "(p \ A\<^sub>Q) \* (p \ \\<^sub>P)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + ultimately show ?case using \A\<^sub>P\<^sub>Q \* \\ \A\<^sub>P\<^sub>Q \* P\ \A\<^sub>P\<^sub>Q \* Q\ \A\<^sub>P\<^sub>Q \* M\ Aeq \eq + by(intro rPar1) (assumption | simp)+ +next + case(cPar2 Q' A\<^sub>P \\<^sub>P) + from \A\<^sub>P \* (A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q)\ have "A\<^sub>P \* A\<^sub>P\<^sub>Q" and "A\<^sub>P \* \\<^sub>P\<^sub>Q" by simp+ + obtain A\<^sub>Q \\<^sub>Q where FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" and "distinct A\<^sub>Q" + "A\<^sub>Q \* (P, Q, \, M, A\<^sub>P, A\<^sub>P\<^sub>Q, \\<^sub>P)" + by(rule freshFrame) + then have "A\<^sub>Q \* P" and "A\<^sub>Q \* Q" and "A\<^sub>Q \* \" and "A\<^sub>Q \* M" and "A\<^sub>Q \* A\<^sub>P" and "A\<^sub>Q \* A\<^sub>P\<^sub>Q" and "A\<^sub>Q \* \\<^sub>P" + by simp+ + + have FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" by fact + + from \A\<^sub>P \* Q\ \A\<^sub>Q \* A\<^sub>P\ FrQ have "A\<^sub>P \* \\<^sub>Q" + by(force dest: extractFrameFreshChain) + + from \extractFrame(P \ Q) = \A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q\\ FrP FrQ \A\<^sub>Q \* A\<^sub>P\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>Q \* \\<^sub>P\ + have "\(A\<^sub>P@A\<^sub>Q), \\<^sub>P \ \\<^sub>Q\ = \A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q\" by simp + moreover from \distinct A\<^sub>P\ \distinct A\<^sub>Q\ \A\<^sub>Q \* A\<^sub>P\ have "distinct(A\<^sub>P@A\<^sub>Q)" + by(auto simp add: fresh_star_def fresh_def name_list_supp) + ultimately obtain p where S: "set p \ set(A\<^sub>P@A\<^sub>Q) \ set((p \ A\<^sub>P)@(p \ A\<^sub>Q))" and "distinctPerm p" + and \eq: "\\<^sub>P\<^sub>Q = (p \ \\<^sub>P) \ (p \ \\<^sub>Q)" and Aeq: "A\<^sub>P\<^sub>Q = (p \ A\<^sub>P)@(p \ A\<^sub>Q)" + using \A\<^sub>P \* A\<^sub>P\<^sub>Q\ \A\<^sub>Q \* A\<^sub>P\<^sub>Q\ \distinct A\<^sub>P\<^sub>Q\ + by(elim frameChainEq') (assumption | simp add: eqvts)+ + + from \\ \ \\<^sub>P \ Q \M\\*xvec\\N\ \ Q'\ S \A\<^sub>P\<^sub>Q \* Q\ \A\<^sub>P \* Q\ \A\<^sub>Q \* Q\ \A\<^sub>P\<^sub>Q \* M\ \A\<^sub>P \* M\ \A\<^sub>Q \* M\ Aeq + have "(p \ (\ \ \\<^sub>P)) \ Q \M\\*xvec\\N\ \ Q'" + by(elim outputPermFrame) (assumption | simp)+ + with S \A\<^sub>P\<^sub>Q \* \\ \A\<^sub>P \* \\ \A\<^sub>Q \* \\ Aeq have "\ \ (p \ \\<^sub>P) \ Q \M\\*xvec\\N\ \ Q'" + by(simp add: eqvts) + moreover from FrP have "(p \ extractFrame P) = p \ \A\<^sub>P, \\<^sub>P\" by simp + with S \A\<^sub>P\<^sub>Q \* P\ \A\<^sub>P \* P\ \A\<^sub>Q \* P\ Aeq have "extractFrame P = \(p \ A\<^sub>P), p \ \\<^sub>P\" + by(simp add: eqvts) + moreover from FrQ have "(p \ extractFrame Q) = p \ \A\<^sub>Q, \\<^sub>Q\" by simp + with S \A\<^sub>P\<^sub>Q \* Q\ \A\<^sub>P \* Q\ \A\<^sub>Q \* Q\ Aeq have "extractFrame Q = \(p \ A\<^sub>Q), p \ \\<^sub>Q\" + by(simp add: eqvts) + moreover from \distinct A\<^sub>P\ \distinct A\<^sub>Q\ have "distinct(p \ A\<^sub>P)" and "distinct(p \ A\<^sub>Q)" + by simp+ + moreover from \A\<^sub>Q \* A\<^sub>P\ have "(p \ A\<^sub>P) \* (p \ A\<^sub>Q)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + moreover from \A\<^sub>P \* \\<^sub>Q\ have "(p \ A\<^sub>P) \* (p \ \\<^sub>Q)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + moreover from \A\<^sub>Q \* \\<^sub>P\ have "(p \ A\<^sub>Q) \* (p \ \\<^sub>P)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + ultimately show ?case using \A\<^sub>P\<^sub>Q \* \\ \A\<^sub>P\<^sub>Q \* P\ \A\<^sub>P\<^sub>Q \* Q\ \A\<^sub>P\<^sub>Q \* M\ Aeq \eq + by(intro rPar2) (assumption | simp)+ +qed + +lemma parCasesBrOutputFrame[consumes 11, case_names cPar1 cPar2 cBrComm1 cBrComm2]: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and Q :: "('a, 'b, 'c) psi" + and M :: 'a + and xvec :: "name list" + and N :: 'a + and T :: "('a, 'b, 'c) psi" + and C :: "'d::fs_name" + +assumes Trans: "\ \ P \ Q \\M\\*xvec\\N\ \ T" + and "xvec \* \" + and "xvec \* P" + and "xvec \* Q" + and "xvec \* M" + and "extractFrame(P \ Q) = \A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q\" + and "distinct A\<^sub>P\<^sub>Q" + and "A\<^sub>P\<^sub>Q \* \" + and "A\<^sub>P\<^sub>Q \* P" + and "A\<^sub>P\<^sub>Q \* Q" + and "A\<^sub>P\<^sub>Q \* M" + and rPar1: "\P' A\<^sub>P \\<^sub>P A\<^sub>Q \\<^sub>Q. \\ \ \\<^sub>Q \ P \\M\\*xvec\\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; + distinct A\<^sub>P; distinct A\<^sub>Q; A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* Q; A\<^sub>P \* M; A\<^sub>Q \* \; A\<^sub>Q \* P; A\<^sub>Q \* Q; A\<^sub>Q \* M; + A\<^sub>P \* \\<^sub>Q; A\<^sub>Q \* \\<^sub>P; A\<^sub>P \* A\<^sub>Q; A\<^sub>P\<^sub>Q = A\<^sub>P@A\<^sub>Q; \\<^sub>P\<^sub>Q = \\<^sub>P \ \\<^sub>Q\ \ Prop (P' \ Q)" + and rPar2: "\Q' A\<^sub>P \\<^sub>P A\<^sub>Q \\<^sub>Q. \\ \ \\<^sub>P \ Q \\M\\*xvec\\N\ \ Q'; extractFrame P = \A\<^sub>P, \\<^sub>P\; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; + distinct A\<^sub>P; distinct A\<^sub>Q; A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* Q; A\<^sub>P \* M; A\<^sub>Q \* \; A\<^sub>Q \* P; A\<^sub>Q \* Q; A\<^sub>Q \* M; + A\<^sub>P \* \\<^sub>Q; A\<^sub>Q \* \\<^sub>P; A\<^sub>P \* A\<^sub>Q; A\<^sub>P\<^sub>Q = A\<^sub>P@A\<^sub>Q; \\<^sub>P\<^sub>Q = \\<^sub>P \ \\<^sub>Q\ \ Prop (P \ Q')" + and rBrComm1: "\\\<^sub>Q P' A\<^sub>P \\<^sub>P Q' A\<^sub>Q. + \\ \ \\<^sub>Q \ P \\M\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \ \ \\<^sub>P \ Q \\M\\*xvec\\N\ \ Q'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + distinct xvec; + A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* Q; A\<^sub>P \* A\<^sub>Q; + A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; A\<^sub>Q \* P; A\<^sub>Q \* Q; + A\<^sub>P \* M; A\<^sub>Q \* M; xvec \* M; + xvec \* \; xvec \* P; xvec \* Q; + A\<^sub>P\<^sub>Q = A\<^sub>P@A\<^sub>Q; \\<^sub>P\<^sub>Q = \\<^sub>P \ \\<^sub>Q\ \ + Prop (P' \ Q')" + and rBrComm2: "\\\<^sub>Q P' A\<^sub>P \\<^sub>P Q' A\<^sub>Q. + \\ \ \\<^sub>Q \ P \\M\\*xvec\\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; + \ \ \\<^sub>P \ Q \\M\N\ \ Q'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; + distinct xvec; + A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* Q; A\<^sub>P \* A\<^sub>Q; + A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; A\<^sub>Q \* P; A\<^sub>Q \* Q; + A\<^sub>P \* M; A\<^sub>Q \* M; xvec \* M; + xvec \* \; xvec \* P; xvec \* Q; + A\<^sub>P\<^sub>Q = A\<^sub>P@A\<^sub>Q; \\<^sub>P\<^sub>Q = \\<^sub>P \ \\<^sub>Q\ \ + Prop (P' \ Q')" +shows "Prop T" + using Trans \xvec \* \\ \xvec \* P\ \xvec \* Q\ \xvec \* M\ +proof(induct rule: parBrOutputCases[of _ _ _ _ _ _ _ "(A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q)"]) + case(cPar1 P' A\<^sub>Q \\<^sub>Q) + from \A\<^sub>Q \* (A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q)\ have "A\<^sub>Q \* A\<^sub>P\<^sub>Q" and "A\<^sub>Q \* \\<^sub>P\<^sub>Q" by simp+ + obtain A\<^sub>P \\<^sub>P where FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" and "distinct A\<^sub>P" + "A\<^sub>P \* (P, Q, \, M, A\<^sub>Q, A\<^sub>P\<^sub>Q, \\<^sub>Q)" + by(rule freshFrame) + then have "A\<^sub>P \* P" and "A\<^sub>P \* Q" and "A\<^sub>P \* \" and "A\<^sub>P \* M" and "A\<^sub>P \* A\<^sub>Q" and "A\<^sub>P \* A\<^sub>P\<^sub>Q" and "A\<^sub>P \* \\<^sub>Q" + by simp+ + + have FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" by fact + + from \A\<^sub>Q \* P\ \A\<^sub>P \* A\<^sub>Q\ FrP have "A\<^sub>Q \* \\<^sub>P" + by(force dest: extractFrameFreshChain) + + from \extractFrame(P \ Q) = \A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q\\ FrP FrQ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>Q \* \\<^sub>P\ + have "\(A\<^sub>P@A\<^sub>Q), \\<^sub>P \ \\<^sub>Q\ = \A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q\" by simp + moreover from \distinct A\<^sub>P\ \distinct A\<^sub>Q\ \A\<^sub>P \* A\<^sub>Q\ have "distinct(A\<^sub>P@A\<^sub>Q)" + by(auto simp add: fresh_star_def fresh_def name_list_supp) + ultimately obtain p where S: "set p \ set(A\<^sub>P@A\<^sub>Q) \ set((p \ A\<^sub>P)@(p \ A\<^sub>Q))" and "distinctPerm p" + and \eq: "\\<^sub>P\<^sub>Q = (p \ \\<^sub>P) \ (p \ \\<^sub>Q)" and Aeq: "A\<^sub>P\<^sub>Q = (p \ A\<^sub>P)@(p \ A\<^sub>Q)" + using \A\<^sub>P \* A\<^sub>P\<^sub>Q\ \A\<^sub>Q \* A\<^sub>P\<^sub>Q\ \distinct A\<^sub>P\<^sub>Q\ + by(elim frameChainEq') (assumption | simp add: eqvts)+ + + from \\ \ \\<^sub>Q \ P \\M\\*xvec\\N\ \ P'\ S \A\<^sub>P\<^sub>Q \* P\ \A\<^sub>P \* P\ \A\<^sub>Q \* P\ \A\<^sub>P\<^sub>Q \* M\ \A\<^sub>P \* M\ \A\<^sub>Q \* M\ Aeq + have "(p \ (\ \ \\<^sub>Q)) \ P \\M\\*xvec\\N\ \ P'" + by(elim broutputPermFrame) (assumption | simp)+ + + with S \A\<^sub>P\<^sub>Q \* \\ \A\<^sub>P \* \\ \A\<^sub>Q \* \\ Aeq have "\ \ (p \ \\<^sub>Q) \ P \\M\\*xvec\\N\ \ P'" + by(simp add: eqvts) + moreover from FrP have "(p \ extractFrame P) = p \ \A\<^sub>P, \\<^sub>P\" by simp + with S \A\<^sub>P\<^sub>Q \* P\ \A\<^sub>P \* P\ \A\<^sub>Q \* P\ Aeq have "extractFrame P = \(p \ A\<^sub>P), p \ \\<^sub>P\" + by(simp add: eqvts) + moreover from FrQ have "(p \ extractFrame Q) = p \ \A\<^sub>Q, \\<^sub>Q\" by simp + with S \A\<^sub>P\<^sub>Q \* Q\ \A\<^sub>P \* Q\ \A\<^sub>Q \* Q\ Aeq have "extractFrame Q = \(p \ A\<^sub>Q), p \ \\<^sub>Q\" + by(simp add: eqvts) + moreover from \distinct A\<^sub>P\ \distinct A\<^sub>Q\ have "distinct(p \ A\<^sub>P)" and "distinct(p \ A\<^sub>Q)" + by simp+ + moreover from \A\<^sub>P \* A\<^sub>Q\ have "(p \ A\<^sub>P) \* (p \ A\<^sub>Q)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + moreover from \A\<^sub>P \* \\<^sub>Q\ have "(p \ A\<^sub>P) \* (p \ \\<^sub>Q)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + moreover from \A\<^sub>Q \* \\<^sub>P\ have "(p \ A\<^sub>Q) \* (p \ \\<^sub>P)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + ultimately show ?case using \A\<^sub>P\<^sub>Q \* \\ \A\<^sub>P\<^sub>Q \* P\ \A\<^sub>P\<^sub>Q \* Q\ \A\<^sub>P\<^sub>Q \* M\ Aeq \eq + by(intro rPar1) (assumption | simp)+ +next + case(cPar2 Q' A\<^sub>P \\<^sub>P) + from \A\<^sub>P \* (A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q)\ have "A\<^sub>P \* A\<^sub>P\<^sub>Q" and "A\<^sub>P \* \\<^sub>P\<^sub>Q" by simp+ + obtain A\<^sub>Q \\<^sub>Q where FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" and "distinct A\<^sub>Q" + "A\<^sub>Q \* (P, Q, \, M, A\<^sub>P, A\<^sub>P\<^sub>Q, \\<^sub>P)" + by(rule freshFrame) + then have "A\<^sub>Q \* P" and "A\<^sub>Q \* Q" and "A\<^sub>Q \* \" and "A\<^sub>Q \* M" and "A\<^sub>Q \* A\<^sub>P" and "A\<^sub>Q \* A\<^sub>P\<^sub>Q" and "A\<^sub>Q \* \\<^sub>P" + by simp+ + + have FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" by fact + + from \A\<^sub>P \* Q\ \A\<^sub>Q \* A\<^sub>P\ FrQ have "A\<^sub>P \* \\<^sub>Q" + by(force dest: extractFrameFreshChain) + + from \extractFrame(P \ Q) = \A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q\\ FrP FrQ \A\<^sub>Q \* A\<^sub>P\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>Q \* \\<^sub>P\ + have "\(A\<^sub>P@A\<^sub>Q), \\<^sub>P \ \\<^sub>Q\ = \A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q\" by simp + moreover from \distinct A\<^sub>P\ \distinct A\<^sub>Q\ \A\<^sub>Q \* A\<^sub>P\ have "distinct(A\<^sub>P@A\<^sub>Q)" + by(auto simp add: fresh_star_def fresh_def name_list_supp) + ultimately obtain p where S: "set p \ set(A\<^sub>P@A\<^sub>Q) \ set((p \ A\<^sub>P)@(p \ A\<^sub>Q))" and "distinctPerm p" + and \eq: "\\<^sub>P\<^sub>Q = (p \ \\<^sub>P) \ (p \ \\<^sub>Q)" and Aeq: "A\<^sub>P\<^sub>Q = (p \ A\<^sub>P)@(p \ A\<^sub>Q)" + using \A\<^sub>P \* A\<^sub>P\<^sub>Q\ \A\<^sub>Q \* A\<^sub>P\<^sub>Q\ \distinct A\<^sub>P\<^sub>Q\ + by(elim frameChainEq') (assumption | simp add: eqvts)+ + + from \\ \ \\<^sub>P \ Q \\M\\*xvec\\N\ \ Q'\ S \A\<^sub>P\<^sub>Q \* Q\ \A\<^sub>P \* Q\ \A\<^sub>Q \* Q\ \A\<^sub>P\<^sub>Q \* M\ \A\<^sub>P \* M\ \A\<^sub>Q \* M\ Aeq + have "(p \ (\ \ \\<^sub>P)) \ Q \\M\\*xvec\\N\ \ Q'" + by(elim broutputPermFrame) (assumption | simp)+ + with S \A\<^sub>P\<^sub>Q \* \\ \A\<^sub>P \* \\ \A\<^sub>Q \* \\ Aeq have "\ \ (p \ \\<^sub>P) \ Q \\M\\*xvec\\N\ \ Q'" + by(simp add: eqvts) + moreover from FrP have "(p \ extractFrame P) = p \ \A\<^sub>P, \\<^sub>P\" by simp + with S \A\<^sub>P\<^sub>Q \* P\ \A\<^sub>P \* P\ \A\<^sub>Q \* P\ Aeq have "extractFrame P = \(p \ A\<^sub>P), p \ \\<^sub>P\" + by(simp add: eqvts) + moreover from FrQ have "(p \ extractFrame Q) = p \ \A\<^sub>Q, \\<^sub>Q\" by simp + with S \A\<^sub>P\<^sub>Q \* Q\ \A\<^sub>P \* Q\ \A\<^sub>Q \* Q\ Aeq have "extractFrame Q = \(p \ A\<^sub>Q), p \ \\<^sub>Q\" + by(simp add: eqvts) + moreover from \distinct A\<^sub>P\ \distinct A\<^sub>Q\ have "distinct(p \ A\<^sub>P)" and "distinct(p \ A\<^sub>Q)" + by simp+ + moreover from \A\<^sub>Q \* A\<^sub>P\ have "(p \ A\<^sub>P) \* (p \ A\<^sub>Q)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + moreover from \A\<^sub>P \* \\<^sub>Q\ have "(p \ A\<^sub>P) \* (p \ \\<^sub>Q)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + moreover from \A\<^sub>Q \* \\<^sub>P\ have "(p \ A\<^sub>Q) \* (p \ \\<^sub>P)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + ultimately show ?case using \A\<^sub>P\<^sub>Q \* \\ \A\<^sub>P\<^sub>Q \* P\ \A\<^sub>P\<^sub>Q \* Q\ \A\<^sub>P\<^sub>Q \* M\ Aeq \eq + by(intro rPar2) (assumption | simp)+ +next + case(cBrComm1 \\<^sub>Q P' A\<^sub>P \\<^sub>P Q' A\<^sub>Q) + then have FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" and FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" + and "A\<^sub>P \* A\<^sub>P\<^sub>Q" and "A\<^sub>Q \* A\<^sub>P\<^sub>Q" + by simp+ + + from \extractFrame(P \ Q) = \A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q\\ FrP FrQ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>Q \* \\<^sub>P\ + have "\(A\<^sub>P@A\<^sub>Q), \\<^sub>P \ \\<^sub>Q\ = \A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q\" by simp + moreover from \distinct A\<^sub>P\ \distinct A\<^sub>Q\ \A\<^sub>P \* A\<^sub>Q\ have "distinct(A\<^sub>P@A\<^sub>Q)" + by(auto simp add: fresh_star_def fresh_def name_list_supp) + ultimately obtain p where S: "set p \ set(A\<^sub>P@A\<^sub>Q) \ set((p \ A\<^sub>P)@(p \ A\<^sub>Q))" and "distinctPerm p" + and \eq: "\\<^sub>P\<^sub>Q = (p \ \\<^sub>P) \ (p \ \\<^sub>Q)" and Aeq: "A\<^sub>P\<^sub>Q = (p \ A\<^sub>P)@(p \ A\<^sub>Q)" + using \A\<^sub>P \* A\<^sub>P\<^sub>Q\ \A\<^sub>Q \* A\<^sub>P\<^sub>Q\ \distinct A\<^sub>P\<^sub>Q\ + by(elim frameChainEq') (assumption | simp add: eqvts)+ + + from \\ \ \\<^sub>Q \ P \\M\N\ \ P'\ S \A\<^sub>P\<^sub>Q \* P\ \A\<^sub>P \* P\ \A\<^sub>Q \* P\ \A\<^sub>P\<^sub>Q \* M\ \A\<^sub>P \* M\ \A\<^sub>Q \* M\ Aeq + have "(p \ (\ \ \\<^sub>Q)) \ P \\M\N\ \ P'" + by(elim brinputPermFrame) (assumption | simp)+ + + from \\ \ \\<^sub>P \ Q \\M\\*xvec\\N\ \ Q'\ S \A\<^sub>P\<^sub>Q \* Q\ \A\<^sub>Q \* Q\ \A\<^sub>P \* Q\ \A\<^sub>P\<^sub>Q \* M\ \A\<^sub>P \* M\ \A\<^sub>Q \* M\ Aeq + have "(p \ (\ \ \\<^sub>P)) \ Q \\M\\*xvec\\N\ \ Q'" + by(elim broutputPermFrame) (assumption | simp)+ + + from \(p \ (\ \ \\<^sub>Q)) \ P \\M\N\ \ P'\ S \A\<^sub>P\<^sub>Q \* \\ \A\<^sub>P \* \\ \A\<^sub>Q \* \\ Aeq have "\ \ (p \ \\<^sub>Q) \ P \\M\N\ \ P'" + by(simp add: eqvts) + moreover from \(p \ (\ \ \\<^sub>P)) \ Q \\M\\*xvec\\N\ \ Q'\ S \A\<^sub>P\<^sub>Q \* \\ \A\<^sub>P \* \\ \A\<^sub>Q \* \\ Aeq have "\ \ (p \ \\<^sub>P) \ Q \\M\\*xvec\\N\ \ Q'" + by(simp add: eqvts) + + moreover from FrP have "(p \ extractFrame P) = p \ \A\<^sub>P, \\<^sub>P\" by simp + with S \A\<^sub>P\<^sub>Q \* P\ \A\<^sub>P \* P\ \A\<^sub>Q \* P\ Aeq have "extractFrame P = \(p \ A\<^sub>P), p \ \\<^sub>P\" + by(simp add: eqvts) + moreover from FrQ have "(p \ extractFrame Q) = p \ \A\<^sub>Q, \\<^sub>Q\" by simp + with S \A\<^sub>P\<^sub>Q \* Q\ \A\<^sub>P \* Q\ \A\<^sub>Q \* Q\ Aeq have "extractFrame Q = \(p \ A\<^sub>Q), p \ \\<^sub>Q\" + by(simp add: eqvts) + moreover from \distinct A\<^sub>P\ \distinct A\<^sub>Q\ have "distinct(p \ A\<^sub>P)" and "distinct(p \ A\<^sub>Q)" + by simp+ + moreover from \A\<^sub>P \* A\<^sub>Q\ have "(p \ A\<^sub>P) \* (p \ A\<^sub>Q)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + moreover from \A\<^sub>P \* \\<^sub>Q\ have "(p \ A\<^sub>P) \* (p \ \\<^sub>Q)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + moreover from \A\<^sub>Q \* \\<^sub>P\ have "(p \ A\<^sub>Q) \* (p \ \\<^sub>P)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + ultimately show ?case using \distinct xvec\ \xvec \* M\ \xvec \* \\ \xvec \* P\ \xvec \* Q\ \A\<^sub>P\<^sub>Q \* \\ \A\<^sub>P\<^sub>Q \* P\ \A\<^sub>P\<^sub>Q \* Q\ \A\<^sub>P\<^sub>Q \* M\ Aeq \eq + by(intro rBrComm1) (assumption | simp)+ +next + case(cBrComm2 \\<^sub>Q P' A\<^sub>P \\<^sub>P Q' A\<^sub>Q) + then have FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" and FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" + and "A\<^sub>P \* A\<^sub>P\<^sub>Q" and "A\<^sub>Q \* A\<^sub>P\<^sub>Q" + by simp+ + + from \extractFrame(P \ Q) = \A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q\\ FrP FrQ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>Q \* \\<^sub>P\ + have "\(A\<^sub>P@A\<^sub>Q), \\<^sub>P \ \\<^sub>Q\ = \A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q\" by simp + moreover from \distinct A\<^sub>P\ \distinct A\<^sub>Q\ \A\<^sub>P \* A\<^sub>Q\ have "distinct(A\<^sub>P@A\<^sub>Q)" + by(auto simp add: fresh_star_def fresh_def name_list_supp) + ultimately obtain p where S: "set p \ set(A\<^sub>P@A\<^sub>Q) \ set((p \ A\<^sub>P)@(p \ A\<^sub>Q))" and "distinctPerm p" + and \eq: "\\<^sub>P\<^sub>Q = (p \ \\<^sub>P) \ (p \ \\<^sub>Q)" and Aeq: "A\<^sub>P\<^sub>Q = (p \ A\<^sub>P)@(p \ A\<^sub>Q)" + using \A\<^sub>P \* A\<^sub>P\<^sub>Q\ \A\<^sub>Q \* A\<^sub>P\<^sub>Q\ \distinct A\<^sub>P\<^sub>Q\ + by(elim frameChainEq') (assumption | simp add: eqvts)+ + + from \\ \ \\<^sub>Q \ P \\M\\*xvec\\N\ \ P'\ S \A\<^sub>P\<^sub>Q \* P\ \A\<^sub>P \* P\ \A\<^sub>Q \* P\ \A\<^sub>P\<^sub>Q \* M\ \A\<^sub>P \* M\ \A\<^sub>Q \* M\ Aeq + have "(p \ (\ \ \\<^sub>Q)) \ P \\M\\*xvec\\N\ \ P'" + by(elim broutputPermFrame) (assumption | simp)+ + + from \\ \ \\<^sub>P \ Q \\M\N\ \ Q'\ S \A\<^sub>P\<^sub>Q \* Q\ \A\<^sub>Q \* Q\ \A\<^sub>P \* Q\ \A\<^sub>P\<^sub>Q \* M\ \A\<^sub>P \* M\ \A\<^sub>Q \* M\ Aeq + have "(p \ (\ \ \\<^sub>P)) \ Q \\M\N\ \ Q'" + by(elim brinputPermFrame) (assumption | simp)+ + + from \(p \ (\ \ \\<^sub>Q)) \ P \\M\\*xvec\\N\ \ P'\ S \A\<^sub>P\<^sub>Q \* \\ \A\<^sub>P \* \\ \A\<^sub>Q \* \\ Aeq have "\ \ (p \ \\<^sub>Q) \ P \\M\\*xvec\\N\ \ P'" + by(simp add: eqvts) + moreover from \(p \ (\ \ \\<^sub>P)) \ Q \\M\N\ \ Q'\ S \A\<^sub>P\<^sub>Q \* \\ \A\<^sub>P \* \\ \A\<^sub>Q \* \\ Aeq have "\ \ (p \ \\<^sub>P) \ Q \\M\N\ \ Q'" + by(simp add: eqvts) + + moreover from FrP have "(p \ extractFrame P) = p \ \A\<^sub>P, \\<^sub>P\" by simp + with S \A\<^sub>P\<^sub>Q \* P\ \A\<^sub>P \* P\ \A\<^sub>Q \* P\ Aeq have "extractFrame P = \(p \ A\<^sub>P), p \ \\<^sub>P\" + by(simp add: eqvts) + moreover from FrQ have "(p \ extractFrame Q) = p \ \A\<^sub>Q, \\<^sub>Q\" by simp + with S \A\<^sub>P\<^sub>Q \* Q\ \A\<^sub>P \* Q\ \A\<^sub>Q \* Q\ Aeq have "extractFrame Q = \(p \ A\<^sub>Q), p \ \\<^sub>Q\" + by(simp add: eqvts) + moreover from \distinct A\<^sub>P\ \distinct A\<^sub>Q\ have "distinct(p \ A\<^sub>P)" and "distinct(p \ A\<^sub>Q)" + by simp+ + moreover from \A\<^sub>P \* A\<^sub>Q\ have "(p \ A\<^sub>P) \* (p \ A\<^sub>Q)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + moreover from \A\<^sub>P \* \\<^sub>Q\ have "(p \ A\<^sub>P) \* (p \ \\<^sub>Q)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + moreover from \A\<^sub>Q \* \\<^sub>P\ have "(p \ A\<^sub>Q) \* (p \ \\<^sub>P)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + ultimately show ?case using \distinct xvec\ \xvec \* M\ \xvec \* \\ \xvec \* P\ \xvec \* Q\ \A\<^sub>P\<^sub>Q \* \\ \A\<^sub>P\<^sub>Q \* P\ \A\<^sub>P\<^sub>Q \* Q\ \A\<^sub>P\<^sub>Q \* M\ Aeq \eq + by(intro rBrComm2) (assumption | simp)+ +qed + +inductive bangPred :: "('a, 'b, 'c) psi \ ('a, 'b, 'c) psi \ bool" + where + aux1: "bangPred P (!P)" + | aux2: "bangPred P (P \ !P)" + +lemma bangInduct[consumes 1, case_names cPar1 cPar2 cComm1 cComm2 cBrMerge cBrComm1 cBrComm2 cBang]: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and Rs :: "('a, 'b, 'c) residual" + and Prop :: "'d::fs_name \ 'b \ ('a, 'b, 'c) psi \ ('a, 'b, 'c) residual \ bool" + and C :: 'd + +assumes "\ \ !P \ Rs" + and rPar1: "\\ P' C. \\ \ P \\ \ P'; bn \ \* \; bn \ \* P; bn \ \* subject \; bn \ \* C; distinct(bn \)\ \ Prop C \ (P \ !P) (\ \ (P' \ !P))" + and rPar2: "\\ P' C. \\ \ !P \\ \ P'; bn \ \* \; bn \ \* P; bn \ \* subject \; bn \ \* C; distinct(bn \); + \C. Prop C \ (!P) (\ \ P')\ \ Prop C \ (P \ !P) (\ \ (P \ P'))" + and rComm1: "\M N P' K xvec P'' C. \\ \ P \M\N\ \ P'; \ \ !P \K\\*xvec\\N\ \ P''; \C. Prop C \ (!P) (K\\*xvec\\N\ \ P''); \ \ M \ K; + xvec \* \; xvec \* P; xvec \* M; xvec \* K; xvec \* C; distinct xvec\ \ Prop C \ (P \ !P) (\ \ \\*xvec\(P' \ P''))" + and rComm2: "\M xvec N P' K P'' C. \\ \ P \M\\*xvec\\N\ \ P'; \ \ !P \K\N\ \ P''; \C. Prop C \ (!P) (K\N\ \ P''); \ \ M \ K; + xvec \* \; xvec \* P; xvec \* M; xvec \* K; xvec \* C; distinct xvec\ \ Prop C \ (P \ !P) (\ \ \\*xvec\(P' \ P''))" + and rBrMerge: "\M N P' P'' C. \\ \ P \\M\N\ \ P'; \ \ !P \\M\N\ \ P''; \C. Prop C \ (!P) (\M\N\ \ P'')\ \ + Prop C \ (P \ !P) (\M\N\ \ (P' \ P''))" + and rBrComm1: "\M N P' xvec P'' C. \\ \ P \\M\N\ \ P'; \ \ !P \\M\\*xvec\\N\ \ P''; \C. Prop C \ (!P) (\M\\*xvec\\N\ \ P''); + xvec \* \; xvec \* P; xvec \* M; xvec \* C; distinct xvec\ \ Prop C \ (P \ !P) (\M\\*xvec\\N\ \ (P' \ P''))" + and rBrComm2: "\M N P' xvec P'' C. \\ \ P \\M\\*xvec\\N\ \ P'; \ \ !P \\M\N\ \ P''; \C. Prop C \ (!P) (\M\N\ \ P''); + xvec \* \; xvec \* P; xvec \* M; xvec \* C; distinct xvec\ \ Prop C \ (P \ !P) (\M\\*xvec\\N\ \ (P' \ P''))" + and rBang: "\Rs C. \\ \ P \ !P \ Rs; \C. Prop C \ (P \ !P) Rs; guarded P\ \ Prop C \ (!P) Rs" +shows "Prop C \ (!P) Rs" +proof - + from \\ \ !P \ Rs\ have "guarded P" + by(nominal_induct \ P=="!P" Rs rule: semantics.strong_induct) (auto simp add: psi.inject) + { + fix Q :: "('a, 'b, 'c) psi" + and \' :: 'b + + assume "\' \ Q \ Rs" + and "guarded Q" + and "bangPred P Q" + and "\ \ \'" + + then have "Prop C \ Q Rs" using rPar1 rPar1 rPar2 rPar2 rComm1 rComm2 rBrMerge rBrComm1 rBrComm2 rBang + proof(nominal_induct avoiding: \ C rule: semantics.strong_induct) + case(cInput \' M K xvec N Tvec Q \ C) + then show ?case by - (ind_cases "bangPred P (M\\*xvec N\.Q)") + next + case(cBrInput \' K M xvec N Tvec Q \ C) + then show ?case by - (ind_cases "bangPred P (M\\*xvec N\.Q)") + next + case(Output \' M K N Q \ C) + then show ?case by - (ind_cases "bangPred P (M\N\.Q)") + next + case(BrOutput \' M K N Q \ C) + then show ?case by - (ind_cases "bangPred P (M\N\.Q)") + next + case(Case \' Q Rs \ Cs \ C) + then show ?case by - (ind_cases "bangPred P (Cases Cs)") + next + case(cPar1 \' \\<^sub>R Q \ P' R A\<^sub>R \ C) + have rPar1: "\\ P' C. \\ \ P \\ \ P'; bn \ \* \; bn \ \* P; bn \ \* subject \; bn \ \* C; distinct(bn \)\ \ Prop C \ (P \ !P) (\ \ (P' \ !P))" + by fact + from \bangPred P (Q \ R)\ have "Q = P" and "R = !P" + by - (ind_cases "bangPred P (Q \ R)", auto simp add: psi.inject)+ + from \R = !P\ \extractFrame R = \A\<^sub>R, \\<^sub>R\\ have "A\<^sub>R = []" and "\\<^sub>R = \" by auto + from \\' \ \\<^sub>R \ Q \\ \ P'\ \Q = P\ \\ \ \'\ \\\<^sub>R = \\ have "\ \ P \\ \ P'" + by(metis statEqTransition Identity AssertionStatEqSym) + then have "Prop C \ (P \ !P) (\ \ (P' \ !P))" using \bn \ \* \\ \bn \ \* Q\ \bn \ \* subject \\ \bn \ \* C\ \Q = P\ \distinct(bn \)\ + by(intro rPar1) auto + with \R = !P\ \Q = P\ show ?case by simp + next + case(cPar2 \' \\<^sub>P R \ P' Q A\<^sub>P \ C) + have rPar2: "\\ P' C. \\ \ !P \\ \ P'; bn \ \* \; bn \ \* P; bn \ \* subject \; bn \ \* C; distinct(bn \); + \C. Prop C \ (!P) (\ \ P')\ \ Prop C \ (P \ !P) (\ \ (P \ P'))" + by fact + from \bangPred P (Q \ R)\ have "Q = P" and "R = !P" + by - (ind_cases "bangPred P (Q \ R)", auto simp add: psi.inject)+ + from \Q = P\ \extractFrame Q = \A\<^sub>P, \\<^sub>P\\ \guarded P\ have "\\<^sub>P \ \" and "supp \\<^sub>P = ({}::name set)" + by(blast dest: guardedStatEq)+ + from \\' \ \\<^sub>P \ R \\ \ P'\ \R = !P\ \\ \ \'\ \\\<^sub>P \ \\ have "\ \ !P \\ \ P'" + by(metis statEqTransition Identity Composition Commutativity AssertionStatEqSym) + moreover + { + fix C + have "bangPred P (!P)" by(rule aux1) + moreover from \\ \ \'\ \\\<^sub>P \ \\ have "\ \ \' \ \\<^sub>P" by(metis Composition Identity Commutativity AssertionStatEqSym AssertionStatEqTrans) + ultimately have "\C. Prop C \ (!P) (\ \ P')" using cPar2 \R = !P\ \guarded P\ by simp + } + ultimately have "Prop C \ (P \ !P) (\ \ (P \ P'))" using \bn \ \* \\ \bn \ \* Q\ \bn \ \* subject \\ \bn \ \* C\ \Q = P\ \distinct(bn \)\ + by(elim rPar2) auto + with \R = !P\ \Q = P\ show ?case by simp + next + case(cComm1 \' \\<^sub>R Q M N P' A\<^sub>P \\<^sub>P R K xvec P'' A\<^sub>R \ C) + have rComm1: "\M N P' K xvec P'' C. \\ \ P \M\N\ \ P'; \ \ !P \K\\*xvec\\N\ \ P''; \C. Prop C \ (!P) (K\\*xvec\\N\ \ P''); \ \ M \ K; + xvec \* \; xvec \* P; xvec \* M; xvec \* K; xvec \* C; distinct xvec\ \ Prop C \ (P \ !P) (\ \ \\*xvec\(P' \ P''))" + by fact + from \bangPred P (Q \ R)\ have "Q = P" and "R = !P" + by - (ind_cases "bangPred P (Q \ R)", auto simp add: psi.inject)+ + from \R = !P\ \extractFrame R = \A\<^sub>R, \\<^sub>R\\ have "A\<^sub>R = []" and "\\<^sub>R = \" by auto + from \\' \ \\<^sub>R \ Q \M\N\ \ P'\ \Q = P\ \\ \ \'\ \\\<^sub>R = \\ have "\ \ P \M\N\ \ P'" + by(metis statEqTransition Identity AssertionStatEqSym) + moreover from \Q = P\ \extractFrame Q = \A\<^sub>P, \\<^sub>P\\ \guarded P\ have "\\<^sub>P \ \" and "supp \\<^sub>P = ({}::name set)" + by(blast dest: guardedStatEq)+ + moreover from \\' \ \\<^sub>P \ R \K\\*xvec\\N\ \ P''\ \R = !P\ \\\<^sub>P \ \\ \\ \ \'\ have "\ \ !P \K\\*xvec\\N\ \ P''" + by(metis statEqTransition Identity Composition Commutativity AssertionStatEqSym) + moreover + { + fix C + have "bangPred P (!P)" by(rule aux1) + moreover from \\ \ \'\ \\\<^sub>P \ \\ have "\ \ \' \ \\<^sub>P" by(metis Composition Identity Commutativity AssertionStatEqSym AssertionStatEqTrans) + ultimately have "\C. Prop C \ (!P) (K\\*xvec\\N\ \ P'')" using cComm1 \R = !P\ \guarded P\ by simp + } + moreover from \\' \ \\<^sub>P \ \\<^sub>R \ M \ K\ \\\<^sub>P \ \\ \\ \ \'\ \\\<^sub>R = \\ have "\ \ M \ K" + by(metis statEqEnt Identity Composition Commutativity AssertionStatEqSym) + ultimately have "Prop C \ (P \ !P) (\ \ \\*xvec\(P' \ P''))" using \xvec \* \\ \xvec \* Q\ \xvec \* M\ \xvec \* K\ \xvec \* C\ \Q = P\ \distinct xvec\ + by(elim rComm1[where K=K and M=M and N=N]) auto + with \R = !P\ \Q = P\ show ?case by simp + next + case(cComm2 \' \\<^sub>R Q M xvec N P' A\<^sub>P \\<^sub>P R K P'' A\<^sub>R \ C) + have rComm2: "\M xvec N P' K P'' C. \\ \ P \M\\*xvec\\N\ \ P'; \ \ !P \K\N\ \ P''; \C. Prop C \ (!P) (K\N\ \ P''); \ \ M \ K; + xvec \* \; xvec \* P; xvec \* M; xvec \* K; xvec \* C; distinct xvec\ \ Prop C \ (P \ !P) (\ \ \\*xvec\(P' \ P''))" + by fact + from \bangPred P (Q \ R)\ have "Q = P" and "R = !P" + by - (ind_cases "bangPred P (Q \ R)", auto simp add: psi.inject)+ + from \R = !P\ \extractFrame R = \A\<^sub>R, \\<^sub>R\\ have "A\<^sub>R = []" and "\\<^sub>R = \" by auto + from \\' \ \\<^sub>R \ Q \M\\*xvec\\N\ \ P'\ \Q = P\ \\ \ \'\ \\\<^sub>R = \\ have "\ \ P \M\\*xvec\\N\ \ P'" + by(metis statEqTransition Identity AssertionStatEqSym) + moreover from \Q = P\ \extractFrame Q = \A\<^sub>P, \\<^sub>P\\ \guarded P\ have "\\<^sub>P \ \" and "supp \\<^sub>P = ({}::name set)" + by(blast dest: guardedStatEq)+ + moreover from \\' \ \\<^sub>P \ R \K\N\ \ P''\ \R = !P\ \\\<^sub>P \ \\ \\ \ \'\ have "\ \ !P \K\N\ \ P''" + by(metis statEqTransition Identity Composition Commutativity AssertionStatEqSym) + moreover + { + fix C + have "bangPred P (!P)" by(rule aux1) + moreover from \\ \ \'\ \\\<^sub>P \ \\ have "\ \ \' \ \\<^sub>P" by(metis Composition Identity Commutativity AssertionStatEqSym AssertionStatEqTrans) + ultimately have "\C. Prop C \ (!P) (K\N\ \ P'')" using cComm2 \R = !P\ \guarded P\ by simp + } + moreover from \\' \ \\<^sub>P \ \\<^sub>R \ M \ K\ \\\<^sub>P \ \\ \\ \ \'\ \\\<^sub>R = \\ have "\ \ M \ K" + by(metis statEqEnt Identity Composition Commutativity AssertionStatEqSym) + ultimately have "Prop C \ (P \ !P) (\ \ \\*xvec\(P' \ P''))" using \xvec \* \\ \xvec \* Q\ \xvec \* M\ \xvec \* K\ \xvec \* C\ \Q = P\ \distinct xvec\ + by(elim rComm2[where K=K and M=M and N=N]) auto + with \R = !P\ \Q = P\ show ?case by simp + next + case(cBrMerge \' \\<^sub>R Q M N P' A\<^sub>P \\<^sub>P R P'' A\<^sub>R \ C) + have rBrMerge: "\M N P' P'' C. \\ \ P \\M\N\ \ P'; \ \ !P \\M\N\ \ P''; \C. Prop C \ (!P) (\M\N\ \ P'')\ \ + Prop C \ (P \ !P) (\M\N\ \ (P' \ P''))" + by fact + from \bangPred P (Q \ R)\ have "Q = P" and "R = !P" + by - (ind_cases "bangPred P (Q \ R)", auto simp add: psi.inject)+ + from \R = !P\ \extractFrame R = \A\<^sub>R, \\<^sub>R\\ have "A\<^sub>R = []" and "\\<^sub>R = \" by auto + from \\' \ \\<^sub>R \ Q \\M\N\ \ P'\ \Q = P\ \\ \ \'\ \\\<^sub>R = \\ have "\ \ P \\M\N\ \ P'" + by(metis statEqTransition Identity AssertionStatEqSym) + moreover from \Q = P\ \extractFrame Q = \A\<^sub>P, \\<^sub>P\\ \guarded P\ have "\\<^sub>P \ \" and "supp \\<^sub>P = ({}::name set)" + by(blast dest: guardedStatEq)+ + from \\' \ \\<^sub>P \ R \\M\N\ \ P''\ \R = !P\ \\ \ \'\ \\\<^sub>P \ \\ have "\ \ !P \\M\N\ \ P''" + by (metis AssertionStatEqSym Identity compositionSym statEqTransition) + moreover + { + fix C + have "bangPred P (!P)" by(rule aux1) + moreover from \\ \ \'\ \\\<^sub>P \ \\ have "\ \ \' \ \\<^sub>P" by(metis Composition Identity Commutativity AssertionStatEqSym AssertionStatEqTrans) + ultimately have "\C. Prop C \ (!P) (\M\N\ \ P'')" using cBrMerge \R = !P\ \guarded P\ by simp + } + ultimately have "Prop C \ (P \ !P) (\M\N\ \ (P' \ P''))" + by(elim rBrMerge) auto + with \R = !P\ \Q = P\ show ?case by simp + next + case(cBrComm1 \' \\<^sub>R Q M N P' A\<^sub>P \\<^sub>P R xvec P'' A\<^sub>R \ C) + have rBrComm1: "\M N P' xvec P'' C. \\ \ P \\M\N\ \ P'; \ \ !P \\M\\*xvec\\N\ \ P''; \C. Prop C \ (!P) (\M\\*xvec\\N\ \ P''); + xvec \* \; xvec \* P; xvec \* M; xvec \* C; distinct xvec\ \ Prop C \ (P \ !P) (\M\\*xvec\\N\ \ (P' \ P''))" + by fact + from \bangPred P (Q \ R)\ have "Q = P" and "R = !P" + by - (ind_cases "bangPred P (Q \ R)", auto simp add: psi.inject)+ + from \R = !P\ \extractFrame R = \A\<^sub>R, \\<^sub>R\\ have "A\<^sub>R = []" and "\\<^sub>R = \" by auto + from \\' \ \\<^sub>R \ Q \\M\N\ \ P'\ \Q = P\ \\ \ \'\ \\\<^sub>R = \\ have "\ \ P \\M\N\ \ P'" + by(metis statEqTransition Identity AssertionStatEqSym) + moreover from \Q = P\ \extractFrame Q = \A\<^sub>P, \\<^sub>P\\ \guarded P\ have "\\<^sub>P \ \" and "supp \\<^sub>P = ({}::name set)" + by(blast dest: guardedStatEq)+ + moreover from \\' \ \\<^sub>P \ R \\M\\*xvec\\N\ \ P''\ \R = !P\ \\\<^sub>P \ \\ \\ \ \'\ have "\ \ !P \\M\\*xvec\\N\ \ P''" + by(metis statEqTransition Identity Composition Commutativity AssertionStatEqSym) + moreover + { + fix C + have "bangPred P (!P)" by(rule aux1) + moreover from \\ \ \'\ \\\<^sub>P \ \\ have "\ \ \' \ \\<^sub>P" by(metis Composition Identity Commutativity AssertionStatEqSym AssertionStatEqTrans) + ultimately have "\C. Prop C \ (!P) (\M\\*xvec\\N\ \ P'')" using cBrComm1 \R = !P\ \guarded P\ by simp + } + ultimately have "Prop C \ (P \ !P) (\M\\*xvec\\N\ \ (P' \ P''))" using \xvec \* \\ \xvec \* Q\ \xvec \* M\ \xvec \* C\ \Q = P\ \distinct xvec\ + by(elim rBrComm1) auto + with \R = !P\ \Q = P\ show ?case by simp + next + case(cBrComm2 \' \\<^sub>R Q M xvec N P' A\<^sub>P \\<^sub>P R P'' A\<^sub>R \ C) + have rBrComm2: "\M N P' xvec P'' C. \\ \ P \\M\\*xvec\\N\ \ P'; \ \ !P \\M\N\ \ P''; \C. Prop C \ (!P) (\M\N\ \ P''); + xvec \* \; xvec \* P; xvec \* M; xvec \* C; distinct xvec\ \ Prop C \ (P \ !P) (\M\\*xvec\\N\ \ (P' \ P''))" + by fact + from \bangPred P (Q \ R)\ have "Q = P" and "R = !P" + by - (ind_cases "bangPred P (Q \ R)", auto simp add: psi.inject)+ + from \R = !P\ \extractFrame R = \A\<^sub>R, \\<^sub>R\\ have "A\<^sub>R = []" and "\\<^sub>R = \" by auto + from \\' \ \\<^sub>R \ Q \\M\\*xvec\\N\ \ P'\ \Q = P\ \\ \ \'\ \\\<^sub>R = \\ have "\ \ P \\M\\*xvec\\N\ \ P'" + by(metis statEqTransition Identity AssertionStatEqSym) + moreover from \Q = P\ \extractFrame Q = \A\<^sub>P, \\<^sub>P\\ \guarded P\ have "\\<^sub>P \ \" and "supp \\<^sub>P = ({}::name set)" + by(blast dest: guardedStatEq)+ + moreover from \\' \ \\<^sub>P \ R \\M\N\ \ P''\ \R = !P\ \\\<^sub>P \ \\ \\ \ \'\ have "\ \ !P \\M\N\ \ P''" + by(metis statEqTransition Identity Composition Commutativity AssertionStatEqSym) + moreover + { + fix C + have "bangPred P (!P)" by(rule aux1) + moreover from \\ \ \'\ \\\<^sub>P \ \\ have "\ \ \' \ \\<^sub>P" by(metis Composition Identity Commutativity AssertionStatEqSym AssertionStatEqTrans) + ultimately have "\C. Prop C \ (!P) (\M\N\ \ P'')" using cBrComm2 \R = !P\ \guarded P\ by simp + } + ultimately have "Prop C \ (P \ !P) (\M\\*xvec\\N\ \ (P' \ P''))" using \xvec \* \\ \xvec \* Q\ \xvec \* M\ \xvec \* C\ \Q = P\ \distinct xvec\ + by(elim rBrComm2) auto + with \R = !P\ \Q = P\ show ?case by simp + next + case(cBrClose \' Q M xvec N P' x \ C) + then show ?case by - (ind_cases "bangPred P (\\x\Q)") + next + case(cOpen \ Q M xvec yvec N P' x C) + then show ?case by - (ind_cases "bangPred P (\\x\Q)") + next + case(cBrOpen \ Q M xvec yvec N P' x C) + then show ?case by - (ind_cases "bangPred P (\\x\Q)") + next + case(cScope \ Q \ P' x C) + then show ?case by - (ind_cases "bangPred P (\\x\Q)") + next + case(Bang \' Q Rs \ C) + have rBang: "\Rs C. \\ \ P \ !P \ Rs; \C. Prop C \ (P \ !P) Rs; guarded P\ \ Prop C \ (!P) Rs" + by fact + from \bangPred P (!Q)\ have "P = Q" + by - (ind_cases "bangPred P (!Q)", auto simp add: psi.inject) + with \\' \ Q \ !Q \ Rs\ \\ \ \'\ have "\ \ P \ !P \ Rs" by(metis statEqTransition AssertionStatEqSym) + moreover + { + fix C + have "bangPred P (P \ !P)" by(rule aux2) + with Bang \P = Q\ have "\C. Prop C \ (P \ !P) Rs" by simp + } + moreover from \guarded Q\ \P = Q\ have "guarded P" by simp + ultimately have "Prop C \ (!P) Rs" by(rule rBang) + with \P = Q\ show ?case by simp + qed + } + with \guarded P\ \\ \ !P \ Rs\ + show ?thesis by(force intro: aux1) +qed + +lemma bangInputInduct[consumes 1, case_names cPar1 cPar2 cBang]: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and M :: 'a + and N :: 'a + and P' :: "('a, 'b, 'c) psi" + and Prop :: "'b \ ('a, 'b, 'c) psi \ 'a \ 'a \ ('a, 'b, 'c) psi \ bool" + +assumes "\ \ !P \M\N\ \ P'" + and rPar1: "\P'. \ \ P \M\N\ \ P' \ Prop \ (P \ !P) M N (P' \ !P)" + and rPar2: "\P'. \\ \ !P \M\N\ \ P'; Prop \ (!P) M N P'\ \ Prop \ (P \ !P) M N (P \ P')" + and rBang: "\P'. \\ \ P \ !P \M\N\ \ P'; Prop \ (P \ !P) M N P'; guarded P\ \ Prop \ (!P) M N P'" +shows "Prop \ (!P) M N P'" + using \\ \ !P \M\N\ \ P'\ + by(nominal_induct \ P Rs=="M\N\ \ P'" arbitrary: P' rule: bangInduct) + (auto simp add: residualInject intro: rPar1 rPar2 rBang) + +lemma brbangInputInduct[consumes 1, case_names cPar1 cPar2 cBrMerge cBang]: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and M :: 'a + and N :: 'a + and P' :: "('a, 'b, 'c) psi" + and Prop :: "'b \ ('a, 'b, 'c) psi \ 'a \ 'a \ ('a, 'b, 'c) psi \ bool" + +assumes "\ \ !P \\M\N\ \ P'" + and rPar1: "\P'. \ \ P \\M\N\ \ P' \ Prop \ (P \ !P) M N (P' \ !P)" + and rPar2: "\P'. \\ \ !P \\M\N\ \ P'; Prop \ (!P) M N P'\ \ Prop \ (P \ !P) M N (P \ P')" + and rBrMerge: "\P' P'' C. \\ \ P \\M\N\ \ P'; \ \ !P \\M\N\ \ P''; \C. Prop \ (!P) M N P''\ \ + Prop \ (P \ !P) M N (P' \ P'')" + and rBang: "\P'. \\ \ P \ !P \\M\N\ \ P'; Prop \ (P \ !P) M N P'; guarded P\ \ Prop \ (!P) M N P'" +shows "Prop \ (!P) M N P'" + using \\ \ !P \\M\N\ \ P'\ + by(nominal_induct \ P Rs=="\M\N\ \ P'" arbitrary: P' rule: bangInduct) + (auto simp add: residualInject intro: rPar1 rPar2 rBrMerge rBang) + +lemma bangOutputInduct[consumes 1, case_names cPar1 cPar2 cBang]: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and M :: 'a + and xvec :: "name list" + and N :: 'a + and P' :: "('a, 'b, 'c) psi" + and Prop :: "'d::fs_name \ 'b \ ('a, 'b, 'c) psi \ 'a \ ('a, 'b, 'c) boundOutput \ bool" + and C :: 'd + +assumes "\ \ !P \M\\*xvec\\N\ \ P'" + and rPar1: "\xvec N P' C. \\ \ P \M\\*xvec\\N\ \ P'; xvec \* \; xvec \* P; xvec \* M; xvec \* C; distinct xvec\ \ Prop C \ (P \ !P) M (\\*xvec\N \' (P' \ !P))" + and rPar2: "\xvec N P' C. \\ \ !P \M\\*xvec\\N\ \ P'; \C. Prop C \ (!P) M (\\*xvec\N \' P'); xvec \* \; xvec \* P; xvec \* M; xvec \* C; distinct xvec\ \ + Prop C \ (P \ !P) M (\\*xvec\N \' (P \ P'))" + and rBang: "\B C. \\ \ P \ !P \(ROut M B); \C. Prop C \ (P \ !P) M B; guarded P\ \ Prop C \ (!P) M B" + +shows "Prop C \ (!P) M (\\*xvec\N \' P')" + using \\ \ !P \M\\*xvec\\N\ \ P'\ + apply(simp add: residualInject) + by(nominal_induct \ P Rs=="ROut M (\\*xvec\N \' P')" avoiding: C arbitrary: xvec N P' rule: bangInduct) + (force simp add: residualInject intro: rPar1 rPar2 rBang)+ + +lemma bangTauInduct[consumes 1, case_names cPar1 cPar2 cComm1 cComm2 cBang]: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and P' :: "('a, 'b, 'c) psi" + and Prop :: "'d::fs_name \ 'b \ ('a, 'b, 'c) psi \ ('a, 'b, 'c) psi \ bool" + and C :: 'd + +assumes "\ \ !P \\ \ P'" + and rPar1: "\P' C. \ \ P \\ \ P' \ Prop C \ (P \ !P) (P' \ !P)" + and rPar2: "\P' C. \\ \ !P \\ \ P'; \C. Prop C \ (!P) P'\ \ Prop C \ (P \ !P) (P \ P')" + and rComm1: "\M N P' K xvec P'' C. \\ \ P \M\N\ \ P'; \ \ !P \K\\*xvec\\N\ \ P''; \ \ M \ K; + xvec \* \; xvec \* P; xvec \* M; xvec \* K; xvec \* C\ \ Prop C \ (P \ !P) (\\*xvec\(P' \ P''))" + and rComm2: "\M N P' K xvec P'' C. \\ \ P \M\\*xvec\\N\ \ P'; \ \ !P \K\N\ \ P''; \ \ M \ K; + xvec \* \; xvec \* P; xvec \* M; xvec \* K; xvec \* C\ \ Prop C \ (P \ !P) (\\*xvec\(P' \ P''))" + and rBang: "\P' C. \\ \ P \ !P \\ \ P'; \C. Prop C \ (P \ !P) P'; guarded P\ \ Prop C \ (!P) P'" + +shows "Prop C \ (!P) P'" + using \\ \ !P \\ \ P'\ + by(nominal_induct \ P Rs=="\ \ P'" avoiding: C arbitrary: P' rule: bangInduct) + (auto simp add: residualInject intro: rPar1 rPar2 rComm1 rComm2 rBang) + +lemma bangInduct'[consumes 2, case_names cAlpha cPar1 cPar2 cComm1 cComm2 cBrMerge cBrComm1 cBrComm2 cBang]: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and \ :: "'a action" + and P' :: "('a, 'b, 'c) psi" + and Prop :: "'d::fs_name \ 'b \ ('a, 'b, 'c) psi \ 'a action \ ('a, 'b, 'c) psi \ bool" + and C :: "'d::fs_name" + +assumes "\ \ !P \\ \ P'" + and "bn \ \* subject \" + and rAlpha: "\\ P' p C. \bn \ \* \; bn \ \* P; bn \ \* subject \; bn \ \* C; + set p \ set(bn \) \ set(bn(p \ \)); distinctPerm p; + bn(p \ \) \* \; bn(p \ \) \* P'; Prop C \ (P \ !P) \ P'\ \ + Prop C \ (P \ !P) (p \ \) (p \ P')" + and rPar1: "\\ P' C. + \\ \ P \\ \ P'; bn \ \* \; bn \ \* P; bn \ \* subject \; bn \ \* C; distinct(bn \)\ \ + Prop C \ (P \ !P) \ (P' \ !P)" + and rPar2: "\\ P' C. + \\ \ !P \\ \ P'; \C. Prop C \ (!P) \ P'; + bn \ \* \; bn \ \* P; bn \ \* subject \; bn \ \* C; distinct(bn \)\ \ + Prop C \ (P \ !P) \ (P \ P')" + and rComm1: "\M N P' K xvec P'' C. \\ \ P \M\N\ \ P'; \ \ !P \K\\*xvec\\N\ \ P''; \C. Prop C \ (!P) (K\\*xvec\\N\) P''; \ \ M \ K; + xvec \* \; xvec \* P; xvec \* M; xvec \* K; xvec \* C; distinct xvec\ \ Prop C \ (P \ !P) (\) (\\*xvec\(P' \ P''))" + and rComm2: "\M xvec N P' K P'' C. \\ \ P \M\\*xvec\\N\ \ P'; \ \ !P \K\N\ \ P''; \C. Prop C \ (!P) (K\N\) P''; \ \ M \ K; + xvec \* \; xvec \* P; xvec \* M; xvec \* K; xvec \* C; distinct xvec\ \ Prop C \ (P \ !P) (\) (\\*xvec\(P' \ P''))" + and rBrMerge: "\M N P' P'' C. \\ \ P \\M\N\ \ P'; \ \ !P \\M\N\ \ P''; \C. Prop C \ (!P) (\M\N\) P''\ \ + Prop C \ (P \ !P) (\M\N\) (P' \ P'')" + and rBrComm1: "\M N P' xvec P'' C. \\ \ P \\M\N\ \ P'; \ \ !P \\M\\*xvec\\N\ \ P''; \C. Prop C \ (!P) (\M\\*xvec\\N\) P''; + xvec \* \; xvec \* P; xvec \* M; xvec \* C; distinct xvec\ \ Prop C \ (P \ !P) (\M\\*xvec\\N\) (P' \ P'')" + and rBrComm2: "\M N P' xvec P'' C. \\ \ P \\M\\*xvec\\N\ \ P'; \ \ !P \\M\N\ \ P''; \C. Prop C \ (!P) (\M\N\) P''; + xvec \* \; xvec \* P; xvec \* M; xvec \* C; distinct xvec\ \ Prop C \ (P \ !P) (\M\\*xvec\\N\) (P' \ P'')" + and rBang: "\\ P' C. + \\ \ P \ !P \\ \ P'; guarded P; \C. Prop C \ (P \ !P) \ P'; guarded P; distinct(bn \)\ \ + Prop C \ (!P) \ P'" +shows "Prop C \ (!P) \ P'" +proof - + from \\ \ !P \\ \ P'\ have "distinct(bn \)" by(rule boundOutputDistinct) + with \\ \ !P \\ \ P'\ \bn \ \* subject \\ show ?thesis + proof(nominal_induct \ P Rs=="\ \ P'" avoiding: C \ P' rule: bangInduct) + case(cPar1 \ P' C \' P'') + note \\ \ (P' \ !P) = \' \ P''\ + moreover from \bn \ \* \'\ have "bn \ \* bn \'" by simp + moreover note \distinct(bn \)\ \distinct(bn \')\ + moreover from \bn \ \* subject \\ have "bn \ \* (\ \ P' \ !P)" by simp + moreover from \bn \' \* subject \'\ have "bn \' \* (\' \ P'')" by simp + ultimately obtain p where S: "set p \ set(bn \) \ set(bn(p \ \))" and "distinctPerm p" and "\' = p \ \" + and P'eq: "P'' = p \ (P' \ !P)" and "bn(p \ \) \* \" and "bn(p \ \) \* (P' \ !P)" + by(elim residualEq) + + from \\ \ P \\ \ P'\ \bn \ \* \\ \bn \ \* P\ \bn \ \* subject \\ \bn \ \* C\ \distinct(bn \)\ + have "Prop C \ (P \ !P) \ (P' \ !P)" + by(rule rPar1) + + with \bn \ \* \\ \bn \ \* P\ \bn \ \* subject \\ \bn \ \* C\ S \distinctPerm p\ \bn \ \* \'\ \bn \ \* P''\ \\' = (p \ \)\ P'eq \bn(p \ \) \* \\ \bn(p \ \) \* (P' \ !P)\ + have "Prop C \ (P \ !P) (p \ \) (p \ (P' \ !P))" + by(elim rAlpha) + with P'eq \\' = p \ \\ \distinctPerm p\ show ?case by simp + next + case(cPar2 \ P' C \' P'') + note \\ \ (P \ P') = \' \ P''\ + moreover from \bn \ \* \'\ have "bn \ \* bn \'" by simp + moreover note \distinct(bn \)\ \distinct(bn \')\ + moreover from \bn \ \* subject \\ have "bn \ \* (\ \ P \ P')" by simp + moreover from \bn \' \* subject \'\ have "bn \' \* (\' \ P'')" by simp + ultimately obtain p where S: "set p \ set(bn \) \ set(bn(p \ \))" and "distinctPerm p" and "\' = p \ \" + and P'eq: "P'' = p \ (P \ P')" and "bn(p \ \) \* \" and "bn(p \ \) \* (P \ P')" + by(elim residualEq) + + note \\ \ !P \\ \ P'\ + moreover from \bn \ \* subject \\ \distinct(bn \)\ have "\C. Prop C \ (!P) \ P'" by(intro cPar2) auto + moreover note \bn \ \* \\ \bn \ \* P\ \bn \ \* subject \\ \bn \ \* C\ \distinct(bn \)\ + ultimately have "Prop C \ (P \ !P) \ (P \ P')" + by(rule rPar2) + with \bn \ \* \\ \bn \ \* P\ \bn \ \* subject \\ \bn \ \* C\ S \distinctPerm p\ \bn \ \* \'\ \bn \ \* P''\ \\' = (p \ \)\ P'eq \bn(p \ \) \* \\ \bn(p \ \) \* (P \ P')\ + have "Prop C \ (P \ !P) (p \ \) (p \ (P \ P'))" + by(elim rAlpha) + with P'eq \\' = p \ \\ show ?case by simp + next + case(cComm1 M N P' K xvec P'' C \ P''') + then have "Prop C \ (P \ !P) (\) (\\*xvec\(P' \ P''))" + by(elim rComm1) (assumption | simp)+ + then show ?case using \\ \ \\*xvec\(P' \ P'') = \ \ P'''\ + by(simp add: residualInject) + next + case(cComm2 M xvec N P' K P'' C \ P''') + then have "Prop C \ (P \ !P) (\) (\\*xvec\(P' \ P''))" + by(elim rComm2) (assumption | simp)+ + then show ?case using \\ \ \\*xvec\(P' \ P'') = \ \ P'''\ + by(simp add: residualInject) + next + case(cBrMerge M N P' P'' C \ P''') + then have "Prop C \ (P \ !P) (\M\N\) (P' \ P'')" + by(elim rBrMerge) (assumption | simp)+ + then show ?case using \\M\N\ \ P' \ P'' = \ \ P'''\ + by(simp add: residualInject) + next + case(cBrComm1 M N P' xvec P'' C \ P''') + note \(\M\\*xvec\\N\) \ (P' \ P'') = \ \ P'''\ + moreover from \xvec \* \\ have "xvec \* bn \" by simp + moreover note \distinct xvec\ \distinct(bn \)\ + moreover from \xvec \* M\ have "xvec \* ((\M\\*xvec\\N\) \ (P' \ P''))" by simp + moreover from \bn \ \* subject \\ have "bn \ \* (\ \ P''')" by simp + ultimately obtain p where S: "set p \ set xvec \ set(p \ xvec)" and "distinctPerm p" and "\ = p \ (\M\\*xvec\\N\)" + and P'eq: "P''' = p \ (P' \ P'')" and "(p \ xvec) \* (\M\\*xvec\\N\)" and "(p \ xvec) \* (P' \ P'')" + using residualEq[where \="(\M\\*xvec\\N\)" and \=\] by (smt (verit, best) Sigma_cong bn.simps(4) bnEqvt) + + from cBrComm1 have "Prop C \ (P \ !P) (\M\\*xvec\\N\) (P' \ P'')" + by(elim rBrComm1) (assumption | simp)+ + + with \xvec \* \\ \xvec \* P\ \xvec \* M\ \xvec \* C\ S \distinctPerm p\ \xvec \* \\ \xvec \* P'''\ \\ = (p \ (\M\\*xvec\\N\))\ P'eq \(p \ xvec) \* (\M\\*xvec\\N\)\ \(p \ xvec) \* (P' \ P'')\ + have "Prop C \ (P \ !P) (p \ (\M\\*xvec\\N\)) (p \ (P' \ P''))" + by(intro rAlpha) simp+ + + with P'eq \\ = p \ (\M\\*xvec\\N\)\ \distinctPerm p\ show ?case by simp + next + case(cBrComm2 M N P' xvec P'' C \ P''') + note \(\M\\*xvec\\N\) \ (P' \ P'') = \ \ P'''\ + moreover from \xvec \* \\ have "xvec \* bn \" by simp + moreover note \distinct xvec\ \distinct(bn \)\ + moreover from \xvec \* M\ have "xvec \* ((\M\\*xvec\\N\) \ (P' \ P''))" by simp + moreover from \bn \ \* subject \\ have "bn \ \* (\ \ P''')" by simp + ultimately obtain p where S: "set p \ set xvec \ set(p \ xvec)" and "distinctPerm p" and "\ = p \ (\M\\*xvec\\N\)" + and P'eq: "P''' = p \ (P' \ P'')" and "(p \ xvec) \* (\M\\*xvec\\N\)" and "(p \ xvec) \* (P' \ P'')" + using residualEq[where \="(\M\\*xvec\\N\)" and \=\] by (smt (verit, best) Sigma_cong bn.simps(4) bnEqvt) + + from cBrComm2 have "Prop C \ (P \ !P) (\M\\*xvec\\N\) (P' \ P'')" + by(elim rBrComm2) (assumption | simp)+ + + with \xvec \* \\ \xvec \* P\ \xvec \* M\ \xvec \* C\ S \distinctPerm p\ \xvec \* \\ \xvec \* P'''\ \\ = (p \ (\M\\*xvec\\N\))\ P'eq \(p \ xvec) \* (\M\\*xvec\\N\)\ \(p \ xvec) \* (P' \ P'')\ + have "Prop C \ (P \ !P) (p \ (\M\\*xvec\\N\)) (p \ (P' \ P''))" + by(intro rAlpha) simp+ + + with P'eq \\ = p \ (\M\\*xvec\\N\)\ \distinctPerm p\ show ?case by simp + next + case(cBang C \ P') + then show ?case by(auto intro: rBang) + qed +qed + +lemma brCommInAuxTooMuch: + fixes \ :: 'b + and \\<^sub>Q :: 'b + and R :: "('a, 'b, 'c) psi" + and M :: 'a + and N :: 'a + and R' :: "('a, 'b, 'c) psi" + and A\<^sub>R :: "name list" + and \\<^sub>R :: 'b + and A\<^sub>P :: "name list" + and \\<^sub>P :: 'b + and A\<^sub>Q :: "name list" + +assumes RTrans: "\ \ \\<^sub>Q \ R \\M\N\ \ R'" + and FrR: "extractFrame R = \A\<^sub>R, \\<^sub>R\" + and "distinct A\<^sub>R" + and QimpP: "\A\<^sub>Q, (\ \ \\<^sub>Q) \ \\<^sub>R\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\<^sub>R\" + and "A\<^sub>R \* A\<^sub>P" + and "A\<^sub>R \* A\<^sub>Q" + and "A\<^sub>R \* \" + and "A\<^sub>R \* \\<^sub>P" + and "A\<^sub>R \* \\<^sub>Q" + and "A\<^sub>R \* (\ \ \\<^sub>Q)" + and "A\<^sub>R \* R" + and "A\<^sub>R \* M" + and "A\<^sub>P \* R" + and "A\<^sub>P \* M" + and "A\<^sub>Q \* R" + and "A\<^sub>Q \* M" + +shows "\ \ \\<^sub>P \ R \\M\N\ \ R'" + using assms +proof(nominal_induct avoiding: A\<^sub>P \\<^sub>P A\<^sub>Q \\<^sub>Q \ rule: brinputFrameInduct) + case(cAlpha \' P M N P' A\<^sub>R \\<^sub>R p A\<^sub>P \\<^sub>P A\<^sub>Q \\<^sub>Q \) + have S: "set p \ set A\<^sub>R \ set (p \ A\<^sub>R)" by fact + from \\A\<^sub>Q, \' \ (p \ \\<^sub>R)\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ (p \ \\<^sub>R)\\ + have "(p \ \A\<^sub>Q, \' \ (p \ \\<^sub>R)\) \\<^sub>F (p \ \A\<^sub>P, (\ \ \\<^sub>P) \ (p \ \\<^sub>R)\)" + by(rule FrameStatImpClosed) + with \A\<^sub>R \* A\<^sub>P\ \(p \ A\<^sub>R) \* A\<^sub>P\ \A\<^sub>R \* \'\ \(p \ A\<^sub>R) \* \'\ \A\<^sub>R \* \\<^sub>P\ \(p \ A\<^sub>R) \* \\<^sub>P\ \A\<^sub>R \* A\<^sub>Q\ + \(p \ A\<^sub>R) \* A\<^sub>Q\ \A\<^sub>R \* \\ \(p \ A\<^sub>R) \* \\ S \distinctPerm p\ + have "\A\<^sub>Q, \' \ \\<^sub>R\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\<^sub>R\" by(simp add: eqvts) + moreover note \A\<^sub>R \* \'\ \A\<^sub>R \* \\ \A\<^sub>R \* P\ \A\<^sub>R \* \\<^sub>P\ \A\<^sub>R \* \\<^sub>Q\ \A\<^sub>R \* M\ \A\<^sub>P \* P\ + \A\<^sub>Q \* P\ \A\<^sub>P \* M\ \A\<^sub>Q \* M\ \A\<^sub>R \* A\<^sub>P\ \A\<^sub>R \* A\<^sub>Q\ + ultimately show ?case + by(elim cAlpha) +next + case(cBrInput \' M K xvec N Tvec P A\<^sub>P \\<^sub>P A\<^sub>Q \\<^sub>Q \) + from \A\<^sub>P \* (M\\*xvec N\.P)\ \A\<^sub>Q \* (M\\*xvec N\.P)\ + have "A\<^sub>P \* M" and "A\<^sub>Q \* M" by simp+ + from \\' \ K \ M\ + have "\' \ \ \ K \ M" + by(blast intro: statEqEnt Identity AssertionStatEqSym) + with \A\<^sub>Q \* K\ \A\<^sub>Q \* M\ + have "(\A\<^sub>Q, \' \ \\) \\<^sub>F K \ M" + by(force intro: frameImpI) + with \\A\<^sub>Q, \' \ \\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\\ + have "(\A\<^sub>P, (\ \ \\<^sub>P) \ \\) \\<^sub>F K \ M" + by(simp add: FrameStatImp_def) + with \A\<^sub>P \* K\ \A\<^sub>P \* M\ have "(\ \ \\<^sub>P) \ \ \ K \ M" + by(force dest: frameImpE) + then have "\ \ \\<^sub>P \ K \ M" by(blast intro: statEqEnt Identity) + then show ?case using \distinct xvec\ \set xvec \ supp N\ \length xvec = length Tvec\ + by(rule BrInput) +next + case(cCase \' P M N P' \ Cs A\<^sub>R \\<^sub>R A\<^sub>Q \\<^sub>Q A\<^sub>P \\<^sub>P \) + from \A\<^sub>P \* (Cases Cs)\ \A\<^sub>Q \* (Cases Cs)\ \(\, P) \ set Cs\ + have "A\<^sub>P \* P" and "A\<^sub>Q \* P" and "A\<^sub>P \* \" and "A\<^sub>Q \* \" + by(auto dest: memFreshChain) + + from \\' \ \\ + have "\' \ \ \ \" + by(blast intro: statEqEnt Identity AssertionStatEqSym) + with \A\<^sub>P \* \\ + have "(\A\<^sub>P, \' \ \\) \\<^sub>F \" + by(force intro: frameImpI) + with \\A\<^sub>P, \' \ \\ \\<^sub>F \A\<^sub>Q, (\ \ \\<^sub>Q) \ \\\ + have "(\A\<^sub>Q, (\ \ \\<^sub>Q) \ \\) \\<^sub>F \" + by(simp add: FrameStatImp_def) + with \A\<^sub>Q \* \\ have "(\ \ \\<^sub>Q) \ \ \ \" + by(force dest: frameImpE) + then have "\ \ \\<^sub>Q \ \" by(blast intro: statEqEnt Identity) + + have "\A\<^sub>P, \' \ \\<^sub>R\ \\<^sub>F \A\<^sub>Q, (\ \ \\<^sub>Q) \ \\<^sub>R\" + proof - + from \\\<^sub>R \ \\ have "\A\<^sub>P, \' \ \\<^sub>R\ \\<^sub>F \A\<^sub>P, \' \ \\" + by(metis Identity Commutativity AssertionStatEqSym Composition frameResChainPres frameNilStatEq AssertionStatEqTrans) + moreover note \\A\<^sub>P, \' \ \\ \\<^sub>F \A\<^sub>Q, (\ \ \\<^sub>Q) \ \\\ + moreover from \\\<^sub>R \ \\ have "\A\<^sub>Q, (\ \ \\<^sub>Q) \ \\ \\<^sub>F \A\<^sub>Q, (\ \ \\<^sub>Q) \ \\<^sub>R\" + by(metis Identity Commutativity AssertionStatEqSym Composition frameResChainPres frameNilStatEq AssertionStatEqTrans) + ultimately show ?thesis by(rule FrameStatEqImpCompose) + qed + moreover note \A\<^sub>R \* \'\ \A\<^sub>R \* \\ \A\<^sub>R \* P\ \A\<^sub>R \* \\<^sub>Q\ \A\<^sub>R \* M\ + \A\<^sub>Q \* P\ \A\<^sub>P \* P\ \A\<^sub>Q \* M\ \A\<^sub>P \* M\ \A\<^sub>R \* A\<^sub>Q\ \A\<^sub>R \* A\<^sub>P\ + ultimately have "\ \ \\<^sub>Q \ P \ \M\N\ \ P'" + by(elim cCase(4)) + moreover note \(\, P) \ set Cs\ \\ \ \\<^sub>Q \ \\ \guarded P\ + ultimately show ?case + by(rule Case) +next + case(cPar1 \' \\<^sub>Q P M N P' A\<^sub>Q Q A\<^sub>P \\<^sub>P A\<^sub>P' \\<^sub>P' A\<^sub>Q' \\<^sub>Q' \) + from \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ \A\<^sub>Q \* A\<^sub>P'\ \A\<^sub>P' \* (P \ Q)\ have "A\<^sub>P' \* \\<^sub>Q" + by(force dest: extractFrameFreshChain) + + have "\A\<^sub>P', (\ \ \\<^sub>P') \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>P', (\ \ \\<^sub>P') \ \\<^sub>Q \ \\<^sub>P\" + by (metis Commutativity FrameStatEq_def frameIntCompositionSym) + moreover have "\A\<^sub>P', (\ \ \\<^sub>P') \ \\<^sub>Q \ \\<^sub>P\ \\<^sub>F \A\<^sub>P', ((\ \ \\<^sub>P') \ \\<^sub>Q) \ \\<^sub>P\" + by (metis FrameStatEq_def frameIntAssociativity) + moreover have "\A\<^sub>P', ((\ \ \\<^sub>P') \ \\<^sub>Q) \ \\<^sub>P\ \\<^sub>F \A\<^sub>P', ((\ \ \\<^sub>Q) \ \\<^sub>P') \ \\<^sub>P\" + by (metis FrameStatEq_def associativitySym frameIntComposition) + ultimately have right: "\A\<^sub>P', (\ \ \\<^sub>P') \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>P', ((\ \ \\<^sub>Q) \ \\<^sub>P') \ \\<^sub>P\" + by (metis FrameStatImpTrans) + have "\A\<^sub>Q', \' \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>Q', \' \ \\<^sub>Q \ \\<^sub>P\" + by (metis AssertionStatEq_def Commutativity compositionSym frameImpNilStatEq frameImpResChainPres) + moreover have "\A\<^sub>Q', \' \ \\<^sub>Q \ \\<^sub>P\ \\<^sub>F \A\<^sub>Q', (\' \ \\<^sub>Q) \ \\<^sub>P\" + by (metis FrameStatEq_def frameIntAssociativity) + ultimately have left: "\A\<^sub>Q', \' \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>Q', (\' \ \\<^sub>Q) \ \\<^sub>P\" + by (metis FrameStatImpTrans) + from \\A\<^sub>Q', \' \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>P', (\ \ \\<^sub>P') \ \\<^sub>P \ \\<^sub>Q\\ left right + have "\A\<^sub>Q', (\' \ \\<^sub>Q) \ \\<^sub>P\ \\<^sub>F \A\<^sub>P', ((\ \ \\<^sub>Q) \ \\<^sub>P') \ \\<^sub>P\" + by (metis AssertionStatEqSym AssertionStatEqTrans AssertionStatEq_def Associativity FrameStatImpTrans associativitySym frameImpNilStatEq frameImpResChainPres) + moreover note \A\<^sub>P \* \'\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>P \* \\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>P \* P\ \A\<^sub>P \* M\ + \A\<^sub>P \* \\<^sub>P'\ \A\<^sub>P' \* (P \ Q)\ \A\<^sub>Q' \* (P \ Q)\ \A\<^sub>P' \* \\<^sub>Q\ \A\<^sub>P' \* M\ \A\<^sub>Q' \* M\ \A\<^sub>P \* A\<^sub>P'\ \A\<^sub>P \* A\<^sub>Q'\ + ultimately have "(\ \ \\<^sub>Q) \ \\<^sub>P' \ P \ \M\N\ \ P'" + by(elim cPar1(6)) (simp | force)+ + then have "(\ \ \\<^sub>P') \ \\<^sub>Q \ P \ \M\N\ \ P'" + by (metis associativitySym statEqTransition) + + moreover note \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ \A\<^sub>Q \* \\ \A\<^sub>Q \* \\<^sub>P'\ + \A\<^sub>Q \* P\ \A\<^sub>Q \* M\ \A\<^sub>Q \* N\ + ultimately show ?case + by(elim Par1) (simp | force)+ +next + case(cPar2 \' \\<^sub>P Q M N Q' A\<^sub>P P A\<^sub>Q \\<^sub>Q A\<^sub>P' \\<^sub>P' A\<^sub>Q' \\<^sub>Q' \) + from \extractFrame P = \A\<^sub>P, \\<^sub>P\\ \A\<^sub>P \* A\<^sub>P'\ \A\<^sub>P' \* (P \ Q)\ have "A\<^sub>P' \* \\<^sub>P" + by(force dest: extractFrameFreshChain) + have "\A\<^sub>P', (\ \ \\<^sub>P') \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>P', ((\ \ \\<^sub>P') \ \\<^sub>P) \ \\<^sub>Q\" + by (metis FrameStatEq_def frameIntAssociativity) + moreover have "\A\<^sub>P', ((\ \ \\<^sub>P') \ \\<^sub>P) \ \\<^sub>Q\ \\<^sub>F \A\<^sub>P', ((\ \ \\<^sub>P) \ \\<^sub>P') \ \\<^sub>Q\" + by (metis FrameStatEq_def associativitySym frameIntComposition) + ultimately have right: "\A\<^sub>P', (\ \ \\<^sub>P') \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>P', ((\ \ \\<^sub>P) \ \\<^sub>P') \ \\<^sub>Q\" + by (metis FrameStatImpTrans) + moreover have left: "\A\<^sub>Q', \' \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>Q', (\' \ \\<^sub>P) \ \\<^sub>Q\" + by (metis FrameStatEq_def frameIntAssociativity) + from \\A\<^sub>Q', \' \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>P', (\ \ \\<^sub>P') \ \\<^sub>P \ \\<^sub>Q\\ left right + have "\A\<^sub>Q', (\' \ \\<^sub>P) \ \\<^sub>Q\ \\<^sub>F \A\<^sub>P', ((\ \ \\<^sub>P) \ \\<^sub>P') \ \\<^sub>Q\" + by (metis FrameStatEq_def FrameStatImpTrans frameIntAssociativity) + moreover note \A\<^sub>Q \* \'\ \A\<^sub>Q \* \\<^sub>P\ \A\<^sub>Q \* \\ \A\<^sub>Q \* \\<^sub>P\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* M\ + \A\<^sub>Q \* \\<^sub>P'\ \A\<^sub>P' \* (P \ Q)\ \A\<^sub>Q' \* (P \ Q)\ \A\<^sub>P' \* \\<^sub>P\ \A\<^sub>P' \* M\ \A\<^sub>Q' \* M\ \A\<^sub>Q \* A\<^sub>Q'\ \A\<^sub>Q \* A\<^sub>P'\ + ultimately have "(\ \ \\<^sub>P) \ \\<^sub>P' \ Q \ \M\N\ \ Q'" + by(elim cPar2(6)) (simp | force)+ + then have "(\ \ \\<^sub>P') \ \\<^sub>P \ Q \ \M\N\ \ Q'" + by (metis associativitySym statEqTransition) + moreover note \extractFrame P = \A\<^sub>P, \\<^sub>P\\ \A\<^sub>P \* \\ \A\<^sub>P \* \\<^sub>P'\ + \A\<^sub>P \* Q\ \A\<^sub>P \* M\ \A\<^sub>P \* N\ + ultimately show ?case + by(elim Par2) (simp | force)+ +next + case(cBrMerge \' \\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q Q' A\<^sub>Q A\<^sub>P' \\<^sub>P' A\<^sub>Q' \\<^sub>Q' \) + from \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ \A\<^sub>Q \* A\<^sub>P'\ \A\<^sub>P' \* (P \ Q)\ have "A\<^sub>P' \* \\<^sub>Q" + by(force dest: extractFrameFreshChain) + from \extractFrame P = \A\<^sub>P, \\<^sub>P\\ \A\<^sub>P \* A\<^sub>P'\ \A\<^sub>P' \* (P \ Q)\ have "A\<^sub>P' \* \\<^sub>P" + by(force dest: extractFrameFreshChain) + have "\A\<^sub>P', (\ \ \\<^sub>P') \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>P', (\ \ \\<^sub>P') \ \\<^sub>Q \ \\<^sub>P\" + by (metis Commutativity FrameStatEq_def frameIntCompositionSym) + moreover have "\A\<^sub>P', (\ \ \\<^sub>P') \ \\<^sub>Q \ \\<^sub>P\ \\<^sub>F \A\<^sub>P', ((\ \ \\<^sub>P') \ \\<^sub>Q) \ \\<^sub>P\" + by (metis FrameStatEq_def frameIntAssociativity) + moreover have "\A\<^sub>P', ((\ \ \\<^sub>P') \ \\<^sub>Q) \ \\<^sub>P\ \\<^sub>F \A\<^sub>P', ((\ \ \\<^sub>Q) \ \\<^sub>P') \ \\<^sub>P\" + by (metis FrameStatEq_def associativitySym frameIntComposition) + ultimately have right: "\A\<^sub>P', (\ \ \\<^sub>P') \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>P', ((\ \ \\<^sub>Q) \ \\<^sub>P') \ \\<^sub>P\" + by (metis FrameStatImpTrans) + have "\A\<^sub>Q', \' \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>Q', \' \ \\<^sub>Q \ \\<^sub>P\" + by (metis AssertionStatEq_def Commutativity compositionSym frameImpNilStatEq frameImpResChainPres) + moreover have "\A\<^sub>Q', \' \ \\<^sub>Q \ \\<^sub>P\ \\<^sub>F \A\<^sub>Q', (\' \ \\<^sub>Q) \ \\<^sub>P\" + by (metis FrameStatEq_def frameIntAssociativity) + ultimately have left: "\A\<^sub>Q', \' \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>Q', (\' \ \\<^sub>Q) \ \\<^sub>P\" + by (metis FrameStatImpTrans) + from \\A\<^sub>Q', \' \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>P', (\ \ \\<^sub>P') \ \\<^sub>P \ \\<^sub>Q\\ left right + have "\A\<^sub>Q', (\' \ \\<^sub>Q) \ \\<^sub>P\ \\<^sub>F \A\<^sub>P', ((\ \ \\<^sub>Q) \ \\<^sub>P') \ \\<^sub>P\" + by (metis AssertionStatEqSym AssertionStatEqTrans AssertionStatEq_def Associativity FrameStatImpTrans associativitySym frameImpNilStatEq frameImpResChainPres) + moreover note \A\<^sub>P \* \'\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>P \* \\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>P \* P\ \A\<^sub>P \* M\ + \A\<^sub>P \* \\<^sub>P'\ \A\<^sub>P' \* (P \ Q)\ \A\<^sub>Q' \* (P \ Q)\ \A\<^sub>P' \* \\<^sub>Q\ \A\<^sub>P' \* M\ \A\<^sub>Q' \* M\ \A\<^sub>P \* A\<^sub>P'\ \A\<^sub>P \* A\<^sub>Q'\ + ultimately have "(\ \ \\<^sub>Q) \ \\<^sub>P' \ P \ \M\N\ \ P'" + by(elim cBrMerge(2)) (simp | force)+ + then have Ptrans: "(\ \ \\<^sub>P') \ \\<^sub>Q \ P \ \M\N\ \ P'" + by (metis associativitySym statEqTransition) + + have left2: "\A\<^sub>Q', \' \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>Q', (\' \ \\<^sub>P) \ \\<^sub>Q\" + by (metis FrameStatEq_def frameIntAssociativity) + have "\A\<^sub>P', (\ \ \\<^sub>P') \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>P', ((\ \ \\<^sub>P') \ \\<^sub>P) \ \\<^sub>Q\" + by (metis FrameStatEq_def frameIntAssociativity) + moreover have "\A\<^sub>P', ((\ \ \\<^sub>P') \ \\<^sub>P) \ \\<^sub>Q\ \\<^sub>F \A\<^sub>P', ((\ \ \\<^sub>P) \ \\<^sub>P') \ \\<^sub>Q\" + by (metis FrameStatEq_def associativitySym frameIntComposition) + ultimately have right2: "\A\<^sub>P', (\ \ \\<^sub>P') \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>P', ((\ \ \\<^sub>P) \ \\<^sub>P') \ \\<^sub>Q\" + by (metis FrameStatImpTrans) + from \\A\<^sub>Q', \' \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>P', (\ \ \\<^sub>P') \ \\<^sub>P \ \\<^sub>Q\\ left2 right2 + have "\A\<^sub>Q', (\' \ \\<^sub>P) \ \\<^sub>Q\ \\<^sub>F \A\<^sub>P', ((\ \ \\<^sub>P) \ \\<^sub>P') \ \\<^sub>Q\" + by (metis FrameStatEq_def FrameStatImpTrans frameIntAssociativity) + moreover note \A\<^sub>Q \* \'\ \A\<^sub>Q \* \\<^sub>P\ \A\<^sub>Q \* \\ \A\<^sub>Q \* \\<^sub>P\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* M\ + \A\<^sub>Q \* \\<^sub>P'\ \A\<^sub>P' \* (P \ Q)\ \A\<^sub>Q' \* (P \ Q)\ \A\<^sub>P' \* \\<^sub>P\ \A\<^sub>P' \* M\ \A\<^sub>Q' \* M\ \A\<^sub>Q \* A\<^sub>Q'\ \A\<^sub>Q \* A\<^sub>P'\ + ultimately have "(\ \ \\<^sub>P) \ \\<^sub>P' \ Q \ \M\N\ \ Q'" + by(elim cBrMerge(6)) (simp | force)+ + then have Qtrans: "(\ \ \\<^sub>P') \ \\<^sub>P \ Q \ \M\N\ \ Q'" + by (metis associativitySym statEqTransition) + + from Ptrans \extractFrame P = \A\<^sub>P, \\<^sub>P\\ Qtrans \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ + \A\<^sub>P \* \\ \A\<^sub>P \* \\<^sub>P'\ \A\<^sub>P \* P\ \A\<^sub>P \* Q\ \A\<^sub>P \* M\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>Q \* \\ \A\<^sub>Q \* \\<^sub>P'\ \A\<^sub>Q \* P\ + \A\<^sub>Q \* Q\ \A\<^sub>Q \* M\ + show ?case + by(elim BrMerge) (simp | force)+ +next + case(cScope \' P M N P' x A\<^sub>P \\<^sub>P A\<^sub>P' \\<^sub>P' A\<^sub>Q \\<^sub>Q \) + then have "\ \ \\<^sub>P' \ P \ \M\N\ \ P'" + by simp + with \x \ \\ \x \ \\<^sub>P'\ \x \ M\ \x \ N\ + show ?case + by(elim Scope) (simp | force)+ +next + case(cBang \' P M N P' A\<^sub>R \\<^sub>R A\<^sub>P \\<^sub>P A\<^sub>Q \\<^sub>Q \) + from \A\<^sub>R \* P\ have "A\<^sub>R \* (P \ !P)" + by simp + from \A\<^sub>P \* !P\ have "A\<^sub>P \* (P \ !P)" + by simp + from \A\<^sub>Q \* !P\ have "A\<^sub>Q \* (P \ !P)" + by simp + + have "\A\<^sub>Q, \' \ \\<^sub>R \ \\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\<^sub>R \ \\" + proof - + from \\\<^sub>R \ \\ have "\A\<^sub>Q, \' \ \\<^sub>R \ \\ \\<^sub>F \A\<^sub>Q, \' \ \\" + by(metis Identity Commutativity AssertionStatEqSym Composition frameResChainPres frameNilStatEq AssertionStatEqTrans) + moreover note \\A\<^sub>Q, \' \ \\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\\ + moreover from \\\<^sub>R \ \\ have "\A\<^sub>P, (\ \ \\<^sub>P) \ \\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\<^sub>R \ \\" + by(metis Identity Commutativity AssertionStatEqSym Composition frameResChainPres frameNilStatEq AssertionStatEqTrans) + ultimately show ?thesis by(rule FrameStatEqImpCompose) + qed + then have "\ \ \\<^sub>P \ P \ !P \ \M\N\ \ P'" + using \A\<^sub>R \* \'\ \A\<^sub>R \* \\ \A\<^sub>R \* (P \ !P)\ \A\<^sub>R \* \\<^sub>P\ + \A\<^sub>P \* (P \ !P)\ \A\<^sub>Q \* (P \ !P)\ \A\<^sub>P \* M\ \A\<^sub>Q \* M\ \A\<^sub>R \* M\ \A\<^sub>R \* A\<^sub>P\ \A\<^sub>R \* A\<^sub>Q\ + by(elim cBang(5)) + then show ?case using \guarded P\ + by(rule Bang) +qed + +lemma brCommInAux: + fixes \ :: 'b + and \\<^sub>Q :: 'b + and R :: "('a, 'b, 'c) psi" + and M :: 'a + and N :: 'a + and R' :: "('a, 'b, 'c) psi" + and A\<^sub>R :: "name list" + and \\<^sub>R :: 'b + and A\<^sub>P :: "name list" + and \\<^sub>P :: 'b + and A\<^sub>Q :: "name list" + +assumes RTrans: "\ \ \\<^sub>Q \ R \\M\N\ \ R'" + and FrR: "extractFrame R = \A\<^sub>R, \\<^sub>R\" + and "distinct A\<^sub>R" + and QimpP: "\A\<^sub>Q, (\ \ \\<^sub>Q) \ \\<^sub>R\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\<^sub>R\" + and "A\<^sub>R \* A\<^sub>P" + and "A\<^sub>R \* A\<^sub>Q" + and "A\<^sub>R \* \" + and "A\<^sub>R \* \\<^sub>P" + and "A\<^sub>R \* \\<^sub>Q" + and "A\<^sub>R \* R" + and "A\<^sub>R \* M" + and "A\<^sub>P \* R" + and "A\<^sub>P \* M" + and "A\<^sub>Q \* R" + and "A\<^sub>Q \* M" + +shows "\ \ \\<^sub>P \ R \\M\N\ \ R'" +proof - + from \A\<^sub>R \* \\ \A\<^sub>R \* \\<^sub>Q\ + have "A\<^sub>R \* (\ \ \\<^sub>Q)" + by auto + with assms + show ?thesis + by(simp add: brCommInAuxTooMuch) +qed + +lemma brCommOutAuxTooMuch: + fixes \ :: 'b + and \\<^sub>Q :: 'b + and R :: "('a, 'b, 'c) psi" + and M :: 'a + and xvec :: "name list" + and N :: 'a + and R' :: "('a, 'b, 'c) psi" + and A\<^sub>R :: "name list" + and \\<^sub>R :: 'b + and A\<^sub>P :: "name list" + and \\<^sub>P :: 'b + and A\<^sub>Q :: "name list" + +assumes RTrans: "\ \ \\<^sub>Q \ R \ RBrOut M (\\*xvec\N \' R')" + and FrR: "extractFrame R = \A\<^sub>R, \\<^sub>R\" + and "distinct A\<^sub>R" + and QimpP: "\A\<^sub>Q, (\ \ \\<^sub>Q) \ \\<^sub>R\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\<^sub>R\" + and "A\<^sub>R \* A\<^sub>P" + and "A\<^sub>R \* A\<^sub>Q" + and "A\<^sub>R \* \" + and "A\<^sub>R \* \\<^sub>P" + and "A\<^sub>R \* \\<^sub>Q" + and "A\<^sub>R \* (\ \ \\<^sub>Q)" + and "A\<^sub>R \* R" + and "A\<^sub>R \* M" + and "A\<^sub>P \* R" + and "A\<^sub>P \* M" + and "A\<^sub>Q \* R" + and "A\<^sub>Q \* M" + +shows "\ \ \\<^sub>P \ R \ \M\\*xvec\\N\ \ R'" + using assms +proof(nominal_induct R M B=="\\*xvec\N \' R'" A\<^sub>R \\<^sub>R avoiding: \ A\<^sub>P \\<^sub>P A\<^sub>Q \\<^sub>Q N R' xvec rule: broutputFrameInduct) + case (cAlpha \' P M A\<^sub>R \\<^sub>R p \ A\<^sub>P \\<^sub>P A\<^sub>Q \\<^sub>Q N P' xvec) + have S: "set p \ set A\<^sub>R \ set (p \ A\<^sub>R)" by fact + from \\A\<^sub>Q, \' \ (p \ \\<^sub>R)\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ (p \ \\<^sub>R)\\ + have "(p \ \A\<^sub>Q, \' \ (p \ \\<^sub>R)\) \\<^sub>F (p \ \A\<^sub>P, (\ \ \\<^sub>P) \ (p \ \\<^sub>R)\)" + by(rule FrameStatImpClosed) + with \A\<^sub>R \* A\<^sub>P\ \(p \ A\<^sub>R) \* A\<^sub>P\ \A\<^sub>R \* \'\ \(p \ A\<^sub>R) \* \'\ \A\<^sub>R \* \\<^sub>P\ \(p \ A\<^sub>R) \* \\<^sub>P\ \A\<^sub>R \* A\<^sub>Q\ + \(p \ A\<^sub>R) \* A\<^sub>Q\ \A\<^sub>R \* \\ \(p \ A\<^sub>R) \* \\ S \distinctPerm p\ + have "\A\<^sub>Q, \' \ \\<^sub>R\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\<^sub>R\" by(simp add: eqvts) + moreover note \A\<^sub>R \* \'\ \A\<^sub>R \* \\ \A\<^sub>R \* P\ \A\<^sub>R \* \\<^sub>P\ \A\<^sub>R \* \\<^sub>Q\ \A\<^sub>R \* M\ \A\<^sub>P \* P\ + \A\<^sub>Q \* P\ \A\<^sub>P \* M\ \A\<^sub>Q \* M\ \A\<^sub>R \* A\<^sub>P\ \A\<^sub>R \* A\<^sub>Q\ + ultimately show ?case + by(elim cAlpha) +next + case(cBrOutput \' M K N P \ A\<^sub>P \\<^sub>P A\<^sub>Q \\<^sub>Q N' R' xvec) + from \A\<^sub>P \* (M\N\.P)\ \A\<^sub>Q \* (M\N\.P)\ + have "A\<^sub>P \* M" and "A\<^sub>Q \* M" by simp+ + from \\' \ M \ K\ + have "\' \ \ \ M \ K" + by(blast intro: statEqEnt Identity AssertionStatEqSym) + with \A\<^sub>Q \* K\ \A\<^sub>Q \* M\ + have "(\A\<^sub>Q, \' \ \\) \\<^sub>F M \ K" + by(force intro: frameImpI) + with \\A\<^sub>Q, \' \ \\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\\ + have "(\A\<^sub>P, (\ \ \\<^sub>P) \ \\) \\<^sub>F M \ K" + by(simp add: FrameStatImp_def) + with \A\<^sub>P \* K\ \A\<^sub>P \* M\ have "(\ \ \\<^sub>P) \ \ \ M \ K" + by(force dest: frameImpE) + then have "\ \ \\<^sub>P \ M \ K" by(blast intro: statEqEnt Identity) + then have "\ \ \\<^sub>P \ M\N\.P \ \K\N\ \ P" + by(rule BrOutput) + with \N \' P = \\*xvec\N' \' R'\ + show ?case + by(simp add: residualInject) +next + case(cCase \' R M \ Cs A\<^sub>R \\<^sub>R \ A\<^sub>P \\<^sub>P A\<^sub>Q \\<^sub>Q N R' xvec) + from \A\<^sub>P \* (Cases Cs)\ \A\<^sub>Q \* (Cases Cs)\ \(\, R) \ set Cs\ + have "A\<^sub>P \* R" and "A\<^sub>Q \* R" and "A\<^sub>P \* \" and "A\<^sub>Q \* \" + by(auto dest: memFreshChain) + + from \\' \ \\ have "\' \ \ \ \" by(blast intro: statEqEnt Identity AssertionStatEqSym) + with \A\<^sub>Q \* \\ have "(\A\<^sub>Q, \' \ \\) \\<^sub>F \" by(force intro: frameImpI) + with \\A\<^sub>Q, \' \ \\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\\ have "(\A\<^sub>P, (\ \ \\<^sub>P) \ \\) \\<^sub>F \" + by(simp add: FrameStatImp_def) + with \A\<^sub>P \* \\ have "(\ \ \\<^sub>P) \ \ \ \" by(force dest: frameImpE) + then have "\ \ \\<^sub>P \ \" by(blast intro: statEqEnt Identity) + + have "\A\<^sub>Q, \' \ \\<^sub>R\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\<^sub>R\" + proof - + from \\\<^sub>R \ \\ have "\A\<^sub>Q, \' \ \\<^sub>R\ \\<^sub>F \A\<^sub>Q, \' \ \\" + by(metis Identity Commutativity AssertionStatEqSym Composition frameResChainPres frameNilStatEq AssertionStatEqTrans) + moreover note \\A\<^sub>Q, \' \ \\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\\ + moreover from \\\<^sub>R \ \\ have "\A\<^sub>P, (\ \ \\<^sub>P) \ \\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\<^sub>R\" + by(metis Identity Commutativity AssertionStatEqSym Composition frameResChainPres frameNilStatEq AssertionStatEqTrans) + ultimately show ?thesis by(rule FrameStatEqImpCompose) + qed + with \A\<^sub>P \* R\ \A\<^sub>Q \* R\ + have "\ \ \\<^sub>P \ R \ \M\\*xvec\\N\ \ R'" using cCase + by(intro cCase(4)) simp+ + then show ?case using \(\, R) \ set Cs\ \\ \ \\<^sub>P \ \\ \guarded R\ + by(rule Case) +next + case(cPar1 \' \\<^sub>Q P M xvec N P' A\<^sub>Q Q A\<^sub>P \\<^sub>P \ A\<^sub>P' \\<^sub>P' A\<^sub>Q' \\<^sub>Q' N' R' yvec) + from \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ \A\<^sub>Q \* A\<^sub>P'\ \A\<^sub>P' \* (P \ Q)\ have "A\<^sub>P' \* \\<^sub>Q" + by(force dest: extractFrameFreshChain) + + have "\A\<^sub>P', (\ \ \\<^sub>P') \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>P', (\ \ \\<^sub>P') \ \\<^sub>Q \ \\<^sub>P\" + by (metis Commutativity FrameStatEq_def frameIntCompositionSym) + moreover have "\A\<^sub>P', (\ \ \\<^sub>P') \ \\<^sub>Q \ \\<^sub>P\ \\<^sub>F \A\<^sub>P', ((\ \ \\<^sub>P') \ \\<^sub>Q) \ \\<^sub>P\" + by (metis FrameStatEq_def frameIntAssociativity) + moreover have "\A\<^sub>P', ((\ \ \\<^sub>P') \ \\<^sub>Q) \ \\<^sub>P\ \\<^sub>F \A\<^sub>P', ((\ \ \\<^sub>Q) \ \\<^sub>P') \ \\<^sub>P\" + by (metis FrameStatEq_def associativitySym frameIntComposition) + ultimately have right: "\A\<^sub>P', (\ \ \\<^sub>P') \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>P', ((\ \ \\<^sub>Q) \ \\<^sub>P') \ \\<^sub>P\" + by (metis FrameStatImpTrans) + have "\A\<^sub>Q', \' \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>Q', \' \ \\<^sub>Q \ \\<^sub>P\" + by (metis AssertionStatEq_def Commutativity compositionSym frameImpNilStatEq frameImpResChainPres) + moreover have "\A\<^sub>Q', \' \ \\<^sub>Q \ \\<^sub>P\ \\<^sub>F \A\<^sub>Q', (\' \ \\<^sub>Q) \ \\<^sub>P\" + by (metis FrameStatEq_def frameIntAssociativity) + ultimately have left: "\A\<^sub>Q', \' \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>Q', (\' \ \\<^sub>Q) \ \\<^sub>P\" + by (metis FrameStatImpTrans) + from \\A\<^sub>Q', \' \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>P', (\ \ \\<^sub>P') \ \\<^sub>P \ \\<^sub>Q\\ left right + have "\A\<^sub>Q', (\' \ \\<^sub>Q) \ \\<^sub>P\ \\<^sub>F \A\<^sub>P', ((\ \ \\<^sub>Q) \ \\<^sub>P') \ \\<^sub>P\" + by (metis AssertionStatEqSym AssertionStatEqTrans AssertionStatEq_def Associativity FrameStatImpTrans associativitySym frameImpNilStatEq frameImpResChainPres) + moreover note \A\<^sub>P \* \'\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>P \* \\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>P \* P\ \A\<^sub>P \* M\ + \A\<^sub>P \* \\<^sub>P'\ \A\<^sub>P' \* (P \ Q)\ \A\<^sub>Q' \* (P \ Q)\ \A\<^sub>P' \* \\<^sub>Q\ \A\<^sub>P' \* M\ \A\<^sub>Q' \* M\ \A\<^sub>P \* A\<^sub>P'\ \A\<^sub>P \* A\<^sub>Q'\ + ultimately have "(\ \ \\<^sub>Q) \ \\<^sub>P' \ P \\M\\*xvec\\N\ \ P'" + by(intro cPar1(6)) (simp | force)+ + then have "(\ \ \\<^sub>P') \ \\<^sub>Q \ P \\M\\*xvec\\N\ \ P'" + by (metis associativitySym statEqTransition) + moreover note \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ \xvec \* Q\ \A\<^sub>Q \* \\ \A\<^sub>Q \* \\<^sub>P'\ + \A\<^sub>Q \* P\ \A\<^sub>Q \* M\ \A\<^sub>Q \* xvec\ \A\<^sub>Q \* N\ + ultimately have "\ \ \\<^sub>P' \ P \ Q \\M\\*xvec\\N\ \ P' \ Q" + by(elim Par1) (simp | force)+ + then show ?case using \\\*xvec\N \' (P' \ Q) = \\*yvec\N' \' R'\ + by(simp add: residualInject) +next + case(cPar2 \' \\<^sub>P Q M xvec N Q' A\<^sub>P P A\<^sub>Q \\<^sub>Q \ A\<^sub>P' \\<^sub>P' A\<^sub>Q' \\<^sub>Q' N' R' yvec) + from \extractFrame P = \A\<^sub>P, \\<^sub>P\\ \A\<^sub>P \* A\<^sub>P'\ \A\<^sub>P' \* (P \ Q)\ have "A\<^sub>P' \* \\<^sub>P" + by(force dest: extractFrameFreshChain) + have "\A\<^sub>P', (\ \ \\<^sub>P') \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>P', ((\ \ \\<^sub>P') \ \\<^sub>P) \ \\<^sub>Q\" + by (metis FrameStatEq_def frameIntAssociativity) + moreover have "\A\<^sub>P', ((\ \ \\<^sub>P') \ \\<^sub>P) \ \\<^sub>Q\ \\<^sub>F \A\<^sub>P', ((\ \ \\<^sub>P) \ \\<^sub>P') \ \\<^sub>Q\" + by (metis FrameStatEq_def associativitySym frameIntComposition) + ultimately have right: "\A\<^sub>P', (\ \ \\<^sub>P') \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>P', ((\ \ \\<^sub>P) \ \\<^sub>P') \ \\<^sub>Q\" + by (metis FrameStatImpTrans) + moreover have left: "\A\<^sub>Q', \' \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>Q', (\' \ \\<^sub>P) \ \\<^sub>Q\" + by (metis FrameStatEq_def frameIntAssociativity) + from \\A\<^sub>Q', \' \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>P', (\ \ \\<^sub>P') \ \\<^sub>P \ \\<^sub>Q\\ left right + have "\A\<^sub>Q', (\' \ \\<^sub>P) \ \\<^sub>Q\ \\<^sub>F \A\<^sub>P', ((\ \ \\<^sub>P) \ \\<^sub>P') \ \\<^sub>Q\" + by (metis FrameStatEq_def FrameStatImpTrans frameIntAssociativity) + moreover note \A\<^sub>Q \* \'\ \A\<^sub>Q \* \\<^sub>P\ \A\<^sub>Q \* \\ \A\<^sub>Q \* \\<^sub>P\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* M\ + \A\<^sub>Q \* \\<^sub>P'\ \A\<^sub>P' \* (P \ Q)\ \A\<^sub>Q' \* (P \ Q)\ \A\<^sub>P' \* \\<^sub>P\ \A\<^sub>P' \* M\ \A\<^sub>Q' \* M\ \A\<^sub>Q \* A\<^sub>Q'\ \A\<^sub>Q \* A\<^sub>P'\ + ultimately have "(\ \ \\<^sub>P) \ \\<^sub>P' \ Q \\M\\*xvec\\N\ \ Q'" + by(intro cPar2(6)) (simp | force)+ + then have "(\ \ \\<^sub>P') \ \\<^sub>P \ Q \\M\\*xvec\\N\ \ Q'" + by (metis associativitySym statEqTransition) + moreover note \extractFrame P = \A\<^sub>P, \\<^sub>P\\ \xvec \* P\ \A\<^sub>P \* \\ \A\<^sub>P \* \\<^sub>P'\ + \A\<^sub>P \* Q\ \A\<^sub>P \* M\ \A\<^sub>P \* xvec\ \A\<^sub>P \* N\ + ultimately have "\ \ \\<^sub>P' \ P \ Q \\M\\*xvec\\N\ \ P \ Q'" + by(elim Par2) (simp | force)+ + then show ?case using \\\*xvec\N \' (P \ Q') = \\*yvec\N' \' R'\ + by(simp add: residualInject) +next + case(cBrComm1 \' \\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q xvec Q' A\<^sub>Q \ A\<^sub>P' \\<^sub>P' A\<^sub>Q' \\<^sub>Q' N' R' yvec) + have right: "\A\<^sub>P', (\ \ \\<^sub>P') \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>P', (\\<^sub>Q \ (\ \ \\<^sub>P')) \ \\<^sub>P\" + by (metis AssertionStatEqTrans AssertionStatEq_def Associativity Commutativity FrameStatImpTrans frameImpNilStatEq frameImpResChainPres) + with \\A\<^sub>Q', \' \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>P', (\ \ \\<^sub>P') \ \\<^sub>P \ \\<^sub>Q\\ + have "\A\<^sub>Q', (\\<^sub>Q \ \') \ \\<^sub>P\ \\<^sub>F \A\<^sub>P', (\\<^sub>Q \ (\ \ \\<^sub>P')) \ \\<^sub>P\" + by (metis AssertionStatEqTrans AssertionStatEq_def Associativity Commutativity FrameStatImpTrans frameImpNilStatEq frameImpResChainPres) + moreover from \\' \ \\<^sub>Q \ P \ \M\N\ \ P'\ + have "\\<^sub>Q \ \' \ P \ \M\N\ \ P'" + by (metis Commutativity statEqTransition) + moreover note \extractFrame P = \A\<^sub>P, \\<^sub>P\\ \distinct A\<^sub>P\ \A\<^sub>P \* A\<^sub>P'\ + \A\<^sub>P \* A\<^sub>Q'\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>P \* \\ \A\<^sub>P \* \\<^sub>P'\ \A\<^sub>P \* \'\ \A\<^sub>P \* P\ \A\<^sub>P \* M\ \A\<^sub>P' \* (P \ Q)\ + \A\<^sub>P' \* M\ \A\<^sub>Q' \* (P \ Q)\ \A\<^sub>Q' \* M\ + ultimately have "\\<^sub>Q \ (\ \ \\<^sub>P') \ P \ \M\N\ \ P'" + by(elim brCommInAux) (simp | force)+ + then have Ptrans: "(\ \ \\<^sub>P') \ \\<^sub>Q \ P \ \M\N\ \ P'" + by (metis Commutativity statEqTransition) + + have right2: "\A\<^sub>P', (\ \ \\<^sub>P') \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>P', ((\ \ \\<^sub>P') \ \\<^sub>P) \ \\<^sub>Q\" + by (metis FrameStatEq_def frameIntAssociativity) + with \\A\<^sub>Q', \' \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>P', (\ \ \\<^sub>P') \ \\<^sub>P \ \\<^sub>Q\\ + have "\A\<^sub>Q', (\' \ \\<^sub>P) \ \\<^sub>Q\ \\<^sub>F \A\<^sub>P', ((\ \ \\<^sub>P') \ \\<^sub>P) \ \\<^sub>Q\" + by (metis FrameStatEq_def FrameStatImpTrans frameIntAssociativity) + then have Qtrans: "(\ \ \\<^sub>P') \ \\<^sub>P \ Q \ \M\\*xvec\\N\ \ Q'" + using \A\<^sub>Q \* A\<^sub>P'\ \A\<^sub>Q \* A\<^sub>Q'\ \A\<^sub>Q \* \\ \A\<^sub>Q \* \\<^sub>P'\ + \A\<^sub>Q \* \\<^sub>P\ \A\<^sub>Q \* \'\ \A\<^sub>Q \* \\<^sub>P\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* M\ \A\<^sub>P' \* (P \ Q)\ \A\<^sub>P' \* M\ + \A\<^sub>Q' \* (P \ Q)\ \A\<^sub>Q' \* M\ + by(intro cBrComm1(8)) (simp | force)+ + from Ptrans \extractFrame P = \A\<^sub>P, \\<^sub>P\\ Qtrans \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ + \A\<^sub>P \* \\ \A\<^sub>P \* \\<^sub>P'\ \A\<^sub>P \* P\ \A\<^sub>P \* Q\ + \A\<^sub>P \* M\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>Q \* \\ \A\<^sub>Q \* \\<^sub>P'\ \A\<^sub>Q \* P\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* M\ \xvec \* P\ + have "\ \ \\<^sub>P' \ P \ Q \ \M\\*xvec\\N\ \ P' \ Q'" + by(elim BrComm1) (simp | force)+ + then show ?case using \\\*xvec\N \' (P' \ Q') = \\*yvec\N' \' R'\ + by(simp add: residualInject) +next + case(cBrComm2 \' \\<^sub>Q P M xvec N P' A\<^sub>P \\<^sub>P Q Q' A\<^sub>Q \ A\<^sub>P' \\<^sub>P' A\<^sub>Q' \\<^sub>Q' N' R' yvec) + have "\A\<^sub>P', (\ \ \\<^sub>P') \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>P', (\\<^sub>P \ (\ \ \\<^sub>P')) \ \\<^sub>Q\" + by (metis AssertionStatEqTrans AssertionStatEq_def Commutativity associativitySym frameImpNilStatEq frameImpResChainPres) + with \\A\<^sub>Q', \' \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>P', (\ \ \\<^sub>P') \ \\<^sub>P \ \\<^sub>Q\\ + have "\A\<^sub>Q', (\\<^sub>P \ \') \ \\<^sub>Q\ \\<^sub>F \A\<^sub>P', (\\<^sub>P \ (\ \ \\<^sub>P')) \ \\<^sub>Q\" + by (metis AssertionStatEqTrans AssertionStatEq_def Commutativity FrameStatImpTrans associativitySym frameImpNilStatEq frameImpResChainPres) + moreover from \\' \ \\<^sub>P \ Q \ \M\N\ \ Q'\ + have "\\<^sub>P \ \' \ Q \ \M\N\ \ Q'" + by (metis Commutativity statEqTransition) + moreover note \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ \distinct A\<^sub>Q\ + \A\<^sub>Q \* A\<^sub>P'\ \A\<^sub>Q \* A\<^sub>Q'\ \A\<^sub>Q \* \\<^sub>P\ \A\<^sub>Q \* \\ \A\<^sub>Q \* \\<^sub>P'\ + \A\<^sub>Q \* \'\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* M\ \A\<^sub>P' \* (P \ Q)\ \A\<^sub>P' \* M\ \A\<^sub>Q' \* (P \ Q)\ \A\<^sub>Q' \* M\ + ultimately have "\\<^sub>P \ (\ \ \\<^sub>P') \ Q \ \M\N\ \ Q'" + by (elim brCommInAux) (simp | force)+ + then have Qtrans: "(\ \ \\<^sub>P') \ \\<^sub>P \ Q \ \M\N\ \ Q'" + by (metis Commutativity statEqTransition) + + have "\A\<^sub>P', (\ \ \\<^sub>P') \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>P', ((\ \ \\<^sub>P') \ \\<^sub>Q) \ \\<^sub>P\" + by (metis AssertionStatEqTrans AssertionStatEq_def Associativity associativitySym frameImpNilStatEq frameImpResChainPres) + with \\A\<^sub>Q', \' \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>P', (\ \ \\<^sub>P') \ \\<^sub>P \ \\<^sub>Q\\ + have "\A\<^sub>Q', (\' \ \\<^sub>Q) \ \\<^sub>P\ \\<^sub>F \A\<^sub>P', ((\ \ \\<^sub>P') \ \\<^sub>Q) \ \\<^sub>P\" + by (metis AssertionStatEqTrans AssertionStatEq_def Associativity FrameStatImpTrans associativitySym frameImpNilStatEq frameImpResChainPres) + with \A\<^sub>P \* A\<^sub>P'\ \A\<^sub>P \* A\<^sub>Q'\ \A\<^sub>P \* \\ \A\<^sub>P \* \\<^sub>P'\ \A\<^sub>P \* \\<^sub>Q\ + \A\<^sub>P \* \'\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>P \* P\ \A\<^sub>P \* M\ \A\<^sub>P' \* (P \ Q)\ \A\<^sub>P' \* M\ + \A\<^sub>Q' \* (P \ Q)\ \A\<^sub>Q' \* M\ + have Ptrans: "(\ \ \\<^sub>P') \ \\<^sub>Q \ P \ \M\\*xvec\\N\ \ P'" + by(intro cBrComm2(4)) (simp | force)+ + from Ptrans \extractFrame P = \A\<^sub>P, \\<^sub>P\\ Qtrans \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ + \A\<^sub>P \* \\ \A\<^sub>P \* \\<^sub>P'\ \A\<^sub>P \* P\ \A\<^sub>P \* Q\ \A\<^sub>P \* M\ \A\<^sub>P \* A\<^sub>Q\ + \A\<^sub>Q \* \\ \A\<^sub>Q \* \\<^sub>P'\ \A\<^sub>Q \* P\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* M\ \xvec \* Q\ + have "\ \ \\<^sub>P' \ P \ Q \ \M\\*xvec\\N\ \ P' \ Q'" + by(elim BrComm2) (simp | force)+ + then show ?case using \\\*xvec\N \' (P' \ Q') = \\*yvec\N' \' R'\ + by(simp add: residualInject) +next + case(cBrOpen \' R M' xvec yvec N' R' x A\<^sub>R \\<^sub>R \ A\<^sub>P \\<^sub>P A\<^sub>Q \\<^sub>Q N R'' zvec) + have "\ \ \\<^sub>P \ R \ \M'\\*(xvec@yvec)\\N'\ \ R'" using cBrOpen + by(intro cBrOpen(4)) (assumption | simp)+ + + then have "\ \ \\<^sub>P \ \\x\R \ \M'\\*(xvec@x#yvec)\\N'\ \ R'" + using \x \ supp N'\ \x \ \\ \x \ \\<^sub>P\ \x \ M'\ \x \ xvec\ \x \ yvec\ + by(elim BrOpen) (simp | force)+ + with \\\*(xvec @ x # yvec)\N' \' R' = \\*zvec\N \' R''\ + show ?case + by(simp add: residualInject) +next + case(cScope \' R M' xvec N' R' x A\<^sub>R \\<^sub>R \ A\<^sub>P \\<^sub>P A\<^sub>Q \\<^sub>Q N R'' yvec) + then have "\ \ \\<^sub>P \ R \ \M'\\*xvec\\N'\ \ R'" + by(intro cScope(4)) (simp | force)+ + with \x \ \\ \x \ \\<^sub>P\ \x \ M'\ \x \ xvec\ \x \ N'\ + have "\ \ \\<^sub>P \ \\x\R \ \M'\\*xvec\\N'\ \ \\x\R'" + by(elim Scope) (simp | force)+ + then show ?case using \\\*xvec\N' \' \\x\R' = \\*yvec\N \' R''\ + by(simp add: residualInject) +next + case(cBang \' R M' A\<^sub>R \\<^sub>R \ A\<^sub>P \\<^sub>P A\<^sub>Q \\<^sub>Q N R' xvec) + have "\A\<^sub>Q, \' \ \\<^sub>R \ \\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\<^sub>R \ \\" + proof - + from \\\<^sub>R \ \\ have "\A\<^sub>Q, \' \ \\<^sub>R \ \\ \\<^sub>F \A\<^sub>Q, \' \ \\" + by(metis Identity Commutativity AssertionStatEqSym Composition frameResChainPres frameNilStatEq AssertionStatEqTrans) + moreover note \\A\<^sub>Q, \' \ \\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\\ + moreover from \\\<^sub>R \ \\ have "\A\<^sub>P, (\ \ \\<^sub>P) \ \\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\<^sub>R \ \\" + by(metis Identity Commutativity AssertionStatEqSym Composition frameResChainPres frameNilStatEq AssertionStatEqTrans) + ultimately show ?thesis by(rule FrameStatEqImpCompose) + qed + then have "\ \ \\<^sub>P \ R \ !R \ \M'\\*xvec\\N\ \ R'" using cBang + by(intro cBang(5)) (simp | force)+ + then show ?case using \guarded R\ + by(rule Bang) +qed + +lemma brCommOutAux: + fixes \ :: 'b + and \\<^sub>Q :: 'b + and R :: "('a, 'b, 'c) psi" + and M :: 'a + and xvec :: "name list" + and N :: 'a + and R' :: "('a, 'b, 'c) psi" + and A\<^sub>R :: "name list" + and \\<^sub>R :: 'b + and A\<^sub>P :: "name list" + and \\<^sub>P :: 'b + and A\<^sub>Q :: "name list" + +assumes RTrans: "\ \ \\<^sub>Q \ R \ \M\\*xvec\\N\ \ R'" + and FrR: "extractFrame R = \A\<^sub>R, \\<^sub>R\" + and "distinct A\<^sub>R" + and QimpP: "\A\<^sub>Q, (\ \ \\<^sub>Q) \ \\<^sub>R\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\<^sub>R\" + and "A\<^sub>R \* A\<^sub>P" + and "A\<^sub>R \* A\<^sub>Q" + and "A\<^sub>R \* \" + and "A\<^sub>R \* \\<^sub>P" + and "A\<^sub>R \* \\<^sub>Q" + and "A\<^sub>R \* R" + and "A\<^sub>R \* M" + and "A\<^sub>P \* R" + and "A\<^sub>P \* M" + and "A\<^sub>Q \* R" + and "A\<^sub>Q \* M" + +shows "\ \ \\<^sub>P \ R \ \M\\*xvec\\N\ \ R'" +proof - + from RTrans have "\ \ \\<^sub>Q \ R \ RBrOut M (\\*xvec\N \' R')" + by(simp add: residualInject) + moreover from \A\<^sub>R \* \\ \A\<^sub>R \* \\<^sub>Q\ have "A\<^sub>R \* (\ \ \\<^sub>Q)" + by force + ultimately show ?thesis using assms + by(elim brCommOutAuxTooMuch) +qed + +lemma comm1Aux: + fixes \ :: 'b + and \\<^sub>Q :: 'b + and R :: "('a, 'b, 'c) psi" + and K :: 'a + and xvec :: "name list" + and N :: 'a + and R' :: "('a, 'b, 'c) psi" + and A\<^sub>R :: "name list" + and \\<^sub>R :: 'b + and P :: "('a, 'b, 'c) psi" + and M :: 'a + and L :: 'a + and P' :: "('a, 'b, 'c) psi" + and A\<^sub>P :: "name list" + and \\<^sub>P :: 'b + and A\<^sub>Q :: "name list" + +assumes RTrans: "\ \ \\<^sub>Q \ R \K\\*xvec\\N\ \ R'" + and FrR: "extractFrame R = \A\<^sub>R, \\<^sub>R\" + and PTrans: "\ \ \\<^sub>R \ P \M\L\ \ P'" + and MeqK: "\ \ \\<^sub>Q \ \\<^sub>R \ M \ K" + and PeqQ: "\A\<^sub>Q, (\ \ \\<^sub>Q) \ \\<^sub>R\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\<^sub>R\" + and FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" + and FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" + and "distinct A\<^sub>P" + and "distinct A\<^sub>R" + and "A\<^sub>R \* A\<^sub>P" + and "A\<^sub>R \* A\<^sub>Q" + and "A\<^sub>R \* \" + and "A\<^sub>R \* P" + and "A\<^sub>R \* Q" + and "A\<^sub>R \* R" + and "A\<^sub>R \* K" + and "A\<^sub>P \* \" + and "A\<^sub>P \* R" + and "A\<^sub>P \* P" + and "A\<^sub>P \* M" + and "A\<^sub>Q \* R" + and "A\<^sub>Q \* M" + +obtains K' where "\ \ \\<^sub>P \ R \K'\\*xvec\\N\ \ R'" and "\ \ \\<^sub>P \ \\<^sub>R \ M \ K'" and "A\<^sub>R \* K'" +proof - + from \A\<^sub>R \* P\ \A\<^sub>R \* Q\ \A\<^sub>R \* A\<^sub>P\ \A\<^sub>R \* A\<^sub>Q\ FrP FrQ have "A\<^sub>R \* \\<^sub>P" and "A\<^sub>R \* \\<^sub>Q" + by(force dest: extractFrameFreshChain)+ + assume Assumptions: "\K'. \\ \ \\<^sub>P \ R \K'\\*xvec\\N\ \ R'; \ \ \\<^sub>P \ \\<^sub>R \ M \ K'; A\<^sub>R \* K'\ \ thesis" + { + fix \' ::'b + and zvec ::"name list" + + assume A: "\ \ \\<^sub>Q \ \'" + assume "\' \ R \K\\*xvec\\N\ \ R'" + then have "\' \ R \ROut K (\\*xvec\N \' R')" by(simp add: residualInject) + moreover note FrR \distinct A\<^sub>R\ PTrans + moreover from \\' \ R \K\\*xvec\\N\ \ R'\ have "distinct xvec" by(auto dest: boundOutputDistinct) + moreover assume "\' \ \\<^sub>R \ M \ K" and "\A\<^sub>Q, \' \ \\<^sub>R\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\<^sub>R\" + and "A\<^sub>R \* zvec" and "A\<^sub>P \* zvec" and "zvec \* R" and "zvec \* P" + and "A\<^sub>R \* \'" + ultimately have "\K'. \ \ \\<^sub>P \ R \ROut K' (\\*xvec\N \' R') \ \ \ \\<^sub>P \ \\<^sub>R \ M \ K' \ zvec \* K' \ A\<^sub>R \* K'" + using FrP \A\<^sub>P \* R\ \A\<^sub>Q \* R\ \A\<^sub>P \* \\ \A\<^sub>P \* P\ \distinct A\<^sub>P\ \A\<^sub>R \* \\ \A\<^sub>R \* P\ \A\<^sub>R \* R\ + \A\<^sub>P \* M\ \A\<^sub>Q \* M\ \A\<^sub>R \* K\ \A\<^sub>R \* A\<^sub>P\ \A\<^sub>R \* A\<^sub>Q\ \A\<^sub>R \* \\<^sub>P\ + proof(nominal_induct \' R K B=="\\*xvec\N \' R'" A\<^sub>R \\<^sub>R avoiding: \ P A\<^sub>P \\<^sub>P A\<^sub>Q zvec xvec N R' arbitrary: M rule: outputFrameInduct) + case(cAlpha \' R K A\<^sub>R \\<^sub>R p \ P A\<^sub>P \\<^sub>P A\<^sub>Q zvec xvec N R' M) + have S: "set p \ set A\<^sub>R \ set (p \ A\<^sub>R)" by fact + from \\' \ (p \ \\<^sub>R) \ M \ K\ have "(p \ (\' \ (p \ \\<^sub>R))) \ (p \ M) \ (p \ K)" + by(rule chanEqClosed) + with \A\<^sub>R \* \'\ \(p \ A\<^sub>R) \* \'\ \A\<^sub>R \* K\ \(p \ A\<^sub>R) \* K\ S \distinctPerm p\ + have "\' \ \\<^sub>R \ (p \ M) \ K" by(simp add: eqvts) + moreover from \\ \ (p \ \\<^sub>R) \ P \M\L\ \ P'\ S \A\<^sub>R \* P\ \(p \ A\<^sub>R) \* P\ have "(p \ (\ \ (p \ \\<^sub>R))) \ P \(p \ M)\L\ \ P'" + by(elim inputPermFrameSubject) auto + with \A\<^sub>R \* \\ \(p \ A\<^sub>R) \* \\ S \distinctPerm p\ have "\ \ \\<^sub>R \ P \(p \ M)\L\ \ P'" + by(simp add: eqvts) + moreover from \A\<^sub>P \* M\ have "(p \ A\<^sub>P) \* (p \ M)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>R \* A\<^sub>P\ \(p \ A\<^sub>R) \* A\<^sub>P\ S have "A\<^sub>P \* (p \ M)" by simp + moreover from \A\<^sub>Q \* M\ have "(p \ A\<^sub>Q) \* (p \ M)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>R \* A\<^sub>Q\ \(p \ A\<^sub>R) \* A\<^sub>Q\ S have "A\<^sub>Q \* (p \ M)" by simp + moreover from \\A\<^sub>Q, \' \ (p \ \\<^sub>R)\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ (p \ \\<^sub>R)\ \ + have "(p \ \A\<^sub>Q, \' \ (p \ \\<^sub>R)\) \\<^sub>F (p \ \A\<^sub>P, (\ \ \\<^sub>P) \ (p \ \\<^sub>R)\)" + by(rule FrameStatImpClosed) + with \A\<^sub>R \* A\<^sub>P\ \(p \ A\<^sub>R) \* A\<^sub>P\ \A\<^sub>R \* \'\ \(p \ A\<^sub>R) \* \'\ \A\<^sub>R \* \\<^sub>P\ \(p \ A\<^sub>R) \* \\<^sub>P\ \A\<^sub>R \* A\<^sub>Q\ + \(p \ A\<^sub>R) \* A\<^sub>Q\ \A\<^sub>R \* \\ \(p \ A\<^sub>R) \* \\ S \distinctPerm p\ + have "\A\<^sub>Q, \' \ \\<^sub>R\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\<^sub>R\" by(simp add: eqvts) + ultimately obtain K' where "\ \ \\<^sub>P \ R \ROut K' (\\*xvec\N \' R')" and "\ \ \\<^sub>P \ \\<^sub>R \ (p \ M) \ K'" and "zvec \* K'" and "A\<^sub>R \* K'" + using cAlpha + by metis + from \\ \ \\<^sub>P \ R \ROut K' (\\*xvec\N \' R')\ S \A\<^sub>R \* R\ \(p \ A\<^sub>R) \* R\ + have "(p \ (\ \ \\<^sub>P)) \ R \(ROut (p \ K') (\\*xvec\N \' R'))" using outputPermFrameSubject + by(auto simp add: residualInject) + with S \A\<^sub>R \* \\ \(p \ A\<^sub>R) \* \\ \A\<^sub>R \* \\<^sub>P\ \(p \ A\<^sub>R) \* \\<^sub>P\ have "\ \ \\<^sub>P \ R \(ROut (p \ K') (\\*xvec\N \' R'))" + by(simp add: eqvts) + moreover from \\ \ \\<^sub>P \ \\<^sub>R \ (p \ M) \ K'\ have "(p \ (\ \ \\<^sub>P \ \\<^sub>R))\ (p \ p \ M) \ (p \ K')" + by(rule chanEqClosed) + with S \A\<^sub>R \* \\ \(p \ A\<^sub>R) \* \\ \A\<^sub>R \* \\<^sub>P\ \(p \ A\<^sub>R) \* \\<^sub>P\ \distinctPerm p\ have "\ \ \\<^sub>P \ (p \ \\<^sub>R) \ M \ (p \ K')" + by(simp add: eqvts) + moreover from \zvec \* K'\ have "(p \ zvec) \* (p \ K')" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>R \* zvec\ \(p \ A\<^sub>R) \* zvec\ S have "zvec \* (p \ K')" by simp + moreover from \A\<^sub>R \* K'\ have "(p \ A\<^sub>R) \* (p \ K')" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + ultimately show ?case by blast + next + case(cOutput \' M' K N R \ P A\<^sub>P \\<^sub>P A\<^sub>Q zvec xvec N' R' M) + from \A\<^sub>P \* (M'\N\.R)\ \A\<^sub>Q \* (M'\N\.R)\ \zvec \* (M'\N\.R)\ + have "A\<^sub>P \* M'" and "A\<^sub>Q \* M'" and "zvec \* M'" by simp+ + + from \\' \ M' \ K\ have "\' \ \ \ M' \ K" by(blast intro: statEqEnt Identity AssertionStatEqSym) + then have "\' \ \ \ M' \ M'" by(blast intro: chanEqSym chanEqTrans) + with \A\<^sub>Q \* M'\ have "(\A\<^sub>Q, \' \ \\) \\<^sub>F M' \ M'" by(force intro: frameImpI) + + with \\A\<^sub>Q, \' \ \\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\\ have "(\A\<^sub>P, (\ \ \\<^sub>P) \ \\) \\<^sub>F M' \ M'" + by(simp add: FrameStatImp_def) + with \A\<^sub>P \* M'\ have "(\ \ \\<^sub>P) \ \ \ M' \ M'" by(force dest: frameImpE) + then have "\ \ \\<^sub>P \ M' \ M'" by(blast intro: statEqEnt Identity) then have "\ \ \\<^sub>P \ M'\N\.R \M'\N\ \ R" + by(rule Output) + + moreover from \\' \ \ \ M \ K\ \\' \ \ \ M' \ K\ + have "\' \ \ \ M \ M'" by(metis chanEqSym chanEqTrans) + with \A\<^sub>Q \* M\ \A\<^sub>Q \* M'\ + have "(\A\<^sub>Q, \' \ \\) \\<^sub>F M \ M'" + by(force intro: frameImpI) + with \\A\<^sub>Q, \' \ \\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\\ + have "(\A\<^sub>P, (\ \ \\<^sub>P) \ \\) \\<^sub>F M \ M'" + by(simp add: FrameStatImp_def) + with \A\<^sub>P \* M\ \A\<^sub>P \* M'\ have "(\ \ \\<^sub>P) \ \ \ M \ M'" + by(force dest: frameImpE) + then have "\ \ \\<^sub>P \ \ \ M \ M'" + by(metis statEqEnt Associativity) + ultimately show ?case using cOutput by(auto simp add: residualInject) + next + case(cCase \' R M' \ Cs A\<^sub>R \\<^sub>R \ P A\<^sub>P \\<^sub>P A\<^sub>Q zvec xvec N R' M) + from \guarded R\ \extractFrame R = \A\<^sub>R, \\<^sub>R\\ have "\\<^sub>R \ \" + by(metis guardedStatEq) + with \\' \ \ \ M \ M'\ have "\' \ \\<^sub>R \ M \ M'" + by(metis Identity Commutativity statEqEnt AssertionStatEqSym Composition) + moreover have "\A\<^sub>Q, \' \ \\<^sub>R\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\<^sub>R\" + proof - + from \\\<^sub>R \ \\ have "\A\<^sub>Q, \' \ \\<^sub>R\ \\<^sub>F \A\<^sub>Q, \' \ \\" + by(metis Identity Commutativity AssertionStatEqSym Composition frameResChainPres frameNilStatEq AssertionStatEqTrans) + moreover note \\A\<^sub>Q, \' \ \\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\\ + moreover from \\\<^sub>R \ \\ have "\A\<^sub>P, (\ \ \\<^sub>P) \ \\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\<^sub>R\" + by(metis Identity Commutativity AssertionStatEqSym Composition frameResChainPres frameNilStatEq AssertionStatEqTrans) + ultimately show ?thesis by(rule FrameStatEqImpCompose) + qed + moreover from \\ \ \ \ P \M\L\ \ P'\ \\\<^sub>R \ \\ + have "\ \ \\<^sub>R \ P \M\L\ \ P'" by(metis statEqTransition Identity Commutativity AssertionStatEqSym Composition) + moreover from \zvec \* (Cases Cs)\ \A\<^sub>P \* (Cases Cs)\ \A\<^sub>Q \* (Cases Cs)\ \(\, R) \ set Cs\ + have "A\<^sub>P \* R" and "A\<^sub>Q \* R" and "zvec \* R" and "A\<^sub>P \* \" and "A\<^sub>Q \* \" + by(auto dest: memFreshChain) + ultimately have "\K'. \ \ \\<^sub>P \ R \ROut K' (\\*xvec\N \' R') \ \ \ \\<^sub>P \ \\<^sub>R \ M \ K' \ zvec \* K' \ A\<^sub>R \* K'" using cCase + by(intro cCase) (assumption | simp)+ + then obtain K' where RTrans: "\ \ \\<^sub>P \ R \ROut K' (\\*xvec\N \' R')" + and MeqK': "\ \ \\<^sub>P \ \\<^sub>R \ M \ K'" and "zvec \* K'" and "A\<^sub>R \* K'" + by metis + note RTrans \(\, R) \ set Cs\ + moreover from \\' \ \\ have "\' \ \ \ \" by(blast intro: statEqEnt Identity AssertionStatEqSym) + with \A\<^sub>Q \* \\ have "(\A\<^sub>Q, \' \ \\) \\<^sub>F \" by(force intro: frameImpI) + with \\A\<^sub>Q, \' \ \\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\\ have "(\A\<^sub>P, (\ \ \\<^sub>P) \ \\) \\<^sub>F \" + by(simp add: FrameStatImp_def) + with \A\<^sub>P \* \\ have "(\ \ \\<^sub>P) \ \ \ \" by(force dest: frameImpE) + then have "\ \ \\<^sub>P \ \" by(blast intro: statEqEnt Identity) + ultimately have "\ \ \\<^sub>P \ Cases Cs \ROut K' (\\*xvec\N \' R')" using \guarded R\ by(rule Case) + moreover from MeqK' \\\<^sub>R \ \\ have "\ \ \\<^sub>P \ \ \ M \ K'" + by(metis Identity Commutativity statEqEnt AssertionStatEqSym Composition AssertionStatEqTrans) + ultimately show ?case using \zvec \* K'\ + by auto + next + case(cPar1 \' \\<^sub>R\<^sub>2 R\<^sub>1 M' xvec N' R\<^sub>1' A\<^sub>R\<^sub>2 R\<^sub>2 A\<^sub>R\<^sub>1 \\<^sub>R\<^sub>1 \ P A\<^sub>P \\<^sub>P A\<^sub>Q zvec yvec N R' M) + have FrR2: "extractFrame R\<^sub>2 = \A\<^sub>R\<^sub>2, \\<^sub>R\<^sub>2\" by fact + from \\' \ \\<^sub>R\<^sub>1 \ \\<^sub>R\<^sub>2 \ M \ M'\ have "(\' \ \\<^sub>R\<^sub>2) \ \\<^sub>R\<^sub>1 \ M \ M'" + by(metis statEqEnt Associativity Composition Commutativity) + moreover have "\A\<^sub>Q, (\' \ \\<^sub>R\<^sub>2) \ \\<^sub>R\<^sub>1\ \\<^sub>F \A\<^sub>P, ((\ \ \\<^sub>R\<^sub>2) \ \\<^sub>P) \ \\<^sub>R\<^sub>1\" + proof - + have "\A\<^sub>Q, (\' \ \\<^sub>R\<^sub>2) \ \\<^sub>R\<^sub>1\ \\<^sub>F \A\<^sub>Q, \' \ \\<^sub>R\<^sub>1 \ \\<^sub>R\<^sub>2\" + by(metis Associativity Composition Commutativity AssertionStatEqTrans AssertionStatEqSym frameNilStatEq frameResChainPres) + moreover note \\A\<^sub>Q, \' \ \\<^sub>R\<^sub>1 \ \\<^sub>R\<^sub>2\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\<^sub>R\<^sub>1 \ \\<^sub>R\<^sub>2\\ + moreover have "\A\<^sub>P, (\ \ \\<^sub>P) \ \\<^sub>R\<^sub>1 \ \\<^sub>R\<^sub>2\ \\<^sub>F \A\<^sub>P, ((\ \ \\<^sub>R\<^sub>2) \ \\<^sub>P) \ \\<^sub>R\<^sub>1\" + by(metis Associativity Composition Commutativity AssertionStatEqTrans AssertionStatEqSym frameNilStatEq frameResChainPres) + ultimately show ?thesis by(rule FrameStatEqImpCompose) + qed + moreover from \\ \ \\<^sub>R\<^sub>1 \ \\<^sub>R\<^sub>2 \ P \M\L\ \ P'\ have "(\ \ \\<^sub>R\<^sub>2) \ \\<^sub>R\<^sub>1 \ P \M\L\ \ P'" + by(metis statEqTransition Associativity Composition Commutativity) + moreover from \A\<^sub>R\<^sub>1 \* A\<^sub>P\ \A\<^sub>R\<^sub>2 \* A\<^sub>P\ \A\<^sub>P \* (R\<^sub>1 \ R\<^sub>2)\ \extractFrame R\<^sub>1 = \A\<^sub>R\<^sub>1, \\<^sub>R\<^sub>1\\ FrR2 have "A\<^sub>P \* \\<^sub>R\<^sub>1" and "A\<^sub>P \* \\<^sub>R\<^sub>2" + by(force dest: extractFrameFreshChain)+ + + moreover from \\\*xvec\N' \' (R\<^sub>1' \ R\<^sub>2) = \\*yvec\N \' R'\ \xvec \* yvec\ + obtain p T where "\\*xvec\N' \' R\<^sub>1' = \\*yvec\N \' T" and "R' = T \ (p \ R\<^sub>2)" and "set p \ set yvec \ set xvec" + apply(drule_tac sym) + by(elim boundOutputPar1Dest') (assumption | simp | blast dest: sym)+ + ultimately have "\K'. (\ \ \\<^sub>R\<^sub>2) \ \\<^sub>P \ R\<^sub>1 \ROut K' (\\*yvec\N \' T) \ (\ \ \\<^sub>R\<^sub>2) \ \\<^sub>P \ \\<^sub>R\<^sub>1 \ M \ K' \ (A\<^sub>R\<^sub>2@zvec) \* K' \ A\<^sub>R\<^sub>1 \* K'" using cPar1 + by(elim cPar1(6)[where ba=P and bb=A\<^sub>P and bd=A\<^sub>Q]) auto + then obtain K' where RTrans: "(\ \ \\<^sub>R\<^sub>2) \ \\<^sub>P \ R\<^sub>1 \K'\\*xvec\\N'\ \ R\<^sub>1'" + and MeqK': "(\ \ \\<^sub>R\<^sub>2) \ \\<^sub>P \ \\<^sub>R\<^sub>1 \ M \ K'" and "A\<^sub>R\<^sub>2 \* K'" and "A\<^sub>R\<^sub>1 \* K'" and "zvec \* K'" + using \\\*xvec\N' \' R\<^sub>1' = \\*yvec\N \' T\ by(auto simp add: residualInject) + + from RTrans have "(\ \ \\<^sub>P) \ \\<^sub>R\<^sub>2 \ R\<^sub>1 \K'\\*xvec\\N'\ \ R\<^sub>1'" + by(metis statEqTransition Associativity Composition Commutativity) + then have "\ \ \\<^sub>P \ (R\<^sub>1 \ R\<^sub>2) \K'\\*xvec\\N'\ \ (R\<^sub>1' \ R\<^sub>2)" using FrR2 \xvec \* R\<^sub>2\ \A\<^sub>R\<^sub>2 \* \\ \A\<^sub>R\<^sub>2 \* \\<^sub>P\ \A\<^sub>R\<^sub>2 \* K'\ \A\<^sub>R\<^sub>2 \* R\<^sub>1\ \A\<^sub>R\<^sub>2 \* xvec\ \A\<^sub>R\<^sub>2 \* N'\ + by(force intro: Par1) + moreover from MeqK' have "\ \ \\<^sub>P \ \\<^sub>R\<^sub>1 \ \\<^sub>R\<^sub>2 \ M \ K'" + by(metis statEqEnt Associativity Composition Commutativity) + ultimately show ?case using \zvec \* K'\ \A\<^sub>R\<^sub>1 \* K'\ \A\<^sub>R\<^sub>2 \* K'\ \\\*xvec\N' \' (R\<^sub>1' \ R\<^sub>2) = \\*yvec\N \' R'\ + by(auto simp add: residualInject) + next + case(cPar2 \' \\<^sub>R\<^sub>1 R\<^sub>2 M' xvec N' R\<^sub>2' A\<^sub>R\<^sub>1 R\<^sub>1 A\<^sub>R\<^sub>2 \\<^sub>R\<^sub>2 \ P A\<^sub>P \\<^sub>P A\<^sub>Q zvec yvec N R' M) + have FrR1: "extractFrame R\<^sub>1 = \A\<^sub>R\<^sub>1, \\<^sub>R\<^sub>1\" by fact + from \\' \ \\<^sub>R\<^sub>1 \ \\<^sub>R\<^sub>2 \ M \ M'\ have "(\' \ \\<^sub>R\<^sub>1) \ \\<^sub>R\<^sub>2 \ M \ M'" + by(metis statEqEnt Associativity Composition Commutativity) + moreover have "\A\<^sub>Q, (\' \ \\<^sub>R\<^sub>1) \ \\<^sub>R\<^sub>2\ \\<^sub>F \A\<^sub>P, ((\ \ \\<^sub>R\<^sub>1) \ \\<^sub>P) \ \\<^sub>R\<^sub>2\" + proof - + have "\A\<^sub>Q, (\' \ \\<^sub>R\<^sub>1) \ \\<^sub>R\<^sub>2\ \\<^sub>F \A\<^sub>Q, \' \ \\<^sub>R\<^sub>1 \ \\<^sub>R\<^sub>2\" + by(metis Associativity Composition Commutativity AssertionStatEqTrans AssertionStatEqSym frameNilStatEq frameResChainPres) + moreover note \\A\<^sub>Q, \' \ \\<^sub>R\<^sub>1 \ \\<^sub>R\<^sub>2\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\<^sub>R\<^sub>1 \ \\<^sub>R\<^sub>2\\ + moreover have "\A\<^sub>P, (\ \ \\<^sub>P) \ \\<^sub>R\<^sub>1 \ \\<^sub>R\<^sub>2\ \\<^sub>F \A\<^sub>P, ((\ \ \\<^sub>R\<^sub>1) \ \\<^sub>P) \ \\<^sub>R\<^sub>2\" + by(metis Associativity Composition Commutativity AssertionStatEqTrans AssertionStatEqSym frameNilStatEq frameResChainPres) + ultimately show ?thesis by(rule FrameStatEqImpCompose) + qed + moreover from \\ \ \\<^sub>R\<^sub>1 \ \\<^sub>R\<^sub>2 \ P \M\L\ \ P'\ have "(\ \ \\<^sub>R\<^sub>1) \ \\<^sub>R\<^sub>2 \ P \M\L\ \ P'" + by(metis statEqTransition Associativity Composition Commutativity) + moreover from \A\<^sub>R\<^sub>1 \* A\<^sub>P\ \A\<^sub>R\<^sub>2 \* A\<^sub>P\ \A\<^sub>P \* (R\<^sub>1 \ R\<^sub>2)\ FrR1 \extractFrame R\<^sub>2 = \A\<^sub>R\<^sub>2, \\<^sub>R\<^sub>2\\ have "A\<^sub>P \* \\<^sub>R\<^sub>1" and "A\<^sub>P \* \\<^sub>R\<^sub>2" + by(force dest: extractFrameFreshChain)+ + moreover from \\\*xvec\N' \' (R\<^sub>1 \ R\<^sub>2') = \\*yvec\N \' R'\ \xvec \* yvec\ + obtain p T where "\\*xvec\N' \' R\<^sub>2' = \\*yvec\N \' T" and "R' = (p \ R\<^sub>1) \ T" and "set p \ set yvec \ set xvec" + apply(drule_tac sym) + by(elim boundOutputPar2Dest') (assumption | simp | blast dest: sym)+ + ultimately have "\K'. (\ \ \\<^sub>R\<^sub>1) \ \\<^sub>P \ R\<^sub>2 \ROut K' (\\*yvec\N \' T) \ (\ \ \\<^sub>R\<^sub>1) \ \\<^sub>P \ \\<^sub>R\<^sub>2 \ M \ K' \ (A\<^sub>R\<^sub>1@zvec) \* K' \ A\<^sub>R\<^sub>2 \* K'" using cPar2 + by(elim cPar2(6)) (assumption | simp | auto)+ + then obtain K' where RTrans: "(\ \ \\<^sub>R\<^sub>1) \ \\<^sub>P \ R\<^sub>2 \K'\\*xvec\\N'\ \ R\<^sub>2'" + and MeqK': "(\ \ \\<^sub>R\<^sub>1) \ \\<^sub>P \ \\<^sub>R\<^sub>2 \ M \ K'" and "A\<^sub>R\<^sub>1 \* K'" and "zvec \* K'" and "A\<^sub>R\<^sub>2 \* K'" + using \\\*xvec\N' \' R\<^sub>2' = \\*yvec\N \' T\ by(auto simp add: residualInject) + + from RTrans have "(\ \ \\<^sub>P) \ \\<^sub>R\<^sub>1 \ R\<^sub>2 \K'\\*xvec\\N'\ \ R\<^sub>2'" + by(metis statEqTransition Associativity Composition Commutativity) + then have "\ \ \\<^sub>P \ (R\<^sub>1 \ R\<^sub>2) \K'\\*xvec\\N'\ \ (R\<^sub>1 \ R\<^sub>2')" using FrR1 \xvec \* R\<^sub>1\ \A\<^sub>R\<^sub>1 \* \\ \A\<^sub>R\<^sub>1 \* \\<^sub>P\ \A\<^sub>R\<^sub>1 \* K'\\A\<^sub>R\<^sub>1 \* xvec\ \A\<^sub>R\<^sub>1 \* N'\ \A\<^sub>R\<^sub>1 \* R\<^sub>2\ + by(force intro: Par2) + moreover from MeqK' have "\ \ \\<^sub>P \ \\<^sub>R\<^sub>1 \ \\<^sub>R\<^sub>2 \ M \ K'" + by(metis statEqEnt Associativity Composition Commutativity) + ultimately show ?case using \zvec \* K'\ \A\<^sub>R\<^sub>1 \* K'\ \A\<^sub>R\<^sub>2 \* K'\ \\\*xvec\N' \' (R\<^sub>1 \ R\<^sub>2') = \\*yvec\N \' R'\ + by(auto simp add: residualInject) + next + case(cOpen \' R M' xvec yvec N' R' x A\<^sub>R \\<^sub>R \ P A\<^sub>P \\<^sub>P A\<^sub>Q zvec zvec2 N R'' M) + from \\\*(xvec @ x # yvec)\N' \' R' = \\*zvec2\N \' R''\ \x \ xvec\ \x \ yvec\ \x \ zvec2\ \x \ R''\ \x \ N\ \distinct zvec2\ + obtain xvec' x' yvec' where A: "\\*(xvec@yvec)\N' \' R' = \\*(xvec'@yvec')\([(x, x')] \ N) \' ([(x, x')] \ R'')" + and B: "zvec2 = (xvec'@x'#yvec')" + by(elim boundOutputOpenDest) auto + then have "\K'. \ \ \\<^sub>P \ R \ROut K' (\\*(xvec'@yvec')\([(x, x')] \ N) \' ([(x, x')] \ R'')) \ \ \ \\<^sub>P \ \\<^sub>R \ M \ K' \ (zvec) \* K' \ A\<^sub>R \* K'" using cOpen + by(elim cOpen(4)) (assumption | simp)+ + then obtain K' where RTrans: "\ \ \\<^sub>P \ R \K'\\*(xvec@yvec)\\N'\ \ R'" + and MeqK': "\ \ \\<^sub>P \ \\<^sub>R \ M \ K'" and "zvec \* K'" and "A\<^sub>R \* K'" + using A by(auto simp add: residualInject) + from \A\<^sub>R \* A\<^sub>P\ \A\<^sub>P \* (\\x\R)\ \x \ A\<^sub>P\ \extractFrame R = \A\<^sub>R, \\<^sub>R\\ have "A\<^sub>P \* \\<^sub>R" + by(force dest: extractFrameFreshChain)+ + + from \\ \ \\<^sub>R \ P \M\L\ \ P'\ \extractFrame P = \A\<^sub>P, \\<^sub>P\\ \distinct A\<^sub>P\ \zvec \* P\ \A\<^sub>P \* \\<^sub>R\ \x \ A\<^sub>P\ \A\<^sub>P \* M\ \A\<^sub>P \* P\ \A\<^sub>P \* zvec\ \A\<^sub>P \* \\ \A\<^sub>P \* zvec\ \A\<^sub>R \* P\ \A\<^sub>R \* A\<^sub>P\ \x \ A\<^sub>P\ \x \ P\ + + obtain K'' where MeqK'': "(\ \ \\<^sub>R) \ \\<^sub>P \ M \ K''" and "A\<^sub>R \* K''" and "zvec \* K''" and "x \ K''" + by(elim inputObtainPrefix[where B="(x#A\<^sub>R@zvec)"]) (assumption | simp | force)+ + + from MeqK'' MeqK' have KeqK'': "(\ \ \\<^sub>P) \ \\<^sub>R \ K' \ K''" + by(metis statEqEnt Associativity Composition Commutativity chanEqSym chanEqTrans) + with RTrans \extractFrame R = \A\<^sub>R, \\<^sub>R\\ \distinct A\<^sub>R\ \A\<^sub>R \* \\ \A\<^sub>R \* \\<^sub>P\ \A\<^sub>R \* K'\ \A\<^sub>R \* K''\ \A\<^sub>R \* R\ + have "\ \ \\<^sub>P \ R \K''\\*(xvec@yvec)\\N'\ \ R'" + by(elim outputRenameSubject) (assumption | force)+ + then have "\ \ \\<^sub>P \ \\x\R \K''\\*(xvec@x#yvec)\\N'\ \ R'" + using \x \ \\ \x \ \\<^sub>P\ \x \ K''\ \x \ xvec\ \ x \ yvec\ \x \ supp N'\ \xvec \* \\ \xvec \* \\<^sub>P\ \xvec \* R\ \x \ K''\ + by(elim Open) (assumption | force)+ + moreover from MeqK'' have "\ \ \\<^sub>P \ \\<^sub>R \ M \ K''" + by(metis statEqEnt Associativity Composition Commutativity) + ultimately show ?case using \zvec \* K''\ \x \ K''\ \A\<^sub>R \* K''\ B \\\*(xvec @ x # yvec)\N' \' R' = \\*zvec2\N \' R''\ + by(auto simp add: residualInject) + next + case(cScope \' R M' xvec N' R' x A\<^sub>R \\<^sub>R \ P A\<^sub>P \\<^sub>P A\<^sub>Q zvec yvec N R'' M) + from \\\*xvec\N' \' \\x\R' = \\*yvec\N \' R''\ \x \ xvec\ \x \ yvec\ + obtain R''' where "R'' = \\x\R'''" and "\\*xvec\N' \' R' = \\*yvec\N \' R'''" + apply(drule_tac sym) + by(metis boundOutputScopeDest) + then have "\K'. \ \ \\<^sub>P \ R \ROut K' (\\*yvec\N \' R''') \ \ \ \\<^sub>P \ \\<^sub>R \ M \ K' \ zvec \* K' \ A\<^sub>R \* K'" using cScope + by(elim cScope(4)) (assumption | simp)+ + then obtain K' where RTrans: "\ \ \\<^sub>P \ R \K'\\*xvec\\N'\ \ R'" + and MeqK': "\ \ \\<^sub>P \ \\<^sub>R \ M \ K'" and "zvec \* K'" and "A\<^sub>R \* K'" + using \\\*xvec\N' \' R' = \\*yvec\N \' R'''\ + by(auto simp add: residualInject) + from \A\<^sub>R \* A\<^sub>P\ \A\<^sub>P \* (\\x\R)\ \x \ A\<^sub>P\ \extractFrame R = \A\<^sub>R, \\<^sub>R\\ have "A\<^sub>P \* \\<^sub>R" + by(force dest: extractFrameFreshChain)+ + from \\ \ \\<^sub>R \ P \M\L\ \ P'\ \extractFrame P = \A\<^sub>P, \\<^sub>P\\ \distinct A\<^sub>P\ \x \ P\ \zvec \* P\ \A\<^sub>P \* \\<^sub>R\ \x \ A\<^sub>P\ \A\<^sub>P \* M\ \A\<^sub>P \* P\ \A\<^sub>P \* zvec\ \A\<^sub>P \* \\ \A\<^sub>P \* zvec\ \A\<^sub>R \* P\ \A\<^sub>R \* A\<^sub>P\ + obtain K'' where MeqK'': "(\ \ \\<^sub>R) \ \\<^sub>P \ M \ K''" and "x \ K''" and "A\<^sub>R \* K''" and "zvec \* K''" + by(elim inputObtainPrefix[where B="(x#A\<^sub>R@zvec)"]) (assumption | force)+ + + from MeqK'' MeqK' have KeqK'': "(\ \ \\<^sub>P) \ \\<^sub>R \ K' \ K''" + by(metis statEqEnt Associativity Composition Commutativity chanEqSym chanEqTrans) + with RTrans \extractFrame R = \A\<^sub>R, \\<^sub>R\\ \distinct A\<^sub>R\ \A\<^sub>R \* \\ \A\<^sub>R \* \\<^sub>P\ \A\<^sub>R \* K'\ \A\<^sub>R \* K''\ \A\<^sub>R \* R\ + have "\ \ \\<^sub>P \ R \K''\\*xvec\\N'\ \ R'" + by(elim outputRenameSubject) (assumption | force)+ + then have "\ \ \\<^sub>P \ \\x\R \K''\\*xvec\\N'\ \ \\x\R'" using \x \ \\ \x \ \\<^sub>P\ \x \ K''\ \x \ xvec\ \x \ N'\ \xvec \* \\ \xvec \* \\<^sub>P\ \xvec \* R\ \x \ K''\ + by(elim Scope) (assumption | force)+ + moreover from MeqK'' have "\ \ \\<^sub>P \ \\<^sub>R \ M \ K''" + by(metis statEqEnt Associativity Composition Commutativity) + ultimately show ?case using \zvec \* K''\ \x \ K''\ \A\<^sub>R \* K''\ \\\*xvec\N' \' \\x\R' = \\*yvec\N \' R''\ + by(auto simp add: residualInject) + next + case(cBang \' R M' A\<^sub>R \\<^sub>R \ P A\<^sub>P \\<^sub>P A\<^sub>Q zvec xvec N R' M) + from \guarded R\ \extractFrame R = \A\<^sub>R, \\<^sub>R\\ have "\\<^sub>R \ \" + by(metis guardedStatEq) + with \\' \ \ \ M \ M'\ have "\' \ \\<^sub>R \ \ \ M \ M'" + by(metis Identity Commutativity statEqEnt AssertionStatEqSym Composition) + moreover have "\A\<^sub>Q, \' \ \\<^sub>R \ \\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\<^sub>R \ \\" + proof - + from \\\<^sub>R \ \\ have "\A\<^sub>Q, \' \ \\<^sub>R \ \\ \\<^sub>F \A\<^sub>Q, \' \ \\" + by(metis Identity Commutativity AssertionStatEqSym Composition frameResChainPres frameNilStatEq AssertionStatEqTrans) + moreover note \\A\<^sub>Q, \' \ \\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\\ + moreover from \\\<^sub>R \ \\ have "\A\<^sub>P, (\ \ \\<^sub>P) \ \\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\<^sub>R \ \\" + by(metis Identity Commutativity AssertionStatEqSym Composition frameResChainPres frameNilStatEq AssertionStatEqTrans) + ultimately show ?thesis by(rule FrameStatEqImpCompose) + qed + moreover from \\ \ \ \ P \M\L\ \ P'\ \\\<^sub>R \ \\ + have "\ \ \\<^sub>R \ \ \ P \M\L\ \ P'" by(metis statEqTransition Identity Commutativity AssertionStatEqSym Composition) + ultimately have "\K'. \ \ \\<^sub>P \ R \ !R \ROut K' (\\*xvec\N \' R') \ \ \ \\<^sub>P \ \\<^sub>R \ \ \ M \ K' \ zvec \* K' \ A\<^sub>R \* K'" using cBang + by(intro cBang(5)) (assumption |simp)+ + then obtain K' where RTrans: "\ \ \\<^sub>P \ R \ !R \ROut K' (\\*xvec\N \' R')" + and MeqK': "\ \ \\<^sub>P \ \\<^sub>R \ \ \ M \ K'" and "zvec \* K'" and "A\<^sub>R \* K'" + by metis + from RTrans \guarded R\ have "\ \ \\<^sub>P \ !R \ROut K' (\\*xvec\N \' R')" by(rule Bang) + moreover from MeqK' \\\<^sub>R \ \\ have "\ \ \\<^sub>P \ \ \ M \ K'" + by(metis Identity Commutativity statEqEnt AssertionStatEqSym Composition AssertionStatEqTrans) + ultimately show ?case using \zvec \* K'\ + by force + qed + } + note Goal = this + have "\ \ \\<^sub>Q \ \ \ \\<^sub>Q" by simp + moreover note RTrans + moreover from MeqK have "(\ \ \\<^sub>Q) \ \\<^sub>R \ M \ K" + by(metis statEqEnt Associativity Commutativity) + moreover note PeqQ \A\<^sub>R \* \\ \A\<^sub>R \* \\<^sub>Q\ + ultimately have "\K'. \ \ \\<^sub>P \ R \ROut K' (\\*xvec\N \' R') \ \ \ \\<^sub>P \ \\<^sub>R \ M \ K' \ ([]::name list) \* K' \ A\<^sub>R \* K'" + by(elim Goal) (assumption | force simp add: residualInject)+ + with Assumptions show ?thesis + by(force simp add: residualInject) +qed + +lemma comm2Aux: + fixes \ :: 'b + and \\<^sub>Q :: 'b + and R :: "('a, 'b, 'c) psi" + and K :: 'a + and N :: 'a + and R' :: "('a, 'b, 'c) psi" + and A\<^sub>R :: "name list" + and \\<^sub>R :: 'b + and P :: "('a, 'b, 'c) psi" + and M :: 'a + and xvec :: "name list" + and P' :: "('a, 'b, 'c) psi" + and A\<^sub>P :: "name list" + and \\<^sub>P :: 'b + and A\<^sub>Q :: "name list" + +assumes RTrans: "\ \ \\<^sub>Q \ R \K\N\ \ R'" + and FrR: "extractFrame R = \A\<^sub>R, \\<^sub>R\" + and PTrans: "\ \ \\<^sub>R \ P \M\\*xvec\\N\ \ P'" + and MeqK: "\ \ \\<^sub>Q \ \\<^sub>R \ M \ K" + and QimpP: "\A\<^sub>Q, (\ \ \\<^sub>Q) \ \\<^sub>R\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\<^sub>R\" + and FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" + and FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" + and "distinct A\<^sub>P" + and "distinct A\<^sub>R" + and "A\<^sub>R \* A\<^sub>P" + and "A\<^sub>R \* A\<^sub>Q" + and "A\<^sub>R \* \" + and "A\<^sub>R \* P" + and "A\<^sub>R \* Q" + and "A\<^sub>R \* R" + and "A\<^sub>R \* K" + and "A\<^sub>P \* \" + and "A\<^sub>P \* R" + and "A\<^sub>P \* P" + and "A\<^sub>P \* M" + and "A\<^sub>Q \* R" + and "A\<^sub>Q \* M" + and "A\<^sub>R \* xvec" + and "xvec \* M" + +obtains K' where "\ \ \\<^sub>P \ R \K'\N\ \ R'" and "\ \ \\<^sub>P \ \\<^sub>R \ M \ K'" and "A\<^sub>R \* K'" +proof - + from \A\<^sub>R \* P\ \A\<^sub>R \* Q\ \A\<^sub>R \* A\<^sub>P\ \A\<^sub>R \* A\<^sub>Q\ FrP FrQ have "A\<^sub>R \* \\<^sub>P" and "A\<^sub>R \* \\<^sub>Q" + by(force dest: extractFrameFreshChain)+ + assume Assumptions: "\K'. \\ \ \\<^sub>P \ R \K'\N\ \ R'; \ \ \\<^sub>P \ \\<^sub>R \ M \ K'; A\<^sub>R \* K'\ \ thesis" + { + fix \'::'b + fix zvec::"name list" + assume "A\<^sub>R \* \'" + assume "A\<^sub>R \* zvec" + assume "A\<^sub>P \* zvec" + assume "zvec \* R" + assume "zvec \* P" + + assume A: "\ \ \\<^sub>Q \ \'" + with RTrans have "\' \ R \K\N\ \ R'" + by(rule statEqTransition) + moreover note FrR \distinct A\<^sub>R\ + moreover from \\ \ \\<^sub>Q \ \\<^sub>R \ M \ K\ have "(\ \ \\<^sub>Q) \ \\<^sub>R \ M \ K" + by(blast intro: statEqEnt Associativity AssertionStatEqSym) + with A have "\' \ \\<^sub>R \ M \ K" by(rule statEqEnt[OF Composition]) + moreover have "\A\<^sub>Q, \' \ \\<^sub>R\ \\<^sub>F \A\<^sub>Q, (\ \ \\<^sub>Q) \ \\<^sub>R\" using A + by(blast dest: frameIntComposition FrameStatEqTrans FrameStatEqSym) + with QimpP have "\A\<^sub>Q, \' \ \\<^sub>R\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\<^sub>R\" + by(force intro: FrameStatEqImpCompose) + moreover from PTrans have "distinct xvec" by(auto dest: boundOutputDistinct) + ultimately have "\K'. \ \ \\<^sub>P \ R \K'\N\ \ R' \ \ \ \\<^sub>P \ \\<^sub>R \ M \ K' \ zvec \* K' \ A\<^sub>R \* K'" + using PTrans FrP \A\<^sub>R \* K\ \A\<^sub>R \* \'\ \A\<^sub>R \* \\ \A\<^sub>R \* P\ \A\<^sub>R \* R\ \A\<^sub>R \* \\<^sub>P\ \A\<^sub>P \* R\ \A\<^sub>Q \* R\ \A\<^sub>P \* \\ + \A\<^sub>P \* P\ \A\<^sub>P \* M\ \A\<^sub>P \* zvec\ \A\<^sub>Q \* M\ \A\<^sub>R \* zvec\ \zvec \* R\ \zvec \* P\ \distinct A\<^sub>P\ + \A\<^sub>R \* A\<^sub>P\ \A\<^sub>R \* A\<^sub>Q\ \A\<^sub>R \* xvec\ \xvec \* M\ + proof(nominal_induct avoiding: \ P A\<^sub>P \\<^sub>P A\<^sub>Q zvec xvec arbitrary: M rule: inputFrameInduct) + case(cAlpha \' R K N R' A\<^sub>R \\<^sub>R p \ P A\<^sub>P \\<^sub>P A\<^sub>Q zvec xvec M) + have S: "set p \ set A\<^sub>R \ set (p \ A\<^sub>R)" by fact + from \\' \ (p \ \\<^sub>R) \ M \ K\ have "(p \ (\' \ (p \ \\<^sub>R))) \ (p \ M) \ (p \ K)" + by(rule chanEqClosed) + with \A\<^sub>R \* \'\ \(p \ A\<^sub>R) \* \'\ \A\<^sub>R \* K\ \(p \ A\<^sub>R) \* K\ S \distinctPerm p\ + have "\' \ \\<^sub>R \ (p \ M) \ K" by(simp add: eqvts) + moreover from \\ \ (p \ \\<^sub>R) \ P \M\\*xvec\\N\ \ P'\ S \A\<^sub>R \* P\ \(p \ A\<^sub>R) \* P\ \A\<^sub>R \* xvec\ \(p \ A\<^sub>R) \* xvec\ \xvec \* M\ + have "(p \ (\ \ (p \ \\<^sub>R))) \ P \(p \ M)\\*xvec\\N\ \ P'" + using outputPermFrameSubject by(auto simp add: residualInject) + with \A\<^sub>R \* \\ \(p \ A\<^sub>R) \* \\ S \distinctPerm p\ have "\ \ \\<^sub>R \ P \(p \ M)\\*xvec\\N\ \ P'" + by(simp add: eqvts) + moreover from \A\<^sub>P \* M\ have "(p \ A\<^sub>P) \* (p \ M)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>R \* A\<^sub>P\ \(p \ A\<^sub>R) \* A\<^sub>P\ S have "A\<^sub>P \* (p \ M)" by simp + moreover from \A\<^sub>Q \* M\ have "(p \ A\<^sub>Q) \* (p \ M)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>R \* A\<^sub>Q\ \(p \ A\<^sub>R) \* A\<^sub>Q\ S have "A\<^sub>Q \* (p \ M)" by simp + + moreover from \\A\<^sub>Q, \' \ (p \ \\<^sub>R)\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ (p \ \\<^sub>R)\\ + have "(p \ \A\<^sub>Q, \' \ (p \ \\<^sub>R)\) \\<^sub>F (p \ \A\<^sub>P, (\ \ \\<^sub>P) \ (p \ \\<^sub>R)\)" + by(rule FrameStatImpClosed) + with \A\<^sub>R \* A\<^sub>P\ \(p \ A\<^sub>R) \* A\<^sub>P\ \A\<^sub>R \* \'\ \(p \ A\<^sub>R) \* \'\ \A\<^sub>R \* \\<^sub>P\ \(p \ A\<^sub>R) \* \\<^sub>P\ \A\<^sub>R \* A\<^sub>Q\ + \(p \ A\<^sub>R) \* A\<^sub>Q\ \A\<^sub>R \* \\ \(p \ A\<^sub>R) \* \\ S \distinctPerm p\ + have "\A\<^sub>Q, \' \ \\<^sub>R\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\<^sub>R\" by(simp add: eqvts) + moreover from \xvec \* M\ have "(p \ xvec) \* (p \ M)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with S \A\<^sub>R \* xvec\ \(p \ A\<^sub>R) \* xvec\ have "xvec \* (p \ M)" by simp + ultimately obtain K' where "\ \ \\<^sub>P \ R \K'\N\ \ R'" and "\ \ \\<^sub>P \ \\<^sub>R \ (p \ M) \ K'" and "zvec \* K'" and "A\<^sub>R \* K'" + using cAlpha + by(auto simp del: freshChainSimps) + from \\ \ \\<^sub>P \ R \K'\N\ \ R'\ S \A\<^sub>R \* R\ \(p \ A\<^sub>R) \* R\ have "(p \ (\ \ \\<^sub>P)) \ R \(p \ K')\N\ \ R'" + by(elim inputPermFrameSubject) auto + with S \A\<^sub>R \* \\ \(p \ A\<^sub>R) \* \\ \A\<^sub>R \* \\<^sub>P\ \(p \ A\<^sub>R) \* \\<^sub>P\ have "\ \ \\<^sub>P \ R \(p \ K')\N\ \ R'" + by(simp add: eqvts) + moreover from \\ \ \\<^sub>P \ \\<^sub>R \ (p \ M) \ K'\ have "(p \ (\ \ \\<^sub>P \ \\<^sub>R))\ (p \ p \ M) \ (p \ K')" + by(rule chanEqClosed) + with S \A\<^sub>R \* \\ \(p \ A\<^sub>R) \* \\ \A\<^sub>R \* \\<^sub>P\ \(p \ A\<^sub>R) \* \\<^sub>P\ \distinctPerm p\ have "\ \ \\<^sub>P \ (p \ \\<^sub>R) \ M \ (p \ K')" + by(simp add: eqvts) + moreover from \zvec \* K'\ have "(p \ zvec) \* (p \ K')" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>R \* zvec\ \(p \ A\<^sub>R) \* zvec\ S have "zvec \* (p \ K')" by simp + moreover from \A\<^sub>R \* K'\ have "(p \ A\<^sub>R) \* (p \ K')" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + ultimately show ?case by blast + next + case(cInput \' M' K xvec N Tvec R \ P A\<^sub>P \\<^sub>P A\<^sub>Q zvec yvec M) + from \A\<^sub>P \* (M'\\*xvec N\.R)\ \A\<^sub>Q \* (M'\\*xvec N\.R)\ \zvec \* (M'\\*xvec N\.R)\ + have "A\<^sub>P \* M'" and "A\<^sub>Q \* M'" and "zvec \* M'" by simp+ + + from \\' \ M' \ K\ + have "\' \ \ \ M' \ K" + by(blast intro: statEqEnt Identity AssertionStatEqSym) + then have "\' \ \ \ M' \ M'" + by(blast intro: chanEqSym chanEqTrans) + with \A\<^sub>Q \* M'\ + have "(\A\<^sub>Q, \' \ \\) \\<^sub>F M' \ M'" + by(force intro: frameImpI) + + with \\A\<^sub>Q, \' \ \\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\\ + have "(\A\<^sub>P, (\ \ \\<^sub>P) \ \\) \\<^sub>F M' \ M'" + by(simp add: FrameStatImp_def) + with \A\<^sub>P \* M'\ have "(\ \ \\<^sub>P) \ \ \ M' \ M'" + by(force dest: frameImpE) + then have "\ \ \\<^sub>P \ M' \ M'" by(blast intro: statEqEnt Identity) + then have "\ \ \\<^sub>P \ M'\\*xvec N\.R \M'\(N[xvec::=Tvec])\ \ R[xvec::=Tvec]" + using \distinct xvec\ \set xvec \ supp N\ \length xvec = length Tvec\ + by(rule Input) + + moreover from \\' \ \ \ M \ K\ \\' \ \ \ M' \ K\ + have "\' \ \ \ M \ M'" by(metis chanEqSym chanEqTrans) + with \A\<^sub>Q \* M\ \A\<^sub>Q \* M'\ + have "(\A\<^sub>Q, \' \ \\) \\<^sub>F M \ M'" + by(force intro: frameImpI) + with \\A\<^sub>Q, \' \ \\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\ \ + have "(\A\<^sub>P, (\ \ \\<^sub>P) \ \\) \\<^sub>F M \ M'" + by(simp add: FrameStatImp_def) + with \A\<^sub>P \* M\ \A\<^sub>P \* M'\ have "(\ \ \\<^sub>P) \ \ \ M \ M'" + by(force dest: frameImpE) + then have "\ \ \\<^sub>P \ \ \ M \ M'" + by(metis statEqEnt Associativity) + ultimately show ?case using \zvec \* M'\ + by force + next + case(cCase \' R M' N R' \ Cs A\<^sub>R \\<^sub>R \ P A\<^sub>P \\<^sub>P A\<^sub>Q zvec xvec M) + from \guarded R\ \extractFrame R = \A\<^sub>R, \\<^sub>R\\ have "\\<^sub>R \ \" + by(metis guardedStatEq) + with \\' \ \ \ M \ M'\ have "\' \ \\<^sub>R \ M \ M'" + by(metis Identity Commutativity statEqEnt AssertionStatEqSym Composition) + moreover have "\A\<^sub>Q, \' \ \\<^sub>R\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\<^sub>R\" + proof - + from \\\<^sub>R \ \\ have "\A\<^sub>Q, \' \ \\<^sub>R\ \\<^sub>F \A\<^sub>Q, \' \ \\" + by(metis Identity Commutativity AssertionStatEqSym Composition frameResChainPres frameNilStatEq AssertionStatEqTrans) + moreover note \\A\<^sub>Q, \' \ \\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\\ + moreover from \\\<^sub>R \ \\ have "\A\<^sub>P, (\ \ \\<^sub>P) \ \\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\<^sub>R\" + by(metis Identity Commutativity AssertionStatEqSym Composition frameResChainPres frameNilStatEq AssertionStatEqTrans) + ultimately show ?thesis by(rule FrameStatEqImpCompose) + qed + moreover from \\ \ \ \ P \M\\*xvec\\N\ \ P'\ \\\<^sub>R \ \\ + have "\ \ \\<^sub>R \ P \M\\*xvec\\N\ \ P'" by(metis statEqTransition Identity Commutativity AssertionStatEqSym Composition) + moreover from \zvec \* (Cases Cs)\ \A\<^sub>P \* (Cases Cs)\ \A\<^sub>Q \* (Cases Cs)\ \(\, R) \ set Cs\ + have "A\<^sub>P \* R" and "A\<^sub>Q \* R" and "zvec \* R" and "A\<^sub>P \* \" and "A\<^sub>Q \* \" + by(auto dest: memFreshChain) + ultimately have "\K'. \ \ \\<^sub>P \ R \K'\N\ \ R'\ \ \ \\<^sub>P \ \\<^sub>R \ M \ K' \ zvec \* K' \ A\<^sub>R \* K'" using cCase + by(elim cCase) (assumption |simp)+ + then obtain K' where RTrans: "\ \ \\<^sub>P \ R \K'\N\ \ R'" + and MeqK': "\ \ \\<^sub>P \ \\<^sub>R \ M \ K'" and "zvec \* K'" and "A\<^sub>R \* K'" + by metis + note RTrans \(\, R) \ set Cs\ + moreover from \\' \ \\ have "\' \ \ \ \" by(blast intro: statEqEnt Identity AssertionStatEqSym) + with \A\<^sub>Q \* \\ have "(\A\<^sub>Q, \' \ \\) \\<^sub>F \" by(force intro: frameImpI) + with \\A\<^sub>Q, \' \ \\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\\ have "(\A\<^sub>P, (\ \ \\<^sub>P) \ \\) \\<^sub>F \" + by(simp add: FrameStatImp_def) + with \A\<^sub>P \* \\ have "(\ \ \\<^sub>P) \ \ \ \" by(force dest: frameImpE) + then have "\ \ \\<^sub>P \ \" by(blast intro: statEqEnt Identity) + ultimately have "\ \ \\<^sub>P \ Cases Cs \K'\N\ \ R'" using \guarded R\ by(rule Case) + moreover from MeqK' \\\<^sub>R \ \\ have "\ \ \\<^sub>P \ \ \ M \ K'" + by(metis Identity Commutativity statEqEnt AssertionStatEqSym Composition AssertionStatEqTrans) + ultimately show ?case using \zvec \* K'\ + by force + next + case(cPar1 \' \\<^sub>R\<^sub>2 R\<^sub>1 M' N R\<^sub>1' A\<^sub>R\<^sub>2 R\<^sub>2 A\<^sub>R\<^sub>1 \\<^sub>R\<^sub>1 \ P A\<^sub>P \\<^sub>P A\<^sub>Q zvec xvec M) + have FrR2: "extractFrame R\<^sub>2 = \A\<^sub>R\<^sub>2, \\<^sub>R\<^sub>2\" by fact + from \\' \ \\<^sub>R\<^sub>1 \ \\<^sub>R\<^sub>2 \ M \ M'\ have "(\' \ \\<^sub>R\<^sub>2) \ \\<^sub>R\<^sub>1 \ M \ M'" + by(metis statEqEnt Associativity Composition Commutativity) + moreover have "\A\<^sub>Q, (\' \ \\<^sub>R\<^sub>2) \ \\<^sub>R\<^sub>1\ \\<^sub>F \A\<^sub>P, ((\ \ \\<^sub>R\<^sub>2) \ \\<^sub>P) \ \\<^sub>R\<^sub>1\" + proof - + have "\A\<^sub>Q, (\' \ \\<^sub>R\<^sub>2) \ \\<^sub>R\<^sub>1\ \\<^sub>F \A\<^sub>Q, \' \ \\<^sub>R\<^sub>1 \ \\<^sub>R\<^sub>2\" + by(metis Associativity Composition Commutativity AssertionStatEqTrans AssertionStatEqSym frameNilStatEq frameResChainPres) + moreover note \\A\<^sub>Q, \' \ \\<^sub>R\<^sub>1 \ \\<^sub>R\<^sub>2\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\<^sub>R\<^sub>1 \ \\<^sub>R\<^sub>2\\ + moreover have "\A\<^sub>P, (\ \ \\<^sub>P) \ \\<^sub>R\<^sub>1 \ \\<^sub>R\<^sub>2\ \\<^sub>F \A\<^sub>P, ((\ \ \\<^sub>R\<^sub>2) \ \\<^sub>P) \ \\<^sub>R\<^sub>1\" + by(metis Associativity Composition Commutativity AssertionStatEqTrans AssertionStatEqSym frameNilStatEq frameResChainPres) + ultimately show ?thesis by(rule FrameStatEqImpCompose) + qed + moreover from \\ \ \\<^sub>R\<^sub>1 \ \\<^sub>R\<^sub>2 \ P \M\\*xvec\\N\ \ P'\ have "(\ \ \\<^sub>R\<^sub>2) \ \\<^sub>R\<^sub>1 \ P \M\\*xvec\\N\ \ P'" + by(metis statEqTransition Associativity Composition Commutativity) + moreover from \A\<^sub>R\<^sub>1 \* A\<^sub>P\ \A\<^sub>R\<^sub>2 \* A\<^sub>P\ \A\<^sub>P \* (R\<^sub>1 \ R\<^sub>2)\ \extractFrame R\<^sub>1 = \A\<^sub>R\<^sub>1, \\<^sub>R\<^sub>1\\ FrR2 have "A\<^sub>P \* \\<^sub>R\<^sub>1" and "A\<^sub>P \* \\<^sub>R\<^sub>2" + by(force dest: extractFrameFreshChain)+ + moreover note \distinct xvec\ + + ultimately have "\K'. (\ \ \\<^sub>R\<^sub>2) \ \\<^sub>P \ R\<^sub>1 \K'\N\ \ R\<^sub>1' \ (\ \ \\<^sub>R\<^sub>2) \ \\<^sub>P \ \\<^sub>R\<^sub>1 \ M \ K' \ (A\<^sub>R\<^sub>2@zvec) \* K' \ A\<^sub>R\<^sub>1 \* K'" using cPar1 + by(elim cPar1(6)[where ba=P and bb=A\<^sub>P and bd=A\<^sub>Q and bf=xvec]) auto + then obtain K' where RTrans: "(\ \ \\<^sub>R\<^sub>2) \ \\<^sub>P \ R\<^sub>1 \K'\N\ \ R\<^sub>1'" + and MeqK': "(\ \ \\<^sub>R\<^sub>2) \ \\<^sub>P \ \\<^sub>R\<^sub>1 \ M \ K'" and "A\<^sub>R\<^sub>2 \* K'" and "zvec \* K'" and "A\<^sub>R\<^sub>1 \* K'" + by force + + from RTrans have "(\ \ \\<^sub>P) \ \\<^sub>R\<^sub>2 \ R\<^sub>1 \K'\N\ \ R\<^sub>1'" + by(metis statEqTransition Associativity Composition Commutativity) + then have "\ \ \\<^sub>P \ (R\<^sub>1 \ R\<^sub>2) \K'\N\ \ (R\<^sub>1' \ R\<^sub>2)" using FrR2 \A\<^sub>R\<^sub>2 \* \\ \A\<^sub>R\<^sub>2 \* \\<^sub>P\ \A\<^sub>R\<^sub>2 \* K'\ \A\<^sub>R\<^sub>2 \* R\<^sub>1\ \A\<^sub>R\<^sub>2 \* N\ + by(force intro: Par1) + moreover from MeqK' have "\ \ \\<^sub>P \ \\<^sub>R\<^sub>1 \ \\<^sub>R\<^sub>2 \ M \ K'" + by(metis statEqEnt Associativity Composition Commutativity) + ultimately show ?case using \zvec \* K'\ \A\<^sub>R\<^sub>1 \* K'\ \A\<^sub>R\<^sub>2 \* K'\ + by force + next + case(cPar2 \' \\<^sub>R\<^sub>1 R\<^sub>2 M' N R\<^sub>2' A\<^sub>R\<^sub>1 R\<^sub>1 A\<^sub>R\<^sub>2 \\<^sub>R\<^sub>2 \ P A\<^sub>P \\<^sub>P A\<^sub>Q zvec xvec M) + have FrR1: "extractFrame R\<^sub>1 = \A\<^sub>R\<^sub>1, \\<^sub>R\<^sub>1\" by fact + from \\' \ \\<^sub>R\<^sub>1 \ \\<^sub>R\<^sub>2 \ M \ M'\ have "(\' \ \\<^sub>R\<^sub>1) \ \\<^sub>R\<^sub>2 \ M \ M'" + by(metis statEqEnt Associativity Composition Commutativity) + moreover have "\A\<^sub>Q, (\' \ \\<^sub>R\<^sub>1) \ \\<^sub>R\<^sub>2\ \\<^sub>F \A\<^sub>P, ((\ \ \\<^sub>R\<^sub>1) \ \\<^sub>P) \ \\<^sub>R\<^sub>2\" + proof - + have "\A\<^sub>Q, (\' \ \\<^sub>R\<^sub>1) \ \\<^sub>R\<^sub>2\ \\<^sub>F \A\<^sub>Q, \' \ \\<^sub>R\<^sub>1 \ \\<^sub>R\<^sub>2\" + by(metis Associativity Composition Commutativity AssertionStatEqTrans AssertionStatEqSym frameNilStatEq frameResChainPres) + moreover note \\A\<^sub>Q, \' \ \\<^sub>R\<^sub>1 \ \\<^sub>R\<^sub>2\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\<^sub>R\<^sub>1 \ \\<^sub>R\<^sub>2\\ + moreover have "\A\<^sub>P, (\ \ \\<^sub>P) \ \\<^sub>R\<^sub>1 \ \\<^sub>R\<^sub>2\ \\<^sub>F \A\<^sub>P, ((\ \ \\<^sub>R\<^sub>1) \ \\<^sub>P) \ \\<^sub>R\<^sub>2\" + by(metis Associativity Composition Commutativity AssertionStatEqTrans AssertionStatEqSym frameNilStatEq frameResChainPres) + ultimately show ?thesis by(rule FrameStatEqImpCompose) + qed + moreover from \\ \ \\<^sub>R\<^sub>1 \ \\<^sub>R\<^sub>2 \ P \M\\*xvec\\N\ \ P'\ have "(\ \ \\<^sub>R\<^sub>1) \ \\<^sub>R\<^sub>2 \ P \M\\*xvec\\N\ \ P'" + by(metis statEqTransition Associativity Composition Commutativity) + moreover from \A\<^sub>R\<^sub>1 \* A\<^sub>P\ \A\<^sub>R\<^sub>2 \* A\<^sub>P\ \A\<^sub>P \* (R\<^sub>1 \ R\<^sub>2)\ FrR1 \extractFrame R\<^sub>2 = \A\<^sub>R\<^sub>2, \\<^sub>R\<^sub>2\\ have "A\<^sub>P \* \\<^sub>R\<^sub>1" and "A\<^sub>P \* \\<^sub>R\<^sub>2" + by(force dest: extractFrameFreshChain)+ + ultimately have "\K'. (\ \ \\<^sub>R\<^sub>1) \ \\<^sub>P \ R\<^sub>2 \K'\N\ \ R\<^sub>2' \ (\ \ \\<^sub>R\<^sub>1) \ \\<^sub>P \ \\<^sub>R\<^sub>2 \ M \ K' \ (A\<^sub>R\<^sub>1@zvec) \* K' \ A\<^sub>R\<^sub>2 \* K'" using \distinct xvec\ cPar2 + by(elim cPar2(6)[where ba=P and bb=A\<^sub>P and bd=A\<^sub>Q and bf=xvec]) auto + then obtain K' where RTrans: "(\ \ \\<^sub>R\<^sub>1) \ \\<^sub>P \ R\<^sub>2 \K'\N\ \ R\<^sub>2'" + and MeqK': "(\ \ \\<^sub>R\<^sub>1) \ \\<^sub>P \ \\<^sub>R\<^sub>2 \ M \ K'" and "A\<^sub>R\<^sub>1 \* K'" and "zvec \* K'" and "A\<^sub>R\<^sub>2 \* K'" + by force + + from RTrans have "(\ \ \\<^sub>P) \ \\<^sub>R\<^sub>1 \ R\<^sub>2 \K'\N\ \ R\<^sub>2'" + by(metis statEqTransition Associativity Composition Commutativity) + then have "\ \ \\<^sub>P \ (R\<^sub>1 \ R\<^sub>2) \K'\N\ \ (R\<^sub>1 \ R\<^sub>2')" using FrR1 \A\<^sub>R\<^sub>1 \* \\ \A\<^sub>R\<^sub>1 \* \\<^sub>P\ \A\<^sub>R\<^sub>1 \* K'\ \A\<^sub>R\<^sub>1 \* R\<^sub>2\ \A\<^sub>R\<^sub>1 \* N\ + by(force intro: Par2) + moreover from MeqK' have "\ \ \\<^sub>P \ \\<^sub>R\<^sub>1 \ \\<^sub>R\<^sub>2 \ M \ K'" + by(metis statEqEnt Associativity Composition Commutativity) + ultimately show ?case using \zvec \* K'\ \A\<^sub>R\<^sub>1 \* K'\ \A\<^sub>R\<^sub>2 \* K'\ + by force + next + case(cScope \' R M' N R' x A\<^sub>R \\<^sub>R \ P A\<^sub>P \\<^sub>P A\<^sub>Q zvec xvec M) + then have "\K'. \ \ \\<^sub>P \ R \K'\N\ \ R' \ \ \ \\<^sub>P \ \\<^sub>R \ M \ K' \ zvec \* K' \ A\<^sub>R \* K'" + by(elim cScope(4)) (assumption | simp del: freshChainSimps)+ + then obtain K' where RTrans: "\ \ \\<^sub>P \ R \K'\N\ \ R'" + and MeqK': "\ \ \\<^sub>P \ \\<^sub>R \ M \ K'" and "zvec \* K'" and "A\<^sub>R \* K'" + by metis + from \A\<^sub>R \* A\<^sub>P\ \A\<^sub>P \* (\\x\R)\ \x \ A\<^sub>P\ \extractFrame R = \A\<^sub>R, \\<^sub>R\\ have "A\<^sub>P \* \\<^sub>R" + by(force dest: extractFrameFreshChain)+ + + from \\ \ \\<^sub>R \ P \M\\*xvec\\N\ \ P'\ have "\ \ \\<^sub>R \ P \ ROut M (\\*xvec\N \' P')" by(simp add: residualInject) + with \extractFrame P = \A\<^sub>P, \\<^sub>P\\ \distinct A\<^sub>P\ \x \ P\ \zvec \* P\ \A\<^sub>P \* \\<^sub>R\ \x \ A\<^sub>P\ \A\<^sub>P \* M\ \A\<^sub>P \* P\ \A\<^sub>P \* zvec\ \A\<^sub>P \* \\ \A\<^sub>P \* zvec\ \A\<^sub>R \* P\ \A\<^sub>R \* A\<^sub>P\ \xvec \* M\ \distinct xvec\ + obtain K'' where MeqK'': "(\ \ \\<^sub>R) \ \\<^sub>P \ M \ K''" and "x \ K''" and "A\<^sub>R \* K''" and "zvec \* K''" + by(elim outputObtainPrefix[where B="(x#A\<^sub>R@zvec)"]) (assumption | simp | force | metis freshChainSym)+ + + from MeqK'' MeqK' have KeqK'': "(\ \ \\<^sub>P) \ \\<^sub>R \ K' \ K''" + by(metis statEqEnt Associativity Composition Commutativity chanEqSym chanEqTrans) + with RTrans \extractFrame R = \A\<^sub>R, \\<^sub>R\\ \distinct A\<^sub>R\ \A\<^sub>R \* \\ \A\<^sub>R \* \\<^sub>P\ \A\<^sub>R \* K'\ \A\<^sub>R \* K''\ \A\<^sub>R \* R\ + have "\ \ \\<^sub>P \ R \K''\N\ \ R'" + by(elim inputRenameSubject) (assumption | force)+ + then have "\ \ \\<^sub>P \ \\x\R \K''\N\ \ \\x\R'" using \x \ \\ \x \ \\<^sub>P\ \x \ K''\ \x \ N\ + by(elim Scope) (assumption | force)+ + moreover from MeqK'' have "\ \ \\<^sub>P \ \\<^sub>R \ M \ K''" + by(metis statEqEnt Associativity Composition Commutativity) + ultimately show ?case using \zvec \* K''\ \x \ K''\ \A\<^sub>R \* K''\ + by force + next + case(cBang \' R M' N R' A\<^sub>R \\<^sub>R \ P A\<^sub>P \\<^sub>P A\<^sub>Q zvec xvec M) + from \guarded R\ \extractFrame R = \A\<^sub>R, \\<^sub>R\\ have "\\<^sub>R \ \" + by(metis guardedStatEq) + with \\' \ \ \ M \ M'\ have "\' \ \\<^sub>R \ \ \ M \ M'" + by(metis Identity Commutativity statEqEnt AssertionStatEqSym Composition) + moreover have "\A\<^sub>Q, \' \ \\<^sub>R \ \\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\<^sub>R \ \\" + proof - + from \\\<^sub>R \ \\ have "\A\<^sub>Q, \' \ \\<^sub>R \ \\ \\<^sub>F \A\<^sub>Q, \' \ \\" + by(metis Identity Commutativity AssertionStatEqSym Composition frameResChainPres frameNilStatEq AssertionStatEqTrans) + moreover note \\A\<^sub>Q, \' \ \\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\\ + moreover from \\\<^sub>R \ \\ have "\A\<^sub>P, (\ \ \\<^sub>P) \ \\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\<^sub>R \ \\" + by(metis Identity Commutativity AssertionStatEqSym Composition frameResChainPres frameNilStatEq AssertionStatEqTrans) + ultimately show ?thesis by(rule FrameStatEqImpCompose) + qed + moreover from \\ \ \ \ P \M\\*xvec\\N\ \ P'\ \\\<^sub>R \ \\ + have "\ \ \\<^sub>R \ \ \ P \M\\*xvec\\N\ \ P'" by(metis statEqTransition Identity Commutativity AssertionStatEqSym Composition) + ultimately have "\K'. \ \ \\<^sub>P \ R \ !R \K'\N\ \ R'\ \ \ \\<^sub>P \ \\<^sub>R \ \ \ M \ K' \ zvec \* K' \ A\<^sub>R \* K'" using cBang + by(elim cBang(5)) (assumption | simp del: freshChainSimps)+ + then obtain K' where RTrans: "\ \ \\<^sub>P \ R \ !R \K'\N\ \ R'" + and MeqK': "\ \ \\<^sub>P \ \\<^sub>R \ \ \ M \ K'" and "zvec \* K'" and "A\<^sub>R \* K'" + by metis + from RTrans \guarded R\ have "\ \ \\<^sub>P \ !R \K'\N\ \ R'" by(rule Bang) + moreover from MeqK' \\\<^sub>R \ \\ have "\ \ \\<^sub>P \ \ \ M \ K'" + by(metis Identity Commutativity statEqEnt AssertionStatEqSym Composition AssertionStatEqTrans) + ultimately show ?case using \zvec \* K'\ + by force + qed + } + note Goal = this + have "\ \ \\<^sub>Q \ \ \ \\<^sub>Q" by(simp add: AssertionStatEqRefl) + moreover from \A\<^sub>R \* \\ \A\<^sub>R \* \\<^sub>Q\ have "A\<^sub>R \* (\ \ \\<^sub>Q)" by force + ultimately have "\K'. \ \ \\<^sub>P \ R \K'\N\ \ R' \ \ \ \\<^sub>P \ \\<^sub>R \ M \ K' \ ([]::name list) \* K' \ A\<^sub>R \* K'" + by(intro Goal) (assumption | force)+ + with Assumptions show ?thesis + by blast +qed + +end + +end diff --git a/thys/Broadcast_Psi/Sim_Pres.thy b/thys/Broadcast_Psi/Sim_Pres.thy new file mode 100644 --- /dev/null +++ b/thys/Broadcast_Psi/Sim_Pres.thy @@ -0,0 +1,1168 @@ +theory Sim_Pres + imports Simulation +begin + +text \This file is a (heavily modified) variant of the theory {\it Psi\_Calculi.Sim\_Pres} +from~\cite{DBLP:journals/afp/Bengtson12}.\ + +context env begin + +lemma inputPres: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and Rel :: "('b \ ('a, 'b, 'c) psi \ ('a, 'b, 'c) psi) set" + and Q :: "('a, 'b, 'c) psi" + and M :: 'a + and xvec :: "name list" + and N :: 'a + +assumes PRelQ: "\Tvec. length xvec = length Tvec \ (\, P[xvec::=Tvec], Q[xvec::=Tvec]) \ Rel" + +shows "\ \ M\\*xvec N\.P \[Rel] M\\*xvec N\.Q" +proof - + { + fix \ Q' + assume "\ \ M\\*xvec N\.Q \\ \ Q'" + then have "\P'. \ \ M\\*xvec N\.P \\ \ P' \ (\, P', Q') \ Rel" + by(induct rule: inputCases) (auto intro: Input BrInput PRelQ) + } + then show ?thesis + by(auto simp add: simulation_def residual.inject psi.inject) +qed + +lemma outputPres: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and Rel :: "('b \ ('a, 'b, 'c) psi \ ('a, 'b, 'c) psi) set" + and Q :: "('a, 'b, 'c) psi" + and M :: 'a + and N :: 'a + +assumes PRelQ: "(\, P, Q) \ Rel" + +shows "\ \ M\N\.P \[Rel] M\N\.Q" +proof - + { + fix \ Q' + assume "\ \ M\N\.Q \\ \ Q'" + then have "\P'. \ \ M\N\.P \\ \ P' \ (\, P', Q') \ Rel" + by(induct rule: outputCases) (auto intro: Output BrOutput PRelQ) + } + then show ?thesis + by(auto simp add: simulation_def residual.inject psi.inject) +qed + +lemma casePres: + fixes \ :: 'b + and CsP :: "('c \ ('a, 'b, 'c) psi) list" + and Rel :: "('b \ ('a, 'b, 'c) psi \ ('a, 'b, 'c) psi) set" + and CsQ :: "('c \ ('a, 'b, 'c) psi) list" + and M :: 'a + and N :: 'a + +assumes PRelQ: "\\ Q. (\, Q) \ set CsQ \ \P. (\, P) \ set CsP \ guarded P \ (\, P, Q) \ Rel" + and Sim: "\\' R S. (\', R, S) \ Rel \ \' \ R \[Rel] S" + and "Rel \ Rel'" + +shows "\ \ Cases CsP \[Rel'] Cases CsQ" +proof - + { + fix \ Q' + assume "\ \ Cases CsQ \\ \ Q'" and "bn \ \* CsP" and "bn \ \* \" + then have "\P'. \ \ Cases CsP \\ \ P' \ (\, P', Q') \ Rel'" + proof(induct rule: caseCases) + case(cCase \ Q) + from \(\, Q) \ set CsQ\ obtain P where "(\, P) \ set CsP" and "guarded P" and "(\, P, Q) \ Rel" + by(metis PRelQ) + from \(\, P, Q) \ Rel\ have "\ \ P \[Rel] Q" + by(rule Sim) + moreover from \bn \ \* CsP\ \(\, P) \ set CsP\ have "bn \ \* P" + by(auto dest: memFreshChain) + moreover note \\ \ Q \\ \ Q'\ \bn \ \* \\ + ultimately obtain P' where PTrans: "\ \ P \\ \ P'" and P'RelQ': "(\, P', Q') \ Rel" + by(blast dest: simE) + from PTrans \(\, P) \ set CsP\ \\ \ \\ \guarded P\ have "\ \ Cases CsP \\ \ P'" + by(rule Case) + moreover from P'RelQ' \Rel \ Rel'\ have "(\, P', Q') \ Rel'" + by blast + ultimately show ?case + by blast + qed + } + then show ?thesis + by(auto simp add: simulation_def residual.inject psi.inject) +qed + +lemma resPres: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and Rel :: "('b \ ('a, 'b, 'c) psi \ ('a, 'b, 'c) psi) set" + and Q :: "('a, 'b, 'c) psi" + and x :: name + and Rel' :: "('b \ ('a, 'b, 'c) psi \ ('a, 'b, 'c) psi) set" + +assumes PSimQ: "\ \ P \[Rel] Q" + and "eqvt Rel'" + and "x \ \" + and "Rel \ Rel'" + and C1: "\\' R S yvec. \(\', R, S) \ Rel; yvec \* \'\ \ (\', \\*yvec\R, \\*yvec\S) \ Rel'" + +shows "\ \ \\x\P \[Rel'] \\x\Q" +proof - + note \eqvt Rel'\ \x \ \\ + moreover have "x \ \\x\P" and "x \ \\x\Q" by(simp add: abs_fresh)+ + ultimately show ?thesis + proof(induct rule: simIFresh[where C="()"]) + case(cSim \ Q') + from \bn \ \* \\x\P\ \bn \ \* \\x\Q\ \x \ \\ have "bn \ \* P" and "bn \ \* Q" by simp+ + from \\ \ \\x\Q \\ \ Q'\ \x \ \\ \x \ \\ \x \ Q'\ \bn \ \* \\ \bn \ \* Q\ \bn \ \* subject \\ + \bn \ \* \\ \bn \ \* P\ \x \ \\ + show ?case + proof(induct rule: resCases[where C="P"]) + case(cOpen M xvec1 xvec2 y N Q') + from \bn (M\\*(xvec1@y#xvec2)\\N\) \* \\ have "xvec1 \* \" and "y \ \" and "xvec2 \* \" by simp+ + from \bn (M\\*(xvec1@y#xvec2)\\N\) \* P\ have "xvec1 \* P" and "y \ P" and "xvec2 \* P" by simp+ + from \x \ (M\\*(xvec1@y#xvec2)\\N\)\ have "x \ xvec1" and "x \ y" and "x \ xvec2" and "x \ M" by simp+ + from PSimQ \\ \ Q \M\\*(xvec1@xvec2)\\([(x, y)] \ N)\ \ ([(x, y)] \ Q')\ + \xvec1 \* \\ \xvec2 \* \\ \xvec1 \* P\ \xvec2 \* P\ + obtain P' where PTrans: "\ \ P \M\\*(xvec1@xvec2)\\([(x, y)] \ N)\ \ P'" and P'RelQ': "(\, P', ([(x, y)] \ Q')) \ Rel" + by(force dest: simE) + from \y \ supp N\ \x \ y\ have "x \ supp([(x, y)] \ N)" + apply - + apply(drule pt_set_bij2[OF pt_name_inst, OF at_name_inst, where pi="[(x, y)]"]) + by(simp add: eqvts calc_atm) + with PTrans \x \ \\ \x \ M\ \x \ xvec1\ \x \ xvec2\ + have "\ \ \\x\P \M\\*(xvec1@x#xvec2)\\([(x, y)] \ N)\ \ P'" + by(metis Open) + then have "([(x, y)] \ \) \ ([(x, y)] \ \\x\P) \([(x, y)] \ (M\\*(xvec1@x#xvec2)\\([(x, y)] \ N)\ \ P'))" + by(rule eqvts) + with \x \ \\ \y \ \\ \y \ P\ \x \ M\ \y \ M\ \x \ xvec1\ \y \ xvec1\ \x \ xvec2\ \y \ xvec2\ \x \ y\ + have "\ \ \\x\P \M\\*(xvec1@y#xvec2)\\N\ \ ([(x, y)] \ P')" by(simp add: eqvts calc_atm alphaRes) + moreover from P'RelQ' \Rel \ Rel'\ \eqvt Rel'\ have "([(y, x)] \ \, [(y, x)] \ P', [(y, x)] \ [(x, y)] \ Q') \ Rel'" + by(force simp add: eqvt_def) + with \x \ \\ \y \ \\ have "(\, [(x, y)] \ P', Q') \ Rel'" by(simp add: name_swap) + ultimately show ?case by blast + next + case(cBrOpen M xvec1 xvec2 y N Q') + from \bn (\M\\*(xvec1@y#xvec2)\\N\) \* \\ have "xvec1 \* \" and "y \ \" and "xvec2 \* \" by simp+ + from \bn (\M\\*(xvec1@y#xvec2)\\N\) \* P\ have "xvec1 \* P" and "y \ P" and "xvec2 \* P" by simp+ + from \x \ (\M\\*(xvec1@y#xvec2)\\N\)\ have "x \ xvec1" and "x \ y" and "x \ xvec2" and "x \ M" by simp+ + from PSimQ \\ \ Q \\M\\*(xvec1@xvec2)\\([(x, y)] \ N)\ \ ([(x, y)] \ Q')\ + \xvec1 \* \\ \xvec2 \* \\ \xvec1 \* P\ \xvec2 \* P\ + obtain P' where PTrans: "\ \ P \\M\\*(xvec1@xvec2)\\([(x, y)] \ N)\ \ P'" and P'RelQ': "(\, P', ([(x, y)] \ Q')) \ Rel" + by(force dest: simE) + from \y \ supp N\ \x \ y\ have "x \ supp([(x, y)] \ N)" + apply - + apply(drule pt_set_bij2[OF pt_name_inst, OF at_name_inst, where pi="[(x, y)]"]) + by(simp add: eqvts calc_atm) + with PTrans \x \ \\ \x \ M\ \x \ xvec1\ \x \ xvec2\ + have "\ \ \\x\P \\M\\*(xvec1@x#xvec2)\\([(x, y)] \ N)\ \ P'" + by(metis BrOpen) + then have "([(x, y)] \ \) \ ([(x, y)] \ \\x\P) \([(x, y)] \ (\M\\*(xvec1@x#xvec2)\\([(x, y)] \ N)\ \ P'))" + by(rule eqvts) + with \x \ \\ \y \ \\ \y \ P\ \x \ M\ \y \ M\ \x \ xvec1\ \y \ xvec1\ \x \ xvec2\ \y \ xvec2\ \x \ y\ + have "\ \ \\x\P \\M\\*(xvec1@y#xvec2)\\N\ \ ([(x, y)] \ P')" by(simp add: eqvts calc_atm alphaRes) + moreover from P'RelQ' \Rel \ Rel'\ \eqvt Rel'\ have "([(y, x)] \ \, [(y, x)] \ P', [(y, x)] \ [(x, y)] \ Q') \ Rel'" + by(force simp add: eqvt_def) + with \x \ \\ \y \ \\ have "(\, [(x, y)] \ P', Q') \ Rel'" by(simp add: name_swap) + ultimately show ?case by blast + next + case(cRes Q') + from PSimQ \\ \ Q \\ \ Q'\ \bn \ \* \\ \bn \ \* P\ + obtain P' where PTrans: "\ \ P \\ \ P'" and P'RelQ': "(\, P', Q') \ Rel" + by(blast dest: simE) + from PTrans \x \ \\ \x \ \\ have "\ \ \\x\P \\ \ \\x\P'" + by(rule Scope) + moreover from P'RelQ' \x \ \\ have "(\, \\*[x]\P', \\*[x]\Q') \ Rel'" + by(intro C1[where yvec="[x]"]) simp+ + moreover then have "(\, \\x\P', \\x\Q') \ Rel'" + by (metis resChain.base resChain.step) + ultimately show ?case by blast + next + case(cBrClose M xvec N Q') + from PSimQ \\ \ Q \ \M\\*xvec\\N\ \ Q'\ \xvec \* \\ \xvec \* P\ + obtain P' where PTrans: "\ \ P \\M\\*xvec\\N\ \ P'" and P'RelQ': "(\, P', Q') \ Rel" + by(force dest: simE) + with \x \ \\ \xvec \* \\ + have "(x#xvec) \* \" by simp + + from P'RelQ' + have "(\, \\*(x#xvec)\P', \\*(x#xvec)\Q') \ Rel'" + using \(x#xvec) \* \\ + by(rule C1) + then have "(\, \\x\(\\*xvec\P'), \\x\(\\*xvec\Q')) \ Rel'" + by simp + + moreover from \\ \ P \ \M\\*xvec\\N\ \ P'\ \x \ supp M\ \x \ \\ + have "\ \ \\x\P \ \ \ \\x\(\\*xvec\P')" + by(rule BrClose) + ultimately show ?case + by blast + qed + qed +qed + +lemma resChainPres: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and Rel :: "('b \ ('a, 'b, 'c) psi \ ('a, 'b, 'c) psi) set" + and Q :: "('a, 'b, 'c) psi" + and xvec :: "name list" + +assumes PSimQ: "\ \ P \[Rel] Q" + and "eqvt Rel" + and "xvec \* \" + and C1: "\\' R S yvec. \(\', R, S) \ Rel; yvec \* \'\ \ (\', \\*yvec\R, \\*yvec\S) \ Rel" + +shows "\ \ \\*xvec\P \[Rel] \\*xvec\Q" + using \xvec \* \\ +proof(induct xvec) + case Nil + from PSimQ show ?case by simp +next + case(Cons x xvec) + from \(x#xvec) \* \\ have "x \ \" and "xvec \* \" by simp+ + from \xvec \* \\ have "\ \ \\*xvec\P \[Rel] \\*xvec\Q" by(rule Cons) + moreover note \eqvt Rel\ \x \ \\ + moreover have "Rel \ Rel" by simp + ultimately have "\ \ \\x\(\\*xvec\P) \[Rel] \\x\(\\*xvec\Q)" using C1 + by(rule resPres) + then show ?case by simp +qed + +lemma parPres: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and Rel :: "('b \ ('a, 'b, 'c) psi \ ('a, 'b, 'c) psi) set" + 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 PRelQ: "\A\<^sub>R \\<^sub>R. \extractFrame R = \A\<^sub>R, \\<^sub>R\; A\<^sub>R \* \; A\<^sub>R \* P; A\<^sub>R \* Q\ \ (\ \ \\<^sub>R, P, Q) \ Rel" + and Eqvt: "eqvt Rel" + and Eqvt': "eqvt Rel'" + +and StatImp: "\\' S T. (\', S, T) \ Rel \ insertAssertion (extractFrame T) \' \\<^sub>F insertAssertion (extractFrame S) \'" +and Sim: "\\' S T. (\', S, T) \ Rel \ \' \ S \[Rel] T" +and Ext: "\\' S T \''. \(\', S, T) \ Rel\ \ (\' \ \'', S, T) \ Rel" + + +and C1: "\\' S T A\<^sub>U \\<^sub>U U. \(\' \ \\<^sub>U, S, T) \ Rel; extractFrame U = \A\<^sub>U, \\<^sub>U\; A\<^sub>U \* \'; A\<^sub>U \* S; A\<^sub>U \* T\ \ (\', S \ U, T \ U) \ Rel'" +and C2: "\\' S T xvec. \(\', S, T) \ Rel'; xvec \* \'\ \ (\', \\*xvec\S, \\*xvec\T) \ Rel'" +and C3: "\\' S T \''. \(\', S, T) \ Rel; \' \ \''\ \ (\'', S, T) \ Rel" + +shows "\ \ P \ R \[Rel'] Q \ R" + using Eqvt' +proof(induct rule: simI[of _ _ _ _ "()"]) + case(cSim \ QR) + from \bn \ \* (P \ R)\ \bn \ \* (Q \ R)\ + have "bn \ \* P" and "bn \ \* Q" and "bn \ \* R" + by simp+ + from \\ \ Q \ R \\ \ QR\ \bn \ \* \\ \bn \ \* Q\ \bn \ \* R\ \bn \ \* subject \\ + show ?case + proof(induct rule: parCases[where C = "(P, R)"]) + case(cPar1 Q' A\<^sub>R \\<^sub>R) + from \A\<^sub>R \* (P, R)\ have "A\<^sub>R \* P" by simp + have FrR: "extractFrame R = \A\<^sub>R, \\<^sub>R\" by fact + from \A\<^sub>R \* \\ \bn \ \* R\ FrR have "bn \ \* \\<^sub>R" + by(auto dest: extractFrameFreshChain) + from FrR \A\<^sub>R \* \\ \A\<^sub>R \* P\ \A\<^sub>R \* Q\ have "\ \ \\<^sub>R \ P \[Rel] Q" + by(blast intro: Sim PRelQ) + moreover have QTrans: "\ \ \\<^sub>R \ Q \\ \ Q'" by fact + ultimately obtain P' where PTrans: "\ \ \\<^sub>R \ P \\ \ P'" + and P'RelQ': "(\ \ \\<^sub>R, P', Q') \ Rel" + using \bn \ \* \\ \bn \ \* \\<^sub>R\ \bn \ \* P\ + by(force dest: simE) + from PTrans QTrans \A\<^sub>R \* P\ \A\<^sub>R \* Q\ \A\<^sub>R \* \\ \bn \ \* subject \\ \distinct(bn \)\ have "A\<^sub>R \* P'" and "A\<^sub>R \* Q'" + by(blast dest: freeFreshChainDerivative)+ + from PTrans \bn \ \* R\ FrR \A\<^sub>R \* \\ \A\<^sub>R \* P\ \A\<^sub>R \* \\ have "\ \ P \ R \\ \ (P' \ R)" + by(metis Par1) + moreover from P'RelQ' FrR \A\<^sub>R \* \\ \A\<^sub>R \* P'\ \A\<^sub>R \* Q'\ have "(\, P' \ R, Q' \ R) \ Rel'" by(rule C1) + ultimately show ?case by blast + next + case(cPar2 R' A\<^sub>Q \\<^sub>Q) + from \A\<^sub>Q \* (P, R)\ have "A\<^sub>Q \* P" and "A\<^sub>Q \* R" by simp+ + obtain A\<^sub>P \\<^sub>P where FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" and "A\<^sub>P \* (\, A\<^sub>Q, \\<^sub>Q, \, R)" + by(rule freshFrame) + then have "A\<^sub>P \* \" and "A\<^sub>P \* A\<^sub>Q" and "A\<^sub>P \* \\<^sub>Q" and "A\<^sub>P \* \" and "A\<^sub>P \* R" + by simp+ + + have FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" by fact + from \A\<^sub>Q \* P\ FrP \A\<^sub>P \* A\<^sub>Q\ have "A\<^sub>Q \* \\<^sub>P" + by(auto dest: extractFrameFreshChain) + + from FrP FrQ \bn \ \* P\ \bn \ \* Q\ \A\<^sub>P \* \\ \A\<^sub>Q \* \\ + have "bn \ \* \\<^sub>P" and "bn \ \* \\<^sub>Q" + by(force dest: extractFrameFreshChain)+ + + obtain A\<^sub>R \\<^sub>R where FrR: "extractFrame R = \A\<^sub>R, \\<^sub>R\" and "A\<^sub>R \* (\, P, Q, A\<^sub>Q, A\<^sub>P, \\<^sub>Q, \\<^sub>P, \, R)" and "distinct A\<^sub>R" + by(rule freshFrame) + then have "A\<^sub>R \* \" and "A\<^sub>R \* P" and "A\<^sub>R \* Q" and "A\<^sub>R \* A\<^sub>Q" and "A\<^sub>R \* A\<^sub>P" and "A\<^sub>R \* \\<^sub>Q" and "A\<^sub>R \* \\<^sub>P" and "A\<^sub>R \* \" and "A\<^sub>R \* R" + by simp+ + + from \A\<^sub>Q \* R\ FrR \A\<^sub>R \* A\<^sub>Q\ have "A\<^sub>Q \* \\<^sub>R" + by(auto dest: extractFrameFreshChain) + from \A\<^sub>P \* R\ \A\<^sub>R \* A\<^sub>P\ FrR have "A\<^sub>P \* \\<^sub>R" + by(auto dest: extractFrameFreshChain) + + have RTrans: "\ \ \\<^sub>Q \ R \\ \ R'" by fact + moreover have "\A\<^sub>Q, (\ \ \\<^sub>Q) \ \\<^sub>R\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\<^sub>R\" + proof - + have "\A\<^sub>Q, (\ \ \\<^sub>Q) \ \\<^sub>R\ \\<^sub>F \A\<^sub>Q, (\ \ \\<^sub>R) \ \\<^sub>Q\" + by(metis frameIntAssociativity Commutativity FrameStatEqTrans frameIntCompositionSym FrameStatEqSym) + moreover from FrR \A\<^sub>R \* \\ \A\<^sub>R \* P\ \A\<^sub>R \* Q\ + have "(insertAssertion (extractFrame Q) (\ \ \\<^sub>R)) \\<^sub>F (insertAssertion (extractFrame P) (\ \ \\<^sub>R))" + by(blast intro: PRelQ StatImp) + with FrP FrQ \A\<^sub>P \* \\ \A\<^sub>Q \* \\ \A\<^sub>P \* \\<^sub>R\ \A\<^sub>Q \* \\<^sub>R\ + have "\A\<^sub>Q, (\ \ \\<^sub>R) \ \\<^sub>Q\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>R) \ \\<^sub>P\" using freshCompChain by auto + moreover have "\A\<^sub>P, (\ \ \\<^sub>R) \ \\<^sub>P\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\<^sub>R\" + by(metis frameIntAssociativity Commutativity FrameStatEqTrans frameIntCompositionSym frameIntAssociativity[THEN FrameStatEqSym]) + ultimately show ?thesis + by(rule FrameStatEqImpCompose) + qed + + ultimately have "\ \ \\<^sub>P \ R \\ \ R'" + using \A\<^sub>P \* \\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>Q \* \\ \A\<^sub>Q \* \\<^sub>P\ \A\<^sub>P \* R\ \A\<^sub>Q \* R\ \A\<^sub>P \* \\ \A\<^sub>Q \* \\ + \A\<^sub>R \* A\<^sub>P\ \A\<^sub>R \* A\<^sub>Q\ \A\<^sub>R \* \\<^sub>P\ \A\<^sub>R \* \\<^sub>Q\ \A\<^sub>R \* \\ FrR \distinct A\<^sub>R\ + by(force intro: transferFrame) + with \bn \ \* P\ \A\<^sub>P \* \\ \A\<^sub>P \* R\ \A\<^sub>P \* \\ FrP have "\ \ P \ R \\ \ (P \ R')" + by(force intro: Par2) + moreover obtain A\<^sub>R' \\<^sub>R' where "extractFrame R' = \A\<^sub>R', \\<^sub>R'\" and "A\<^sub>R' \* \" and "A\<^sub>R' \* P" and "A\<^sub>R' \* Q" + by(rule freshFrame[where C="(\, P, Q)"]) auto + + moreover from RTrans FrR \distinct A\<^sub>R\ \A\<^sub>R \* \\ \A\<^sub>R \* P\ \A\<^sub>R \* Q\ \A\<^sub>R \* R\ \A\<^sub>R \* \\ \bn \ \* \\ \bn \ \* P\ \bn \ \* Q\ \bn \ \* R\ \bn \ \* subject \\ \distinct(bn \)\ + obtain p \' A\<^sub>R' \\<^sub>R' where S: "set p \ set(bn \) \ set(bn(p \ \))" and "(p \ \\<^sub>R) \ \' \ \\<^sub>R'" and FrR': "extractFrame R' = \A\<^sub>R', \\<^sub>R'\" + and "bn(p \ \) \* R" and "bn(p \ \) \* \" and "bn(p \ \) \* P" and "bn(p \ \) \* Q" and "bn(p \ \) \* R" + and "A\<^sub>R' \* \" and "A\<^sub>R' \* P" and "A\<^sub>R' \* Q" + apply - + apply (rule expandFrame[where C="(\, P, Q, R)" and C'="(\, P, Q, R)"]) + apply simp+ + done + + from \A\<^sub>R \* \\ have "(p \ A\<^sub>R) \* (p \ \)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \bn \ \* \\ \bn(p \ \) \* \\ S have "(p \ A\<^sub>R) \* \" by simp + from \A\<^sub>R \* P\ have "(p \ A\<^sub>R) \* (p \ P)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \bn \ \* P\ \bn(p \ \) \* P\ S have "(p \ A\<^sub>R) \* P" by simp + from \A\<^sub>R \* Q\ have "(p \ A\<^sub>R) \* (p \ Q)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \bn \ \* Q\ \bn(p \ \) \* Q\ S have "(p \ A\<^sub>R) \* Q" by simp + + from FrR have "(p \ extractFrame R) = p \ \A\<^sub>R, \\<^sub>R\" by simp + with \bn \ \* R\ \bn(p \ \) \* R\ S have "extractFrame R = \(p \ A\<^sub>R), (p \ \\<^sub>R)\" + by(simp add: eqvts) + + with \(p \ A\<^sub>R) \* \\ \(p \ A\<^sub>R) \* P\ \(p \ A\<^sub>R) \* Q\ have "(\ \ (p \ \\<^sub>R), P, Q) \ Rel" + by(metis PRelQ) + + then have "((\ \ (p \ \\<^sub>R)) \ \', P, Q) \ Rel" by(rule Ext) + with \(p \ \\<^sub>R) \ \' \ \\<^sub>R'\ have "(\ \ \\<^sub>R', P, Q) \ Rel" by(blast intro: C3 Associativity compositionSym) + with FrR' \A\<^sub>R' \* \\ \A\<^sub>R' \* P\ \A\<^sub>R' \* Q\ have "(\, P \ R', Q \ R') \ Rel'" + by(metis C1) + ultimately show ?case by blast + next + case(cComm1 \\<^sub>R M N Q' A\<^sub>Q \\<^sub>Q K xvec R' A\<^sub>R) + have FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" by fact + from \A\<^sub>Q \* (P, R)\ have "A\<^sub>Q \* P" and "A\<^sub>Q \* R" by simp+ + + have FrR: "extractFrame R = \A\<^sub>R, \\<^sub>R\" by fact + from \A\<^sub>R \* (P, R)\ have "A\<^sub>R \* P" and "A\<^sub>R \* R" by simp+ + + from \xvec \* (P, R)\ have "xvec \* P" and "xvec \* R" by simp+ + + obtain A\<^sub>P \\<^sub>P where FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" and "A\<^sub>P \* (\, A\<^sub>Q, \\<^sub>Q, A\<^sub>R, M, N, K, R, P, xvec)" and "distinct A\<^sub>P" + by(rule freshFrame) + then have "A\<^sub>P \* \" and "A\<^sub>P \* A\<^sub>Q" and "A\<^sub>P \* \\<^sub>Q" and "A\<^sub>P \* M" and "A\<^sub>P \* R" + and "A\<^sub>P \* N" and "A\<^sub>P \* K" and "A\<^sub>P \* A\<^sub>R" and "A\<^sub>P \* P" and "A\<^sub>P \* xvec" + by simp+ + + have QTrans: "\ \ \\<^sub>R \ Q \M\N\ \ Q'" and RTrans: "\ \ \\<^sub>Q \ R \K\\*xvec\\N\ \ R'" + and MeqK: "\ \ \\<^sub>Q \ \\<^sub>R \ M \K" by fact+ + + from FrP FrR \A\<^sub>Q \* P\ \A\<^sub>P \* R\ \A\<^sub>R \* P\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* A\<^sub>R\ \A\<^sub>P \* xvec\ \xvec \* P\ + have "A\<^sub>P \* \\<^sub>R" and "A\<^sub>Q \* \\<^sub>P" and "A\<^sub>R \* \\<^sub>P" and "xvec \* \\<^sub>P" + by(auto dest!: extractFrameFreshChain)+ + + from RTrans FrR \distinct A\<^sub>R\ \A\<^sub>R \* R\ \A\<^sub>R \* xvec\ \xvec \* R\ \xvec \* Q\ \xvec \* \\ \xvec \* \\<^sub>Q\ \A\<^sub>R \* Q\ + \A\<^sub>R \* \\ \A\<^sub>R \* \\<^sub>Q\ \xvec \* K\ \A\<^sub>R \* K\ \A\<^sub>R \* N\ \A\<^sub>R \* R\ \xvec \* R\ \A\<^sub>R \* P\ \xvec \* P\ \A\<^sub>P \* A\<^sub>R\ \A\<^sub>P \* xvec\ + \A\<^sub>Q \* A\<^sub>R\ \A\<^sub>Q \* xvec\ \A\<^sub>R \* \\<^sub>P\ \xvec \* \\<^sub>P\ \distinct xvec\ \xvec \* M\ + obtain p \' A\<^sub>R' \\<^sub>R' where S: "set p \ set xvec \ set(p \ xvec)" and FrR': "extractFrame R' = \A\<^sub>R', \\<^sub>R'\" + and "(p \ \\<^sub>R) \ \' \ \\<^sub>R'" and "A\<^sub>R' \* Q" and "A\<^sub>R' \* \" and "(p \ xvec) \* \" + and "(p \ xvec) \* Q" and "(p \ xvec) \* \\<^sub>Q" and "(p \ xvec) \* K" and "(p \ xvec) \* R" + and "(p \ xvec) \* P" and "(p \ xvec) \* A\<^sub>P" and "(p \ xvec) \* A\<^sub>Q" and "(p \ xvec) \* \\<^sub>P" + and "A\<^sub>R' \* P" and "A\<^sub>R' \* N" + by (subst expandFrame[where C="(\, Q, \\<^sub>Q, K, R, P, A\<^sub>P, A\<^sub>Q, \\<^sub>P)" and C'="(\, Q, \\<^sub>Q, K, R, P, A\<^sub>P, A\<^sub>Q, \\<^sub>P)"]) simp+ + + from \A\<^sub>R \* \\ have "(p \ A\<^sub>R) \* (p \ \)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \xvec \* \\ \(p \ xvec) \* \\ S have "(p \ A\<^sub>R) \* \" by simp + from \A\<^sub>R \* P\ have "(p \ A\<^sub>R) \* (p \ P)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \xvec \* P\ \(p \ xvec) \* P\ S have "(p \ A\<^sub>R) \* P" by simp + from \A\<^sub>R \* Q\ have "(p \ A\<^sub>R) \* (p \ Q)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \xvec \* Q\ \(p \ xvec) \* Q\ S have "(p \ A\<^sub>R) \* Q" by simp + from \A\<^sub>R \* R\ have "(p \ A\<^sub>R) \* (p \ R)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \xvec \* R\ \(p \ xvec) \* R\ S have "(p \ A\<^sub>R) \* R" by simp + from \A\<^sub>R \* K\ have "(p \ A\<^sub>R) \* (p \ K)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \xvec \* K\ \(p \ xvec) \* K\ S have "(p \ A\<^sub>R) \* K" by simp + + from \A\<^sub>P \* xvec\ \(p \ xvec) \* A\<^sub>P\ \A\<^sub>P \* M\ S have "A\<^sub>P \* (p \ M)" by(simp add: freshChainSimps) + from \A\<^sub>Q \* xvec\ \(p \ xvec) \* A\<^sub>Q\ \A\<^sub>Q \* M\ S have "A\<^sub>Q \* (p \ M)" by(simp add: freshChainSimps) + from \A\<^sub>P \* xvec\ \(p \ xvec) \* A\<^sub>P\ \A\<^sub>P \* A\<^sub>R\ S have "(p \ A\<^sub>R) \* A\<^sub>P" by(simp add: freshChainSimps) + from \A\<^sub>Q \* xvec\ \(p \ xvec) \* A\<^sub>Q\ \A\<^sub>Q \* A\<^sub>R\ S have "(p \ A\<^sub>R) \* A\<^sub>Q" by(simp add: freshChainSimps) + + from QTrans S \xvec \* Q\ \(p \ xvec) \* Q\ have "(p \ (\ \ \\<^sub>R)) \ Q \ (p \ M)\N\ \ Q'" + using inputPermFrameSubject name_list_set_fresh by blast + with \xvec \* \\ \(p \ xvec) \* \\ S have QTrans: "(\ \ (p \ \\<^sub>R)) \ Q \ (p \ M)\N\ \ Q'" + by(simp add: eqvts) + + from FrR have "(p \ extractFrame R) = p \ \A\<^sub>R, \\<^sub>R\" by simp + with \xvec \* R\ \(p \ xvec) \* R\ S have FrR: "extractFrame R = \(p \ A\<^sub>R), (p \ \\<^sub>R)\" + by(simp add: eqvts) + + note RTrans FrR + moreover from FrR \(p \ A\<^sub>R) \* \\ \(p \ A\<^sub>R) \* P\ \(p \ A\<^sub>R) \* Q\ have "\ \ (p \ \\<^sub>R) \ P \[Rel] Q" + by(metis Sim PRelQ) + with QTrans obtain P' where PTrans: "\ \ (p \ \\<^sub>R) \ P \(p \ M)\N\ \ P'" and P'RelQ': "(\ \ (p \ \\<^sub>R), P', Q') \ Rel" + by(force dest: simE) + from PTrans QTrans \A\<^sub>R' \* P\ \A\<^sub>R' \* Q\ \A\<^sub>R' \* N\ have "A\<^sub>R' \* P'" and "A\<^sub>R' \* Q'" + by(blast dest: inputFreshChainDerivative)+ + + note PTrans + moreover from MeqK have "(p \ (\ \ \\<^sub>Q \ \\<^sub>R)) \ (p \ M) \ (p \ K)" by(rule chanEqClosed) + with \xvec \* \\ \(p \ xvec) \* \\ \xvec \* \\<^sub>Q\ \(p \ xvec) \* \\<^sub>Q\ \xvec \* K\ \(p \ xvec) \* K\ S + have MeqK: "\ \ \\<^sub>Q \ (p \ \\<^sub>R) \ (p \ M) \ K" by(simp add: eqvts) + + moreover have "\A\<^sub>Q, (\ \ \\<^sub>Q) \ (p \ \\<^sub>R)\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ (p \ \\<^sub>R)\" + proof - + have "\A\<^sub>P, (\ \ (p \ \\<^sub>R)) \ \\<^sub>P\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ (p \ \\<^sub>R)\" + by(metis frameResChainPres frameNilStatEq Commutativity AssertionStatEqTrans Composition Associativity) + moreover from FrR \(p \ A\<^sub>R) \* \\ \(p \ A\<^sub>R) \* P\ \(p \ A\<^sub>R) \* Q\ + have "(insertAssertion (extractFrame Q) (\ \ (p \ \\<^sub>R))) \\<^sub>F (insertAssertion (extractFrame P) (\ \ (p \ \\<^sub>R)))" + by(metis PRelQ StatImp) + with FrP FrQ \A\<^sub>P \* \\ \A\<^sub>Q \* \\ \A\<^sub>P \* \\<^sub>R\ \A\<^sub>Q \* \\<^sub>R\ \A\<^sub>P \* xvec\ \(p \ xvec) \* A\<^sub>P\ \A\<^sub>Q \* xvec\ \(p \ xvec) \* A\<^sub>Q\ S + have "\A\<^sub>Q, (\ \ (p \ \\<^sub>R)) \ \\<^sub>Q\ \\<^sub>F \A\<^sub>P, (\ \ (p \ \\<^sub>R)) \ \\<^sub>P\" using freshCompChain + by(simp add: freshChainSimps) + moreover have "\A\<^sub>Q, (\ \ \\<^sub>Q) \ (p \ \\<^sub>R)\ \\<^sub>F \A\<^sub>Q, (\ \ (p \ \\<^sub>R)) \ \\<^sub>Q\" + by(metis frameResChainPres frameNilStatEq Commutativity AssertionStatEqTrans Composition Associativity) + ultimately show ?thesis + by(metis FrameStatEqImpCompose) + qed + moreover note FrP FrQ \distinct A\<^sub>P\ + moreover from \distinct A\<^sub>R\ have "distinct(p \ A\<^sub>R)" by simp + moreover note \(p \ A\<^sub>R) \* A\<^sub>P\ \(p \ A\<^sub>R) \* A\<^sub>Q\ \(p \ A\<^sub>R) \* \\ \(p \ A\<^sub>R) \* P\ \(p \ A\<^sub>R) \* Q\ \(p \ A\<^sub>R) \* R\ \(p \ A\<^sub>R) \* K\ + \A\<^sub>P \* \\ \A\<^sub>P \* R\ \A\<^sub>P \* P\ \A\<^sub>P \* (p \ M)\ \A\<^sub>Q \* R\ \A\<^sub>Q \* (p \ M)\ \A\<^sub>P \* xvec\ \xvec \* P\ \A\<^sub>P \* R\ + ultimately obtain K' where "\ \ \\<^sub>P \ R \K'\\*xvec\\N\ \ R'" and "\ \ \\<^sub>P \ (p \ \\<^sub>R) \ (p \ M) \ K'" and "(p \ A\<^sub>R) \* K'" + using comm1Aux by blast + + with PTrans FrP have "\ \ P \ R \\ \ \\*xvec\(P' \ R')" using FrR \(p \ A\<^sub>R) \* \\ \(p \ A\<^sub>R) \* P\ \(p \ A\<^sub>R) \* R\ + \xvec \* P\ \A\<^sub>P \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* R\ \A\<^sub>P \* (p \ M)\ \(p \ A\<^sub>R) \* K'\ \(p \ A\<^sub>R) \* A\<^sub>P\ + by (intro Comm1) (assumption | simp)+ + + moreover from P'RelQ' have "((\ \ (p \ \\<^sub>R)) \ \', P', Q') \ Rel" by(rule Ext) + with \(p \ \\<^sub>R) \ \' \ \\<^sub>R'\ have "(\ \ \\<^sub>R', P', Q') \ Rel" by(metis C3 Associativity compositionSym) + with FrR' \A\<^sub>R' \* P'\ \A\<^sub>R' \* Q'\ \A\<^sub>R' \* \\ have "(\, P' \ R', Q' \ R') \ Rel'" + by(metis C1) + with \xvec \* \\ have "(\, \\*xvec\(P' \ R'), \\*xvec\(Q' \ R')) \ Rel'" + by(metis C2) + ultimately show ?case by blast + next + case(cComm2 \\<^sub>R M xvec N Q' A\<^sub>Q \\<^sub>Q K R' A\<^sub>R) + have FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" by fact + from \A\<^sub>Q \* (P, R)\ have "A\<^sub>Q \* P" and "A\<^sub>Q \* R" by simp+ + + have FrR: "extractFrame R = \A\<^sub>R, \\<^sub>R\" by fact + from \A\<^sub>R \* (P, R)\ have "A\<^sub>R \* P" and "A\<^sub>R \* R" by simp+ + + from \xvec \* (P, R)\ have "xvec \* P" and "xvec \* R" by simp+ + + obtain A\<^sub>P \\<^sub>P where FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" and "A\<^sub>P \* (\, A\<^sub>Q, \\<^sub>Q, A\<^sub>R, M, N, K, R, P, xvec)" and "distinct A\<^sub>P" + by(rule freshFrame) + then have "A\<^sub>P \* \" and "A\<^sub>P \* A\<^sub>Q" and "A\<^sub>P \* \\<^sub>Q" and "A\<^sub>P \* M" and "A\<^sub>P \* R" + and "A\<^sub>P \* N" "A\<^sub>P \* K" and "A\<^sub>P \* A\<^sub>R" and "A\<^sub>P \* P" and "A\<^sub>P \* xvec" + by simp+ + + from FrP FrR \A\<^sub>Q \* P\ \A\<^sub>P \* R\ \A\<^sub>R \* P\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* A\<^sub>R\ \A\<^sub>P \* xvec\ \xvec \* P\ + have "A\<^sub>P \* \\<^sub>R" and "A\<^sub>Q \* \\<^sub>P" and "A\<^sub>R \* \\<^sub>P" and "xvec \* \\<^sub>P" + by(auto dest!: extractFrameFreshChain)+ + + have QTrans: "\ \ \\<^sub>R \ Q \M\\*xvec\\N\ \ Q'" by fact + + note \\ \ \\<^sub>Q \ R \K\N\ \ R'\ FrR \\ \ \\<^sub>Q \ \\<^sub>R \ M \ K\ + moreover from FrR \A\<^sub>R \* \\ \A\<^sub>R \* P\ \A\<^sub>R \* Q\ have "\ \ \\<^sub>R \ P \[Rel] Q" by(metis PRelQ Sim) + with QTrans obtain P' where PTrans: "\ \ \\<^sub>R \ P \M\\*xvec\\N\ \ P'" and P'RelQ': "(\ \ \\<^sub>R, P', Q') \ Rel" + using \xvec \* \\ \xvec \* \\<^sub>R\ \xvec \* P\ + by(force dest: simE) + from PTrans QTrans \A\<^sub>R \* P\ \A\<^sub>R \* Q\ \A\<^sub>R \* xvec\ \xvec \* M\ \distinct xvec\ have "A\<^sub>R \* P'" and "A\<^sub>R \* Q'" + by(blast dest: outputFreshChainDerivative)+ + note PTrans \\ \ \\<^sub>Q \ \\<^sub>R \ M \ K\ + moreover have "\A\<^sub>Q, (\ \ \\<^sub>Q) \ \\<^sub>R\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\<^sub>R\" + proof - + have "\A\<^sub>P, (\ \ \\<^sub>R) \ \\<^sub>P\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\<^sub>R\" + by(metis frameResChainPres frameNilStatEq Commutativity AssertionStatEqTrans Composition Associativity) + moreover from FrR \A\<^sub>R \* \\ \A\<^sub>R \* P\ \A\<^sub>R \* Q\ + have "(insertAssertion (extractFrame Q) (\ \ \\<^sub>R)) \\<^sub>F (insertAssertion (extractFrame P) (\ \ \\<^sub>R))" + by(metis PRelQ StatImp) + with FrP FrQ \A\<^sub>P \* \\ \A\<^sub>Q \* \\ \A\<^sub>P \* \\<^sub>R\ \A\<^sub>Q \* \\<^sub>R\ + have "\A\<^sub>Q, (\ \ \\<^sub>R) \ \\<^sub>Q\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>R) \ \\<^sub>P\" using freshCompChain by simp + moreover have "\A\<^sub>Q, (\ \ \\<^sub>Q) \ \\<^sub>R\ \\<^sub>F \A\<^sub>Q, (\ \ \\<^sub>R) \ \\<^sub>Q\" + by(metis frameResChainPres frameNilStatEq Commutativity AssertionStatEqTrans Composition Associativity) + ultimately show ?thesis + by(metis FrameStatEqImpCompose) + qed + moreover note FrP FrQ \distinct A\<^sub>P\ \distinct A\<^sub>R\ + moreover from \A\<^sub>P \* A\<^sub>R\ \A\<^sub>Q \* A\<^sub>R\ have "A\<^sub>R \* A\<^sub>P" and "A\<^sub>R \* A\<^sub>Q" by simp+ + moreover note \A\<^sub>R \* \\ \A\<^sub>R \* P\ \A\<^sub>R \* Q\ \A\<^sub>R \* R\ \A\<^sub>R \* K\ \A\<^sub>P \* \\ \A\<^sub>P \* P\ + \A\<^sub>P \* R\ \A\<^sub>P \* M\ \A\<^sub>Q \* R\ \A\<^sub>Q \* M\ \A\<^sub>R \* xvec\ \xvec \* M\ + ultimately obtain K' where "\ \ \\<^sub>P \ R \K'\N\ \ R'" and "\ \ \\<^sub>P \ \\<^sub>R \ M \ K'" and "A\<^sub>R \* K'" + using comm2Aux by blast + + with PTrans FrP have "\ \ P \ R \\ \ \\*xvec\(P' \ R')" using FrR \A\<^sub>R \* \\ \A\<^sub>R \* P\ \A\<^sub>R \* R\ + \A\<^sub>R \* \\ \A\<^sub>R \* P\ \A\<^sub>R \* R\ and \xvec \* R\ \A\<^sub>P \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* R\ \A\<^sub>P \* A\<^sub>R\ \A\<^sub>P \* M\ \A\<^sub>R \* K'\ + by(force intro: Comm2) + + moreover from \\ \ \\<^sub>P \ R \K'\N\ \ R'\ FrR \distinct A\<^sub>R\ \A\<^sub>R \* \\ \A\<^sub>R \* R\ \A\<^sub>R \* P'\ \A\<^sub>R \* Q'\ \A\<^sub>R \* N\ \A\<^sub>R \* K'\ + obtain \' A\<^sub>R' \\<^sub>R' where ReqR': "\\<^sub>R \ \' \ \\<^sub>R'" and FrR': "extractFrame R' = \A\<^sub>R', \\<^sub>R'\" + and "A\<^sub>R' \* \" and "A\<^sub>R' \* P'" and "A\<^sub>R' \* Q'" + by(auto intro: expandFrame[where C="(\, P', Q')" and C'="\"]) + + from P'RelQ' have "((\ \ \\<^sub>R) \ \', P', Q') \ Rel" by(rule Ext) + with ReqR' have "(\ \ \\<^sub>R', P', Q') \ Rel" by(metis C3 Associativity compositionSym) + with FrR' \A\<^sub>R' \* P'\ \A\<^sub>R' \* Q'\ \A\<^sub>R' \* \\ have "(\, P' \ R', Q' \ R') \ Rel'" + by(metis C1) + with \xvec \* \\ have "(\, \\*xvec\(P' \ R'), \\*xvec\(Q' \ R')) \ Rel'" + by(metis C2) + ultimately show ?case by blast + next + case(cBrMerge \\<^sub>R M N Q' A\<^sub>Q \\<^sub>Q R' A\<^sub>R) + have FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" by fact + from \A\<^sub>Q \* (P, R)\ have "A\<^sub>Q \* P" and "A\<^sub>Q \* R" by simp+ + + have FrR: "extractFrame R = \A\<^sub>R, \\<^sub>R\" by fact + from \A\<^sub>R \* (P, R)\ have "A\<^sub>R \* P" and "A\<^sub>R \* R" by simp+ + + obtain A\<^sub>P \\<^sub>P where FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" and "A\<^sub>P \* (\, A\<^sub>Q, \\<^sub>Q, A\<^sub>R, M, N, R, P)" and "distinct A\<^sub>P" + by(rule freshFrame) + then have "A\<^sub>P \* \" and "A\<^sub>P \* A\<^sub>Q" and "A\<^sub>P \* \\<^sub>Q" and "A\<^sub>P \* M" and "A\<^sub>P \* R" + and "A\<^sub>P \* N" and "A\<^sub>P \* A\<^sub>R" and "A\<^sub>P \* P" + by simp+ + + from FrP FrR \A\<^sub>Q \* P\ \A\<^sub>P \* R\ \A\<^sub>R \* P\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* A\<^sub>R\ + have "A\<^sub>P \* \\<^sub>R" and "A\<^sub>Q \* \\<^sub>P" and "A\<^sub>R \* \\<^sub>P" + by(auto dest: extractFrameFreshChain)+ + + have QTrans: "\ \ \\<^sub>R \ Q \\M\N\ \ Q'" by fact + + have RTrans: "\ \ \\<^sub>Q \ R \\M\N\ \ R'" by fact + + from FrR \A\<^sub>R \* \\ \A\<^sub>R \* P\ \A\<^sub>R \* Q\ have "\ \ \\<^sub>R \ P \[Rel] Q" by(metis PRelQ Sim) + with QTrans obtain P' where PTrans: "\ \ \\<^sub>R \ P \\M\N\ \ P'" and P'RelQ': "(\ \ \\<^sub>R, P', Q') \ Rel" + by(force dest: simE) + from PTrans QTrans \A\<^sub>R \* P\ \A\<^sub>R \* Q\ \A\<^sub>R \* N\ have "A\<^sub>R \* P'" and "A\<^sub>R \* Q'" + by(blast dest: brinputFreshChainDerivative)+ + + have "\A\<^sub>Q, (\ \ \\<^sub>Q) \ \\<^sub>R\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\<^sub>R\" + proof - + have "\A\<^sub>P, (\ \ \\<^sub>R) \ \\<^sub>P\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\<^sub>R\" + by(metis frameResChainPres frameNilStatEq Commutativity AssertionStatEqTrans Composition Associativity) + moreover from FrR \A\<^sub>R \* \\ \A\<^sub>R \* P\ \A\<^sub>R \* Q\ + have "(insertAssertion (extractFrame Q) (\ \ \\<^sub>R)) \\<^sub>F (insertAssertion (extractFrame P) (\ \ \\<^sub>R))" + by(metis PRelQ StatImp) + with FrP FrQ \A\<^sub>P \* \\ \A\<^sub>Q \* \\ \A\<^sub>P \* \\<^sub>R\ \A\<^sub>Q \* \\<^sub>R\ + have "\A\<^sub>Q, (\ \ \\<^sub>R) \ \\<^sub>Q\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>R) \ \\<^sub>P\" using freshCompChain by simp + moreover have "\A\<^sub>Q, (\ \ \\<^sub>Q) \ \\<^sub>R\ \\<^sub>F \A\<^sub>Q, (\ \ \\<^sub>R) \ \\<^sub>Q\" + by(metis frameResChainPres frameNilStatEq Commutativity AssertionStatEqTrans Composition Associativity) + ultimately show ?thesis + by(metis FrameStatEqImpCompose) + qed + with RTrans \extractFrame R = \A\<^sub>R, \\<^sub>R\\ \distinct A\<^sub>R\ \A\<^sub>P \* A\<^sub>R\ \A\<^sub>Q \* A\<^sub>R\ + \A\<^sub>R \* \\ \A\<^sub>R \* \\<^sub>P\ \A\<^sub>R \* \\<^sub>Q\ \A\<^sub>R \* R\ \A\<^sub>R \* M\ \A\<^sub>P \* R\ \A\<^sub>P \* M\ \A\<^sub>Q \* R\ \A\<^sub>Q \* M\ + have "\ \ \\<^sub>P \ R \\M\N\ \ R'" + using brCommInAux freshChainSym by blast + + with PTrans FrP have Transition: "\ \ P \ R \\M\N\ \ (P' \ R')" using FrR \A\<^sub>R \* \\ \A\<^sub>R \* P\ \A\<^sub>R \* R\ + \A\<^sub>R \* \\ \A\<^sub>R \* P\ \A\<^sub>R \* R\ \A\<^sub>P \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* R\ \A\<^sub>P \* A\<^sub>R\ \A\<^sub>P \* M\ \A\<^sub>R \* M\ + by(force intro: BrMerge) + + from \\ \ \\<^sub>P \ R \\M\N\ \ R'\ FrR \distinct A\<^sub>R\ \A\<^sub>R \* \\ \A\<^sub>R \* R\ \A\<^sub>R \* P'\ \A\<^sub>R \* Q'\ \A\<^sub>R \* N\ \A\<^sub>R \* M\ + obtain \' A\<^sub>R' \\<^sub>R' where ReqR': "\\<^sub>R \ \' \ \\<^sub>R'" and FrR': "extractFrame R' = \A\<^sub>R', \\<^sub>R'\" + and "A\<^sub>R' \* \" and "A\<^sub>R' \* P'" and "A\<^sub>R' \* Q'" + by(auto intro: expandFrame[where C="(\, P', Q')" and C'="\"]) + + from P'RelQ' have "((\ \ \\<^sub>R) \ \', P', Q') \ Rel" by(rule Ext) + with ReqR' have "(\ \ \\<^sub>R', P', Q') \ Rel" by(metis C3 Associativity compositionSym) + with FrR' \A\<^sub>R' \* P'\ \A\<^sub>R' \* Q'\ \A\<^sub>R' \* \\ have Relation: "(\, P' \ R', Q' \ R') \ Rel'" + by(metis C1) + show ?case using Transition Relation + by blast + next + case(cBrComm1 \\<^sub>R M N Q' A\<^sub>Q \\<^sub>Q xvec R' A\<^sub>R) + have FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" by fact + from \A\<^sub>Q \* (P, R)\ have "A\<^sub>Q \* P" and "A\<^sub>Q \* R" by simp+ + + have FrR: "extractFrame R = \A\<^sub>R, \\<^sub>R\" by fact + from \A\<^sub>R \* (P, R)\ have "A\<^sub>R \* P" and "A\<^sub>R \* R" by simp+ + from \\M\\*xvec\\N\ = \\ have "xvec = bn \" + by(auto simp add: action.inject) + from \xvec = bn \\ \bn \ \* P\ \bn \ \* R\ + have "xvec \* P" and "xvec \* R" by simp+ + + obtain A\<^sub>P \\<^sub>P where FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" and "A\<^sub>P \* (\, A\<^sub>Q, \\<^sub>Q, A\<^sub>R, M, N, R, P, xvec)" and "distinct A\<^sub>P" + by(rule freshFrame) + then have "A\<^sub>P \* \" and "A\<^sub>P \* A\<^sub>Q" and "A\<^sub>P \* \\<^sub>Q" and "A\<^sub>P \* M" and "A\<^sub>P \* R" + and "A\<^sub>P \* N" and "A\<^sub>P \* A\<^sub>R" and "A\<^sub>P \* P" and "A\<^sub>P \* xvec" + by simp+ + + from FrP FrR \A\<^sub>Q \* P\ \A\<^sub>P \* R\ \A\<^sub>R \* P\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* A\<^sub>R\ + have "A\<^sub>P \* \\<^sub>R" and "A\<^sub>Q \* \\<^sub>P" and "A\<^sub>R \* \\<^sub>P" + by(auto dest: extractFrameFreshChain)+ + + from \A\<^sub>P \* xvec\ \xvec \* P\ FrP have "xvec \* \\<^sub>P" + by(auto dest: extractFrameFreshChain) + + have QTrans: "\ \ \\<^sub>R \ Q \\M\N\ \ Q'" by fact + + have RTrans: "\ \ \\<^sub>Q \ R \ \M\\*xvec\\N\ \ R'" by fact + + from RTrans FrR \distinct A\<^sub>R\ \A\<^sub>R \* R\ \A\<^sub>R \* xvec\ \xvec \* R\ \xvec \* Q\ \xvec \* \\ \xvec \* \\<^sub>Q\ \A\<^sub>R \* Q\ + \A\<^sub>R \* \\ \A\<^sub>R \* \\<^sub>Q\ \A\<^sub>R \* M\ \A\<^sub>R \* N\ \A\<^sub>R \* R\ \xvec \* R\ \A\<^sub>R \* P\ \xvec \* P\ \A\<^sub>P \* A\<^sub>R\ \A\<^sub>P \* xvec\ + \A\<^sub>Q \* A\<^sub>R\ \A\<^sub>Q \* xvec\ \A\<^sub>R \* \\<^sub>P\ \xvec \* \\<^sub>P\ \distinct xvec\ \xvec \* M\ + obtain p \' A\<^sub>R' \\<^sub>R' where S: "set p \ set xvec \ set(p \ xvec)" and FrR': "extractFrame R' = \A\<^sub>R', \\<^sub>R'\" + and "(p \ \\<^sub>R) \ \' \ \\<^sub>R'" and "A\<^sub>R' \* Q" and "A\<^sub>R' \* \" and "(p \ xvec) \* \" + and "(p \ xvec) \* Q" and "(p \ xvec) \* \\<^sub>Q" and "(p \ xvec) \* R" and "(p \ xvec) \* M" + and "(p \ xvec) \* P" and "(p \ xvec) \* A\<^sub>P" and "(p \ xvec) \* A\<^sub>Q" and "(p \ xvec) \* \\<^sub>P" + and "A\<^sub>R' \* P" and "A\<^sub>R' \* N" + by(auto intro: expandFrame[where C="(\, Q, \\<^sub>Q, R, P, M, A\<^sub>P, A\<^sub>Q, \\<^sub>P)" and C'="(\, Q, \\<^sub>Q, R, P, M, A\<^sub>P, A\<^sub>Q, \\<^sub>P)"]) + + from \A\<^sub>R \* \\ have "(p \ A\<^sub>R) \* (p \ \)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \xvec \* \\ \(p \ xvec) \* \\ S have "(p \ A\<^sub>R) \* \" by simp + from \A\<^sub>R \* \\<^sub>P\ have "(p \ A\<^sub>R) \* (p \ \\<^sub>P)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \xvec \* \\<^sub>P\ \(p \ xvec) \* \\<^sub>P\ S have "(p \ A\<^sub>R) \* \\<^sub>P" by simp + from \A\<^sub>R \* \\<^sub>Q\ have "(p \ A\<^sub>R) \* (p \ \\<^sub>Q)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \xvec \* \\<^sub>Q\ \(p \ xvec) \* \\<^sub>Q\ S have "(p \ A\<^sub>R) \* \\<^sub>Q" by simp + from \A\<^sub>R \* M\ have "(p \ A\<^sub>R) \* (p \ M)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \xvec \* M\ \(p \ xvec) \* M\ S have "(p \ A\<^sub>R) \* M" by simp + from \A\<^sub>R \* P\ have "(p \ A\<^sub>R) \* (p \ P)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \xvec \* P\ \(p \ xvec) \* P\ S have "(p \ A\<^sub>R) \* P" by simp + from \A\<^sub>R \* Q\ have "(p \ A\<^sub>R) \* (p \ Q)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \xvec \* Q\ \(p \ xvec) \* Q\ S have "(p \ A\<^sub>R) \* Q" by simp + from \A\<^sub>R \* R\ have "(p \ A\<^sub>R) \* (p \ R)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \xvec \* R\ \(p \ xvec) \* R\ S have "(p \ A\<^sub>R) \* R" by simp + + from \A\<^sub>P \* xvec\ \(p \ xvec) \* A\<^sub>P\ \A\<^sub>P \* M\ S have "A\<^sub>P \* (p \ M)" by(simp add: freshChainSimps) + from \A\<^sub>Q \* xvec\ \(p \ xvec) \* A\<^sub>Q\ \A\<^sub>Q \* M\ S have "A\<^sub>Q \* (p \ M)" by(simp add: freshChainSimps) + from \A\<^sub>P \* xvec\ \(p \ xvec) \* A\<^sub>P\ \A\<^sub>P \* A\<^sub>R\ S have "(p \ A\<^sub>R) \* A\<^sub>P" by(simp add: freshChainSimps) + from \A\<^sub>Q \* xvec\ \(p \ xvec) \* A\<^sub>Q\ \A\<^sub>Q \* A\<^sub>R\ S have "(p \ A\<^sub>R) \* A\<^sub>Q" by(simp add: freshChainSimps) + + from QTrans S \xvec \* Q\ \(p \ xvec) \* Q\ have "(p \ (\ \ \\<^sub>R)) \ Q \ \(p \ M)\N\ \ Q'" + using brinputPermFrameSubject name_list_set_fresh by blast + with \xvec \* \\ \(p \ xvec) \* \\ S have QTrans: "(\ \ (p \ \\<^sub>R)) \ Q \ \(p \ M)\N\ \ Q'" + by(simp add: eqvts) + + from FrR have "(p \ extractFrame R) = p \ \A\<^sub>R, \\<^sub>R\" by simp + with \xvec \* R\ \(p \ xvec) \* R\ S have FrR: "extractFrame R = \(p \ A\<^sub>R), (p \ \\<^sub>R)\" + by(simp add: eqvts) + + from FrR \(p \ A\<^sub>R) \* \\ \(p \ A\<^sub>R) \* P\ \(p \ A\<^sub>R) \* Q\ have "\ \ (p \ \\<^sub>R) \ P \[Rel] Q" + by(metis Sim PRelQ) + with QTrans obtain P' where PTrans: "\ \ (p \ \\<^sub>R) \ P \\(p \ M)\N\ \ P'" and P'RelQ': "(\ \ (p \ \\<^sub>R), P', Q') \ Rel" + by(force dest: simE) + with QTrans \A\<^sub>R' \* P\ \A\<^sub>R' \* Q\ \A\<^sub>R' \* N\ have "A\<^sub>R' \* P'" and "A\<^sub>R' \* Q'" + by(blast dest: brinputFreshChainDerivative)+ + + have "\A\<^sub>Q, (\ \ \\<^sub>Q) \ (p \ \\<^sub>R)\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ (p \ \\<^sub>R)\" + proof - + have "\A\<^sub>P, (\ \ (p \ \\<^sub>R)) \ \\<^sub>P\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ (p \ \\<^sub>R)\" + by(metis frameResChainPres frameNilStatEq Commutativity AssertionStatEqTrans Composition Associativity) + moreover from FrR \(p \ A\<^sub>R) \* \\ \(p \ A\<^sub>R) \* P\ \(p \ A\<^sub>R) \* Q\ + have "(insertAssertion (extractFrame Q) (\ \ (p \ \\<^sub>R))) \\<^sub>F (insertAssertion (extractFrame P) (\ \ (p \ \\<^sub>R)))" + by(metis PRelQ StatImp) + with FrP FrQ \A\<^sub>P \* \\ \A\<^sub>Q \* \\ \A\<^sub>P \* \\<^sub>R\ \A\<^sub>Q \* \\<^sub>R\ \A\<^sub>P \* xvec\ \(p \ xvec) \* A\<^sub>P\ \A\<^sub>Q \* xvec\ \(p \ xvec) \* A\<^sub>Q\ S + have "\A\<^sub>Q, (\ \ (p \ \\<^sub>R)) \ \\<^sub>Q\ \\<^sub>F \A\<^sub>P, (\ \ (p \ \\<^sub>R)) \ \\<^sub>P\" using freshCompChain + by(simp add: freshChainSimps) + moreover have "\A\<^sub>Q, (\ \ \\<^sub>Q) \ (p \ \\<^sub>R)\ \\<^sub>F \A\<^sub>Q, (\ \ (p \ \\<^sub>R)) \ \\<^sub>Q\" + by(metis frameResChainPres frameNilStatEq Commutativity AssertionStatEqTrans Composition Associativity) + ultimately show ?thesis + by(metis FrameStatEqImpCompose) + qed + moreover note RTrans FrR FrP FrQ \distinct A\<^sub>P\ + moreover from \distinct A\<^sub>R\ have "distinct(p \ A\<^sub>R)" by simp + moreover note \(p \ A\<^sub>R) \* A\<^sub>P\ \(p \ A\<^sub>R) \* A\<^sub>Q\ \(p \ A\<^sub>R) \* \\ \(p \ A\<^sub>R) \* \\<^sub>P\ \(p \ A\<^sub>R) \* \\<^sub>Q\ \(p \ A\<^sub>R) \* M\ \(p \ A\<^sub>R) \* P\ \(p \ A\<^sub>R) \* Q\ \(p \ A\<^sub>R) \* R\ + \A\<^sub>P \* \\ \A\<^sub>P \* R\ \A\<^sub>P \* P\ \A\<^sub>P \* M\ \A\<^sub>Q \* M\ \A\<^sub>P \* (p \ M)\ \A\<^sub>Q \* R\ \A\<^sub>P \* xvec\ \xvec \* P\ \A\<^sub>P \* R\ + ultimately have "\ \ \\<^sub>P \ R \\M\\*xvec\\N\ \ R'" + using brCommOutAux by blast + with S \xvec \* M\ \(p \ xvec) \* M\ have "\ \ \\<^sub>P \ R \\(p \ M)\\*xvec\\N\ \ R'" by simp + + with PTrans FrP S \xvec \* M\ \(p \ xvec) \* M\ \(p \ A\<^sub>R) \* M\ have "\ \ P \ R \\(p \ M)\\*xvec\\N\ \ (P' \ R')" using FrR \(p \ A\<^sub>R) \* \\ \(p \ A\<^sub>R) \* P\ \(p \ A\<^sub>R) \* R\ + \xvec \* P\ \A\<^sub>P \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* R\ \A\<^sub>P \* (p \ M)\ \(p \ A\<^sub>R) \* A\<^sub>P\ + by(intro BrComm1) (assumption | simp)+ + + with S \xvec \* M\ \(p \ xvec) \* M\ + have Transition: "\ \ P \ R \\M\\*xvec\\N\ \ (P' \ R')" by simp + + from P'RelQ' have "((\ \ (p \ \\<^sub>R)) \ \', P', Q') \ Rel" by(rule Ext) + with \(p \ \\<^sub>R) \ \' \ \\<^sub>R'\ have "(\ \ \\<^sub>R', P', Q') \ Rel" by(metis C3 Associativity compositionSym) + with FrR' \A\<^sub>R' \* P'\ \A\<^sub>R' \* Q'\ \A\<^sub>R' \* \\ have Relation: "(\, P' \ R', Q' \ R') \ Rel'" + by(metis C1) + + show ?case using Transition Relation + by blast + next + case(cBrComm2 \\<^sub>R M xvec N Q' A\<^sub>Q \\<^sub>Q R' A\<^sub>R) + have FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" by fact + from \A\<^sub>Q \* (P, R)\ have "A\<^sub>Q \* P" and "A\<^sub>Q \* R" by simp+ + + have FrR: "extractFrame R = \A\<^sub>R, \\<^sub>R\" by fact + from \A\<^sub>R \* (P, R)\ have "A\<^sub>R \* P" and "A\<^sub>R \* R" by simp+ + from \\M\\*xvec\\N\ = \\ have "xvec = bn \" + by(auto simp add: action.inject) + from \xvec = bn \\ \bn \ \* P\ \bn \ \* R\ + have "xvec \* P" and "xvec \* R" by simp+ + + obtain A\<^sub>P \\<^sub>P where FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" and "A\<^sub>P \* (\, A\<^sub>Q, \\<^sub>Q, A\<^sub>R, M, N, R, P)" and "distinct A\<^sub>P" + by(rule freshFrame) + then have "A\<^sub>P \* \" and "A\<^sub>P \* A\<^sub>Q" and "A\<^sub>P \* \\<^sub>Q" and "A\<^sub>P \* M" and "A\<^sub>P \* R" + and "A\<^sub>P \* N" and "A\<^sub>P \* A\<^sub>R" and "A\<^sub>P \* P" + by simp+ + + from FrP FrR \A\<^sub>Q \* P\ \A\<^sub>P \* R\ \A\<^sub>R \* P\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* A\<^sub>R\ + have "A\<^sub>P \* \\<^sub>R" and "A\<^sub>Q \* \\<^sub>P" and "A\<^sub>R \* \\<^sub>P" + by(auto dest: extractFrameFreshChain)+ + + have QTrans: "\ \ \\<^sub>R \ Q \\M\\*xvec\\N\ \ Q'" by fact + + have RTrans: "\ \ \\<^sub>Q \ R \\M\N\ \ R'" by fact + + from FrR \A\<^sub>R \* \\ \A\<^sub>R \* P\ \A\<^sub>R \* Q\ have "\ \ \\<^sub>R \ P \[Rel] Q" by(metis PRelQ Sim) + from \\ \ \\<^sub>R \ P \[Rel] Q\ QTrans \xvec \* \\ \xvec \* \\<^sub>R\ \xvec \* P\ + obtain P' where PTrans: "\ \ \\<^sub>R \ P \\M\\*xvec\\N\ \ P'" and P'RelQ': "(\ \ \\<^sub>R, P', Q') \ Rel" + by(force dest: simE) + from PTrans QTrans \A\<^sub>R \* P\ \A\<^sub>R \* Q\ \A\<^sub>R \* N\ \A\<^sub>R \* xvec\ \xvec \* M\ \distinct xvec\ have "A\<^sub>R \* P'" and "A\<^sub>R \* Q'" + by(blast dest: broutputFreshChainDerivative)+ + + have "\A\<^sub>Q, (\ \ \\<^sub>Q) \ \\<^sub>R\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\<^sub>R\" + proof - + have "\A\<^sub>P, (\ \ \\<^sub>R) \ \\<^sub>P\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\<^sub>R\" + by(metis frameResChainPres frameNilStatEq Commutativity AssertionStatEqTrans Composition Associativity) + moreover from FrR \A\<^sub>R \* \\ \A\<^sub>R \* P\ \A\<^sub>R \* Q\ + have "(insertAssertion (extractFrame Q) (\ \ \\<^sub>R)) \\<^sub>F (insertAssertion (extractFrame P) (\ \ \\<^sub>R))" + by(metis PRelQ StatImp) + with FrP FrQ \A\<^sub>P \* \\ \A\<^sub>Q \* \\ \A\<^sub>P \* \\<^sub>R\ \A\<^sub>Q \* \\<^sub>R\ + have "\A\<^sub>Q, (\ \ \\<^sub>R) \ \\<^sub>Q\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>R) \ \\<^sub>P\" using freshCompChain by simp + moreover have "\A\<^sub>Q, (\ \ \\<^sub>Q) \ \\<^sub>R\ \\<^sub>F \A\<^sub>Q, (\ \ \\<^sub>R) \ \\<^sub>Q\" + by(metis frameResChainPres frameNilStatEq Commutativity AssertionStatEqTrans Composition Associativity) + ultimately show ?thesis + by(metis FrameStatEqImpCompose) + qed + with RTrans \extractFrame R = \A\<^sub>R, \\<^sub>R\\ \distinct A\<^sub>R\ \A\<^sub>P \* A\<^sub>R\ \A\<^sub>Q \* A\<^sub>R\ + \A\<^sub>R \* \\ \A\<^sub>R \* \\<^sub>P\ \A\<^sub>R \* \\<^sub>Q\ \A\<^sub>R \* R\ \A\<^sub>R \* M\ \A\<^sub>P \* R\ \A\<^sub>P \* M\ \A\<^sub>Q \* R\ \A\<^sub>Q \* M\ + have "\ \ \\<^sub>P \ R \\M\N\ \ R'" + using brCommInAux freshChainSym by blast + + with PTrans FrP have Transition: "\ \ P \ R \\M\\*xvec\\N\ \ (P' \ R')" using FrR \A\<^sub>R \* \\ \A\<^sub>R \* P\ \A\<^sub>R \* R\ + \A\<^sub>R \* \\ \A\<^sub>R \* P\ \A\<^sub>R \* R\ \A\<^sub>P \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* R\ \A\<^sub>P \* A\<^sub>R\ \A\<^sub>P \* M\ \A\<^sub>R \* M\ \xvec \* R\ + by(force intro: BrComm2) + + from \\ \ \\<^sub>P \ R \\M\N\ \ R'\ FrR \distinct A\<^sub>R\ \A\<^sub>R \* \\ \A\<^sub>R \* R\ \A\<^sub>R \* P'\ \A\<^sub>R \* Q'\ \A\<^sub>R \* N\ \A\<^sub>R \* M\ + obtain \' A\<^sub>R' \\<^sub>R' where ReqR': "\\<^sub>R \ \' \ \\<^sub>R'" and FrR': "extractFrame R' = \A\<^sub>R', \\<^sub>R'\" + and "A\<^sub>R' \* \" and "A\<^sub>R' \* P'" and "A\<^sub>R' \* Q'" + by(auto intro: expandFrame[where C="(\, P', Q')" and C'="\"]) + + from P'RelQ' have "((\ \ \\<^sub>R) \ \', P', Q') \ Rel" by(rule Ext) + with ReqR' have "(\ \ \\<^sub>R', P', Q') \ Rel" by(metis C3 Associativity compositionSym) + with FrR' \A\<^sub>R' \* P'\ \A\<^sub>R' \* Q'\ \A\<^sub>R' \* \\ have Relation: "(\, P' \ R', Q' \ R') \ Rel'" + by(metis C1) + show ?case using Transition Relation + by blast + qed +qed + +no_notation relcomp (infixr "O" 75) + +lemma bangPres: + 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" + and Rel' :: "('b \ ('a, 'b, 'c) psi \ ('a, 'b, 'c) psi) set" + +assumes "(\, P, Q) \ Rel" + and "eqvt Rel'" + and "guarded P" + and "guarded Q" + and cSim: "\\' S T. (\', S, T) \ Rel \ \' \ S \[Rel] T" + and cExt: "\\' S T \''. (\', S, T) \ Rel \ (\' \ \'', S, T) \ Rel" + and cSym: "\\' S T. (\', S, T) \ Rel \ (\', T, S) \ Rel" + and StatEq: "\\' S T \''. \(\', S, T) \ Rel; \' \ \''\ \ (\'', S, T) \ Rel" + and Closed: "\\' S T p. (\', S, T) \ Rel \ ((p::name prm) \ \', p \ S, p \ T) \ Rel" + and Assoc: "\\' S T U. (\', S \ (T \ U), (S \ T) \ U) \ Rel" + and ParPres: "\\' S T U. (\', S, T) \ Rel \ (\', S \ U, T \ U) \ Rel" + and FrameParPres: "\\' \\<^sub>U S T U A\<^sub>U. \(\' \ \\<^sub>U, S, T) \ Rel; extractFrame U = \A\<^sub>U, \\<^sub>U\; A\<^sub>U \* \'; A\<^sub>U \* S; A\<^sub>U \* T\ \ + (\', U \ S, U \ T) \ Rel" + and ResPres: "\\' S T xvec. \(\', S, T) \ Rel; xvec \* \'\ \ (\', \\*xvec\S, \\*xvec\T) \ Rel" + and ScopeExt: "\xvec \' S T. \xvec \* \'; xvec \* T\ \ (\', \\*xvec\(S \ T), (\\*xvec\S) \ T) \ Rel" + and Trans: "\\' S T U. \(\', S, T) \ Rel; (\', T, U) \ Rel\ \ (\', S, U) \ Rel" + and Compose: "\\' S T U O. \(\', S, T) \ Rel; (\', T, U) \ Rel'; (\', U, O) \ Rel\ \ (\', S, O) \ Rel'" + and C1: "\\ S T U. \(\, S, T) \ Rel; guarded S; guarded T\ \ (\, U \ !S, U \ !T) \ Rel'" + and Der: "\\' S \ S' T. \\' \ !S \\ \ S'; (\', S, T) \ Rel; bn \ \* \'; bn \ \* S; bn \ \* T; guarded T; bn \ \* subject \\ \ + \T' U O. \' \ !T \\ \ T' \ (\', S', U \ !S) \ Rel \ (\', T', O \ !T) \ Rel \ + (\', U, O) \ Rel \ ((supp U)::name set) \ supp S' \ + ((supp O)::name set) \ supp T'" + +shows "\ \ R \ !P \[Rel'] R \ !Q" + using \eqvt Rel'\ +proof(induct rule: simI[of _ _ _ _ "()"]) + case(cSim \ RQ') + from \bn \ \* (R \ !P)\ \bn \ \* (R \ !Q)\ have "bn \ \* P" and "bn \ \* (!Q)" and "bn \ \* Q" and "bn \ \* R" by simp+ + from \\ \ R \ !Q \\ \ RQ'\ \bn \ \* \\ \bn \ \* R\ \bn \ \* !Q\ \bn \ \* subject \\ show ?case + proof(induct rule: parCases[where C=P]) + case(cPar1 R' A\<^sub>Q \\<^sub>Q) + from \extractFrame (!Q) = \A\<^sub>Q, \\<^sub>Q\\ have "A\<^sub>Q = []" and "\\<^sub>Q = SBottom'" by simp+ + with \\ \ \\<^sub>Q \ R \\ \ R'\ \bn \ \* P\ have "\ \ R \ !P \\ \ (R' \ !P)" + by(intro Par1) (assumption | simp)+ + moreover from \(\, P, Q) \ Rel\ \guarded P\ \guarded Q\ have "(\, R' \ !P, R' \ !Q) \ Rel'" + by(rule C1) + ultimately show ?case by blast + next + case(cPar2 Q' A\<^sub>R \\<^sub>R) + have QTrans: "\ \ \\<^sub>R \ !Q \\ \ Q'" and FrR: "extractFrame R = \A\<^sub>R, \\<^sub>R\" by fact+ + with \bn \ \* R\ \A\<^sub>R \* \\ have "bn \ \* \\<^sub>R" by(force dest: extractFrameFreshChain) + with QTrans \(\, P, Q) \ Rel\ \bn \ \* \\\bn \ \* P\ \bn \ \* Q\ \guarded P\ \bn \ \* subject \\ + obtain P' S T where PTrans: "\ \ \\<^sub>R \ !P \\ \ P'" and "(\ \ \\<^sub>R, P', T \ !P) \ Rel" + and "(\ \ \\<^sub>R, Q', S \ !Q) \ Rel" and "(\ \ \\<^sub>R, S, T) \ Rel" + and suppT: "((supp T)::name set) \ supp P'" and suppS: "((supp S)::name set) \ supp Q'" + by(drule_tac cSym) (auto dest: Der cExt) + from PTrans FrR \A\<^sub>R \* \\ \A\<^sub>R \* P\ \A\<^sub>R \* \\ \bn \ \* R\ have "\ \ R \ !P \\ \ (R \ P')" + by(intro Par2) auto + moreover + { + from \A\<^sub>R \* P\ \A\<^sub>R \* (!Q)\ \A\<^sub>R \* \\ PTrans QTrans \bn \ \* subject \\ \distinct(bn \)\ have "A\<^sub>R \* P'" and "A\<^sub>R \* Q'" + by(force dest: freeFreshChainDerivative)+ + from \(\ \ \\<^sub>R, P', T \ !P) \ Rel\ FrR \A\<^sub>R \* \\ \A\<^sub>R \* P'\ \A\<^sub>R \* P\ suppT have "(\, R \ P', R \ (T \ !P)) \ Rel" + by(intro FrameParPres) (auto simp add: fresh_star_def fresh_def psi.supp) + then have "(\, R \ P', (R \ T) \ !P) \ Rel" by(blast intro: Assoc Trans) + moreover from \(\, P, Q) \ Rel\ \guarded P\ \guarded Q\ have "(\, (R \ T) \ !P, (R \ T) \ !Q) \ Rel'" + by(rule C1) + moreover from \(\ \ \\<^sub>R, Q', S \ !Q) \ Rel\ \(\ \ \\<^sub>R, S, T) \ Rel\ have "(\ \ \\<^sub>R, Q', T \ !Q) \ Rel" + by(blast intro: ParPres Trans) + with FrR \A\<^sub>R \* \\ \A\<^sub>R \* P'\ \A\<^sub>R \* Q'\ \A\<^sub>R \* (!Q)\ suppT suppS have "(\, R \ Q', R \ (T \ !Q)) \ Rel" + by(intro FrameParPres) (auto simp add: fresh_star_def fresh_def psi.supp) + then have "(\, R \ Q', (R \ T) \ !Q) \ Rel" by(blast intro: Assoc Trans) + ultimately have "(\, R \ P', R \ Q') \ Rel'" by(blast intro: cSym Compose) + } + ultimately show ?case by blast + next + case(cComm1 \\<^sub>Q M N R' A\<^sub>R \\<^sub>R K xvec Q' A\<^sub>Q) + from \extractFrame (!Q) = \A\<^sub>Q, \\<^sub>Q\\ have "A\<^sub>Q = []" and "\\<^sub>Q = SBottom'" by simp+ + have RTrans: "\ \ \\<^sub>Q \ R \M\N\ \ R'" and FrR: "extractFrame R = \A\<^sub>R, \\<^sub>R\" by fact+ + moreover have QTrans: "\ \ \\<^sub>R \ !Q \K\\*xvec\\N\ \ Q'" by fact + from FrR \xvec \* R\ \A\<^sub>R \* xvec\ have "xvec \* \\<^sub>R" by(force dest: extractFrameFreshChain) + with QTrans \(\, P, Q) \ Rel\ \xvec \* \\\xvec \* P\ \xvec \* (!Q)\ \guarded P\ \xvec \* K\ + obtain P' S T where PTrans: "\ \ \\<^sub>R \ !P \K\\*xvec\\N\ \ P'" and "(\ \ \\<^sub>R, P', T \ !P) \ Rel" + and "(\ \ \\<^sub>R, Q', S \ !Q) \ Rel" and "(\ \ \\<^sub>R, S, T) \ Rel" + and suppT: "((supp T)::name set) \ supp P'" and suppS: "((supp S)::name set) \ supp Q'" + apply(drule_tac cSym) + by (metis Der bn.simps(3) cExt freshCompChain(1) optionFreshChain(1) psiFreshVec(7) subject.simps(3)) + note \\ \ \\<^sub>R \ \\<^sub>Q \ M \ K\ + ultimately have "\ \ R \ !P \\ \ \\*xvec\(R' \ P')" + using PTrans \\\<^sub>Q = SBottom'\ \xvec \* R\ \A\<^sub>R \* \\ \A\<^sub>R \* R\ \A\<^sub>R \* M\ \A\<^sub>R \* P\ + by(intro Comm1) (assumption | simp)+ + + moreover from \A\<^sub>R \* P\ \A\<^sub>R \* (!Q)\ \A\<^sub>R \* xvec\ PTrans QTrans \xvec \* K\ \distinct xvec\ + have "A\<^sub>R \* P'" and "A\<^sub>R \* Q'" by(force dest: outputFreshChainDerivative)+ + moreover with RTrans FrR \distinct A\<^sub>R\ \A\<^sub>R \* R\ \A\<^sub>R \* N\ \A\<^sub>R \* \\ \A\<^sub>R \* P\ \A\<^sub>R \* (!Q)\ \A\<^sub>R \* M\ + obtain \' A\<^sub>R' \\<^sub>R' where FrR': "extractFrame R' = \A\<^sub>R', \\<^sub>R'\" and "\\<^sub>R \ \' \ \\<^sub>R'" and "A\<^sub>R' \* \" + and "A\<^sub>R' \* P'" and "A\<^sub>R' \* Q'" and "A\<^sub>R' \* P" and "A\<^sub>R' \* Q" + by(auto intro: expandFrame[where C="(\, P, P', Q, Q')" and C'=\]) + + moreover + { + from \(\ \ \\<^sub>R, P', T \ !P) \ Rel\ have "((\ \ \\<^sub>R) \ \', P', T \ !P) \ Rel" by(rule cExt) + with \\\<^sub>R \ \' \ \\<^sub>R'\ have "(\ \ \\<^sub>R', P', T \ !P) \ Rel" + by(metis Associativity StatEq compositionSym) + with FrR' \A\<^sub>R' \* \\ \A\<^sub>R' \* P'\ \A\<^sub>R' \* P\ suppT have "(\, R' \ P', R' \ (T \ !P)) \ Rel" + by(intro FrameParPres) (auto simp add: fresh_star_def fresh_def psi.supp) + then have "(\, R' \ P', (R' \ T) \ !P) \ Rel" by(blast intro: Assoc Trans) + with \xvec \* \\ \xvec \* P\ have "(\, \\*xvec\(R' \ P'), (\\*xvec\(R' \ T)) \ !P) \ Rel" + by(metis ResPres psiFreshVec ScopeExt Trans) + moreover from \(\, P, Q) \ Rel\ \guarded P\ \guarded Q\ have "(\, (\\*xvec\(R' \ T)) \ !P, (\\*xvec\(R' \ T)) \ !Q) \ Rel'" + by(rule C1) + moreover from \(\ \ \\<^sub>R, Q', S \ !Q) \ Rel\ \(\ \ \\<^sub>R, S, T) \ Rel\ have "(\ \ \\<^sub>R, Q', T \ !Q) \ Rel" + by(blast intro: ParPres Trans) + then have "((\ \ \\<^sub>R) \ \', Q', T \ !Q) \ Rel" by(rule cExt) + with \\\<^sub>R \ \' \ \\<^sub>R'\ have "(\ \ \\<^sub>R', Q', T \ !Q) \ Rel" + by(metis Associativity StatEq compositionSym) + with FrR' \A\<^sub>R' \* \\ \A\<^sub>R' \* P'\ \A\<^sub>R' \* Q'\ \A\<^sub>R' \* Q\ suppT suppS have "(\, R' \ Q', R' \ (T \ !Q)) \ Rel" + by(intro FrameParPres) (auto simp add: fresh_star_def fresh_def psi.supp) + then have "(\, R' \ Q', (R' \ T) \ !Q) \ Rel" by(blast intro: Assoc Trans) + with \xvec \* \\ \xvec \* (!Q)\ have "(\, \\*xvec\(R' \ Q'), (\\*xvec\(R' \ T)) \ !Q) \ Rel" + by(metis ResPres psiFreshVec ScopeExt Trans) + ultimately have "(\, \\*xvec\(R' \ P'), \\*xvec\(R' \ Q')) \ Rel'" by(blast intro: cSym Compose) + } + ultimately show ?case by blast + next + case(cComm2 \\<^sub>Q M xvec N R' A\<^sub>R \\<^sub>R K Q' A\<^sub>Q) + from \extractFrame (!Q) = \A\<^sub>Q, \\<^sub>Q\\ have "A\<^sub>Q = []" and "\\<^sub>Q = SBottom'" by simp+ + have RTrans: "\ \ \\<^sub>Q \ R \M\\*xvec\\N\ \ R'" and FrR: "extractFrame R = \A\<^sub>R, \\<^sub>R\" by fact+ + then obtain p \' A\<^sub>R' \\<^sub>R' where S: "set p \ set xvec \ set(p \ xvec)" + and FrR': "extractFrame R' = \A\<^sub>R', \\<^sub>R'\" and "(p \ \\<^sub>R) \ \' \ \\<^sub>R'" and "A\<^sub>R' \* \" + and "A\<^sub>R' \* N" and "A\<^sub>R' \* R'" and "A\<^sub>R' \* P" and "A\<^sub>R' \* Q" and "(p \ xvec) \* \" + and "(p \ xvec) \* P" and "(p \ xvec) \* Q" and "xvec \* A\<^sub>R'" and "(p \ xvec) \* A\<^sub>R'" + and "distinctPerm p" and "(p \ xvec) \* R'" and "(p \ xvec) \* N" + using \distinct A\<^sub>R\ \A\<^sub>R \* R\ \A\<^sub>R \* M\ \A\<^sub>R \* xvec\ \A\<^sub>R \* N\ \A\<^sub>R \* \\ \A\<^sub>R \* P\ \A\<^sub>R \* (!Q)\ + \xvec \* \\ \xvec \* P\ \xvec \* (!Q)\ \xvec \* R\ \xvec \* M\ \distinct xvec\ + by(auto intro: expandFrame[where C="(\, P, Q)" and C'="(\, P, Q)"]) + + from RTrans S \(p \ xvec) \* N\ \(p \ xvec) \* R'\ have "\ \ \\<^sub>Q \ R \M\\*(p \ xvec)\\(p \ N)\ \ (p \ R')" + apply(simp add: residualInject) + by(subst boundOutputChainAlpha''[symmetric]) auto + + moreover have QTrans: "\ \ \\<^sub>R \ !Q \K\N\ \ Q'" by fact + with QTrans S \(p \ xvec) \* N\ have "\ \ \\<^sub>R \ !Q \K\(p \ N)\ \ (p \ Q')" using \distinctPerm p\ \xvec \* (!Q)\ \(p \ xvec) \* Q\ + by(intro inputAlpha) auto + with \(\, P, Q) \ Rel\ \guarded P\ + obtain P' S T where PTrans: "\ \ \\<^sub>R \ !P \K\(p \ N)\ \ P'" and "(\ \ \\<^sub>R, P', T \ !P) \ Rel" + and "(\ \ \\<^sub>R, (p \ Q'), S \ !Q) \ Rel" and "(\ \ \\<^sub>R, S, T) \ Rel" + and suppT: "((supp T)::name set) \ supp P'" and suppS: "((supp S)::name set) \ supp(p \ Q')" + by(drule_tac cSym) (auto dest: Der cExt) + note \\ \ \\<^sub>R \ \\<^sub>Q \ M \ K\ + ultimately have "\ \ R \ !P \\ \ \\*(p \ xvec)\((p \ R') \ P')" + using PTrans FrR \\\<^sub>Q = SBottom'\ \(p \ xvec) \* P\ \A\<^sub>R \* \\ \A\<^sub>R \* R\ \A\<^sub>R \* M\ \A\<^sub>R \* P\ + by(intro Comm2) (assumption | simp)+ + + moreover from \A\<^sub>R' \* P\ \A\<^sub>R' \* Q\ \A\<^sub>R' \* N\ S \xvec \* A\<^sub>R'\ \(p \ xvec) \* A\<^sub>R'\ PTrans QTrans \distinctPerm p\ have "A\<^sub>R' \* P'" and "A\<^sub>R' \* Q'" + apply(drule_tac inputFreshChainDerivative) + apply simp + apply(subst pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst, symmetric, of _ _ p], simp) + apply fastforce + using QTrans \A\<^sub>R' \* N\ \A\<^sub>R' \* Q\ inputFreshChainDerivative by force + from \xvec \* P\ \(p \ xvec) \* N\ PTrans \distinctPerm p\ have "(p \ xvec) \* (p \ P')" + apply(drule_tac inputFreshChainDerivative) + apply simp + by(subst pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst, symmetric, of _ _ p], simp)+ + + { + from \(\ \ \\<^sub>R, P', T \ !P) \ Rel\ have "(p \ (\ \ \\<^sub>R), (p \ P'), p \ (T \ !P)) \ Rel" + by(rule Closed) + with \xvec \* \\ \(p \ xvec) \* \\ \xvec \* P\ \(p \ xvec) \* P\ S have "(\ \ (p \ \\<^sub>R), p \ P', (p \ T) \ !P) \ Rel" + by(simp add: eqvts) + then have "((\ \ (p \ \\<^sub>R)) \ \', p \ P', (p \ T) \ !P) \ Rel" by(rule cExt) + with \(p \ \\<^sub>R) \ \' \ \\<^sub>R'\ have "(\ \ \\<^sub>R', (p \ P'), (p \ T) \ !P) \ Rel" + by(metis Associativity StatEq compositionSym) + with FrR' \A\<^sub>R' \* \\ \A\<^sub>R' \* P'\ \A\<^sub>R' \* P\ \xvec \* A\<^sub>R'\ \(p \ xvec) \* A\<^sub>R'\ S \distinctPerm p\ suppT + have "(\, R' \ (p \ P'), R' \ ((p \ T) \ !P)) \ Rel" + apply(intro FrameParPres) + apply(assumption | simp add: freshChainSimps)+ + by(auto simp add: fresh_star_def fresh_def) + then have "(\, R' \ (p \ P'), (R' \ (p \ T)) \ !P) \ Rel" by(blast intro: Assoc Trans) + with \xvec \* \\ \xvec \* P\ have "(\, \\*xvec\(R' \ (p \ P')), (\\*xvec\(R' \ (p \ T))) \ !P) \ Rel" + by(metis ResPres psiFreshVec ScopeExt Trans) + then have "(\, \\*(p \ xvec)\((p \ R') \ P'), (\\*xvec\(R' \ (p \ T))) \ !P) \ Rel" + using \(p \ xvec) \* R'\ \(p \ xvec) \* (p \ P')\ S \distinctPerm p\ + apply(erule_tac rev_mp) + by(subst resChainAlpha[of p]) auto + moreover from \(\, P, Q) \ Rel\ \guarded P\ \guarded Q\ have "(\, (\\*xvec\(R' \ (p \ T))) \ !P, (\\*xvec\(R' \ (p \ T))) \ !Q) \ Rel'" + by(rule C1) + moreover from \(\ \ \\<^sub>R, (p \ Q'), S \ !Q) \ Rel\ \(\ \ \\<^sub>R, S, T) \ Rel\ have "(\ \ \\<^sub>R, (p \ Q'), T \ !Q) \ Rel" + by(blast intro: ParPres Trans) + then have "(p \ (\ \ \\<^sub>R), p \ p \ Q', p \ (T \ !Q)) \ Rel" by(rule Closed) + with S \xvec \* \\ \(p \ xvec) \* \\ \xvec \* (!Q)\ \(p \ xvec) \* Q\ \distinctPerm p\ + have "(\ \ (p \ \\<^sub>R), Q', (p \ T) \ !Q) \ Rel" by(simp add: eqvts) + then have "((\ \ (p \ \\<^sub>R)) \ \', Q', (p \ T) \ !Q) \ Rel" by(rule cExt) + with \(p \ \\<^sub>R) \ \' \ \\<^sub>R'\ have "(\ \ \\<^sub>R', Q', (p \ T) \ !Q) \ Rel" + by(metis Associativity StatEq compositionSym) + with FrR' \A\<^sub>R' \* \\ \A\<^sub>R' \* P'\ \A\<^sub>R' \* Q'\ \A\<^sub>R' \* Q\ suppT suppS \xvec \* A\<^sub>R'\ \(p \ xvec) \* A\<^sub>R'\ S \distinctPerm p\ + have "(\, R' \ Q', R' \ ((p \ T) \ !Q)) \ Rel" + apply(intro FrameParPres) + apply(assumption | simp)+ + apply(simp add: freshChainSimps) + by(auto simp add: fresh_star_def fresh_def) + then have "(\, R' \ Q', (R' \ (p \ T)) \ !Q) \ Rel" by(blast intro: Assoc Trans) + with \xvec \* \\ \xvec \* (!Q)\ have "(\, \\*xvec\(R' \ Q'), (\\*xvec\(R' \ (p \ T))) \ !Q) \ Rel" + by(metis ResPres psiFreshVec ScopeExt Trans) + ultimately have "(\, \\*(p \ xvec)\((p \ R') \ P'), \\*xvec\(R' \ Q')) \ Rel'" by(blast intro: cSym Compose) + } + ultimately show ?case by blast + next + case(cBrMerge \\<^sub>Q M N R' A\<^sub>R \\<^sub>R Q' A\<^sub>Q) + from \extractFrame (!Q) = \A\<^sub>Q, \\<^sub>Q\\ have "A\<^sub>Q = []" and "\\<^sub>Q = SBottom'" by simp+ + have RTrans: "\ \ \\<^sub>Q \ R \\M\N\ \ R'" and FrR: "extractFrame R = \A\<^sub>R, \\<^sub>R\" by fact+ + have QTrans: "\ \ \\<^sub>R \ !Q \\M\N\ \ Q'" by fact + with \(\, P, Q) \ Rel\ \guarded P\ + obtain P' S T where PTrans: "\ \ \\<^sub>R \ !P \\M\N\ \ P'" and "(\ \ \\<^sub>R, P', T \ !P) \ Rel" + and "(\ \ \\<^sub>R, Q', S \ !Q) \ Rel" and "(\ \ \\<^sub>R, S, T) \ Rel" + and suppT: "((supp T)::name set) \ supp P'" and suppS: "((supp S)::name set) \ supp Q'" + by(drule_tac cSym) (auto dest: Der intro: cExt) + with RTrans QTrans FrR have "\ \ R \ !P \\M\N\ \ (R' \ P')" + using PTrans \\\<^sub>Q = SBottom'\ \A\<^sub>R \* \\ \A\<^sub>R \* R\ \A\<^sub>R \* M\ \A\<^sub>R \* P\ + by(intro BrMerge) (assumption | simp)+ + + moreover from \A\<^sub>R \* P\ \A\<^sub>R \* (!Q)\ \A\<^sub>R \* N\ PTrans QTrans + have "A\<^sub>R \* P'" and "A\<^sub>R \* Q'" by(force dest: brinputFreshChainDerivative)+ + moreover with RTrans FrR \distinct A\<^sub>R\ \A\<^sub>R \* R\ \A\<^sub>R \* N\ \A\<^sub>R \* \\ \A\<^sub>R \* P\ \A\<^sub>R \* (!Q)\ \A\<^sub>R \* M\ + obtain \' A\<^sub>R' \\<^sub>R' where FrR': "extractFrame R' = \A\<^sub>R', \\<^sub>R'\" and "\\<^sub>R \ \' \ \\<^sub>R'" and "A\<^sub>R' \* \" + and "A\<^sub>R' \* P'" and "A\<^sub>R' \* Q'" and "A\<^sub>R' \* P" and "A\<^sub>R' \* Q" + by(auto intro: expandFrame[where C="(\, P, P', Q, Q')" and C'=\]) + + moreover + { + from \(\ \ \\<^sub>R, P', T \ !P) \ Rel\ have "((\ \ \\<^sub>R) \ \', P', T \ !P) \ Rel" by(rule cExt) + with \\\<^sub>R \ \' \ \\<^sub>R'\ have "(\ \ \\<^sub>R', P', T \ !P) \ Rel" + by(metis Associativity StatEq compositionSym) + with FrR' \A\<^sub>R' \* \\ \A\<^sub>R' \* P'\ \A\<^sub>R' \* P\ suppT have "(\, R' \ P', R' \ (T \ !P)) \ Rel" + by(intro FrameParPres) (auto simp add: fresh_star_def fresh_def psi.supp) + then have one: "(\, R' \ P', (R' \ T) \ !P) \ Rel" by(blast intro: Assoc Trans) + from \(\, P, Q) \ Rel\ \guarded P\ \guarded Q\ have two: "(\, (R' \ T) \ !P, (R' \ T) \ !Q) \ Rel'" + by(rule C1) + from \(\ \ \\<^sub>R, Q', S \ !Q) \ Rel\ \(\ \ \\<^sub>R, S, T) \ Rel\ have "(\ \ \\<^sub>R, Q', T \ !Q) \ Rel" + by(blast intro: ParPres Trans) + then have "((\ \ \\<^sub>R) \ \', Q', T \ !Q) \ Rel" by(rule cExt) + with \\\<^sub>R \ \' \ \\<^sub>R'\ have "(\ \ \\<^sub>R', Q', T \ !Q) \ Rel" + by(metis Associativity StatEq compositionSym) + with FrR' \A\<^sub>R' \* \\ \A\<^sub>R' \* P'\ \A\<^sub>R' \* Q'\ \A\<^sub>R' \* Q\ suppT suppS have "(\, R' \ Q', R' \ (T \ !Q)) \ Rel" + by(intro FrameParPres) (auto simp add: fresh_star_def fresh_def psi.supp) + then have three: "(\, R' \ Q', (R' \ T) \ !Q) \ Rel" by(blast intro: Assoc Trans) + from one two three have "(\, (R' \ P'), (R' \ Q')) \ Rel'" by(blast intro: cSym Compose) + } + ultimately show ?case by blast + next + case(cBrComm1 \\<^sub>Q M N R' A\<^sub>R \\<^sub>R xvec Q' A\<^sub>Q) + from \\M\\*xvec\\N\ = \\ have "xvec=bn \" + by(auto simp add: action.inject) + from \xvec = bn \\ \bn \ \* P\ \bn \ \* R\ + have "xvec \* P" and "xvec \* R" by simp+ + + from \extractFrame (!Q) = \A\<^sub>Q, \\<^sub>Q\\ have "A\<^sub>Q = []" and "\\<^sub>Q = SBottom'" by simp+ + have RTrans: "\ \ \\<^sub>Q \ R \\M\N\ \ R'" and FrR: "extractFrame R = \A\<^sub>R, \\<^sub>R\" by fact+ + moreover have QTrans: "\ \ \\<^sub>R \ !Q \\M\\*xvec\\N\ \ Q'" by fact + from FrR \xvec \* R\ \A\<^sub>R \* xvec\ have "xvec \* \\<^sub>R" by(force dest: extractFrameFreshChain) + with QTrans \(\, P, Q) \ Rel\ \xvec \* \\\xvec \* P\ \xvec \* (!Q)\ \guarded P\ \xvec \* M\ + obtain P' S T where PTrans: "\ \ \\<^sub>R \ !P \\M\\*xvec\\N\ \ P'" and "(\ \ \\<^sub>R, P', T \ !P) \ Rel" + and "(\ \ \\<^sub>R, Q', S \ !Q) \ Rel" and "(\ \ \\<^sub>R, S, T) \ Rel" + and suppT: "((supp T)::name set) \ supp P'" and suppS: "((supp S)::name set) \ supp Q'" + apply(drule_tac cSym) + by(metis Der \bn \ \* Q\ \xvec = bn \\ cBrComm1.hyps(30) cExt cSim.hyps(6) freshCompChain(1)) + + ultimately have "\ \ R \ !P \\M\\*xvec\\N\ \ (R' \ P')" + using PTrans \\\<^sub>Q = SBottom'\ \xvec \* R\ \A\<^sub>R \* \\ \A\<^sub>R \* R\ \A\<^sub>R \* M\ \A\<^sub>R \* P\ + by(intro BrComm1) (assumption | simp)+ + + moreover from \A\<^sub>R \* P\ \A\<^sub>R \* (!Q)\ \A\<^sub>R \* xvec\ PTrans QTrans \xvec \* M\ \distinct xvec\ + have "A\<^sub>R \* P'" and "A\<^sub>R \* Q'" by(force dest: broutputFreshChainDerivative)+ + moreover with RTrans FrR \distinct A\<^sub>R\ \A\<^sub>R \* R\ \A\<^sub>R \* N\ \A\<^sub>R \* \\ \A\<^sub>R \* P\ \A\<^sub>R \* (!Q)\ \A\<^sub>R \* M\ + obtain \' A\<^sub>R' \\<^sub>R' where FrR': "extractFrame R' = \A\<^sub>R', \\<^sub>R'\" and "\\<^sub>R \ \' \ \\<^sub>R'" and "A\<^sub>R' \* \" + and "A\<^sub>R' \* P'" and "A\<^sub>R' \* Q'" and "A\<^sub>R' \* P" and "A\<^sub>R' \* Q" + by(auto intro: expandFrame[where C="(\, P, P', Q, Q')" and C'=\]) + + moreover + { + from \(\ \ \\<^sub>R, P', T \ !P) \ Rel\ have "((\ \ \\<^sub>R) \ \', P', T \ !P) \ Rel" by(rule cExt) + with \\\<^sub>R \ \' \ \\<^sub>R'\ have "(\ \ \\<^sub>R', P', T \ !P) \ Rel" + by(metis Associativity StatEq compositionSym) + with FrR' \A\<^sub>R' \* \\ \A\<^sub>R' \* P'\ \A\<^sub>R' \* P\ suppT have "(\, R' \ P', R' \ (T \ !P)) \ Rel" + by(intro FrameParPres) (auto simp add: fresh_star_def fresh_def psi.supp) + then have one: "(\, R' \ P', (R' \ T) \ !P) \ Rel" by(blast intro: Assoc Trans) + from \(\, P, Q) \ Rel\ \guarded P\ \guarded Q\ have two: "(\, (R' \ T) \ !P, (R' \ T) \ !Q) \ Rel'" + by(rule C1) + from \(\ \ \\<^sub>R, Q', S \ !Q) \ Rel\ \(\ \ \\<^sub>R, S, T) \ Rel\ have "(\ \ \\<^sub>R, Q', T \ !Q) \ Rel" + by(blast intro: ParPres Trans) + then have "((\ \ \\<^sub>R) \ \', Q', T \ !Q) \ Rel" by(rule cExt) + with \\\<^sub>R \ \' \ \\<^sub>R'\ have "(\ \ \\<^sub>R', Q', T \ !Q) \ Rel" + by(metis Associativity StatEq compositionSym) + with FrR' \A\<^sub>R' \* \\ \A\<^sub>R' \* P'\ \A\<^sub>R' \* Q'\ \A\<^sub>R' \* Q\ suppT suppS have "(\, R' \ Q', R' \ (T \ !Q)) \ Rel" + by(intro FrameParPres) (auto simp add: fresh_star_def fresh_def psi.supp) + then have three: "(\, R' \ Q', (R' \ T) \ !Q) \ Rel" by(blast intro: Assoc Trans) + from one two three have "(\, (R' \ P'), (R' \ Q')) \ Rel'" by(blast intro: cSym Compose) + } + ultimately show ?case by blast + next + case(cBrComm2 \\<^sub>Q M xvec N R' A\<^sub>R \\<^sub>R Q' A\<^sub>Q) + from \\M\\*xvec\\N\ = \\ have "xvec = bn \" + by(auto simp add: action.inject) + from \xvec = bn \\ \bn \ \* P\ \bn \ \* R\ + have "xvec \* P" and "xvec \* R" by simp+ + + from \extractFrame (!Q) = \A\<^sub>Q, \\<^sub>Q\\ have "A\<^sub>Q = []" and "\\<^sub>Q = SBottom'" by simp+ + have RTrans: "\ \ \\<^sub>Q \ R \\M\\*xvec\\N\ \ R'" and FrR: "extractFrame R = \A\<^sub>R, \\<^sub>R\" by fact+ + then obtain p \' A\<^sub>R' \\<^sub>R' where S: "set p \ set xvec \ set(p \ xvec)" + and FrR': "extractFrame R' = \A\<^sub>R', \\<^sub>R'\" and "(p \ \\<^sub>R) \ \' \ \\<^sub>R'" and "A\<^sub>R' \* \" + and "A\<^sub>R' \* N" and "A\<^sub>R' \* M" and "A\<^sub>R' \* R" and "A\<^sub>R' \* R'" and "A\<^sub>R' \* P" and "A\<^sub>R' \* Q" and "(p \ xvec) \* \" + and "(p \ xvec) \* P" and "(p \ xvec) \* Q" and "xvec \* A\<^sub>R'" and "(p \ xvec) \* A\<^sub>R'" + and "distinctPerm p" and "(p \ xvec) \* R" and "(p \ xvec) \* R'" and "(p \ xvec) \* N" and "(p \ xvec) \* M" + using \distinct A\<^sub>R\ \A\<^sub>R \* R\ \A\<^sub>R \* M\ \A\<^sub>R \* xvec\ \A\<^sub>R \* N\ \A\<^sub>R \* \\ \A\<^sub>R \* P\ \A\<^sub>R \* (!Q)\ + \xvec \* \\ \xvec \* P\ \xvec \* (!Q)\ \xvec \* R\ \xvec \* M\ \distinct xvec\ + by(auto intro: expandFrame[where C="(\, P, R, Q, M)" and C'="(\, P, R, Q, M)"]) + + from RTrans S \(p \ xvec) \* N\ \(p \ xvec) \* R'\ have "\ \ \\<^sub>Q \ R \\M\\*(p \ xvec)\\(p \ N)\ \ (p \ R')" + apply(simp add: residualInject) + by(subst boundOutputChainAlpha''[symmetric]) auto + + moreover have QTrans: "\ \ \\<^sub>R \ !Q \\M\N\ \ Q'" by fact + with QTrans S \(p \ xvec) \* N\ have "\ \ \\<^sub>R \ !Q \\M\(p \ N)\ \ (p \ Q')" using \distinctPerm p\ \xvec \* (!Q)\ \(p \ xvec) \* Q\ + by(intro brinputAlpha) auto + with \(\, P, Q) \ Rel\ \guarded P\ + obtain P' S T where PTrans: "\ \ \\<^sub>R \ !P \\M\(p \ N)\ \ P'" and "(\ \ \\<^sub>R, P', T \ !P) \ Rel" + and "(\ \ \\<^sub>R, (p \ Q'), S \ !Q) \ Rel" and "(\ \ \\<^sub>R, S, T) \ Rel" + and suppT: "((supp T)::name set) \ supp P'" and suppS: "((supp S)::name set) \ supp(p \ Q')" + by(drule_tac cSym) (auto dest: Der cExt) + ultimately have "\ \ R \ !P \\M\\*(p \ xvec)\\(p \ N)\ \ ((p \ R') \ P')" + using PTrans FrR \\\<^sub>Q = SBottom'\ \(p \ xvec) \* P\ \A\<^sub>R \* \\ \A\<^sub>R \* R\ \A\<^sub>R \* M\ \A\<^sub>R \* P\ + by(intro BrComm2) (assumption | simp)+ + then have "p \ (\ \ R \ !P \\M\\*(p \ xvec)\\(p \ N)\ \ ((p \ R') \ P'))" by simp + with \distinctPerm p\ have "(p \ \) \ (p \ R) \ !(p \ P) \\(p \ M)\\*xvec\\N\ \ (R' \ (p \ P'))" + by(simp add: eqvts) + with S \xvec \* \\ \(p \ xvec) \* \\ \xvec \* R\ \(p \ xvec) \* R\ + \xvec \* P\ \(p \ xvec) \* P\ \xvec \* M\ \(p \ xvec) \* M\ + have "\ \ R \ !P \\M\\*xvec\\N\ \ (R' \ (p \ P'))" + by simp + + moreover from \A\<^sub>R' \* P\ \A\<^sub>R' \* Q\ \A\<^sub>R' \* N\ S \xvec \* A\<^sub>R'\ \(p \ xvec) \* A\<^sub>R'\ PTrans QTrans \distinctPerm p\ have "A\<^sub>R' \* P'" and "A\<^sub>R' \* Q'" + apply(drule_tac brinputFreshChainDerivative, simp) + apply(subst pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst, symmetric, of _ _ p], simp) + apply force + using QTrans \A\<^sub>R' \* N\ \A\<^sub>R' \* Q\ brinputFreshChainDerivative by force + from \xvec \* P\ \(p \ xvec) \* N\ PTrans \distinctPerm p\ have "(p \ xvec) \* (p \ P')" + apply(drule_tac brinputFreshChainDerivative, simp) + apply(subst pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst, symmetric, of _ _ p], simp) + by(subst pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst, symmetric, of _ _ p], simp) + + { + from \(\ \ \\<^sub>R, P', T \ !P) \ Rel\ have "(p \ (\ \ \\<^sub>R), (p \ P'), p \ (T \ !P)) \ Rel" + by(rule Closed) + with \xvec \* \\ \(p \ xvec) \* \\ \xvec \* P\ \(p \ xvec) \* P\ S have "(\ \ (p \ \\<^sub>R), p \ P', (p \ T) \ !P) \ Rel" + by(simp add: eqvts) + then have "((\ \ (p \ \\<^sub>R)) \ \', p \ P', (p \ T) \ !P) \ Rel" + by(rule cExt) + with \(p \ \\<^sub>R) \ \' \ \\<^sub>R'\ have "(\ \ \\<^sub>R', (p \ P'), (p \ T) \ !P) \ Rel" + by(metis Associativity StatEq compositionSym) + with FrR' \A\<^sub>R' \* \\ \A\<^sub>R' \* P'\ \A\<^sub>R' \* P\ \xvec \* A\<^sub>R'\ \(p \ xvec) \* A\<^sub>R'\ S \distinctPerm p\ suppT + have "(\, R' \ (p \ P'), R' \ ((p \ T) \ !P)) \ Rel" + apply(intro FrameParPres) + apply(assumption | simp add: freshChainSimps)+ + by(auto simp add: fresh_star_def fresh_def) + then have one: "(\, R' \ (p \ P'), (R' \ (p \ T)) \ !P) \ Rel" by(blast intro: Assoc Trans) + from \(\, P, Q) \ Rel\ \guarded P\ \guarded Q\ have two: "(\, (R' \ (p \ T)) \ !P, (R' \ (p \ T)) \ !Q) \ Rel'" + by(rule C1) + from \(\ \ \\<^sub>R, (p \ Q'), S \ !Q) \ Rel\ \(\ \ \\<^sub>R, S, T) \ Rel\ have "(\ \ \\<^sub>R, (p \ Q'), T \ !Q) \ Rel" + by(blast intro: ParPres Trans) + then have "(p \ (\ \ \\<^sub>R), p \ p \ Q', p \ (T \ !Q)) \ Rel" by(rule Closed) + with S \xvec \* \\ \(p \ xvec) \* \\ \xvec \* (!Q)\ \(p \ xvec) \* Q\ \distinctPerm p\ + have "(\ \ (p \ \\<^sub>R), Q', (p \ T) \ !Q) \ Rel" by(simp add: eqvts) + then have "((\ \ (p \ \\<^sub>R)) \ \', Q', (p \ T) \ !Q) \ Rel" by(rule cExt) + with \(p \ \\<^sub>R) \ \' \ \\<^sub>R'\ have "(\ \ \\<^sub>R', Q', (p \ T) \ !Q) \ Rel" + by(metis Associativity StatEq compositionSym) + with FrR' \A\<^sub>R' \* \\ \A\<^sub>R' \* P'\ \A\<^sub>R' \* Q'\ \A\<^sub>R' \* Q\ suppT suppS \xvec \* A\<^sub>R'\ \(p \ xvec) \* A\<^sub>R'\ S \distinctPerm p\ + have "(\, R' \ Q', R' \ ((p \ T) \ !Q)) \ Rel" + apply(intro FrameParPres) + apply(assumption | simp)+ + apply(simp add: freshChainSimps) + by(auto simp add: fresh_star_def fresh_def) + then have three: "(\, R' \ Q', (R' \ (p \ T)) \ !Q) \ Rel" by(blast intro: Assoc Trans) + from one two three have "(\, (R' \ (p \ P')), (R' \ Q')) \ Rel'" by(blast intro: cSym Compose) + } + ultimately show ?case by blast + qed +qed + +notation relcomp (infixr "O" 75) + +end + +end diff --git a/thys/Broadcast_Psi/Sim_Struct_Cong.thy b/thys/Broadcast_Psi/Sim_Struct_Cong.thy new file mode 100644 --- /dev/null +++ b/thys/Broadcast_Psi/Sim_Struct_Cong.thy @@ -0,0 +1,3457 @@ +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 \ \\<^sub>Q P M xvec N P' A\<^sub>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 \ \\<^sub>Q P M N P' A\<^sub>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\<^sub>Q\<^sub>R \\<^sub>Q\<^sub>R) + from \A\<^sub>Q\<^sub>R \* (\, P, Q, R, \)\ have "A\<^sub>Q\<^sub>R \* Q" and "A\<^sub>Q\<^sub>R \* R" + by simp+ + with \extractFrame(Q \ R) = \A\<^sub>Q\<^sub>R, \\<^sub>Q\<^sub>R\\ \distinct A\<^sub>Q\<^sub>R\ + obtain A\<^sub>Q \\<^sub>Q A\<^sub>R \\<^sub>R where "A\<^sub>Q\<^sub>R = A\<^sub>Q@A\<^sub>R" and "\\<^sub>Q\<^sub>R = \\<^sub>Q \ \\<^sub>R" and FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" and FrR: "extractFrame R = \A\<^sub>R, \\<^sub>R\" + and "A\<^sub>Q \* \\<^sub>R" and "A\<^sub>R \* \\<^sub>Q" + by(auto intro: mergeFrameE dest: extractFrameFreshChain) + + from \A\<^sub>Q\<^sub>R = A\<^sub>Q@A\<^sub>R\ \A\<^sub>Q\<^sub>R \* \\ \A\<^sub>Q\<^sub>R \* P\ \A\<^sub>Q\<^sub>R \* Q\ \A\<^sub>Q\<^sub>R \* \\ + have "A\<^sub>Q \* \" and "A\<^sub>R \* \" and "A\<^sub>Q \* P" and "A\<^sub>R \* P" and "A\<^sub>Q \* Q" and "A\<^sub>R \* Q" and "A\<^sub>Q \* \" and "A\<^sub>R \* \" + by simp+ + + from \\ \ \\<^sub>Q\<^sub>R \ P \\ \ P'\ \\\<^sub>Q\<^sub>R = \\<^sub>Q \ \\<^sub>R\ have "(\ \ \\<^sub>R) \ \\<^sub>Q \ P \\ \ P'" + by(metis statEqTransition Associativity Commutativity Composition) + then have "\ \ \\<^sub>R \ P \ Q \\ \ (P' \ Q)" using FrQ \bn \ \* Q\ \A\<^sub>Q \* \\ \A\<^sub>Q \* \\<^sub>R\ \A\<^sub>Q \* P\ \A\<^sub>Q \* \\ + by(intro Par1) auto + then have "\ \ (P \ Q) \ R \\ \ ((P' \ Q) \ R)" using FrR \bn \ \* R\ \A\<^sub>R \* \\ \A\<^sub>R \* P\ \A\<^sub>R \* Q\ \A\<^sub>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\<^sub>P \\<^sub>P) + from \A\<^sub>P \* (\, P, Q, R, \)\ have "A\<^sub>P \* Q" and "A\<^sub>P \* R" and "A\<^sub>P \* \" + by simp+ + have FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" by fact + with \bn \ \* P\ \A\<^sub>P \* \\ have "bn \ \* \\<^sub>P" by(auto dest: extractFrameFreshChain) + with \bn \ \* \\ have "bn \ \* (\ \ \\<^sub>P)" by force + with \\ \ \\<^sub>P \ Q \ R \\ \ QR\ + show ?case using \bn \ \* Q\ \bn \ \* R\ \bn \ \* subject \\ \A\<^sub>P \* Q\ \A\<^sub>P \* R\ + proof(induct rule: parCasesSubject[where C = "(A\<^sub>P, \\<^sub>P, P, Q, R, \)"]) + case(cPar1 Q' A\<^sub>R \\<^sub>R) + from \A\<^sub>R \* (A\<^sub>P, \\<^sub>P, P, Q, R, \)\ have "A\<^sub>R \* A\<^sub>P" and "A\<^sub>R \* P" and "A\<^sub>R \* Q" and "A\<^sub>R \* \\<^sub>P" and "A\<^sub>R \* \" + by simp+ + from \A\<^sub>P \* R\ \extractFrame R = \A\<^sub>R, \\<^sub>R\\ \A\<^sub>R \* A\<^sub>P\ have "A\<^sub>P \* \\<^sub>R" + by(auto dest: extractFrameFreshChain) + from \(\ \ \\<^sub>P) \ \\<^sub>R \ Q \\ \ Q'\ have "(\ \ \\<^sub>R) \ \\<^sub>P \ Q \\ \ Q'" + by(metis statEqTransition Associativity Commutativity Composition) + then have "\ \ \\<^sub>R \ P \ Q \\ \ (P \ Q')" + using FrP \bn \ \* P\ \A\<^sub>P \* \\ \A\<^sub>P \* \\<^sub>R\ \A\<^sub>P \* Q\ \A\<^sub>P \* \\ + by(intro Par2) (assumption | force)+ + then have "\ \ (P \ Q) \ R \\ \ ((P \ Q') \ R)" + using \extractFrame R = \A\<^sub>R, \\<^sub>R\\ \bn \ \* R\ \A\<^sub>R \* \\ \A\<^sub>R \* P\ \A\<^sub>R \* Q\ \A\<^sub>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\<^sub>Q \\<^sub>Q) + from \A\<^sub>Q \* (A\<^sub>P, \\<^sub>P, P, Q, R, \)\ have "A\<^sub>Q \* A\<^sub>P" and "A\<^sub>Q \* R" and "A\<^sub>Q \* \\<^sub>P" and "A\<^sub>Q \* \" + by simp+ + have FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" by fact + from \A\<^sub>P \* Q\ FrQ \A\<^sub>Q \* A\<^sub>P\ have "A\<^sub>P \* \\<^sub>Q" + by(auto dest: extractFrameFreshChain) + from \(\ \ \\<^sub>P) \ \\<^sub>Q \ R \\ \ R'\ + have "\ \ (\\<^sub>P \ \\<^sub>Q) \ R \\ \ R'" + by(blast intro: statEqTransition Associativity) + moreover from FrP FrQ \A\<^sub>Q \* A\<^sub>P\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>Q \* \\<^sub>P\ + have "extractFrame(P \ Q) = \(A\<^sub>P@A\<^sub>Q), \\<^sub>P \ \\<^sub>Q\ " by simp + moreover from \bn \ \* P\ \bn \ \* Q\ have "bn \ \* (P \ Q)" by simp + moreover from \A\<^sub>P \* \\ \A\<^sub>Q \* \\ have "(A\<^sub>P@A\<^sub>Q) \* \" by simp + moreover from \A\<^sub>P \* R\ \A\<^sub>Q \* R\ have "(A\<^sub>P@A\<^sub>Q) \* R" by simp + moreover from \A\<^sub>P \* \\ \A\<^sub>Q \* \\ have "(A\<^sub>P@A\<^sub>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 \\<^sub>R M N Q' A\<^sub>Q \\<^sub>Q K xvec R' A\<^sub>R) + from \A\<^sub>Q \* (A\<^sub>P, \\<^sub>P, P, Q, R, \)\ + have "A\<^sub>Q \* P" and "A\<^sub>Q \* Q" and "A\<^sub>Q \* R" and "A\<^sub>Q \* A\<^sub>P" and "A\<^sub>Q \* \\<^sub>P" and "A\<^sub>Q \* \" by simp+ + from \A\<^sub>R \* (A\<^sub>P, \\<^sub>P, P, Q, R, \)\ have "A\<^sub>R \* P" and "A\<^sub>R \* Q" and "A\<^sub>R \* R" and "A\<^sub>R \* A\<^sub>P" and "A\<^sub>R \* \" by simp+ + from \xvec \* (A\<^sub>P, \\<^sub>P, P, Q, R, \)\ have "xvec \* A\<^sub>P" and "xvec \* P" and "xvec \* Q" and "xvec \* \" by simp+ + + have FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" by fact + with \A\<^sub>P \* Q\ \A\<^sub>Q \* A\<^sub>P\ have "A\<^sub>P \* \\<^sub>Q" + by(auto dest: extractFrameFreshChain) + have FrR: "extractFrame R = \A\<^sub>R, \\<^sub>R\" by fact + with \A\<^sub>P \* R\ \A\<^sub>R \* A\<^sub>P\ have "A\<^sub>P \* \\<^sub>R" + by(auto dest: extractFrameFreshChain) + from \(\ \ \\<^sub>P) \ \\<^sub>Q \ R \K\\*xvec\\N\ \ R'\ \A\<^sub>P \* R\ \xvec \* A\<^sub>P\ \xvec \* K\ \distinct xvec\ have "A\<^sub>P \* N" + by(auto intro: outputFreshChainDerivative) + + from \(\ \ \\<^sub>P) \ \\<^sub>R \ Q \M\N\ \ Q'\ have "(\ \ \\<^sub>R) \ \\<^sub>P \ Q \M\N\ \ Q'" + by(metis statEqTransition Associativity Commutativity Composition) + then have "\ \ \\<^sub>R \ P \ Q \M\N\ \ (P \ Q')" using FrP \A\<^sub>P \* \\ \A\<^sub>P \* \\<^sub>R\ \A\<^sub>P \* Q\ \A\<^sub>P \* M\ \A\<^sub>P \* N\ + by(intro Par2) auto + moreover from FrP FrQ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>Q \* A\<^sub>P\ \A\<^sub>Q \* \\<^sub>P\ have "extractFrame(P \ Q) = \(A\<^sub>P@A\<^sub>Q), \\<^sub>P \ \\<^sub>Q\" + by simp + moreover from \(\ \ \\<^sub>P) \ \\<^sub>Q \ R \K\\*xvec\\N\ \ R'\ have "\ \ \\<^sub>P \ \\<^sub>Q \ R \K\\*xvec\\N\ \ R'" + by(metis statEqTransition Associativity) + moreover note \extractFrame R = \A\<^sub>R, \\<^sub>R\\ + moreover from \(\ \ \\<^sub>P) \ \\<^sub>Q \ \\<^sub>R \ M \ K\ have "\ \ (\\<^sub>P \ \\<^sub>Q) \ \\<^sub>R \ M \ K" + by(metis statEqEnt Associativity Commutativity Composition) + ultimately have "\ \ (P \ Q) \ R \\ \ \\*xvec\((P \ Q') \ R')" + using \A\<^sub>P \* \\ \A\<^sub>Q \* \\ \A\<^sub>R \* \\ \A\<^sub>P \* P\ \A\<^sub>Q \* P\ \A\<^sub>R \* P\ \A\<^sub>P \* Q\ \A\<^sub>Q \* Q\ \A\<^sub>R \* Q\ \A\<^sub>P \* R\ \A\<^sub>Q \* R\ \A\<^sub>R \* R\ + \A\<^sub>P \* M\ \A\<^sub>Q \* M\ \A\<^sub>R \* K\ \A\<^sub>R \* A\<^sub>P\ \A\<^sub>Q \* A\<^sub>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 \\<^sub>R M xvec N Q' A\<^sub>Q \\<^sub>Q K R' A\<^sub>R) + from \A\<^sub>Q \* (A\<^sub>P, \\<^sub>P, P, Q, R, \)\ + have "A\<^sub>Q \* P" and "A\<^sub>Q \* Q" and "A\<^sub>Q \* R" and "A\<^sub>Q \* A\<^sub>P" and "A\<^sub>Q \* \" and "A\<^sub>Q \* \\<^sub>P" by simp+ + from \A\<^sub>R \* (A\<^sub>P, \\<^sub>P, P, Q, R, \)\ have "A\<^sub>R \* P" and "A\<^sub>R \* Q" and "A\<^sub>R \* R" and "A\<^sub>R \* A\<^sub>P"and "A\<^sub>R \* \" by simp+ + from \xvec \* (A\<^sub>P, \\<^sub>P, P, Q, R, \)\ have "xvec \* A\<^sub>P" and "xvec \* P" and "xvec \* Q" and "xvec \* \" by simp+ + + have FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" by fact + with \A\<^sub>P \* Q\ \A\<^sub>Q \* A\<^sub>P\ have "A\<^sub>P \* \\<^sub>Q" + by(auto dest: extractFrameFreshChain) + have FrR: "extractFrame R = \A\<^sub>R, \\<^sub>R\" by fact + with \A\<^sub>P \* R\ \A\<^sub>R \* A\<^sub>P\ have "A\<^sub>P \* \\<^sub>R" + by(auto dest: extractFrameFreshChain) + + from \(\ \ \\<^sub>P) \ \\<^sub>R \ Q \M\\*xvec\\N\ \ Q'\ \A\<^sub>P \* Q\ \xvec \* A\<^sub>P\ \xvec \* M\ \distinct xvec\ have "A\<^sub>P \* N" + by(auto intro: outputFreshChainDerivative) + + from \(\ \ \\<^sub>P) \ \\<^sub>R \ Q \M\\*xvec\\N\ \ Q'\ have "(\ \ \\<^sub>R) \ \\<^sub>P \ Q \M\\*xvec\\N\ \ Q'" + by(metis statEqTransition Associativity Commutativity Composition) + then have "\ \ \\<^sub>R \ P \ Q \M\\*xvec\\N\ \ (P \ Q')" using FrP \A\<^sub>P \* \\ \A\<^sub>P \* \\<^sub>R\ \A\<^sub>P \* Q\ \A\<^sub>P \* M\ \A\<^sub>P \* N\ \xvec \* P\ \xvec \* A\<^sub>P\ + by(intro Par2) auto + moreover from FrP FrQ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>Q \* A\<^sub>P\ \A\<^sub>Q \* \\<^sub>P\ have "extractFrame(P \ Q) = \(A\<^sub>P@A\<^sub>Q), \\<^sub>P \ \\<^sub>Q\" + by simp+ + moreover from \(\ \ \\<^sub>P) \ \\<^sub>Q \ R \K\N\ \ R'\ have "\ \ \\<^sub>P \ \\<^sub>Q \ R \K\N\ \ R'" + by(metis statEqTransition Associativity) + moreover note \extractFrame R = \A\<^sub>R, \\<^sub>R\\ + moreover from \(\ \ \\<^sub>P) \ \\<^sub>Q \ \\<^sub>R \ M \ K\ have "\ \ (\\<^sub>P \ \\<^sub>Q) \ \\<^sub>R \ M \ K" + by(metis statEqEnt Associativity Commutativity Composition) + ultimately have "\ \ (P \ Q) \ R \\ \ \\*xvec\((P \ Q') \ R')" + using \A\<^sub>P \* \\ \A\<^sub>Q \* \\ \A\<^sub>R \* \\ \A\<^sub>P \* P\ \A\<^sub>Q \* P\ \A\<^sub>R \* P\ \A\<^sub>P \* Q\ \A\<^sub>Q \* Q\ \A\<^sub>R \* Q\ \A\<^sub>P \* R\ \A\<^sub>Q \* R\ \A\<^sub>R \* R\ + \A\<^sub>P \* M\ \A\<^sub>Q \* M\ \A\<^sub>R \* K\ \A\<^sub>R \* A\<^sub>P\ \A\<^sub>Q \* A\<^sub>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 \\<^sub>R M N Q' A\<^sub>Q \\<^sub>Q R' A\<^sub>R) + from \A\<^sub>P \* \\ \\ = \M\N\\ have "A\<^sub>P \* M" and "A\<^sub>P \* N" by simp+ + from \A\<^sub>Q \* (A\<^sub>P, \\<^sub>P, P, Q, R, \)\ + have "A\<^sub>Q \* P" and "A\<^sub>Q \* Q" and "A\<^sub>Q \* R" and "A\<^sub>Q \* A\<^sub>P" and "A\<^sub>Q \* \\<^sub>P" and "A\<^sub>Q \* \" by simp+ + from \A\<^sub>R \* (A\<^sub>P, \\<^sub>P, P, Q, R, \)\ have "A\<^sub>R \* P" and "A\<^sub>R \* Q" and "A\<^sub>R \* R" and "A\<^sub>R \* A\<^sub>P" and "A\<^sub>R \* \" by simp+ + + have FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" by fact + with \A\<^sub>P \* Q\ \A\<^sub>Q \* A\<^sub>P\ have "A\<^sub>P \* \\<^sub>Q" + by(auto dest: extractFrameFreshChain) + have FrR: "extractFrame R = \A\<^sub>R, \\<^sub>R\" by fact + with \A\<^sub>P \* R\ \A\<^sub>R \* A\<^sub>P\ have "A\<^sub>P \* \\<^sub>R" + by(auto dest: extractFrameFreshChain) + + from \(\ \ \\<^sub>P) \ \\<^sub>R \ Q \\M\N\ \ Q'\ have "(\ \ \\<^sub>R) \ \\<^sub>P \ Q \\M\N\ \ Q'" + by(metis statEqTransition Associativity Commutativity Composition) + then have "\ \ \\<^sub>R \ P \ Q \\M\N\ \ (P \ Q')" using FrP \A\<^sub>P \* \\ \A\<^sub>P \* \\<^sub>R\ \A\<^sub>P \* Q\ \A\<^sub>P \* M\ \A\<^sub>P \* N\ + by(intro Par2) auto + moreover from FrP FrQ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>Q \* A\<^sub>P\ \A\<^sub>Q \* \\<^sub>P\ have "extractFrame(P \ Q) = \(A\<^sub>P@A\<^sub>Q), \\<^sub>P \ \\<^sub>Q\" + by simp + moreover from \(\ \ \\<^sub>P) \ \\<^sub>Q \ R \\M\N\ \ R'\ have "\ \ \\<^sub>P \ \\<^sub>Q \ R \\M\N\ \ R'" + by(metis statEqTransition Associativity) + moreover note \extractFrame R = \A\<^sub>R, \\<^sub>R\\ + ultimately have "\ \ (P \ Q) \ R \\M\N\ \ (P \ Q') \ R'" + using \A\<^sub>P \* \\ \A\<^sub>Q \* \\ \A\<^sub>R \* \\ \A\<^sub>P \* P\ \A\<^sub>Q \* P\ \A\<^sub>R \* P\ \A\<^sub>P \* Q\ \A\<^sub>Q \* Q\ \A\<^sub>R \* Q\ \A\<^sub>P \* R\ \A\<^sub>Q \* R\ \A\<^sub>R \* R\ + \A\<^sub>P \* M\ \A\<^sub>Q \* M\ \A\<^sub>R \* M\ \A\<^sub>R \* A\<^sub>P\ \A\<^sub>Q \* A\<^sub>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 \\<^sub>R M N Q' A\<^sub>Q \\<^sub>Q xvec R' A\<^sub>R) + from \\M\\*xvec\\N\ = \\ have "xvec = bn \" + by(auto simp add: action.inject) + from \\M\\*xvec\\N\ = \\ \A\<^sub>P \* \\ + have "A\<^sub>P \* xvec" and "A\<^sub>P \* M" and "A\<^sub>P \* N" by auto + from \xvec = bn \\ \bn \ \* P\ \bn \ \* \\ \bn \ \* \\<^sub>P\ have "xvec \* P" and "xvec \* \" and "xvec \* \\<^sub>P" + by simp+ + from \A\<^sub>Q \* (A\<^sub>P, \\<^sub>P, P, Q, R, \)\ + have "A\<^sub>Q \* P" and "A\<^sub>Q \* Q" and "A\<^sub>Q \* R" and "A\<^sub>Q \* A\<^sub>P" and "A\<^sub>Q \* \\<^sub>P" and "A\<^sub>Q \* \" by simp+ + from \A\<^sub>R \* (A\<^sub>P, \\<^sub>P, P, Q, R, \)\ have "A\<^sub>R \* P" and "A\<^sub>R \* Q" and "A\<^sub>R \* R" and "A\<^sub>R \* A\<^sub>P" and "A\<^sub>R \* \" by simp+ + + have FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" by fact + with \A\<^sub>P \* Q\ \A\<^sub>Q \* A\<^sub>P\ have "A\<^sub>P \* \\<^sub>Q" + by(auto dest: extractFrameFreshChain) + have FrR: "extractFrame R = \A\<^sub>R, \\<^sub>R\" by fact + with \A\<^sub>P \* R\ \A\<^sub>R \* A\<^sub>P\ have "A\<^sub>P \* \\<^sub>R" + by(auto dest: extractFrameFreshChain) + + from \(\ \ \\<^sub>P) \ \\<^sub>R \ Q \\M\N\ \ Q'\ have "(\ \ \\<^sub>R) \ \\<^sub>P \ Q \\M\N\ \ Q'" + by(metis statEqTransition Associativity Commutativity Composition) + then have "\ \ \\<^sub>R \ P \ Q \\M\N\ \ (P \ Q')" using FrP \A\<^sub>P \* \\ \A\<^sub>P \* \\<^sub>R\ \A\<^sub>P \* Q\ \A\<^sub>P \* M\ \A\<^sub>P \* N\ + by(intro Par2) auto + moreover from FrP FrQ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>Q \* A\<^sub>P\ \A\<^sub>Q \* \\<^sub>P\ have "extractFrame(P \ Q) = \(A\<^sub>P@A\<^sub>Q), \\<^sub>P \ \\<^sub>Q\" + by simp + moreover from \(\ \ \\<^sub>P) \ \\<^sub>Q \ R \\M\\*xvec\\N\ \ R'\ have "\ \ \\<^sub>P \ \\<^sub>Q \ R \\M\\*xvec\\N\ \ R'" + by(metis statEqTransition Associativity) + moreover note \extractFrame R = \A\<^sub>R, \\<^sub>R\\ + ultimately have "\ \ (P \ Q) \ R \\M\\*xvec\\N\ \ ((P \ Q') \ R')" + using \A\<^sub>P \* \\ \A\<^sub>Q \* \\ \A\<^sub>R \* \\ \A\<^sub>P \* P\ \A\<^sub>Q \* P\ \A\<^sub>R \* P\ \A\<^sub>P \* Q\ \A\<^sub>Q \* Q\ \A\<^sub>R \* Q\ \A\<^sub>P \* R\ \A\<^sub>Q \* R\ \A\<^sub>R \* R\ + \A\<^sub>P \* M\ \A\<^sub>Q \* M\ \A\<^sub>R \* M\ \A\<^sub>R \* A\<^sub>P\ \A\<^sub>Q \* A\<^sub>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 \\<^sub>R M xvec N Q' A\<^sub>Q \\<^sub>Q R' A\<^sub>R) + from \\M\\*xvec\\N\ = \\ have "xvec = bn \" + by(auto simp add: action.inject) + from \\M\\*xvec\\N\ = \\ \A\<^sub>P \* \\ + have "A\<^sub>P \* xvec" and "A\<^sub>P \* M" and "A\<^sub>P \* N" by auto + from \xvec = bn \\ \bn \ \* P\ \bn \ \* \\ \bn \ \* \\<^sub>P\ have "xvec \* P" and "xvec \* \" and "xvec \* \\<^sub>P" + by simp+ + from \A\<^sub>Q \* (A\<^sub>P, \\<^sub>P, P, Q, R, \)\ + have "A\<^sub>Q \* P" and "A\<^sub>Q \* Q" and "A\<^sub>Q \* R" and "A\<^sub>Q \* A\<^sub>P" and "A\<^sub>Q \* \" and "A\<^sub>Q \* \\<^sub>P" by simp+ + from \A\<^sub>R \* (A\<^sub>P, \\<^sub>P, P, Q, R, \)\ have "A\<^sub>R \* P" and "A\<^sub>R \* Q" and "A\<^sub>R \* R" and "A\<^sub>R \* A\<^sub>P"and "A\<^sub>R \* \" by simp+ + + have FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" by fact + with \A\<^sub>P \* Q\ \A\<^sub>Q \* A\<^sub>P\ have "A\<^sub>P \* \\<^sub>Q" + by(auto dest: extractFrameFreshChain) + have FrR: "extractFrame R = \A\<^sub>R, \\<^sub>R\" by fact + with \A\<^sub>P \* R\ \A\<^sub>R \* A\<^sub>P\ have "A\<^sub>P \* \\<^sub>R" + by(auto dest: extractFrameFreshChain) + + from \(\ \ \\<^sub>P) \ \\<^sub>R \ Q \\M\\*xvec\\N\ \ Q'\ have "(\ \ \\<^sub>R) \ \\<^sub>P \ Q \\M\\*xvec\\N\ \ Q'" + by(metis statEqTransition Associativity Commutativity Composition) + then have "\ \ \\<^sub>R \ P \ Q \\M\\*xvec\\N\ \ (P \ Q')" using FrP \A\<^sub>P \* \\ \A\<^sub>P \* \\<^sub>R\ \A\<^sub>P \* Q\ \A\<^sub>P \* M\ \A\<^sub>P \* N\ \xvec \* P\ \A\<^sub>P \* xvec\ + by(intro Par2) auto + moreover from FrP FrQ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>Q \* A\<^sub>P\ \A\<^sub>Q \* \\<^sub>P\ have "extractFrame(P \ Q) = \(A\<^sub>P@A\<^sub>Q), \\<^sub>P \ \\<^sub>Q\" + by simp+ + moreover from \(\ \ \\<^sub>P) \ \\<^sub>Q \ R \\M\N\ \ R'\ have "\ \ \\<^sub>P \ \\<^sub>Q \ R \\M\N\ \ R'" + by(metis statEqTransition Associativity) + moreover note \extractFrame R = \A\<^sub>R, \\<^sub>R\\ + ultimately have "\ \ (P \ Q) \ R \\M\\*xvec\\N\ \ ((P \ Q') \ R')" + using \A\<^sub>P \* \\ \A\<^sub>Q \* \\ \A\<^sub>R \* \\ \A\<^sub>P \* P\ \A\<^sub>Q \* P\ \A\<^sub>R \* P\ \A\<^sub>P \* Q\ \A\<^sub>Q \* Q\ \A\<^sub>R \* Q\ \A\<^sub>P \* R\ \A\<^sub>Q \* R\ \A\<^sub>R \* R\ + \A\<^sub>P \* M\ \A\<^sub>Q \* M\ \A\<^sub>R \* M\ \A\<^sub>R \* A\<^sub>P\ \A\<^sub>Q \* A\<^sub>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 \\<^sub>Q\<^sub>R M N P' A\<^sub>P \\<^sub>P K xvec QR' A\<^sub>Q\<^sub>R) + from \xvec \* (\, P, Q, R, \)\ have "xvec \* Q" and "xvec \* R" by simp+ + from \A\<^sub>Q\<^sub>R \* (\, P, Q, R, \)\ have "A\<^sub>Q\<^sub>R \* Q" and "A\<^sub>Q\<^sub>R \* R" and "A\<^sub>Q\<^sub>R \* \" by simp+ + from \A\<^sub>P \* (Q \ R)\ have "A\<^sub>P \* Q" and "A\<^sub>P \* R" by simp+ + have PTrans: "\ \ \\<^sub>Q\<^sub>R \ P \M\N\ \ P'" and FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" and MeqK: "\ \ \\<^sub>P \ \\<^sub>Q\<^sub>R \ M \ K" by fact+ + note \\ \ \\<^sub>P \ Q \ R \K\\*xvec\\N\ \ QR'\ + moreover from \xvec \* \\ \xvec \* \\<^sub>P\ have "xvec \* (\ \ \\<^sub>P)" by force + moreover note \xvec \* Q\\xvec \* R\ \xvec \* K\ + \extractFrame(Q \ R) = \A\<^sub>Q\<^sub>R, \\<^sub>Q\<^sub>R\\ \distinct A\<^sub>Q\<^sub>R\ + moreover from \A\<^sub>Q\<^sub>R \* \\ \A\<^sub>Q\<^sub>R \* \\<^sub>P\ have "A\<^sub>Q\<^sub>R \* (\ \ \\<^sub>P)" by force + ultimately show ?case using \A\<^sub>Q\<^sub>R \* Q\ \A\<^sub>Q\<^sub>R \* R\ \A\<^sub>Q\<^sub>R \* K\ + proof(induct rule: parCasesOutputFrame) + case(cPar1 Q' A\<^sub>Q \\<^sub>Q A\<^sub>R \\<^sub>R) + have Aeq: "A\<^sub>Q\<^sub>R = A\<^sub>Q@A\<^sub>R" and \eq: "\\<^sub>Q\<^sub>R = \\<^sub>Q \ \\<^sub>R" by fact+ + from PTrans \eq have "(\ \ \\<^sub>R) \ \\<^sub>Q \ P \M\N\ \ P'" + by(metis statEqTransition Associativity Commutativity Composition) + moreover note FrP + moreover from \(\ \ \\<^sub>P) \ \\<^sub>R \ Q \K\\*xvec\\N\ \ Q'\ have "(\ \ \\<^sub>R) \ \\<^sub>P \ Q \K\\*xvec\\N\ \ Q'" + by(metis statEqTransition Associativity Commutativity Composition) + moreover note \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ + moreover from MeqK \eq have "(\ \ \\<^sub>R) \ \\<^sub>P \ \\<^sub>Q \ M \ K" + by(metis statEqEnt Associativity Commutativity Composition) + moreover from \A\<^sub>P \* R\ \A\<^sub>P \* A\<^sub>Q\<^sub>R\ Aeq \extractFrame R = \A\<^sub>R, \\<^sub>R\\ have "A\<^sub>P \* A\<^sub>Q" and "A\<^sub>P \* \\<^sub>R" + by(auto dest: extractFrameFreshChain) + moreover from \A\<^sub>Q\<^sub>R \* P\ \A\<^sub>Q\<^sub>R \* \\ Aeq have "A\<^sub>Q \* P" and "A\<^sub>Q \* \" by simp+ + ultimately have "\ \ \\<^sub>R \ P \ Q \\ \ \\*xvec\(P' \ Q')" + using \A\<^sub>P \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* Q\ \A\<^sub>P \* M\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>Q \* \\ \A\<^sub>Q \* \\<^sub>R\ \A\<^sub>Q \* P\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* K\ \xvec \* P\ + by(intro Comm1) (assumption | force)+ + moreover from \A\<^sub>Q\<^sub>R \* \\ Aeq have "A\<^sub>R \* \" by simp + moreover from \A\<^sub>Q\<^sub>R \* P\ Aeq have "A\<^sub>R \* P" by simp + ultimately have "\ \ (P \ Q) \ R \\ \ (\\*xvec\(P' \ Q')) \ R" using \extractFrame R = \A\<^sub>R, \\<^sub>R\\ \A\<^sub>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\<^sub>Q \\<^sub>Q A\<^sub>R \\<^sub>R) + have Aeq: "A\<^sub>Q\<^sub>R = A\<^sub>Q@A\<^sub>R" and \eq: "\\<^sub>Q\<^sub>R = \\<^sub>Q \ \\<^sub>R" by fact+ + from \A\<^sub>Q\<^sub>R \* \\ \A\<^sub>Q\<^sub>R \* P\ \A\<^sub>Q\<^sub>R \* \\<^sub>P\ \A\<^sub>P \* A\<^sub>Q\<^sub>R\ Aeq have "A\<^sub>R \* \" and "A\<^sub>R \* \\<^sub>P" and "A\<^sub>P \* A\<^sub>R" and "A\<^sub>P \* A\<^sub>Q" and "A\<^sub>R \* P" by simp+ + from \A\<^sub>Q\<^sub>R \* \\ Aeq have "A\<^sub>Q \* \" by simp + from \A\<^sub>Q\<^sub>R \* P\ \A\<^sub>P \* A\<^sub>Q\<^sub>R\ Aeq FrP have "A\<^sub>Q \* \\<^sub>P" by(auto dest: extractFrameFreshChain) + from \A\<^sub>P \* A\<^sub>Q\<^sub>R\ \extractFrame R = \A\<^sub>R, \\<^sub>R\\ \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ Aeq \A\<^sub>P \* Q\ \A\<^sub>P \* R\ have "A\<^sub>P \* \\<^sub>Q" and "A\<^sub>P \* \\<^sub>R" by(auto dest: extractFrameFreshChain) + have RTrans: "(\ \ \\<^sub>P) \ \\<^sub>Q \ R \K\\*xvec\\N\ \ R'" and FrR: "extractFrame R = \A\<^sub>R, \\<^sub>R\" by fact+ + then have "(\ \ \\<^sub>P) \ \\<^sub>Q \ R \ ROut K (\\*xvec\N \' R')" by(simp add: residualInject) + then obtain K' where KeqK': "((\ \ \\<^sub>P) \ \\<^sub>Q) \ \\<^sub>R \ K \ K'" and "A\<^sub>P \* K'" and "A\<^sub>Q \* K'" + using \A\<^sub>P \* R\ \A\<^sub>Q \* R\ \A\<^sub>R \* \\ \A\<^sub>R \* \\<^sub>P\ \A\<^sub>R \* \\<^sub>Q\ \A\<^sub>Q \* A\<^sub>R\ \A\<^sub>P \* A\<^sub>R\ \A\<^sub>R \* R\ \A\<^sub>R \* R\ \A\<^sub>R \* K\ \distinct A\<^sub>R\ \xvec \* K\ \distinct xvec\ FrR + using outputObtainPrefix[where B="A\<^sub>P@A\<^sub>Q"] + by (smt (verit, ccfv_threshold) freshChainAppend freshChainSym freshCompChain(1)) + from PTrans \eq have "(\ \ \\<^sub>R) \ \\<^sub>Q \ P \M\N\ \ P'" + by(metis statEqTransition Associativity Commutativity Composition) + moreover from MeqK KeqK' \eq have MeqK': "((\ \ \\<^sub>R) \ \\<^sub>Q) \ \\<^sub>P \ M \ K'" + by(metis statEqEnt Associativity Commutativity Composition chanEqTrans) + ultimately have "(\ \ \\<^sub>R) \ \\<^sub>Q \ P \K'\N\ \ P'" using FrP \distinct A\<^sub>P\ \A\<^sub>P \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* M\ \A\<^sub>P \* K'\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>P \* \\<^sub>R\ + by(auto intro: inputRenameSubject) + moreover from \A\<^sub>Q\<^sub>R \* P\ \A\<^sub>Q\<^sub>R \* N\ Aeq have "A\<^sub>Q \* P" and "A\<^sub>Q \* N" by simp+ + ultimately have "\ \ \\<^sub>R \ P \ Q \K'\N\ \ P' \ Q" using \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ \A\<^sub>Q \* \\<^sub>R\ \A\<^sub>Q \* K'\ \A\<^sub>Q \* \\ + by(intro Par1) (assumption | force)+ + moreover from FrP \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>Q \* \\<^sub>P\ + have "extractFrame(P \ Q) = \(A\<^sub>P@A\<^sub>Q), \\<^sub>P \ \\<^sub>Q\" by simp+ + moreover from RTrans have "\ \ (\\<^sub>P \ \\<^sub>Q) \ R \K\\*xvec\\N\ \ R'" by(metis Associativity statEqTransition) + moreover note FrR + moreover from MeqK' KeqK' have "\ \ (\\<^sub>P \ \\<^sub>Q) \ \\<^sub>R \ K' \ K" + by(metis statEqEnt Associativity Commutativity Composition chanEqTrans chanEqSym) + ultimately have "\ \ (P \ Q) \ R \\ \ \\*xvec\((P' \ Q) \ R')" + using \A\<^sub>P \* \\ \A\<^sub>Q \* \\ \A\<^sub>P \* P\ \A\<^sub>Q \* P\ \A\<^sub>P \* Q\ \A\<^sub>Q \* Q\ \A\<^sub>P \* R\ \A\<^sub>Q \* R\ \A\<^sub>P \* K'\ \A\<^sub>Q \* K'\ \A\<^sub>P \* A\<^sub>R\ \A\<^sub>Q \* A\<^sub>R\ + \A\<^sub>R \* \\ \A\<^sub>R \* P\ \A\<^sub>R \* Q\ \A\<^sub>R \* R\ \A\<^sub>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 \\<^sub>Q\<^sub>R M xvec N P' A\<^sub>P \\<^sub>P K QR' A\<^sub>Q\<^sub>R) + from \A\<^sub>Q\<^sub>R \* (\, P, Q, R, \)\ have "A\<^sub>Q\<^sub>R \* Q" and "A\<^sub>Q\<^sub>R \* R" and "A\<^sub>Q\<^sub>R \* \" by simp+ + from \A\<^sub>P \* (Q \ R)\ \xvec \* (Q \ R)\ have "A\<^sub>P \* Q" and "A\<^sub>P \* R" and "xvec \* Q" and "xvec \* R" by simp+ + have PTrans: "\ \ \\<^sub>Q\<^sub>R \ P \M\\*xvec\\N\ \ P'" and FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" and MeqK: "\ \ \\<^sub>P \ \\<^sub>Q\<^sub>R \ M \ K" by fact+ + note \\ \ \\<^sub>P \ Q \ R \K\N\ \ QR'\ \extractFrame(Q \ R) = \A\<^sub>Q\<^sub>R, \\<^sub>Q\<^sub>R\\ \distinct A\<^sub>Q\<^sub>R\ + moreover from \A\<^sub>Q\<^sub>R \* \\ \A\<^sub>Q\<^sub>R \* \\<^sub>P\ have "A\<^sub>Q\<^sub>R \* (\ \ \\<^sub>P)" by force + ultimately show ?case using \A\<^sub>Q\<^sub>R \* Q\ \A\<^sub>Q\<^sub>R \* R\ \A\<^sub>Q\<^sub>R \* K\ + proof(induct rule: parCasesInputFrame) + case(cPar1 Q' A\<^sub>Q \\<^sub>Q A\<^sub>R \\<^sub>R) + have Aeq: "A\<^sub>Q\<^sub>R = A\<^sub>Q@A\<^sub>R" and \eq: "\\<^sub>Q\<^sub>R = \\<^sub>Q \ \\<^sub>R" by fact+ + from PTrans \eq have "(\ \ \\<^sub>R) \ \\<^sub>Q \ P \M\\*xvec\\N\ \ P'" + by(metis statEqTransition Associativity Commutativity Composition) + moreover note FrP + moreover from \(\ \ \\<^sub>P) \ \\<^sub>R \ Q \K\N\ \ Q'\ have "(\ \ \\<^sub>R) \ \\<^sub>P \ Q \K\N\ \ Q'" + by(metis statEqTransition Associativity Commutativity Composition) + moreover note \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ + moreover from MeqK \eq have "(\ \ \\<^sub>R) \ \\<^sub>P \ \\<^sub>Q \ M \ K" + by(metis statEqEnt Associativity Commutativity Composition) + moreover from \A\<^sub>P \* Q\ \A\<^sub>P \* R\ \A\<^sub>P \* A\<^sub>Q\<^sub>R\ Aeq \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ \extractFrame R = \A\<^sub>R, \\<^sub>R\\ + have "A\<^sub>P \* A\<^sub>Q" and "A\<^sub>P \* \\<^sub>R" by(auto dest: extractFrameFreshChain) + moreover from \A\<^sub>Q\<^sub>R \* \\ \A\<^sub>Q\<^sub>R \* P\ Aeq have "A\<^sub>Q \* \" and "A\<^sub>Q \* P" by simp+ + ultimately have "\ \ \\<^sub>R \ P \ Q \\ \ \\*xvec\(P' \ Q')" + using \A\<^sub>P \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* Q\ \A\<^sub>P \* M\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>Q \* \\ \A\<^sub>Q \* \\<^sub>R\ \A\<^sub>Q \* P\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* K\ \xvec \* Q\ + by(intro Comm2) (assumption | force)+ + moreover from \A\<^sub>Q\<^sub>R \* \\ \A\<^sub>Q\<^sub>R \* P\ Aeq have "A\<^sub>R \* \" and "A\<^sub>R \* P" by simp+ + ultimately have "\ \ (P \ Q) \ R \\ \ (\\*xvec\(P' \ Q')) \ R" using \extractFrame R = \A\<^sub>R, \\<^sub>R\\ \A\<^sub>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\<^sub>Q \\<^sub>Q A\<^sub>R \\<^sub>R) + have Aeq: "A\<^sub>Q\<^sub>R = A\<^sub>Q@A\<^sub>R" and \eq: "\\<^sub>Q\<^sub>R = \\<^sub>Q \ \\<^sub>R" by fact+ + from \A\<^sub>Q\<^sub>R \* \\ \A\<^sub>Q\<^sub>R \* P\ \A\<^sub>Q\<^sub>R \* \\<^sub>P\ \A\<^sub>P \* A\<^sub>Q\<^sub>R\ Aeq + have "A\<^sub>R \* \" and "A\<^sub>R \* \\<^sub>P" and "A\<^sub>P \* A\<^sub>R" and "A\<^sub>P \* A\<^sub>Q" and "A\<^sub>R \* P" and "A\<^sub>Q \* \" and "A\<^sub>Q \* \\<^sub>P" by simp+ + have RTrans: "(\ \ \\<^sub>P) \ \\<^sub>Q \ R \K\N\ \ R'" and FrR: "extractFrame R = \A\<^sub>R, \\<^sub>R\" by fact+ + then obtain K' where KeqK': "((\ \ \\<^sub>P) \ \\<^sub>Q) \ \\<^sub>R \ K \ K'" and "A\<^sub>P \* K'" and "A\<^sub>Q \* K'" + using \A\<^sub>P \* R\ \A\<^sub>Q \* R\ \A\<^sub>R \* \\ \A\<^sub>R \* \\<^sub>P\ \A\<^sub>R \* \\<^sub>Q\ \A\<^sub>Q \* A\<^sub>R\ \A\<^sub>P \* A\<^sub>R\ \A\<^sub>R \* R\ \A\<^sub>R \* R\ \A\<^sub>R \* K\ \distinct A\<^sub>R\ + using inputObtainPrefix[where B="A\<^sub>P@A\<^sub>Q"] + by (smt (verit, ccfv_threshold) freshChainAppend freshChainSym freshCompChain(1)) + from PTrans \eq have "(\ \ \\<^sub>R) \ \\<^sub>Q \ P \M\\*xvec\\N\ \ P'" + by(metis statEqTransition Associativity Commutativity Composition) + moreover from MeqK KeqK' \eq have MeqK': "((\ \ \\<^sub>R) \ \\<^sub>Q) \ \\<^sub>P \ M \ K'" + by(metis statEqEnt Associativity Commutativity Composition chanEqTrans) + moreover from \A\<^sub>P \* R\ \A\<^sub>P \* Q\ \A\<^sub>P \* A\<^sub>Q\<^sub>R\ FrR \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ Aeq have "A\<^sub>P \* \\<^sub>Q" and "A\<^sub>P \* \\<^sub>R" + by(auto dest: extractFrameFreshChain) + ultimately have "(\ \ \\<^sub>R) \ \\<^sub>Q \ P \K'\\*xvec\\N\ \ P'" using FrP \distinct A\<^sub>P\ \A\<^sub>P \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* M\ \A\<^sub>P \* K'\ + by(auto intro: outputRenameSubject) + moreover from \A\<^sub>Q\<^sub>R \* P\ \A\<^sub>Q\<^sub>R \* N\ \A\<^sub>Q\<^sub>R \* xvec\ Aeq have "A\<^sub>Q \* P" and "A\<^sub>Q \* N" and "A\<^sub>Q \* xvec" by simp+ + ultimately have "\ \ \\<^sub>R \ P \ Q \K'\\*xvec\\N\ \ (P' \ Q)" using \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ \A\<^sub>Q \* \\<^sub>R\ \A\<^sub>Q \* K'\ \xvec \* Q\ \A\<^sub>Q \* \\ + by(intro Par1) (assumption | force)+ + moreover from FrP \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>Q \* \\<^sub>P\ + have "extractFrame(P \ Q) = \(A\<^sub>P@A\<^sub>Q), \\<^sub>P \ \\<^sub>Q\" by simp+ + moreover from RTrans have "\ \ (\\<^sub>P \ \\<^sub>Q) \ R \K\N\ \ R'" by(metis Associativity statEqTransition) + moreover note FrR + moreover from MeqK' KeqK' have "\ \ (\\<^sub>P \ \\<^sub>Q) \ \\<^sub>R \ K' \ K" + by(metis statEqEnt Associativity Commutativity Composition chanEqTrans chanEqSym) + ultimately have "\ \ (P \ Q) \ R \\ \ \\*xvec\((P' \ Q) \ R')" + using \A\<^sub>P \* \\ \A\<^sub>Q \* \\ \A\<^sub>P \* P\ \A\<^sub>Q \* P\ \A\<^sub>P \* Q\ \A\<^sub>Q \* Q\ \A\<^sub>P \* R\ \A\<^sub>Q \* R\ \A\<^sub>P \* K'\ \A\<^sub>Q \* K'\ \A\<^sub>P \* A\<^sub>R\ \A\<^sub>Q \* A\<^sub>R\ + \A\<^sub>R \* \\ \A\<^sub>R \* P\ \A\<^sub>R \* Q\ \A\<^sub>R \* R\ \A\<^sub>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 \\<^sub>Q\<^sub>R M N P' A\<^sub>P \\<^sub>P QR' A\<^sub>Q\<^sub>R) + from \A\<^sub>Q\<^sub>R \* (\, P, Q, R, \)\ have "A\<^sub>Q\<^sub>R \* Q" and "A\<^sub>Q\<^sub>R \* R" and "A\<^sub>Q\<^sub>R \* \" by simp+ + from \A\<^sub>P \* (Q \ R)\ have "A\<^sub>P \* Q" and "A\<^sub>P \* R" by simp+ + have PTrans: "\ \ \\<^sub>Q\<^sub>R \ P \\M\N\ \ P'" and FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" by fact+ + note \\ \ \\<^sub>P \ Q \ R \\M\N\ \ QR'\ \extractFrame(Q \ R) = \A\<^sub>Q\<^sub>R, \\<^sub>Q\<^sub>R\\ \distinct A\<^sub>Q\<^sub>R\ + moreover from \A\<^sub>Q\<^sub>R \* \\ \A\<^sub>Q\<^sub>R \* \\<^sub>P\ have "A\<^sub>Q\<^sub>R \* (\ \ \\<^sub>P)" by force + ultimately show ?case using \A\<^sub>Q\<^sub>R \* Q\ \A\<^sub>Q\<^sub>R \* R\ \A\<^sub>Q\<^sub>R \* M\ \A\<^sub>Q\<^sub>R \* N\ \A\<^sub>Q\<^sub>R \* P\ \A\<^sub>Q\<^sub>R \* M\ \A\<^sub>Q\<^sub>R \* \\ + proof(induct rule: parCasesBrInputFrame) + case(cPar1 Q' A\<^sub>Q \\<^sub>Q A\<^sub>R \\<^sub>R) + from \A\<^sub>Q\<^sub>R \* N\ \A\<^sub>Q\<^sub>R = A\<^sub>Q @ A\<^sub>R\ have "A\<^sub>Q \* N" and "A\<^sub>R \* N" + by simp+ + have Aeq: "A\<^sub>Q\<^sub>R = A\<^sub>Q@A\<^sub>R" and \eq: "\\<^sub>Q\<^sub>R = \\<^sub>Q \ \\<^sub>R" by fact+ + from PTrans \eq have "(\ \ \\<^sub>R) \ \\<^sub>Q \ P \\M\N\ \ P'" + by(metis statEqTransition Associativity Commutativity Composition) + moreover note FrP + moreover from \(\ \ \\<^sub>P) \ \\<^sub>R \ Q \\M\N\ \ Q'\ have "(\ \ \\<^sub>R) \ \\<^sub>P \ Q \\M\N\ \ Q'" + by(metis statEqTransition Associativity Commutativity Composition) + moreover note \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ + moreover from \A\<^sub>P \* Q\ \A\<^sub>P \* R\ \A\<^sub>P \* A\<^sub>Q\<^sub>R\ Aeq \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ \extractFrame R = \A\<^sub>R, \\<^sub>R\\ + have "A\<^sub>P \* A\<^sub>Q" and "A\<^sub>P \* \\<^sub>R" by(auto dest: extractFrameFreshChain) + moreover from \A\<^sub>Q\<^sub>R \* \\ \A\<^sub>Q\<^sub>R \* P\ Aeq have "A\<^sub>Q \* \" and "A\<^sub>Q \* P" by simp+ + ultimately have "\ \ \\<^sub>R \ P \ Q \\M\N\ \ (P' \ Q')" + using \A\<^sub>P \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* Q\ \A\<^sub>P \* M\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>Q \* \\ \A\<^sub>Q \* \\<^sub>R\ \A\<^sub>Q \* P\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* M\ + by(intro BrMerge) (assumption | force)+ + moreover from \A\<^sub>Q\<^sub>R \* \\ \A\<^sub>Q\<^sub>R \* P\ Aeq have "A\<^sub>R \* \" and "A\<^sub>R \* P" by simp+ + ultimately have "\ \ (P \ Q) \ R \\M\N\ \ (P' \ Q') \ R" using \extractFrame R = \A\<^sub>R, \\<^sub>R\\ \A\<^sub>R \* Q\ \A\<^sub>R \* M\ \A\<^sub>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\<^sub>Q \\<^sub>Q A\<^sub>R \\<^sub>R) + from \A\<^sub>Q\<^sub>R \* N\ \A\<^sub>Q\<^sub>R = A\<^sub>Q @ A\<^sub>R\ have "A\<^sub>Q \* N" and "A\<^sub>R \* N" + by simp+ + have Aeq: "A\<^sub>Q\<^sub>R = A\<^sub>Q@A\<^sub>R" and \eq: "\\<^sub>Q\<^sub>R = \\<^sub>Q \ \\<^sub>R" by fact+ + from \A\<^sub>Q\<^sub>R \* \\ \A\<^sub>Q\<^sub>R \* P\ \A\<^sub>Q\<^sub>R \* \\<^sub>P\ \A\<^sub>P \* A\<^sub>Q\<^sub>R\ Aeq + have "A\<^sub>R \* \" and "A\<^sub>R \* \\<^sub>P" and "A\<^sub>P \* A\<^sub>R" and "A\<^sub>P \* A\<^sub>Q" and "A\<^sub>R \* P" and "A\<^sub>Q \* \" and "A\<^sub>Q \* \\<^sub>P" by simp+ + have RTrans: "(\ \ \\<^sub>P) \ \\<^sub>Q \ R \\M\N\ \ R'" and FrR: "extractFrame R = \A\<^sub>R, \\<^sub>R\" by fact+ + from PTrans \eq have "(\ \ \\<^sub>R) \ \\<^sub>Q \ P \\M\N\ \ P'" + by(metis statEqTransition Associativity Commutativity Composition) + moreover from \A\<^sub>P \* R\ \A\<^sub>P \* Q\ \A\<^sub>P \* A\<^sub>Q\<^sub>R\ FrR \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ Aeq have "A\<^sub>P \* \\<^sub>Q" and "A\<^sub>P \* \\<^sub>R" + by(auto dest: extractFrameFreshChain) + moreover from \A\<^sub>Q\<^sub>R \* P\ \A\<^sub>Q\<^sub>R \* N\ Aeq have "A\<^sub>Q \* P" and "A\<^sub>Q \* N" by simp+ + ultimately have "\ \ \\<^sub>R \ P \ Q \\M\N\ \ (P' \ Q)" using \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ \A\<^sub>Q \* \\<^sub>R\ \A\<^sub>Q \* M\ \A\<^sub>Q \* \\ + by(intro Par1) (assumption | force)+ + moreover from FrP \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>Q \* \\<^sub>P\ + have "extractFrame(P \ Q) = \(A\<^sub>P@A\<^sub>Q), \\<^sub>P \ \\<^sub>Q\" by simp+ + moreover from RTrans have "\ \ (\\<^sub>P \ \\<^sub>Q) \ R \\M\N\ \ R'" by(metis Associativity statEqTransition) + moreover note FrR + ultimately have "\ \ (P \ Q) \ R \\M\N\ \ ((P' \ Q) \ R')" + using \A\<^sub>P \* \\ \A\<^sub>Q \* \\ \A\<^sub>P \* P\ \A\<^sub>Q \* P\ \A\<^sub>P \* Q\ \A\<^sub>Q \* Q\ \A\<^sub>P \* R\ \A\<^sub>Q \* R\ \A\<^sub>P \* M\ \A\<^sub>Q \* M\ \A\<^sub>P \* A\<^sub>R\ \A\<^sub>Q \* A\<^sub>R\ + \A\<^sub>R \* \\ \A\<^sub>R \* P\ \A\<^sub>R \* Q\ \A\<^sub>R \* R\ \A\<^sub>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 \\<^sub>R Q' A\<^sub>Q \\<^sub>Q R' A\<^sub>R) + have Aeq: "A\<^sub>Q\<^sub>R = A\<^sub>Q@A\<^sub>R" and \eq: "\\<^sub>Q\<^sub>R = \\<^sub>Q \ \\<^sub>R" by fact+ + from \A\<^sub>Q\<^sub>R \* N\ \A\<^sub>Q\<^sub>R \* M\ \A\<^sub>Q\<^sub>R \* P\ \A\<^sub>Q\<^sub>R \* \\ \A\<^sub>Q\<^sub>R \* \\<^sub>P\ \A\<^sub>P \* A\<^sub>Q\<^sub>R\ Aeq \eq + have "A\<^sub>Q \* N" and "A\<^sub>R \* N" and "A\<^sub>Q \* M" and "A\<^sub>R \* M" + and "A\<^sub>Q \* \" and "A\<^sub>R \* \" and "A\<^sub>P \* A\<^sub>Q" and "A\<^sub>P \* A\<^sub>R" + and "A\<^sub>Q \* \\<^sub>P" and "A\<^sub>R \* \\<^sub>P" + and "A\<^sub>Q \* P" and "A\<^sub>R \* P" + by simp+ + + from \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ \A\<^sub>P \* Q\ \A\<^sub>P \* A\<^sub>Q\ + have "A\<^sub>P \* \\<^sub>Q" + by (metis extractFrameFreshChain freshFrameDest) + from Aeq \eq PTrans have "\ \ (\\<^sub>Q \ \\<^sub>R) \ P \ \M\N\ \ P'" by simp + then have "\ \ (\\<^sub>R \ \\<^sub>Q) \ P \ \M\N\ \ P'" + by (metis Commutativity \eq compositionSym statEqTransition) + then have "(\ \ \\<^sub>R) \ \\<^sub>Q \ P \ \M\N\ \ P'" + by (metis AssertionStatEqSym Associativity statEqTransition) + moreover note FrP + moreover from \(\ \ \\<^sub>P) \ \\<^sub>R \ Q \\M\N\ \ Q'\ have "(\ \ \\<^sub>R) \ \\<^sub>P \ Q \\M\N\ \ Q'" + by (metis associativitySym statEqTransition) + moreover note \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ + moreover note \A\<^sub>P \* \\ \A\<^sub>Q \* \\ \A\<^sub>Q \* \\<^sub>R\ + moreover from \extractFrame R = \A\<^sub>R, \\<^sub>R\\ \A\<^sub>P \* A\<^sub>R\ \A\<^sub>P \* R\ have "A\<^sub>P \* \\<^sub>R" + by (metis extractFrameFreshChain freshFrameDest) + moreover note \A\<^sub>P \* P\ \A\<^sub>P \* Q\ \A\<^sub>P \* M\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>Q \* P\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* M\ + ultimately have "(\ \ \\<^sub>R) \ (P \ Q) \\M\N\ \ (P' \ Q')" + by(intro BrMerge) (assumption | force)+ + + from \(\ \ \\<^sub>P) \ \\<^sub>Q \ R \\M\N\ \ R'\ have "\ \ (\\<^sub>P \ \\<^sub>Q) \ R \\M\N\ \ R'" + by (metis Associativity statEqTransition) + + note \(\ \ \\<^sub>R) \ (P \ Q) \\M\N\ \ (P' \ Q')\ + moreover from FrP \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>Q \* \\<^sub>P\ + have "extractFrame(P \ Q) = \(A\<^sub>P@A\<^sub>Q), (\\<^sub>P \ \\<^sub>Q)\" by simp + moreover note \\ \ (\\<^sub>P \ \\<^sub>Q) \ R \\M\N\ \ R'\ \extractFrame R = \A\<^sub>R, \\<^sub>R\\ + moreover note \A\<^sub>P \* \\ \A\<^sub>Q \* \\ \A\<^sub>R \* \\ \A\<^sub>P \* P\ \A\<^sub>Q \* P\ \A\<^sub>P \* Q\ \A\<^sub>Q \* Q\ + \A\<^sub>P \* R\ \A\<^sub>Q \* R\ \A\<^sub>P \* M\ \A\<^sub>Q \* M\ \A\<^sub>P \* A\<^sub>R\ \A\<^sub>Q \* A\<^sub>R\ \A\<^sub>R \* P\ \A\<^sub>R \* Q\ \A\<^sub>R \* R\ \A\<^sub>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 \\<^sub>Q\<^sub>R M N P' A\<^sub>P \\<^sub>P xvec QR' A\<^sub>Q\<^sub>R) + from \\M\\*xvec\\N\ = \\ \bn \ \* Q\ \bn \ \* R\ have "xvec \* Q" and "xvec \* R" by auto + from \A\<^sub>Q\<^sub>R \* (\, P, Q, R, \)\ have "A\<^sub>Q\<^sub>R \* Q" and "A\<^sub>Q\<^sub>R \* R" and "A\<^sub>Q\<^sub>R \* \" by simp+ + from \A\<^sub>P \* (Q \ R)\ have "A\<^sub>P \* Q" and "A\<^sub>P \* R" by simp+ + have PTrans: "\ \ \\<^sub>Q\<^sub>R \ P \\M\N\ \ P'" and FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" by fact+ + note \\ \ \\<^sub>P \ Q \ R \\M\\*xvec\\N\ \ QR'\ + moreover from \xvec \* \\ \xvec \* \\<^sub>P\ have "xvec \* (\ \ \\<^sub>P)" by force + moreover note \xvec \* Q\ \xvec \* R\ \xvec \* M\ + \extractFrame(Q \ R) = \A\<^sub>Q\<^sub>R, \\<^sub>Q\<^sub>R\\ \distinct A\<^sub>Q\<^sub>R\ + moreover from \A\<^sub>Q\<^sub>R \* \\ \A\<^sub>Q\<^sub>R \* \\<^sub>P\ have "A\<^sub>Q\<^sub>R \* (\ \ \\<^sub>P)" by force + ultimately show ?case using \A\<^sub>Q\<^sub>R \* Q\ \A\<^sub>Q\<^sub>R \* R\ \A\<^sub>Q\<^sub>R \* M\ + proof(induct rule: parCasesBrOutputFrame) + case(cPar1 Q' A\<^sub>Q \\<^sub>Q A\<^sub>R \\<^sub>R) + from \A\<^sub>Q\<^sub>R \* N\ \A\<^sub>Q\<^sub>R = A\<^sub>Q @ A\<^sub>R\ have "A\<^sub>Q \* N" and "A\<^sub>R \* N" + by simp+ + from \A\<^sub>Q\<^sub>R \* xvec\ \A\<^sub>Q\<^sub>R = (A\<^sub>Q@A\<^sub>R)\ have "A\<^sub>Q \* xvec" and "A\<^sub>R \* xvec" + by simp+ + have Aeq: "A\<^sub>Q\<^sub>R = A\<^sub>Q@A\<^sub>R" and \eq: "\\<^sub>Q\<^sub>R = \\<^sub>Q \ \\<^sub>R" by fact+ + from PTrans \eq have "(\ \ \\<^sub>R) \ \\<^sub>Q \ P \\M\N\ \ P'" + by(metis statEqTransition Associativity Commutativity Composition) + moreover note FrP + moreover from \(\ \ \\<^sub>P) \ \\<^sub>R \ Q \\M\\*xvec\\N\ \ Q'\ have "(\ \ \\<^sub>R) \ \\<^sub>P \ Q \\M\\*xvec\\N\ \ Q'" + by(metis statEqTransition Associativity Commutativity Composition) + moreover note \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ + moreover from \A\<^sub>P \* Q\ \A\<^sub>P \* R\ \A\<^sub>P \* A\<^sub>Q\<^sub>R\ Aeq \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ \extractFrame R = \A\<^sub>R, \\<^sub>R\\ + have "A\<^sub>P \* A\<^sub>Q" and "A\<^sub>P \* \\<^sub>R" by(auto dest: extractFrameFreshChain) + moreover from \A\<^sub>Q\<^sub>R \* \\ \A\<^sub>Q\<^sub>R \* P\ Aeq have "A\<^sub>Q \* \" and "A\<^sub>Q \* P" by simp+ + ultimately have "\ \ \\<^sub>R \ P \ Q \\M\\*xvec\\N\ \ (P' \ Q')" + using \A\<^sub>P \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* Q\ \A\<^sub>P \* M\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>Q \* \\ \A\<^sub>Q \* \\<^sub>R\ \A\<^sub>Q \* P\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* M\ \xvec \* P\ + by(intro BrComm1) (assumption | force)+ + moreover from \A\<^sub>Q\<^sub>R \* \\ \A\<^sub>Q\<^sub>R \* P\ Aeq have "A\<^sub>R \* \" and "A\<^sub>R \* P" by simp+ + ultimately have "\ \ (P \ Q) \ R \\M\\*xvec\\N\ \ (P' \ Q') \ R" using \extractFrame R = \A\<^sub>R, \\<^sub>R\\ \A\<^sub>R \* Q\ \A\<^sub>R \* M\ \A\<^sub>R \* N\ \xvec \* R\ \A\<^sub>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\<^sub>Q \\<^sub>Q A\<^sub>R \\<^sub>R) + from \A\<^sub>Q\<^sub>R \* N\ \A\<^sub>Q\<^sub>R = A\<^sub>Q @ A\<^sub>R\ have "A\<^sub>Q \* N" and "A\<^sub>R \* N" + by simp+ + have Aeq: "A\<^sub>Q\<^sub>R = A\<^sub>Q@A\<^sub>R" and \eq: "\\<^sub>Q\<^sub>R = \\<^sub>Q \ \\<^sub>R" by fact+ + from \A\<^sub>Q\<^sub>R \* \\ \A\<^sub>Q\<^sub>R \* P\ \A\<^sub>Q\<^sub>R \* \\<^sub>P\ \A\<^sub>P \* A\<^sub>Q\<^sub>R\ Aeq + have "A\<^sub>R \* \" and "A\<^sub>R \* \\<^sub>P" and "A\<^sub>P \* A\<^sub>R" and "A\<^sub>P \* A\<^sub>Q" and "A\<^sub>R \* P" and "A\<^sub>Q \* \" and "A\<^sub>Q \* \\<^sub>P" by simp+ + have RTrans: "(\ \ \\<^sub>P) \ \\<^sub>Q \ R \\M\\*xvec\\N\ \ R'" and FrR: "extractFrame R = \A\<^sub>R, \\<^sub>R\" by fact+ + from PTrans \eq have "(\ \ \\<^sub>R) \ \\<^sub>Q \ P \\M\N\ \ P'" + by(metis statEqTransition Associativity Commutativity Composition) + moreover from \A\<^sub>P \* R\ \A\<^sub>P \* Q\ \A\<^sub>P \* A\<^sub>Q\<^sub>R\ FrR \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ Aeq have "A\<^sub>P \* \\<^sub>Q" and "A\<^sub>P \* \\<^sub>R" + by(auto dest: extractFrameFreshChain) + moreover from \A\<^sub>Q\<^sub>R \* P\ \A\<^sub>Q\<^sub>R \* N\ Aeq have "A\<^sub>Q \* P" and "A\<^sub>Q \* N" by simp+ + ultimately have "\ \ \\<^sub>R \ P \ Q \\M\N\ \ (P' \ Q)" using \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ \A\<^sub>Q \* \\<^sub>R\ \A\<^sub>Q \* M\ \A\<^sub>Q \* \\ + by(intro Par1) (assumption | force)+ + moreover from FrP \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>Q \* \\<^sub>P\ + have "extractFrame(P \ Q) = \(A\<^sub>P@A\<^sub>Q), \\<^sub>P \ \\<^sub>Q\" by simp+ + moreover from RTrans have "\ \ (\\<^sub>P \ \\<^sub>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\<^sub>P \* \\ \A\<^sub>Q \* \\ \A\<^sub>P \* P\ \A\<^sub>Q \* P\ \A\<^sub>P \* Q\ \A\<^sub>Q \* Q\ \A\<^sub>P \* R\ \A\<^sub>Q \* R\ \A\<^sub>P \* M\ \A\<^sub>Q \* M\ \A\<^sub>P \* A\<^sub>R\ \A\<^sub>Q \* A\<^sub>R\ + \A\<^sub>R \* \\ \A\<^sub>R \* P\ \A\<^sub>R \* Q\ \A\<^sub>R \* R\ \A\<^sub>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 \\<^sub>R Q' A\<^sub>Q \\<^sub>Q R' A\<^sub>R) + from \A\<^sub>Q\<^sub>R \* N\ \A\<^sub>Q\<^sub>R = A\<^sub>Q @ A\<^sub>R\ have "A\<^sub>Q \* N" and "A\<^sub>R \* N" + by simp+ + from \A\<^sub>Q\<^sub>R \* xvec\ \A\<^sub>Q\<^sub>R = (A\<^sub>Q@A\<^sub>R)\ have "A\<^sub>Q \* xvec" and "A\<^sub>R \* xvec" + by simp+ + have Aeq: "A\<^sub>Q\<^sub>R = A\<^sub>Q@A\<^sub>R" and \eq: "\\<^sub>Q\<^sub>R = \\<^sub>Q \ \\<^sub>R" by fact+ + from Aeq \A\<^sub>Q\<^sub>R \* P\ have "A\<^sub>Q \* P" and "A\<^sub>R \* P" by simp+ + from Aeq \A\<^sub>Q\<^sub>R \* \\ have "A\<^sub>Q \* \" and "A\<^sub>R \* \" by simp+ + from Aeq \A\<^sub>P \* A\<^sub>Q\<^sub>R\ have "A\<^sub>P \* A\<^sub>Q" and "A\<^sub>P \* A\<^sub>R" by simp+ + with \A\<^sub>P \* Q\ \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ have "A\<^sub>P \* \\<^sub>Q" + by (metis extractFrameFreshChain freshChainSym freshFrameDest) + from PTrans \eq have "(\ \ \\<^sub>R) \ \\<^sub>Q \ P \\M\N\ \ P'" + by(metis statEqTransition Associativity Commutativity Composition) + moreover note FrP + moreover from \(\ \ \\<^sub>P) \ \\<^sub>R \ Q \\M\N\ \ Q'\ have "(\ \ \\<^sub>R) \ \\<^sub>P \ Q \\M\N\ \ Q'" + by(metis statEqTransition Associativity Commutativity Composition) + moreover note \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ + moreover from \A\<^sub>P \* Q\ \A\<^sub>P \* R\ \A\<^sub>P \* A\<^sub>Q\<^sub>R\ Aeq \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ \extractFrame R = \A\<^sub>R, \\<^sub>R\\ + have "A\<^sub>P \* A\<^sub>Q" and "A\<^sub>P \* \\<^sub>R" by(auto dest: extractFrameFreshChain) + moreover from \A\<^sub>Q\<^sub>R \* \\ \A\<^sub>Q\<^sub>R \* P\ Aeq have "A\<^sub>Q \* \" and "A\<^sub>Q \* P" by simp+ + ultimately have PQTrans: "\ \ \\<^sub>R \ P \ Q \\M\N\ \ (P' \ Q')" + using \A\<^sub>P \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* Q\ \A\<^sub>P \* M\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>Q \* \\ \A\<^sub>Q \* \\<^sub>R\ \A\<^sub>Q \* P\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* M\ + by(intro BrMerge) (assumption | force)+ + have RTrans: "(\ \ \\<^sub>P) \ \\<^sub>Q \ R \\M\\*xvec\\N\ \ R'" and FrR: "extractFrame R = \A\<^sub>R, \\<^sub>R\" by fact+ + note PQTrans + moreover from FrP \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>Q\<^sub>R \* \\<^sub>P\ Aeq + have "extractFrame(P \ Q) = \(A\<^sub>P@A\<^sub>Q), \\<^sub>P \ \\<^sub>Q\" by simp+ + moreover from RTrans have "\ \ (\\<^sub>P \ \\<^sub>Q) \ R \\M\\*xvec\\N\ \ R'" + by (metis Associativity statEqTransition) + moreover note FrR + moreover note \A\<^sub>P \* \\ \A\<^sub>Q \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* Q\ \A\<^sub>Q \* P\ \A\<^sub>Q \* Q\ \A\<^sub>P \* R\ \A\<^sub>Q \* R\ + \A\<^sub>P \* M\ \A\<^sub>Q \* M\ \A\<^sub>P \* A\<^sub>R\ \A\<^sub>Q \* A\<^sub>R\ \A\<^sub>R \* \\ \A\<^sub>R \* P\ \A\<^sub>R \* Q\ \A\<^sub>R \* R\ \A\<^sub>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 \\<^sub>R Q' A\<^sub>Q \\<^sub>Q R' A\<^sub>R) + from \A\<^sub>Q\<^sub>R \* N\ \A\<^sub>Q\<^sub>R = A\<^sub>Q @ A\<^sub>R\ have "A\<^sub>Q \* N" and "A\<^sub>R \* N" + by simp+ + from \A\<^sub>Q\<^sub>R \* xvec\ \A\<^sub>Q\<^sub>R = (A\<^sub>Q@A\<^sub>R)\ have "A\<^sub>Q \* xvec" and "A\<^sub>R \* xvec" + by simp+ + have Aeq: "A\<^sub>Q\<^sub>R = A\<^sub>Q@A\<^sub>R" and \eq: "\\<^sub>Q\<^sub>R = \\<^sub>Q \ \\<^sub>R" by fact+ + from Aeq \A\<^sub>Q\<^sub>R \* P\ have "A\<^sub>Q \* P" and "A\<^sub>R \* P" by simp+ + from Aeq \A\<^sub>Q\<^sub>R \* \\ have "A\<^sub>Q \* \" and "A\<^sub>R \* \" by simp+ + from Aeq \A\<^sub>P \* A\<^sub>Q\<^sub>R\ have "A\<^sub>P \* A\<^sub>Q" and "A\<^sub>P \* A\<^sub>R" by simp+ + with \A\<^sub>P \* Q\ \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ have "A\<^sub>P \* \\<^sub>Q" + by (metis extractFrameFreshChain freshChainSym freshFrameDest) + from PTrans \eq have "(\ \ \\<^sub>R) \ \\<^sub>Q \ P \\M\N\ \ P'" + by(metis statEqTransition Associativity Commutativity Composition) + moreover note FrP + moreover from \(\ \ \\<^sub>P) \ \\<^sub>R \ Q \\M\\*xvec\\N\ \ Q'\ have "(\ \ \\<^sub>R) \ \\<^sub>P \ Q \\M\\*xvec\\N\ \ Q'" + by(metis statEqTransition Associativity Commutativity Composition) + moreover note \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ + moreover from \A\<^sub>P \* Q\ \A\<^sub>P \* R\ \A\<^sub>P \* A\<^sub>Q\<^sub>R\ Aeq \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ \extractFrame R = \A\<^sub>R, \\<^sub>R\\ + have "A\<^sub>P \* A\<^sub>Q" and "A\<^sub>P \* \\<^sub>R" by(auto dest: extractFrameFreshChain) + moreover from \A\<^sub>Q\<^sub>R \* \\ \A\<^sub>Q\<^sub>R \* P\ Aeq have "A\<^sub>Q \* \" and "A\<^sub>Q \* P" by simp+ + ultimately have PQTrans: "\ \ \\<^sub>R \ P \ Q \\M\\*xvec\\N\ \ (P' \ Q')" + using \A\<^sub>P \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* Q\ \A\<^sub>P \* M\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>Q \* \\ \A\<^sub>Q \* \\<^sub>R\ \A\<^sub>Q \* P\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* M\ \xvec \* P\ + by(intro BrComm1) (assumption | force)+ + have RTrans: "(\ \ \\<^sub>P) \ \\<^sub>Q \ R \\M\N\ \ R'" and FrR: "extractFrame R = \A\<^sub>R, \\<^sub>R\" by fact+ + note PQTrans + moreover from FrP \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>Q\<^sub>R \* \\<^sub>P\ Aeq + have "extractFrame(P \ Q) = \(A\<^sub>P@A\<^sub>Q), \\<^sub>P \ \\<^sub>Q\" by simp+ + moreover from RTrans have "\ \ (\\<^sub>P \ \\<^sub>Q) \ R \\M\N\ \ R'" + by (metis Associativity statEqTransition) + moreover note FrR + moreover note \A\<^sub>P \* \\ \A\<^sub>Q \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* Q\ \A\<^sub>Q \* P\ \A\<^sub>Q \* Q\ \A\<^sub>P \* R\ \A\<^sub>Q \* R\ + \A\<^sub>P \* M\ \A\<^sub>Q \* M\ \A\<^sub>P \* A\<^sub>R\ \A\<^sub>Q \* A\<^sub>R\ \A\<^sub>R \* \\ \A\<^sub>R \* P\ \A\<^sub>R \* Q\ \A\<^sub>R \* R\ \A\<^sub>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 \\<^sub>Q\<^sub>R M xvec N P' A\<^sub>P \\<^sub>P QR' A\<^sub>Q\<^sub>R) + from \\M\\*xvec\\N\ = \\ \bn \ \* Q\ \bn \ \* R\ have "xvec \* Q" and "xvec \* R" by auto + from \A\<^sub>Q\<^sub>R \* (\, P, Q, R, \)\ have "A\<^sub>Q\<^sub>R \* Q" and "A\<^sub>Q\<^sub>R \* R" and "A\<^sub>Q\<^sub>R \* \" by simp+ + from \A\<^sub>P \* (Q \ R)\ have "A\<^sub>P \* Q" and "A\<^sub>P \* R" by simp+ + have PTrans: "\ \ \\<^sub>Q\<^sub>R \ P \\M\\*xvec\\N\ \ P'" and FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" by fact+ + note \\ \ \\<^sub>P \ Q \ R \\M\N\ \ QR'\ \extractFrame(Q \ R) = \A\<^sub>Q\<^sub>R, \\<^sub>Q\<^sub>R\\ \distinct A\<^sub>Q\<^sub>R\ + moreover from \A\<^sub>Q\<^sub>R \* \\ \A\<^sub>Q\<^sub>R \* \\<^sub>P\ have "A\<^sub>Q\<^sub>R \* (\ \ \\<^sub>P)" by force + ultimately show ?case using \A\<^sub>Q\<^sub>R \* Q\ \A\<^sub>Q\<^sub>R \* R\ \A\<^sub>Q\<^sub>R \* M\ + proof(induct rule: parCasesBrInputFrame) + case(cPar1 Q' A\<^sub>Q \\<^sub>Q A\<^sub>R \\<^sub>R) + from \A\<^sub>Q\<^sub>R \* N\ \A\<^sub>Q\<^sub>R = A\<^sub>Q @ A\<^sub>R\ have "A\<^sub>Q \* N" and "A\<^sub>R \* N" + by simp+ + from \A\<^sub>Q\<^sub>R \* xvec\ \A\<^sub>Q\<^sub>R = (A\<^sub>Q@A\<^sub>R)\ have "A\<^sub>Q \* xvec" and "A\<^sub>R \* xvec" + by simp+ + have Aeq: "A\<^sub>Q\<^sub>R = A\<^sub>Q@A\<^sub>R" and \eq: "\\<^sub>Q\<^sub>R = \\<^sub>Q \ \\<^sub>R" by fact+ + from PTrans \eq have "(\ \ \\<^sub>R) \ \\<^sub>Q \ P \\M\\*xvec\\N\ \ P'" + by(metis statEqTransition Associativity Commutativity Composition) + moreover note FrP + moreover from \(\ \ \\<^sub>P) \ \\<^sub>R \ Q \\M\N\ \ Q'\ have "(\ \ \\<^sub>R) \ \\<^sub>P \ Q \\M\N\ \ Q'" + by(metis statEqTransition Associativity Commutativity Composition) + moreover note \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ + moreover from \A\<^sub>P \* Q\ \A\<^sub>P \* R\ \A\<^sub>P \* A\<^sub>Q\<^sub>R\ Aeq \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ \extractFrame R = \A\<^sub>R, \\<^sub>R\\ + have "A\<^sub>P \* A\<^sub>Q" and "A\<^sub>P \* \\<^sub>R" by(auto dest: extractFrameFreshChain) + moreover from \A\<^sub>Q\<^sub>R \* \\ \A\<^sub>Q\<^sub>R \* P\ Aeq have "A\<^sub>Q \* \" and "A\<^sub>Q \* P" by simp+ + ultimately have "\ \ \\<^sub>R \ P \ Q \\M\\*xvec\\N\ \ (P' \ Q')" + using \A\<^sub>P \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* Q\ \A\<^sub>P \* M\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>Q \* \\ \A\<^sub>Q \* \\<^sub>R\ \A\<^sub>Q \* P\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* M\ \xvec \* Q\ + by(intro BrComm2) (assumption | force)+ + moreover from \A\<^sub>Q\<^sub>R \* \\ \A\<^sub>Q\<^sub>R \* P\ Aeq have "A\<^sub>R \* \" and "A\<^sub>R \* P" by simp+ + ultimately have "\ \ (P \ Q) \ R \\M\\*xvec\\N\ \ (P' \ Q') \ R" using \extractFrame R = \A\<^sub>R, \\<^sub>R\\ \A\<^sub>R \* Q\ \A\<^sub>R \* M\ \A\<^sub>R \* N\ \xvec \* R\ \A\<^sub>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\<^sub>Q \\<^sub>Q A\<^sub>R \\<^sub>R) + from \A\<^sub>Q\<^sub>R \* N\ \A\<^sub>Q\<^sub>R \* xvec\ \A\<^sub>Q\<^sub>R = A\<^sub>Q @ A\<^sub>R\ have "A\<^sub>Q \* N" and "A\<^sub>R \* N" and "A\<^sub>Q \* xvec" and "A\<^sub>R \* xvec" + by simp+ + have Aeq: "A\<^sub>Q\<^sub>R = A\<^sub>Q@A\<^sub>R" and \eq: "\\<^sub>Q\<^sub>R = \\<^sub>Q \ \\<^sub>R" by fact+ + from \A\<^sub>Q\<^sub>R \* \\ \A\<^sub>Q\<^sub>R \* P\ \A\<^sub>Q\<^sub>R \* \\<^sub>P\ \A\<^sub>P \* A\<^sub>Q\<^sub>R\ Aeq + have "A\<^sub>R \* \" and "A\<^sub>R \* \\<^sub>P" and "A\<^sub>P \* A\<^sub>R" and "A\<^sub>P \* A\<^sub>Q" and "A\<^sub>R \* P" and "A\<^sub>Q \* \" and "A\<^sub>Q \* \\<^sub>P" by simp+ + have RTrans: "(\ \ \\<^sub>P) \ \\<^sub>Q \ R \\M\N\ \ R'" and FrR: "extractFrame R = \A\<^sub>R, \\<^sub>R\" by fact+ + from PTrans \eq have "(\ \ \\<^sub>R) \ \\<^sub>Q \ P \\M\\*xvec\\N\ \ P'" + by(metis statEqTransition Associativity Commutativity Composition) + moreover from \A\<^sub>P \* R\ \A\<^sub>P \* Q\ \A\<^sub>P \* A\<^sub>Q\<^sub>R\ FrR \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ Aeq have "A\<^sub>P \* \\<^sub>Q" and "A\<^sub>P \* \\<^sub>R" + by(auto dest: extractFrameFreshChain) + moreover from \A\<^sub>Q\<^sub>R \* P\ \A\<^sub>Q\<^sub>R \* N\ Aeq have "A\<^sub>Q \* P" and "A\<^sub>Q \* N" by simp+ + ultimately have "\ \ \\<^sub>R \ P \ Q \\M\\*xvec\\N\ \ (P' \ Q)" using \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ \A\<^sub>Q \* \\<^sub>R\ + \A\<^sub>Q \* M\ \A\<^sub>Q \* \\ \xvec \* Q\ \A\<^sub>Q \* xvec\ + by(intro Par1) (assumption | force)+ + moreover from FrP \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>Q \* \\<^sub>P\ + have "extractFrame(P \ Q) = \(A\<^sub>P@A\<^sub>Q), \\<^sub>P \ \\<^sub>Q\" by simp+ + moreover from RTrans have "\ \ (\\<^sub>P \ \\<^sub>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\<^sub>P \* \\ \A\<^sub>Q \* \\ \A\<^sub>P \* P\ \A\<^sub>Q \* P\ \A\<^sub>P \* Q\ \A\<^sub>Q \* Q\ \A\<^sub>P \* R\ \A\<^sub>Q \* R\ \A\<^sub>P \* M\ \A\<^sub>Q \* M\ \A\<^sub>P \* A\<^sub>R\ \A\<^sub>Q \* A\<^sub>R\ + \A\<^sub>R \* \\ \A\<^sub>R \* P\ \A\<^sub>R \* Q\ \A\<^sub>R \* R\ \A\<^sub>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 \\<^sub>R Q' A\<^sub>Q \\<^sub>Q R' A\<^sub>R) + from \A\<^sub>Q\<^sub>R \* N\ \A\<^sub>Q\<^sub>R = A\<^sub>Q @ A\<^sub>R\ have "A\<^sub>Q \* N" and "A\<^sub>R \* N" + by simp+ + from \A\<^sub>Q\<^sub>R \* xvec\ \A\<^sub>Q\<^sub>R = (A\<^sub>Q@A\<^sub>R)\ have "A\<^sub>Q \* xvec" and "A\<^sub>R \* xvec" + by simp+ + have Aeq: "A\<^sub>Q\<^sub>R = A\<^sub>Q@A\<^sub>R" and \eq: "\\<^sub>Q\<^sub>R = \\<^sub>Q \ \\<^sub>R" by fact+ + from Aeq \A\<^sub>Q\<^sub>R \* P\ have "A\<^sub>Q \* P" and "A\<^sub>R \* P" by simp+ + from Aeq \A\<^sub>Q\<^sub>R \* M\ have "A\<^sub>Q \* M" and "A\<^sub>R \* M" by simp+ + from Aeq \A\<^sub>Q\<^sub>R \* \\ have "A\<^sub>Q \* \" and "A\<^sub>R \* \" by simp+ + from Aeq \A\<^sub>P \* A\<^sub>Q\<^sub>R\ have "A\<^sub>P \* A\<^sub>Q" and "A\<^sub>P \* A\<^sub>R" by simp+ + with \A\<^sub>P \* Q\ \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ have "A\<^sub>P \* \\<^sub>Q" + by (metis extractFrameFreshChain freshChainSym freshFrameDest) + from PTrans \eq have "(\ \ \\<^sub>R) \ \\<^sub>Q \ P \\M\\*xvec\\N\ \ P'" + by(metis statEqTransition Associativity Commutativity Composition) + moreover note FrP + moreover from \(\ \ \\<^sub>P) \ \\<^sub>R \ Q \\M\N\ \ Q'\ have "(\ \ \\<^sub>R) \ \\<^sub>P \ Q \\M\N\ \ Q'" + by(metis statEqTransition Associativity Commutativity Composition) + moreover note \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ + moreover from \A\<^sub>P \* Q\ \A\<^sub>P \* R\ \A\<^sub>P \* A\<^sub>Q\<^sub>R\ Aeq \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ \extractFrame R = \A\<^sub>R, \\<^sub>R\\ + have "A\<^sub>P \* A\<^sub>Q" and "A\<^sub>P \* \\<^sub>R" by(auto dest: extractFrameFreshChain) + moreover from \A\<^sub>Q\<^sub>R \* \\ \A\<^sub>Q\<^sub>R \* P\ Aeq have "A\<^sub>Q \* \" and "A\<^sub>Q \* P" by simp+ + ultimately have PQTrans: "\ \ \\<^sub>R \ P \ Q \\M\\*xvec\\N\ \ (P' \ Q')" + using \A\<^sub>P \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* Q\ \A\<^sub>P \* M\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>Q \* \\ \A\<^sub>Q \* \\<^sub>R\ \A\<^sub>Q \* P\ \A\<^sub>Q \* Q\ \xvec \* Q\ \A\<^sub>Q \* M\ + by(intro BrComm2) (assumption | force)+ + have RTrans: "(\ \ \\<^sub>P) \ \\<^sub>Q \ R \\M\N\ \ R'" and FrR: "extractFrame R = \A\<^sub>R, \\<^sub>R\" by fact+ + note PQTrans + moreover from FrP \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>Q\<^sub>R \* \\<^sub>P\ Aeq + have "extractFrame(P \ Q) = \(A\<^sub>P@A\<^sub>Q), \\<^sub>P \ \\<^sub>Q\" by simp+ + moreover from RTrans have "\ \ (\\<^sub>P \ \\<^sub>Q) \ R \\M\N\ \ R'" + by (metis Associativity statEqTransition) + moreover note FrR + moreover note \A\<^sub>P \* \\ \A\<^sub>Q \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* Q\ \A\<^sub>Q \* P\ \A\<^sub>Q \* Q\ \A\<^sub>P \* R\ \A\<^sub>Q \* R\ + \A\<^sub>P \* M\ \A\<^sub>Q \* M\ \A\<^sub>P \* A\<^sub>R\ \A\<^sub>Q \* A\<^sub>R\ \A\<^sub>R \* \\ \A\<^sub>R \* P\ \A\<^sub>R \* Q\ \A\<^sub>R \* R\ \A\<^sub>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\<^sub>Q \\<^sub>Q) + from \extractFrame(\) = \A\<^sub>Q, \\<^sub>Q\\ have "\\<^sub>Q = SBottom'" by auto + with \\ \ \\<^sub>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\<^sub>P \\<^sub>P) + from \\ \ \\<^sub>P \ \ \\ \ Q'\ have False + by auto + then show ?case by simp + next + case(cComm1 \\<^sub>Q M N P' A\<^sub>P \\<^sub>P K xvec Q' A\<^sub>Q) + from \\ \ \\<^sub>P \ \ \K\\*xvec\\N\ \ Q'\ have False by auto + then show ?case by simp + next + case(cComm2 \\<^sub>Q M xvec N P' A\<^sub>P \\<^sub>P K Q' A\<^sub>Q) + from \\ \ \\<^sub>P \ \ \K\N\ \ Q'\ have False + by auto + then show ?case by simp + next + case(cBrMerge \\<^sub>Q M N P' A\<^sub>P \\<^sub>P Q' A\<^sub>Q) + from \\ \ \\<^sub>P \ \ \ \M\N\ \ Q'\ have False + by auto + then show ?case by simp + next + case(cBrComm1 \\<^sub>Q M N P' A\<^sub>P \\<^sub>P xvec Q' A\<^sub>Q) + from \\ \ \\<^sub>P \ \ \ \M\\*xvec\\N\ \ Q'\ have False + by auto + then show ?case by simp + next + case(cBrComm2 \\<^sub>Q M xvec N P' A\<^sub>P \\<^sub>P Q' A\<^sub>Q) + from \\ \ \\<^sub>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" + \ \Addition for Broadcast\ + 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\<^sub>Q \\<^sub>Q) + from \x \ P' \ \\x\Q\ have "x \ P'" by simp + have PTrans: "\ \ \\<^sub>Q \ P \\ \ P'" by fact + from \extractFrame(\\x\Q) = \A\<^sub>Q, \\<^sub>Q\\ have FrxQ: "\\x\(extractFrame Q) = \A\<^sub>Q, \\<^sub>Q\" by simp + then obtain y A\<^sub>Q' where A: "A\<^sub>Q = y#A\<^sub>Q'" by(cases A\<^sub>Q) auto + with \A\<^sub>Q \* \\ \A\<^sub>Q \* P\ \A\<^sub>Q \* \\ + have "A\<^sub>Q' \* \" and "A\<^sub>Q' \* P" and "A\<^sub>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\<^sub>Q \* x\ FrxQ have "extractFrame([(y, x)] \ Q) = \A\<^sub>Q', \\<^sub>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\<^sub>Q \* \\ A have "bn \ \* ([(y, x)] \ Q)" by simp + ultimately have "\ \ P \ ([(y, x)] \ Q) \\ \ (P' \ ([(y, x)] \ Q))" + using \A\<^sub>Q' \* \\ \A\<^sub>Q' \* P\ \A\<^sub>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\<^sub>P \\<^sub>P) + from \A\<^sub>P \* \\x\Q\ \A\<^sub>P \* x\ have "A\<^sub>P \* Q" by simp + note \\ \ \\<^sub>P \ \\x\Q \\ \ xQ'\ + moreover have FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" by fact + with \x \ P\ \A\<^sub>P \* x\ have "x \ \\<^sub>P" and "x \ A\<^sub>P" + by(force dest: extractFrameFresh)+ + with \x \ \\ have "x \ \ \ \\<^sub>P" by force + moreover note \x \ \\ + moreover from \x \ P \ xQ'\ have "x \ xQ'" by simp + moreover from FrP \bn \ \* P\ \A\<^sub>P \* \\ have "bn \ \* \\<^sub>P" + by(auto dest: extractFrameFreshChain) + with \bn \ \* \\ have "bn \ \* (\ \ \\<^sub>P)" by force + ultimately show ?case using \bn \ \* Q\ \bn \ \* subject \\ \x \ \\ \bn \ \* P\ \A\<^sub>P \* \\ \bn \ \* \\ \x \ P\ + proof(induct rule: resCases'[where C="(P, A\<^sub>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 \\ \ \\<^sub>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\<^sub>P \* (M\\*(xvec1@y#xvec2)\\N\)\ have "A\<^sub>P \* (M\\*(xvec1@xvec2)\\N\)" and "y \ A\<^sub>P" by simp+ + moreover from \A\<^sub>P \* Q\ \x \ A\<^sub>P\ \y \ A\<^sub>P\ have "A\<^sub>P \* ([(x, y)] \ Q)" by simp + ultimately have "\ \ P \ ([(x, y)] \ Q) \M\\*(xvec1@xvec2)\\N\ \ (P \ Q')" + using \A\<^sub>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 \\ \ \\<^sub>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\<^sub>P \* (\M\\*(xvec1@y#xvec2)\\N\)\ have "A\<^sub>P \* (\M\\*(xvec1@xvec2)\\N\)" and "y \ A\<^sub>P" by simp+ + moreover from \A\<^sub>P \* Q\ \x \ A\<^sub>P\ \y \ A\<^sub>P\ have "A\<^sub>P \* ([(x, y)] \ Q)" by simp + ultimately have "\ \ P \ ([(x, y)] \ Q) \\M\\*(xvec1@xvec2)\\N\ \ (P \ Q')" + using \A\<^sub>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 \\ \ \\<^sub>P \ Q \\ \ Q'\ FrP \bn \ \* P\ + have "\ \ P \ Q \\ \ (P \ Q')" using \A\<^sub>P \* \\ \A\<^sub>P \* Q\ \A\<^sub>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\<^sub>P, \)\ + have "xvec \* P" and "A\<^sub>P \* xvec" and "xvec \* \" + by simp+ + from \\ \ \\<^sub>P \ Q \ \M\\*xvec\\N\ \ Q'\ \A\<^sub>P \* Q\ + \xvec \* M\ + have "A\<^sub>P \* M" + by(simp add: brOutputFreshSubjectChain) + + from \\ \ \\<^sub>P \ Q \ \M\\*xvec\\N\ \ Q'\ \xvec \* M\ \distinct xvec\ \A\<^sub>P \* Q\ \A\<^sub>P \* xvec\ + have "A\<^sub>P \* N" and "A\<^sub>P \* Q'" by(simp add: broutputFreshChainDerivative)+ + from \A\<^sub>P \* M\ \A\<^sub>P \* xvec\ \A\<^sub>P \* N\ have "A\<^sub>P \* (\M\\*xvec\\N\)" + by simp + + from \\ \ \\<^sub>P \ Q \ \M\\*xvec\\N\ \ Q'\ \extractFrame P = \A\<^sub>P, \\<^sub>P\\ \xvec \* P\ \A\<^sub>P \* \\ \A\<^sub>P \* Q\ \A\<^sub>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 \\<^sub>Q M N P' A\<^sub>P \\<^sub>P K xvec xQ' A\<^sub>Q) + have QTrans: "\ \ \\<^sub>P \ \\x\Q \K\\*xvec\\N\ \ xQ'" and FrQ: "extractFrame(\\x\Q) = \A\<^sub>Q, \\<^sub>Q\" by fact+ + have PTrans: "\ \ \\<^sub>Q \ P \M\N\ \ P'" and FrP: "extractFrame P = \A\<^sub>P, \\<^sub>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\<^sub>Q \* x\ have "x \ \\<^sub>Q" + by(drule_tac extractFrameFresh) auto + from \x \ P\ FrP \A\<^sub>P \* x\ have "x \ \\<^sub>P" + by(drule_tac extractFrameFresh) auto + from \A\<^sub>P \* \\x\Q\ \A\<^sub>P \* x\ have "A\<^sub>P \* Q" by simp + from \A\<^sub>Q \* \\x\Q\ \A\<^sub>Q \* x\ have "A\<^sub>Q \* Q" by simp + from PTrans FrP \distinct A\<^sub>P\ \x \ P\ \A\<^sub>Q \* P\ \xvec \* P\ \A\<^sub>P \* \\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>P \* x\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* P\ \A\<^sub>P \* M\ \A\<^sub>P \* xvec\ + obtain M' where "(\ \ \\<^sub>Q) \ \\<^sub>P \ M \ M'" and "x \ M'" and "A\<^sub>Q \* M'" and "xvec \* M'" + by(elim inputObtainPrefix[where B="x#xvec@A\<^sub>Q"]) (assumption | force simp add: fresh_star_list_cons)+ + then have MeqM': "\ \ \\<^sub>P \ \\<^sub>Q \ M \ M'" by(metis statEqEnt Associativity Commutativity Composition) + with \\ \ \\<^sub>P \ \\<^sub>Q \ M \ K\ have "\ \ \\<^sub>P \ \\<^sub>Q \ K \ M'" + by(blast intro: chanEqTrans chanEqSym) + then have "(\ \ \\<^sub>P) \ \\<^sub>Q \ K \ M'" by(metis statEqEnt Associativity Commutativity Composition) + with QTrans FrQ \distinct A\<^sub>Q\ \A\<^sub>Q \* \\ \A\<^sub>Q \* \\<^sub>P\ \A\<^sub>Q \* (\\x\Q)\ \A\<^sub>Q \* K\ \A\<^sub>Q \* M'\ + have "\ \ \\<^sub>P \ \\x\Q \M'\\*xvec\\N\ \ xQ'" + by(force intro: outputRenameSubject) + moreover from \x \ \\ \x \ \\<^sub>P\ have "x \ \ \ \\<^sub>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 \* \\<^sub>P\ have "xvec \* (\ \ \\<^sub>P)" by force + moreover from \xvec \* \\x\Q\ \x \ xvec\ have "xvec \* Q" by simp + moreover note \xvec \* M'\ + + moreover from \xvec \* \\ \xvec \* \\<^sub>P\ have "bn(M'\\*xvec\\N\) \* (\ \ \\<^sub>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\<^sub>Q \* xvec\ \A\<^sub>Q \* M'\ \A\<^sub>Q \* N\ have "A\<^sub>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\<^sub>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\<^sub>Q \* (M'\\*xvec\\N\)\ Eq have "(xvec1@xvec2) \* A\<^sub>Q" and "y \ A\<^sub>Q" and "A\<^sub>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\ \\ \ \\<^sub>Q \ P \M\N\ \ P'\ + have "\ \ \\<^sub>Q \ P \M\([(y, x)] \ N)\ \ ([(y, x)] \ P')" + by(elim inputAlpha[where xvec="[y]"]) (auto simp add: calc_atm) + then have PTrans: "\ \ \\<^sub>Q \ P \M\([(x, y)] \ N)\ \ ([(x, y)] \ P')" + by(simp add: name_swap) + + from \\ \ \\<^sub>P \ [(x, y)] \ Q \ M''\\*(xvec1 @ xvec2)\\N'\ \ Q'\ + have "[(x, y)] \ (\ \ \\<^sub>P \ [(x, y)] \ Q \ M''\\*(xvec1 @ xvec2)\\N'\ \ Q')" + by simp + then have "[(x, y)] \ (\ \ \\<^sub>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 \ (\ \ \\<^sub>P)\ \y \ (\ \ \\<^sub>P)\ \x \ xvec1\ \y \ xvec1\ \x \ xvec2\ \y \ xvec2\ \x \ M''\ \y \ M''\ + have QTrans: "\ \ \\<^sub>P \ Q \ M''\\*(xvec1 @ xvec2)\\([(x, y)] \ N')\ \ [(x, y)] \ Q'" + by simp + with \A\<^sub>Q \* x\ \y \ A\<^sub>Q\ \distinct xvec1\ \distinct xvec2\ \xvec1 \* xvec2\ \xvec1 \* M''\ \xvec2 \* M''\ + \(xvec1@xvec2) \* A\<^sub>Q\ + have "A\<^sub>Q \* ([(x, y)] \ Q')" using \A\<^sub>Q \* Q\ + by(elim outputFreshChainDerivative(2)) (assumption | simp)+ + + from \extractFrame(\\x\Q) = \A\<^sub>Q, \\<^sub>Q\\ have FrxQ: "\\x\(extractFrame Q) = \A\<^sub>Q, \\<^sub>Q\" by simp + then obtain z A\<^sub>Q' where A: "A\<^sub>Q = z#A\<^sub>Q'" by(cases A\<^sub>Q) auto + with \A\<^sub>Q \* \\ \A\<^sub>Q \* P\ \A\<^sub>Q \* P'\ \A\<^sub>Q \* \\<^sub>P\ \A\<^sub>Q \* Q\ \(xvec1@xvec2) \* A\<^sub>Q\ \A\<^sub>Q \* M''\ \A\<^sub>Q \* ([(x, y)] \ Q')\ \y \ A\<^sub>Q\ \A\<^sub>Q \* N\ + have "A\<^sub>Q' \* \" and "A\<^sub>Q' \* P" and "A\<^sub>Q' \* \\<^sub>P" and "A\<^sub>Q' \* Q" + and "z \ \" and "z \ P" and "z \ P'" and "z \ \\<^sub>P" and "z \ Q" and "z \ xvec1" and "z \ xvec2" + and "z \ M''" and "z \ ([(x, y)] \ Q')" and "A\<^sub>Q' \* M''" and "z \ y" and "z \ (xvec1@xvec2)" + by auto + from A \A\<^sub>P \* A\<^sub>Q\ have "A\<^sub>P \* A\<^sub>Q'" and "z \ A\<^sub>P" by simp+ + from A \A\<^sub>Q \* x\ have "x \ z" and "x \ A\<^sub>Q'" by simp+ + + from \distinct A\<^sub>Q\ A have "z \ A\<^sub>Q'" + by(induct A\<^sub>Q') (auto simp add: fresh_list_nil fresh_list_cons) + from PTrans \x \ P\ \z \ P\ \x \ z\ have "\ \ \\<^sub>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)] \ (\ \ \\<^sub>P)) \ ([(x, z)] \ Q) \([(x, z)] \ (M''\\*(xvec1@xvec2)\\([(x, y)] \ N')\ \ ([(x, y)] \ Q')))" + by(rule semantics.eqvt) + with \x \ \\ \z \ \\ \x \ \\<^sub>P\ \z \ \\<^sub>P\ \x \ M''\ \z \ M''\ \x \ xvec1\ \x \ xvec2\ \z \ xvec1\ \z \ xvec2\ + have "\ \ \\<^sub>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\<^sub>Q \* x\ FrxQ have "extractFrame([(x, z)] \ Q) = \A\<^sub>Q', \\<^sub>Q\" + by(clarsimp simp add: alpha' eqvts frame.inject fresh_list_cons name_swap) + moreover from \A\<^sub>P \* Q\ have "([(x, z)] \ A\<^sub>P) \* ([(x, z)] \ Q)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>P \* x\ \z \ A\<^sub>P\ have "A\<^sub>P \* ([(x, z)] \ Q)" by simp + moreover from \A\<^sub>Q' \* Q\ have "([(x, z)] \ A\<^sub>Q') \* ([(x, z)] \ Q)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \x \ A\<^sub>Q'\ \z \ A\<^sub>Q'\ have "A\<^sub>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\<^sub>P \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* A\<^sub>Q'\ \A\<^sub>Q' \* \\ \A\<^sub>Q' \* P\ \(xvec1@xvec2) \* P\ \A\<^sub>P \* M\ \A\<^sub>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: "\ \ \\<^sub>P \ Q \M'\\*xvec\\N\ \ Q'" by fact + with \A\<^sub>Q \* Q\ \A\<^sub>Q \* xvec\ \xvec \* M'\ \distinct xvec\ have "A\<^sub>Q \* Q'" + by(force dest: outputFreshChainDerivative) + + with \extractFrame(\\x\Q) = \A\<^sub>Q, \\<^sub>Q\\ have FrxQ: "\\x\(extractFrame Q) = \A\<^sub>Q, \\<^sub>Q\" by simp + then obtain y A\<^sub>Q' where A: "A\<^sub>Q = y#A\<^sub>Q'" by(cases A\<^sub>Q) auto + with \A\<^sub>Q \* \\ \A\<^sub>Q \* P\ \A\<^sub>Q \* P'\ \A\<^sub>Q \* \\<^sub>P\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* xvec\ \A\<^sub>Q \* M'\ \A\<^sub>Q \* Q'\ + have "A\<^sub>Q' \* \" and "A\<^sub>Q' \* P" and "A\<^sub>Q' \* \\<^sub>P" and "A\<^sub>Q' \* Q" and "A\<^sub>Q \* xvec" and "A\<^sub>Q \* Q'" + and "y \ \" and "y \ P" and "y \ P'" and "y \ \\<^sub>P" and "y \ Q" and "y \ xvec" and "y \ M'" and "y \ Q'" + and "A\<^sub>Q' \* M'" + by(simp)+ + from A \A\<^sub>P \* A\<^sub>Q\ have "A\<^sub>P \* A\<^sub>Q'" and "y \ A\<^sub>P" by(simp add: fresh_star_list_cons)+ + from A \A\<^sub>Q \* x\ have "x \ y" and "x \ A\<^sub>Q'" by(simp add: fresh_list_cons)+ + + with A \distinct A\<^sub>Q\ have "y \ A\<^sub>Q'" + by(induct A\<^sub>Q') (auto simp add: fresh_list_nil fresh_list_cons) + + from \x \ P\ \y \ P\ \x \ y\ \\ \ \\<^sub>Q \ P \M\N\ \ P'\ + have "\ \ \\<^sub>Q \ P \M\([(y, x)] \ N)\ \ [(y, x)] \ P'" + by(intro inputAlpha[where xvec="[y]"]) (auto simp add: calc_atm) + then have "\ \ \\<^sub>Q \ P \M\([(x, y)] \ N)\ \ [(x, y)] \ P'" + by(simp add: name_swap) + moreover note FrP + moreover from \\ \ \\<^sub>P \ Q \M'\\*xvec\\N\ \ Q'\ have "([(x, y)] \ (\ \ \\<^sub>P)) \ ([(x, y)] \ Q) \([(x, y)] \ (M'\\*xvec\\N\ \ Q'))" + by(rule semantics.eqvt) + with \x \ \\ \y \ \\ \x \ \\<^sub>P\ \y \ \\<^sub>P\ \x \ M'\ \y \ M'\ \x \ xvec\ \y \ xvec\ + have "\ \ \\<^sub>P \ ([(x, y)] \ Q) \M'\\*xvec\\([(x, y)] \ N)\ \ ([(x, y)] \ Q')" + by(simp add: eqvts) + moreover from A \A\<^sub>Q \* x\ FrxQ have FrQ: "extractFrame([(x, y)] \ Q) = \A\<^sub>Q', \\<^sub>Q\" + by(clarsimp simp add: alpha' eqvts frame.inject fresh_list_cons name_swap) + moreover from \A\<^sub>P \* Q\ have "([(x, y)] \ A\<^sub>P) \* ([(x, y)] \ Q)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>P \* x\ \y \ A\<^sub>P\ have "A\<^sub>P \* ([(x, y)] \ Q)" by simp + moreover from \A\<^sub>Q' \* Q\ have "([(x, y)] \ A\<^sub>Q') \* ([(x, y)] \ Q)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \x \ A\<^sub>Q'\ \y \ A\<^sub>Q'\ have "A\<^sub>Q' \* ([(x, y)] \ Q)" by simp + ultimately have "\ \ (P \ ([(x, y)] \ Q)) \\ \ \\*xvec\([(x, y)] \ P') \ ([(x, y)] \ Q')" + using MeqM' \A\<^sub>P \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* A\<^sub>Q'\ \A\<^sub>Q' \* \\ \A\<^sub>Q' \* P\ \xvec \* P\ \A\<^sub>P \* M\ \A\<^sub>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 \\<^sub>Q M xvec N P' A\<^sub>P \\<^sub>P K xQ' A\<^sub>Q) + have QTrans: "\ \ \\<^sub>P \ \\x\Q \K\N\ \ xQ'" and FrQ: "extractFrame(\\x\Q) = \A\<^sub>Q, \\<^sub>Q\" by fact+ + have PTrans: "\ \ \\<^sub>Q \ P \M\\*xvec\\N\ \ P'" and FrP: "extractFrame P = \A\<^sub>P, \\<^sub>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\<^sub>Q \* x\ have "x \ \\<^sub>Q" + by(drule_tac extractFrameFresh) auto + from \x \ P\ FrP \A\<^sub>P \* x\ have "x \ \\<^sub>P" + by(drule_tac extractFrameFresh) auto + from \A\<^sub>P \* \\x\Q\ \A\<^sub>P \* x\ have "A\<^sub>P \* Q" by simp + from \A\<^sub>Q \* \\x\Q\ \A\<^sub>Q \* x\ have "A\<^sub>Q \* Q" by simp + from \xvec \* x\ \xvec \* \\x\Q\ have "xvec \* Q" by simp + + from \\ \ \\<^sub>Q \ P \ M\\*xvec\\N\ \ P'\ have PResTrans: "\ \ \\<^sub>Q \ P \ ROut M (\\*xvec\N \' P')" + by(simp add: residualInject) + + from PResTrans FrP \distinct A\<^sub>P\ \x \ P\ \A\<^sub>Q \* P\ \A\<^sub>P \* \\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>P \* x\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* P\ \A\<^sub>P \* M\ \xvec \* M\ \distinct xvec\ + obtain M' where "(\ \ \\<^sub>Q) \ \\<^sub>P \ M \ M'" and "x \ M'" and "A\<^sub>Q \* M'" + by(elim outputObtainPrefix[where B="x#A\<^sub>Q"]) (assumption | force simp add: fresh_star_list_cons)+ + then have MeqM': "\ \ \\<^sub>P \ \\<^sub>Q \ M \ M'" + by(metis statEqEnt Associativity Commutativity Composition) + with \\ \ \\<^sub>P \ \\<^sub>Q \ M \ K\ have "\ \ \\<^sub>P \ \\<^sub>Q \ K \ M'" + by(blast intro: chanEqTrans chanEqSym) + then have "(\ \ \\<^sub>P) \ \\<^sub>Q \ K \ M'" + by(metis statEqEnt Associativity Commutativity Composition) + with QTrans FrQ \distinct A\<^sub>Q\ \A\<^sub>Q \* \\ \A\<^sub>Q \* \\<^sub>P\ \A\<^sub>Q \* (\\x\Q)\ \A\<^sub>Q \* K\ \A\<^sub>Q \* M'\ + have "\ \ \\<^sub>P \ \\x\Q \M'\N\ \ xQ'" by(force intro: inputRenameSubject) + + moreover from \x \ \\ \x \ \\<^sub>P\ have "x \ \ \ \\<^sub>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: "\ \ \\<^sub>P \ Q \M'\N\ \ Q'" by fact + with \A\<^sub>Q \* Q\ \A\<^sub>Q \* N\ have "A\<^sub>Q \* Q'" + by(elim inputFreshChainDerivative) + + with \extractFrame(\\x\Q) = \A\<^sub>Q, \\<^sub>Q\\ have FrxQ: "\\x\(extractFrame Q) = \A\<^sub>Q, \\<^sub>Q\" by simp + then obtain y A\<^sub>Q' where A: "A\<^sub>Q = y#A\<^sub>Q'" by(cases A\<^sub>Q) auto + with \A\<^sub>Q \* \\ \A\<^sub>Q \* P\ \A\<^sub>Q \* P'\ \A\<^sub>Q \* \\<^sub>P\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* xvec\ \A\<^sub>Q \* M'\ \A\<^sub>Q \* Q'\ \A\<^sub>Q \* N\ + have "A\<^sub>Q' \* \" and "A\<^sub>Q' \* P" and "A\<^sub>Q' \* \\<^sub>P" and "A\<^sub>Q' \* Q" and "A\<^sub>Q \* xvec" and "A\<^sub>Q \* Q'" + and "y \ \" and "y \ P" and "y \ P'" and "y \ \\<^sub>P" and "y \ Q" and "y \ xvec" and "y \ M'" and "y \ Q'" and "y \ N" + and "A\<^sub>Q' \* M'" + by(simp)+ + from A \A\<^sub>P \* A\<^sub>Q\ have "A\<^sub>P \* A\<^sub>Q'" and "y \ A\<^sub>P" by(simp add: fresh_star_list_cons)+ + from A \A\<^sub>Q \* x\ have "x \ y" and "x \ A\<^sub>Q'" by(simp add: fresh_list_cons)+ + + with A \distinct A\<^sub>Q\ have "y \ A\<^sub>Q'" + by(induct A\<^sub>Q') (auto simp add: fresh_list_nil fresh_list_cons) + + note PTrans FrP + moreover from \\ \ \\<^sub>P \ Q \M'\N\ \ Q'\ have "([(x, y)] \ (\ \ \\<^sub>P)) \ ([(x, y)] \ Q) \([(x, y)] \ (M'\N\ \ Q'))" + by(rule semantics.eqvt) + with \x \ \\ \y \ \\ \x \ \\<^sub>P\ \y \ \\<^sub>P\ \x \ M'\ \y \ M'\ \x \ N\ \y \ N\ + have "\ \ \\<^sub>P \ ([(x, y)] \ Q) \M'\N\ \ ([(x, y)] \ Q')" + by(simp add: eqvts) + moreover from A \A\<^sub>Q \* x\ FrxQ have FrQ: "extractFrame([(x, y)] \ Q) = \A\<^sub>Q', \\<^sub>Q\" and "y \ extractFrame Q" + by(clarsimp simp add: alpha' eqvts frame.inject fresh_list_cons name_swap)+ + moreover from \A\<^sub>P \* Q\ have "([(x, y)] \ A\<^sub>P) \* ([(x, y)] \ Q)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>P \* x\ \y \ A\<^sub>P\ have "A\<^sub>P \* ([(x, y)] \ Q)" by simp + moreover from \A\<^sub>Q' \* Q\ have "([(x, y)] \ A\<^sub>Q') \* ([(x, y)] \ Q)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \x \ A\<^sub>Q'\ \y \ A\<^sub>Q'\ have "A\<^sub>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\<^sub>P \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* A\<^sub>Q'\ \A\<^sub>Q' \* \\ \A\<^sub>Q' \* P\ \A\<^sub>P \* M\ \A\<^sub>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 \\<^sub>Q M N P' A\<^sub>P \\<^sub>P xQ' A\<^sub>Q) + have QTrans: "\ \ \\<^sub>P \ \\x\Q \\M\N\ \ xQ'" and FrQ: "extractFrame(\\x\Q) = \A\<^sub>Q, \\<^sub>Q\" by fact+ + have PTrans: "\ \ \\<^sub>Q \ P \\M\N\ \ P'" and FrP: "extractFrame P = \A\<^sub>P, \\<^sub>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\<^sub>P \* x\ \x \ P\ have "x \ \\<^sub>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\<^sub>Q \* x\ have "x \ \\<^sub>Q" + by(drule_tac extractFrameFresh) auto + from \A\<^sub>P \* \\x\Q\ \A\<^sub>P \* x\ have "A\<^sub>P \* Q" by simp + from \A\<^sub>Q \* \\x\Q\ \A\<^sub>Q \* x\ have "A\<^sub>Q \* Q" by simp + from \x \ \\ \x \ \\<^sub>P\ + have "x \ (\ \ \\<^sub>P)" by force + from \\ \ \\<^sub>P \ \\x\Q \ \M\N\ \ xQ'\ \x \ (\ \ \\<^sub>P)\ \x \ M\ \x \ N\ \x \ xQ'\ + show ?case + proof(induct rule: resBrInputCases) + case(cRes Q') + have QTrans: "\ \ \\<^sub>P \ Q \\M\N\ \ Q'" by fact + with \A\<^sub>Q \* Q\ \A\<^sub>Q \* N\ have "A\<^sub>Q \* Q'" + by(elim brinputFreshChainDerivative) + + with \extractFrame(\\x\Q) = \A\<^sub>Q, \\<^sub>Q\\ have FrxQ: "\\x\(extractFrame Q) = \A\<^sub>Q, \\<^sub>Q\" by simp + then obtain y A\<^sub>Q' where A: "A\<^sub>Q = y#A\<^sub>Q'" by(cases A\<^sub>Q) auto + with \A\<^sub>Q \* \\ \A\<^sub>Q \* P\ \A\<^sub>Q \* P'\ \A\<^sub>Q \* \\<^sub>P\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* M\ \A\<^sub>Q \* Q'\ \A\<^sub>Q \* N\ + have "A\<^sub>Q' \* \" and "A\<^sub>Q' \* P" and "A\<^sub>Q' \* \\<^sub>P" and "A\<^sub>Q' \* Q" and "A\<^sub>Q \* Q'" + and "y \ \" and "y \ P" and "y \ P'" and "y \ \\<^sub>P" and "y \ Q" and "y \ M" and "y \ Q'" and "y \ N" + and "A\<^sub>Q' \* M" + by(simp)+ + from A \A\<^sub>P \* A\<^sub>Q\ have "A\<^sub>P \* A\<^sub>Q'" and "y \ A\<^sub>P" by(simp add: fresh_star_list_cons)+ + from A \A\<^sub>Q \* x\ have "x \ y" and "x \ A\<^sub>Q'" by(simp add: fresh_list_cons)+ + + with A \distinct A\<^sub>Q\ have "y \ A\<^sub>Q'" + by(induct A\<^sub>Q') (auto simp add: fresh_list_nil fresh_list_cons) + + note PTrans FrP + moreover from \\ \ \\<^sub>P \ Q \\M\N\ \ Q'\ have "([(x, y)] \ (\ \ \\<^sub>P)) \ ([(x, y)] \ Q) \([(x, y)] \ (\M\N\ \ Q'))" + by(rule semantics.eqvt) + with \x \ \\ \y \ \\ \x \ \\<^sub>P\ \y \ \\<^sub>P\ \x \ M\ \y \ M\ \x \ N\ \y \ N\ + have "\ \ \\<^sub>P \ ([(x, y)] \ Q) \\M\N\ \ ([(x, y)] \ Q')" + by(simp add: eqvts) + moreover from A \A\<^sub>Q \* x\ FrxQ have FrQ: "extractFrame([(x, y)] \ Q) = \A\<^sub>Q', \\<^sub>Q\" and "y \ extractFrame Q" + by(clarsimp simp add: alpha' eqvts frame.inject fresh_list_cons name_swap)+ + moreover from \A\<^sub>P \* Q\ have "([(x, y)] \ A\<^sub>P) \* ([(x, y)] \ Q)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>P \* x\ \y \ A\<^sub>P\ have "A\<^sub>P \* ([(x, y)] \ Q)" by simp + moreover from \A\<^sub>Q' \* Q\ have "([(x, y)] \ A\<^sub>Q') \* ([(x, y)] \ Q)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \x \ A\<^sub>Q'\ \y \ A\<^sub>Q'\ have "A\<^sub>Q' \* ([(x, y)] \ Q)" by simp + ultimately have "\ \ (P \ ([(x, y)] \ Q)) \\M\N\ \ (P' \ ([(x, y)] \ Q'))" + using \A\<^sub>P \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* A\<^sub>Q'\ \A\<^sub>Q' \* \\ \A\<^sub>Q' \* P\ \A\<^sub>P \* M\ \A\<^sub>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 \\<^sub>Q M N P' A\<^sub>P \\<^sub>P xvec xQ' A\<^sub>Q) + have QTrans: "\ \ \\<^sub>P \ \\x\Q \\M\\*xvec\\N\ \ xQ'" and FrQ: "extractFrame(\\x\Q) = \A\<^sub>Q, \\<^sub>Q\" by fact+ + have PTrans: "\ \ \\<^sub>Q \ P \\M\N\ \ P'" and FrP: "extractFrame P = \A\<^sub>P, \\<^sub>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\<^sub>Q \* x\ have "x \ \\<^sub>Q" + by(drule_tac extractFrameFresh) auto + from \x \ P\ FrP \A\<^sub>P \* x\ have "x \ \\<^sub>P" + by(drule_tac extractFrameFresh) auto + from \A\<^sub>P \* \\x\Q\ \A\<^sub>P \* x\ have "A\<^sub>P \* Q" by simp + from \A\<^sub>Q \* \\x\Q\ \A\<^sub>Q \* x\ have "A\<^sub>Q \* Q" by simp + note QTrans + moreover from \x \ \\ \x \ \\<^sub>P\ have "x \ \ \ \\<^sub>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 \* \\<^sub>P\ have "xvec \* (\ \ \\<^sub>P)" by force + moreover from \xvec \* \\x\Q\ \x \ xvec\ have "xvec \* Q" by simp + moreover note \xvec \* M\ + + moreover from \xvec \* \\ \xvec \* \\<^sub>P\ have "bn(\M\\*xvec\\N\) \* (\ \ \\<^sub>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\<^sub>Q \* xvec\ \A\<^sub>Q \* M\ \A\<^sub>Q \* N\ have "A\<^sub>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\<^sub>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\<^sub>Q \* (\M\\*xvec\\N\)\ Eq have "(xvec1@xvec2) \* A\<^sub>Q" and "y \ A\<^sub>Q" and "A\<^sub>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\ \\ \ \\<^sub>Q \ P \\M\N\ \ P'\ + have "\ \ \\<^sub>Q \ P \\M\([(y, x)] \ N)\ \ ([(y, x)] \ P')" + by(intro brinputAlpha[where xvec="[y]"]) (auto simp add: calc_atm) + then have PTrans: "\ \ \\<^sub>Q \ P \\M\([(x, y)] \ N)\ \ ([(x, y)] \ P')" + by(simp add: name_swap) + + from \\ \ \\<^sub>P \ [(x, y)] \ Q \ \M''\\*(xvec1 @ xvec2)\\N'\ \ Q'\ + have "[(x, y)] \ (\ \ \\<^sub>P \ [(x, y)] \ Q \ \M''\\*(xvec1 @ xvec2)\\N'\ \ Q')" + by simp + then have "[(x, y)] \ (\ \ \\<^sub>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 \ (\ \ \\<^sub>P)\ \y \ (\ \ \\<^sub>P)\ \x \ xvec1\ \y \ xvec1\ \x \ xvec2\ \y \ xvec2\ \x \ M''\ \y \ M''\ + have QTrans: "\ \ \\<^sub>P \ Q \ \M''\\*(xvec1 @ xvec2)\\([(x, y)] \ N')\ \ [(x, y)] \ Q'" + by simp + with \A\<^sub>Q \* x\ \y \ A\<^sub>Q\ \distinct xvec1\ \distinct xvec2\ \xvec1 \* xvec2\ \xvec1 \* M''\ \xvec2 \* M''\ + \(xvec1@xvec2) \* A\<^sub>Q\ + have "A\<^sub>Q \* ([(x, y)] \ Q')" using \A\<^sub>Q \* Q\ + by(elim broutputFreshChainDerivative(2)) (assumption | simp)+ + + from \extractFrame(\\x\Q) = \A\<^sub>Q, \\<^sub>Q\\ have FrxQ: "\\x\(extractFrame Q) = \A\<^sub>Q, \\<^sub>Q\" by simp + then obtain z A\<^sub>Q' where A: "A\<^sub>Q = z#A\<^sub>Q'" by(cases A\<^sub>Q) auto + with \A\<^sub>Q \* \\ \A\<^sub>Q \* P\ \A\<^sub>Q \* P'\ \A\<^sub>Q \* \\<^sub>P\ \A\<^sub>Q \* Q\ \(xvec1@xvec2) \* A\<^sub>Q\ \A\<^sub>Q \* M''\ \A\<^sub>Q \* ([(x, y)] \ Q')\ \y \ A\<^sub>Q\ \A\<^sub>Q \* N\ + have "A\<^sub>Q' \* \" and "A\<^sub>Q' \* P" and "A\<^sub>Q' \* \\<^sub>P" and "A\<^sub>Q' \* Q" + and "z \ \" and "z \ P" and "z \ P'" and "z \ \\<^sub>P" and "z \ Q" and "z \ xvec1" and "z \ xvec2" + and "z \ M''" and "z \ ([(x, y)] \ Q')" and "A\<^sub>Q' \* M''" and "z \ y" and "z \ (xvec1@xvec2)" + by auto + from A \A\<^sub>P \* A\<^sub>Q\ have "A\<^sub>P \* A\<^sub>Q'" and "z \ A\<^sub>P" by simp+ + from A \A\<^sub>Q \* x\ have "x \ z" and "x \ A\<^sub>Q'" by simp+ + + from \A\<^sub>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\<^sub>Q\ A have "z \ A\<^sub>Q'" + by(induct A\<^sub>Q') (auto simp add: fresh_list_nil fresh_list_cons) + from PTrans \x \ P\ \z \ P\ \x \ z\ have "\ \ \\<^sub>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)] \ (\ \ \\<^sub>P)) \ ([(x, z)] \ Q) \([(x, z)] \ (\M''\\*(xvec1@xvec2)\\([(x, y)] \ N')\ \ ([(x, y)] \ Q')))" + by(rule semantics.eqvt) + with \x \ \\ \z \ \\ \x \ \\<^sub>P\ \z \ \\<^sub>P\ \x \ M''\ \z \ M''\ \x \ xvec1\ \x \ xvec2\ \z \ xvec1\ \z \ xvec2\ + have "\ \ \\<^sub>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\<^sub>Q \* x\ FrxQ have "extractFrame([(x, z)] \ Q) = \A\<^sub>Q', \\<^sub>Q\" + by(clarsimp simp add: alpha' eqvts frame.inject fresh_list_cons name_swap) + moreover from \A\<^sub>P \* Q\ have "([(x, z)] \ A\<^sub>P) \* ([(x, z)] \ Q)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>P \* x\ \z \ A\<^sub>P\ have "A\<^sub>P \* ([(x, z)] \ Q)" by simp + moreover from \A\<^sub>Q' \* Q\ have "([(x, z)] \ A\<^sub>Q') \* ([(x, z)] \ Q)" + by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \x \ A\<^sub>Q'\ \z \ A\<^sub>Q'\ have "A\<^sub>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\<^sub>P \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* A\<^sub>Q'\ \A\<^sub>Q' \* \\ \A\<^sub>Q' \* P\ \(xvec1@xvec2) \* P\ \A\<^sub>P \* M\ \A\<^sub>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: "\ \ \\<^sub>P \ Q \\M\\*xvec\\N\ \ Q'" by fact + with \A\<^sub>Q \* Q\ \A\<^sub>Q \* xvec\ \xvec \* M\ \distinct xvec\ have "A\<^sub>Q \* Q'" + by(force dest: broutputFreshChainDerivative) + + with \extractFrame(\\x\Q) = \A\<^sub>Q, \\<^sub>Q\\ have FrxQ: "\\x\(extractFrame Q) = \A\<^sub>Q, \\<^sub>Q\" by simp + then obtain y A\<^sub>Q' where A: "A\<^sub>Q = y#A\<^sub>Q'" by(cases A\<^sub>Q) auto + with \A\<^sub>Q \* \\ \A\<^sub>Q \* P\ \A\<^sub>Q \* P'\ \A\<^sub>Q \* \\<^sub>P\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* xvec\ \A\<^sub>Q \* M\ \A\<^sub>Q \* Q'\ \A\<^sub>Q \* N\ + have "A\<^sub>Q' \* \" and "A\<^sub>Q' \* P" and "A\<^sub>Q' \* \\<^sub>P" and "A\<^sub>Q' \* Q" and "A\<^sub>Q \* xvec" and "A\<^sub>Q \* Q'" + and "y \ \" and "y \ P" and "y \ P'" and "y \ \\<^sub>P" and "y \ Q" and "y \ xvec" and "y \ M" and "y \ Q'" + and "A\<^sub>Q' \* M" and "y \ N" + by(simp)+ + from A \A\<^sub>P \* A\<^sub>Q\ have "A\<^sub>P \* A\<^sub>Q'" and "y \ A\<^sub>P" by(simp add: fresh_star_list_cons)+ + from A \A\<^sub>Q \* x\ have "x \ y" and "x \ A\<^sub>Q'" by(simp add: fresh_list_cons)+ + + with A \distinct A\<^sub>Q\ have "y \ A\<^sub>Q'" + by(induct A\<^sub>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\ \\ \ \\<^sub>Q \ P \\M\N\ \ P'\ + have "\ \ \\<^sub>Q \ P \\M\([(y, x)] \ N)\ \ [(y, x)] \ P'" + by(elim brinputAlpha[where xvec="[y]"]) (auto simp add: calc_atm) + then have "\ \ \\<^sub>Q \ P \\M\([(x, y)] \ N)\ \ [(x, y)] \ P'" + by(simp add: name_swap) + moreover note FrP + moreover from \\ \ \\<^sub>P \ Q \\M\\*xvec\\N\ \ Q'\ have "([(x, y)] \ (\ \ \\<^sub>P)) \ ([(x, y)] \ Q) \([(x, y)] \ (\M\\*xvec\\N\ \ Q'))" + by(rule semantics.eqvt) + with \x \ \\ \y \ \\ \x \ \\<^sub>P\ \y \ \\<^sub>P\ \x \ M\ \y \ M\ \x \ xvec\ \y \ xvec\ + have "\ \ \\<^sub>P \ ([(x, y)] \ Q) \\M\\*xvec\\([(x, y)] \ N)\ \ ([(x, y)] \ Q')" + by(simp add: eqvts) + moreover from A \A\<^sub>Q \* x\ FrxQ have FrQ: "extractFrame([(x, y)] \ Q) = \A\<^sub>Q', \\<^sub>Q\" + by(clarsimp simp add: alpha' eqvts frame.inject fresh_list_cons name_swap) + moreover from \A\<^sub>P \* Q\ have "([(x, y)] \ A\<^sub>P) \* ([(x, y)] \ Q)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>P \* x\ \y \ A\<^sub>P\ have "A\<^sub>P \* ([(x, y)] \ Q)" by simp + moreover from \A\<^sub>Q' \* Q\ have "([(x, y)] \ A\<^sub>Q') \* ([(x, y)] \ Q)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \x \ A\<^sub>Q'\ \y \ A\<^sub>Q'\ have "A\<^sub>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\<^sub>P \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* A\<^sub>Q'\ \A\<^sub>Q' \* \\ \A\<^sub>Q' \* P\ \xvec \* P\ \A\<^sub>P \* M\ \A\<^sub>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 \\<^sub>Q M xvec N P' A\<^sub>P \\<^sub>P xQ' A\<^sub>Q) + from \x \ \\ \\M\\*xvec\\N\ = \\ + have "x \ M" "x \ xvec" "x \ N" + by force+ + have QTrans: "\ \ \\<^sub>P \ \\x\Q \\M\N\ \ xQ'" and FrQ: "extractFrame(\\x\Q) = \A\<^sub>Q, \\<^sub>Q\" by fact+ + have PTrans: "\ \ \\<^sub>Q \ P \\M\\*xvec\\N\ \ P'" and FrP: "extractFrame P = \A\<^sub>P, \\<^sub>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\<^sub>Q \* x\ have "x \ \\<^sub>Q" + by(drule_tac extractFrameFresh) auto + from \x \ P\ FrP \A\<^sub>P \* x\ have "x \ \\<^sub>P" + by(drule_tac extractFrameFresh) auto + from \A\<^sub>P \* \\x\Q\ \A\<^sub>P \* x\ have "A\<^sub>P \* Q" by simp + from \A\<^sub>Q \* \\x\Q\ \A\<^sub>Q \* x\ have "A\<^sub>Q \* Q" by simp + from \x \ xvec\ \xvec \* \\x\Q\ have "xvec \* Q" by simp + + from \\ \ \\<^sub>Q \ P \ \M\\*xvec\\N\ \ P'\ have PResTrans: "\ \ \\<^sub>Q \ P \ RBrOut M (\\*xvec\N \' P')" + by(simp add: residualInject) + + note QTrans + moreover from \x \ \\ \x \ \\<^sub>P\ have "x \ \ \ \\<^sub>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: "\ \ \\<^sub>P \ Q \\M\N\ \ Q'" by fact + with \A\<^sub>Q \* Q\ \A\<^sub>Q \* N\ have "A\<^sub>Q \* Q'" + by(elim brinputFreshChainDerivative) + + with \extractFrame(\\x\Q) = \A\<^sub>Q, \\<^sub>Q\\ have FrxQ: "\\x\(extractFrame Q) = \A\<^sub>Q, \\<^sub>Q\" by simp + then obtain y A\<^sub>Q' where A: "A\<^sub>Q = y#A\<^sub>Q'" by(cases A\<^sub>Q) auto + with \A\<^sub>Q \* \\ \A\<^sub>Q \* P\ \A\<^sub>Q \* P'\ \A\<^sub>Q \* \\<^sub>P\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* xvec\ \A\<^sub>Q \* M\ \A\<^sub>Q \* Q'\ \A\<^sub>Q \* N\ + have "A\<^sub>Q' \* \" and "A\<^sub>Q' \* P" and "A\<^sub>Q' \* \\<^sub>P" and "A\<^sub>Q' \* Q" and "A\<^sub>Q \* xvec" and "A\<^sub>Q \* Q'" + and "y \ \" and "y \ P" and "y \ P'" and "y \ \\<^sub>P" and "y \ Q" and "y \ xvec" and "y \ M" and "y \ Q'" and "y \ N" + and "A\<^sub>Q' \* M" + by(simp)+ + from A \A\<^sub>P \* A\<^sub>Q\ have "A\<^sub>P \* A\<^sub>Q'" and "y \ A\<^sub>P" by(simp add: fresh_star_list_cons)+ + from A \A\<^sub>Q \* x\ have "x \ y" and "x \ A\<^sub>Q'" by(simp add: fresh_list_cons)+ + + with A \distinct A\<^sub>Q\ have "y \ A\<^sub>Q'" + by(induct A\<^sub>Q') (auto simp add: fresh_list_nil fresh_list_cons) + + note PTrans FrP + moreover from \\ \ \\<^sub>P \ Q \\M\N\ \ Q'\ have "([(x, y)] \ (\ \ \\<^sub>P)) \ ([(x, y)] \ Q) \([(x, y)] \ (\M\N\ \ Q'))" + by(rule semantics.eqvt) + with \x \ \\ \y \ \\ \x \ \\<^sub>P\ \y \ \\<^sub>P\ \x \ M\ \y \ M\ \x \ N\ \y \ N\ + have "\ \ \\<^sub>P \ ([(x, y)] \ Q) \\M\N\ \ ([(x, y)] \ Q')" + by(simp add: eqvts) + moreover from A \A\<^sub>Q \* x\ FrxQ have FrQ: "extractFrame([(x, y)] \ Q) = \A\<^sub>Q', \\<^sub>Q\" and "y \ extractFrame Q" + by(clarsimp simp add: alpha' eqvts frame.inject fresh_list_cons name_swap)+ + moreover from \A\<^sub>P \* Q\ have "([(x, y)] \ A\<^sub>P) \* ([(x, y)] \ Q)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \A\<^sub>P \* x\ \y \ A\<^sub>P\ have "A\<^sub>P \* ([(x, y)] \ Q)" by simp + moreover from \A\<^sub>Q' \* Q\ have "([(x, y)] \ A\<^sub>Q') \* ([(x, y)] \ Q)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \x \ A\<^sub>Q'\ \y \ A\<^sub>Q'\ have "A\<^sub>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\<^sub>P \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* A\<^sub>Q'\ \A\<^sub>Q' \* \\ \A\<^sub>Q' \* P\ \A\<^sub>P \* M\ \A\<^sub>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" + \ \Addition for Broadcast\ + 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\<^sub>Q \\<^sub>Q) + from \y \ xvec1\ \y \ xvec2\ have "y \ (xvec1@xvec2)" by(auto simp add: fresh_list_append) + with \\ \ \\<^sub>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\<^sub>P \\<^sub>P) + have FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" by fact + with \y \ P\ \A\<^sub>P \* y\ have "y \ \\<^sub>P" + apply(drule_tac extractFrameFresh) + by(simp add: frameResChainFresh) (simp add: fresh_def name_list_supp) + from \\ \ \\<^sub>P \ ([(x, y)] \ Q) \M\\*(xvec1@xvec2)\\N\ \ Q'\ \y \ supp N\ \y \ \\ \y \ \\<^sub>P\ \y \ M\ \y \ xvec1\ \y \ xvec2\ + have "\ \ \\<^sub>P \ \\y\(([(x, y)] \ Q)) \M\\*(xvec1@y#xvec2)\\N\ \ Q'" by(force intro: Open) + with \y \ Q\ have "\ \ \\<^sub>P \ \\x\Q \M\\*(xvec1@y#xvec2)\\N\ \ Q'" + by(simp add: alphaRes) + moreover from \A\<^sub>P \* ([(x, y)] \ Q)\ have "A\<^sub>P \* \\y\([(x, y)] \ Q)" by(auto simp add: fresh_star_def abs_fresh) + with \y \ Q\ have "A\<^sub>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\<^sub>P \* \\ \A\<^sub>P \* M\ \y \ P\ \A\<^sub>P \* (xvec1@xvec2)\ \A\<^sub>P \* y\ \A\<^sub>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\<^sub>Q \\<^sub>Q) + from \y \ xvec1\ \y \ xvec2\ have "y \ (xvec1@xvec2)" by(auto simp add: fresh_list_append) + with \\ \ \\<^sub>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\<^sub>P \\<^sub>P) + have FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" by fact + with \y \ P\ \A\<^sub>P \* y\ have "y \ \\<^sub>P" + apply(drule_tac extractFrameFresh) + by(simp add: frameResChainFresh) (simp add: fresh_def name_list_supp) + from \\ \ \\<^sub>P \ ([(x, y)] \ Q) \\M\\*(xvec1@xvec2)\\N\ \ Q'\ \y \ supp N\ \y \ \\ \y \ \\<^sub>P\ \y \ M\ \y \ xvec1\ \y \ xvec2\ + have "\ \ \\<^sub>P \ \\y\(([(x, y)] \ Q)) \\M\\*(xvec1@y#xvec2)\\N\ \ Q'" by(force intro: BrOpen) + with \y \ Q\ have "\ \ \\<^sub>P \ \\x\Q \\M\\*(xvec1@y#xvec2)\\N\ \ Q'" + by(simp add: alphaRes) + moreover from \A\<^sub>P \* ([(x, y)] \ Q)\ have "A\<^sub>P \* \\y\([(x, y)] \ Q)" by(auto simp add: fresh_star_def abs_fresh) + with \y \ Q\ have "A\<^sub>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\<^sub>P \* \\ \A\<^sub>P \* M\ \y \ P\ \A\<^sub>P \* (xvec1@xvec2)\ \A\<^sub>P \* y\ \A\<^sub>P \* N\ + by(intro Par2) auto + moreover have "(\, P \ Q', P \ Q') \ Rel" by(rule C1) + ultimately show ?case by blast + next + case(cBrComm1 \\<^sub>Q P' A\<^sub>P \\<^sub>P Q' A\<^sub>Q) + have FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" by fact + with \y \ P\ \A\<^sub>P \* y\ have "y \ \\<^sub>P" + apply(drule_tac extractFrameFresh) + by(simp add: frameResChainFresh) (simp add: fresh_def name_list_supp) + + from \extractFrame ([(x, y)] \ Q) = \A\<^sub>Q, \\<^sub>Q\\ + have "extractFrame (\\y\([(x, y)] \ Q)) = \(y#A\<^sub>Q), \\<^sub>Q\" + by simp + with \y \ Q\ + have FrQ: "extractFrame (\\x\Q) = \(y#A\<^sub>Q), \\<^sub>Q\" + by(simp add: alphaRes) + from \\ \ \\<^sub>P \ [(x, y)] \ Q \ \M\\*(xvec1 @ xvec2)\\N\ \ Q'\ \y \ supp N\ \y \ \\ \y \ \\<^sub>P\ \y \ M\ \y \ xvec1\ \y \ xvec2\ + have "\ \ \\<^sub>P \ \\y\([(x, y)] \ Q) \ \M\\*(xvec1@y#xvec2)\\N\ \ Q'" + by(force intro: BrOpen) + with \y \ Q\ have QTrans: "\ \ \\<^sub>P \ \\x\Q \\M\\*(xvec1@y#xvec2)\\N\ \ Q'" + by(simp add: alphaRes) + from \A\<^sub>P \* ([(x, y)] \ Q)\ have "A\<^sub>P \* \\y\([(x, y)] \ Q)" by(auto simp add: fresh_star_def abs_fresh) + with \y \ Q\ have "A\<^sub>P \* \\x\Q" by(simp add: alphaRes) + from \A\<^sub>Q \* ([(x, y)] \ Q)\ have "A\<^sub>Q \* \\y\([(x, y)] \ Q)" by(auto simp add: fresh_star_def abs_fresh) + with \y \ Q\ have "A\<^sub>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\<^sub>Q) \* (\\x\Q)" by simp + from \\ \ \\<^sub>Q \ P \ \M\N\ \ P'\ FrP QTrans FrQ + \A\<^sub>P \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* \\x\Q\ \A\<^sub>P \* M\ \A\<^sub>P \* y\ \A\<^sub>P \* A\<^sub>Q\ + \y \ \\ \A\<^sub>Q \* \\ \y \ P\ \A\<^sub>Q \* P\ \(y#A\<^sub>Q) \* \\x\Q\ + \y \ M\ \A\<^sub>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 \\<^sub>Q P' A\<^sub>P \\<^sub>P Q' A\<^sub>Q) + from \y \ xvec1\ \y \ xvec2\ have "y \ (xvec1@xvec2)" by(auto simp add: fresh_list_append) + with \\ \ \\<^sub>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\<^sub>Q \\<^sub>Q) + note \\ \ \\<^sub>Q \ P \\ \ P'\ + moreover with \x \ P\ \x \ \\ \bn \ \* subject \\ \distinct(bn \)\ + have "x \ P'" by(force dest: freeFreshDerivative) + with \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ have "extractFrame(\\x\Q) = \(x#A\<^sub>Q), \\<^sub>Q\" + by simp + moreover from \bn \ \* Q\ have "bn \ \* (\\x\Q)" by(simp add: fresh_star_def abs_fresh) + moreover from \x \ \\ \A\<^sub>Q \* \\ have "(x#A\<^sub>Q) \* \" by simp + moreover from \x \ P\ \A\<^sub>Q \* P\ have "(x#A\<^sub>Q) \* P" by simp + moreover from \x \ \\ \A\<^sub>Q \* \\ have "(x#A\<^sub>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\<^sub>P \\<^sub>P) + have FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" by fact + with \x \ P\ \A\<^sub>P \* x\ have "x \ \\<^sub>P" + apply(drule_tac extractFrameFresh) + by(simp add: frameResChainFresh) (simp add: fresh_def name_list_supp) + from \\ \ \\<^sub>P \ Q \\ \ Q'\ \x \ \\ \x \ \\<^sub>P\ \x \ \\ + have "\ \ \\<^sub>P \ \\x\Q \\ \ \\x\Q'" + by(intro Scope) auto + moreover note FrP \bn \ \* P\ \A\<^sub>P \* \\ + moreover from \A\<^sub>P \* Q\ have "A\<^sub>P \* \\x\Q" by(simp add: fresh_star_def abs_fresh) + ultimately have "\ \ P \ \\x\Q \\ \ (P \ \\x\Q')" using \A\<^sub>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 \\<^sub>Q M N P' A\<^sub>P \\<^sub>P K xvec Q' A\<^sub>Q) + have PTrans: "\ \ \\<^sub>Q \ P \M\N\ \ P'" and FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" by fact+ + have QTrans: "\ \ \\<^sub>P \ Q \K\\*xvec\\N\ \ Q'" and FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" by fact+ + from FrP \x \ P\ have "x \ \A\<^sub>P, \\<^sub>P\" + by(auto dest: extractFrameFresh) + with \A\<^sub>P \* x\ have "x \ \\<^sub>P" by(simp add: frameResChainFresh) (simp add: fresh_def name_list_supp) + from PTrans FrP \distinct A\<^sub>P\ \x \ P\ \A\<^sub>P \* \\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>P \* x\ \A\<^sub>P \* P\ \A\<^sub>P \* M\ \A\<^sub>Q \* P\ \A\<^sub>P \* A\<^sub>Q\ + obtain M' where MeqM': "(\ \ \\<^sub>Q) \ \\<^sub>P \ M \ M'" and "x \ M'" and "A\<^sub>Q \* M'" + by(elim inputObtainPrefix[where B="x#A\<^sub>Q"]) (assumption | force simp add: fresh_star_list_cons)+ + from MeqM' have MeqM': "\ \ \\<^sub>P \ \\<^sub>Q \ M \ M'" + by(metis statEqEnt Associativity Composition Commutativity) + with \\ \ \\<^sub>P \ \\<^sub>Q \ M \ K\ have "\ \ \\<^sub>P \ \\<^sub>Q \ K \ M'" + by(blast intro: chanEqTrans chanEqSym) + then have "(\ \ \\<^sub>P) \ \\<^sub>Q \ K \ M'" + by(metis statEqEnt Associativity Composition Commutativity) + with QTrans FrQ \distinct A\<^sub>Q\ have QTrans: "\ \ \\<^sub>P \ Q \M'\\*xvec\\N\ \ Q'" using \A\<^sub>Q \* \\ \A\<^sub>Q \* \\<^sub>P\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* K\ \A\<^sub>Q \* M'\ + by(elim outputRenameSubject) (assumption | force)+ + show ?case + proof(cases "x \ supp N") + note PTrans FrP + moreover assume "x \ supp N" + then have "\ \ \\<^sub>P \ \\x\Q \M'\\*([]@(x#xvec))\\N\ \ Q'" using QTrans \x \ \\ \x \ \\<^sub>P\ \x \ M'\ \xvec \* x\ + by(intro Open) (assumption | force simp add: fresh_list_nil)+ + then have QTrans: "\ \ \\<^sub>P \ \\x\Q \M'\\*(x#xvec)\\N\ \ Q'" by simp + moreover from \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ have "extractFrame(\\x\Q) = \(x#A\<^sub>Q), \\<^sub>Q\" + by simp + moreover note MeqM' \A\<^sub>P \* \\ \A\<^sub>P \* P\ + moreover from \A\<^sub>P \* Q\ have "A\<^sub>P \* (\\x\Q)" by(simp add: fresh_star_def abs_fresh) + moreover from \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* x\ have "A\<^sub>P \* (x#A\<^sub>Q)" + by(simp add: fresh_star_def fresh_list_cons) (auto simp add: fresh_def name_list_supp) + moreover from \x \ \\ \A\<^sub>Q \* \\ have "(x#A\<^sub>Q) \* \" by simp + moreover from \x \ P\ \A\<^sub>Q \* P\ have "(x#A\<^sub>Q) \* P" by simp + moreover from \x \ M'\ \A\<^sub>Q \* M'\ have "(x#A\<^sub>Q) \* M'" by simp + moreover from \A\<^sub>Q \* Q\ have "(x#A\<^sub>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\<^sub>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 \ \\<^sub>P\ \x \ M'\ \xvec \* x\ + have QTrans: "\ \ \\<^sub>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\<^sub>Q, \\<^sub>Q\\ have "extractFrame(\\x\Q) = \(x#A\<^sub>Q), \\<^sub>Q\" + by simp + moreover note MeqM' \A\<^sub>P \* \\ \A\<^sub>P \* P\ + moreover from \A\<^sub>P \* Q\ have "A\<^sub>P \* (\\x\Q)" by(simp add: fresh_star_def abs_fresh) + moreover from \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* x\ have "A\<^sub>P \* (x#A\<^sub>Q)" + by(simp add: fresh_star_def fresh_list_cons) (auto simp add: fresh_def name_list_supp) + moreover from \x \ \\ \A\<^sub>Q \* \\ have "(x#A\<^sub>Q) \* \" by simp + moreover from \x \ P\ \A\<^sub>Q \* P\ have "(x#A\<^sub>Q) \* P" by simp + moreover from \x \ M'\ \A\<^sub>Q \* M'\ have "(x#A\<^sub>Q) \* M'" by simp + moreover from \A\<^sub>Q \* Q\ have "(x#A\<^sub>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\<^sub>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 \\<^sub>Q M xvec N P' A\<^sub>P \\<^sub>P K Q' A\<^sub>Q) + have PTrans: "\ \ \\<^sub>Q \ P \M\\*xvec\\N\ \ P'" and FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" by fact+ + have QTrans: "\ \ \\<^sub>P \ Q \K\N\ \ Q'" and FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" by fact+ + from FrP \x \ P\ have "x \ \A\<^sub>P, \\<^sub>P\" + by(auto dest: extractFrameFresh) + with \A\<^sub>P \* x\ have "x \ \\<^sub>P" by(simp add: frameResChainFresh) (simp add: fresh_def name_list_supp) + from PTrans + have "\ \ \\<^sub>Q \ P \ ROut M (\\*xvec\N \' P')" + by(simp add: residualInject) + with FrP \distinct A\<^sub>P\ \x \ P\ \A\<^sub>P \* \\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>P \* x\ \A\<^sub>P \* P\ \A\<^sub>P \* M\ \A\<^sub>Q \* P\ \A\<^sub>P \* A\<^sub>Q\ \xvec \* M\ \distinct xvec\ + obtain M' where MeqM': "(\ \ \\<^sub>Q) \ \\<^sub>P \ M \ M'" and "x \ M'" and "A\<^sub>Q \* M'" + by(elim outputObtainPrefix[where B="x#A\<^sub>Q"]) (assumption | force simp add: fresh_star_list_cons)+ + from MeqM' have MeqM': "\ \ \\<^sub>P \ \\<^sub>Q \ M \ M'" + by(metis statEqEnt Associativity Commutativity Composition) + with \\ \ \\<^sub>P \ \\<^sub>Q \ M \ K\ have "\ \ \\<^sub>P \ \\<^sub>Q \ K \ M'" + by(blast intro: chanEqTrans chanEqSym) + then have "(\ \ \\<^sub>P) \ \\<^sub>Q \ K \ M'" + by(metis statEqEnt Associativity Commutativity Composition) + with QTrans FrQ \distinct A\<^sub>Q\ have QTrans: "\ \ \\<^sub>P \ Q \M'\N\ \ Q'" using \A\<^sub>Q \* \\ \A\<^sub>Q \* \\<^sub>P\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* K\ \A\<^sub>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 \ \\<^sub>P\ \x \ M'\ \x \ N\ have QTrans: "\ \ \\<^sub>P \ \\x\Q \M'\N\ \ \\x\Q'" + by(intro Scope) (assumption | force)+ + + note PTrans FrP QTrans + moreover with \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ have "extractFrame(\\x\Q) = \(x#A\<^sub>Q), \\<^sub>Q\" + by simp + moreover note MeqM' \A\<^sub>P \* \\ \A\<^sub>P \* P\ + moreover from \A\<^sub>P \* Q\ have "A\<^sub>P \* (\\x\Q)" by(simp add: fresh_star_def abs_fresh) + moreover from \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* x\ have "A\<^sub>P \* (x#A\<^sub>Q)" + by(simp add: fresh_star_def fresh_list_cons) (auto simp add: fresh_def name_list_supp) + moreover from \x \ \\ \A\<^sub>Q \* \\ have "(x#A\<^sub>Q) \* \" by simp + moreover from \x \ P\ \A\<^sub>Q \* P\ have "(x#A\<^sub>Q) \* P" by simp + moreover from \x \ M'\ \A\<^sub>Q \* M'\ have "(x#A\<^sub>Q) \* M'" by simp + moreover from \A\<^sub>Q \* Q\ have "(x#A\<^sub>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\<^sub>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 \\<^sub>Q M N P' A\<^sub>P \\<^sub>P Q' A\<^sub>Q) + have PTrans: "\ \ \\<^sub>Q \ P \\M\N\ \ P'" and FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" by fact+ + have QTrans: "\ \ \\<^sub>P \ Q \\M\N\ \ Q'" and FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" by fact+ + from FrP \x \ P\ have "x \ \A\<^sub>P, \\<^sub>P\" + by(auto dest: extractFrameFresh) + with \A\<^sub>P \* x\ have "x \ \\<^sub>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 \ \\<^sub>P\ \x \ M\ \x \ N\ have QTrans: "\ \ \\<^sub>P \ \\x\Q \\M\N\ \ \\x\Q'" + by(intro Scope) (assumption | force)+ + + note PTrans FrP QTrans + moreover with \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ have "extractFrame(\\x\Q) = \(x#A\<^sub>Q), \\<^sub>Q\" + by simp + moreover note \A\<^sub>P \* \\ \A\<^sub>P \* P\ + moreover from \A\<^sub>P \* Q\ have "A\<^sub>P \* (\\x\Q)" by(simp add: fresh_star_def abs_fresh) + moreover from \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* x\ have "A\<^sub>P \* (x#A\<^sub>Q)" + by(simp add: fresh_star_def fresh_list_cons) (auto simp add: fresh_def name_list_supp) + moreover from \x \ \\ \A\<^sub>Q \* \\ have "(x#A\<^sub>Q) \* \" by simp + moreover from \x \ P\ \A\<^sub>Q \* P\ have "(x#A\<^sub>Q) \* P" by simp + moreover from \x \ M\ \A\<^sub>Q \* M\ have "(x#A\<^sub>Q) \* M" by simp + moreover from \A\<^sub>Q \* Q\ have "(x#A\<^sub>Q) \* (\\x\Q)" by(simp add: fresh_star_def abs_fresh) + ultimately have Trans: "\ \ P \ \\x\Q \\M\N\ \ (P' \ \\x\Q')" using \A\<^sub>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 \\<^sub>Q M N P' A\<^sub>P \\<^sub>P xvec Q' A\<^sub>Q) + from \x \ \\ \\M\\*xvec\\N\ = \\ have "x \ M" and "x \ xvec" and "x \ N" by force+ + have PTrans: "\ \ \\<^sub>Q \ P \\M\N\ \ P'" and FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" by fact+ + have QTrans: "\ \ \\<^sub>P \ Q \\M\\*xvec\\N\ \ Q'" and FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" by fact+ + from FrP \x \ P\ have "x \ \A\<^sub>P, \\<^sub>P\" + by(auto dest: extractFrameFresh) + with \A\<^sub>P \* x\ have "x \ \\<^sub>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 \ \\<^sub>P\ \x \ M\ \x \ N\ \x \ xvec\ + have QTrans: "\ \ \\<^sub>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\<^sub>Q, \\<^sub>Q\\ have "extractFrame(\\x\Q) = \(x#A\<^sub>Q), \\<^sub>Q\" + by simp + moreover note \A\<^sub>P \* \\ \A\<^sub>P \* P\ + moreover from \A\<^sub>P \* Q\ have "A\<^sub>P \* (\\x\Q)" by(simp add: fresh_star_def abs_fresh) + moreover from \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* x\ have "A\<^sub>P \* (x#A\<^sub>Q)" + by(simp add: fresh_star_def fresh_list_cons) (auto simp add: fresh_def name_list_supp) + moreover from \x \ \\ \A\<^sub>Q \* \\ have "(x#A\<^sub>Q) \* \" by simp + moreover from \x \ P\ \A\<^sub>Q \* P\ have "(x#A\<^sub>Q) \* P" by simp + moreover from \x \ M\ \A\<^sub>Q \* M\ have "(x#A\<^sub>Q) \* M" by simp + moreover from \A\<^sub>Q \* Q\ have "(x#A\<^sub>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\<^sub>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 \\<^sub>Q M xvec N P' A\<^sub>P \\<^sub>P Q' A\<^sub>Q) + from \x \ \\ \\M\\*xvec\\N\ = \\ have "x \ M" and "x \ xvec" and "x \ N" by force+ + have PTrans: "\ \ \\<^sub>Q \ P \\M\\*xvec\\N\ \ P'" and FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" by fact+ + have QTrans: "\ \ \\<^sub>P \ Q \\M\N\ \ Q'" and FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" by fact+ + from FrP \x \ P\ have "x \ \A\<^sub>P, \\<^sub>P\" + by(auto dest: extractFrameFresh) + with \A\<^sub>P \* x\ have "x \ \\<^sub>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 \ \\<^sub>P\ \x \ M\ \x \ N\ have QTrans: "\ \ \\<^sub>P \ \\x\Q \\M\N\ \ \\x\Q'" + by(intro Scope) (assumption | force)+ + + note PTrans FrP QTrans + moreover with \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ have "extractFrame(\\x\Q) = \(x#A\<^sub>Q), \\<^sub>Q\" + by simp + moreover note \A\<^sub>P \* \\ \A\<^sub>P \* P\ + moreover from \A\<^sub>P \* Q\ have "A\<^sub>P \* (\\x\Q)" by(simp add: fresh_star_def abs_fresh) + moreover from \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* x\ have "A\<^sub>P \* (x#A\<^sub>Q)" + by(simp add: fresh_star_def fresh_list_cons) (auto simp add: fresh_def name_list_supp) + moreover from \x \ \\ \A\<^sub>Q \* \\ have "(x#A\<^sub>Q) \* \" by simp + moreover from \x \ P\ \A\<^sub>Q \* P\ have "(x#A\<^sub>Q) \* P" by simp + moreover from \x \ M\ \A\<^sub>Q \* M\ have "(x#A\<^sub>Q) \* M" by simp + moreover from \A\<^sub>Q \* Q\ have "(x#A\<^sub>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\<^sub>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\<^sub>Q \\<^sub>Q) + from \\ \ \\<^sub>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\<^sub>P \\<^sub>P) + from \A\<^sub>P \* x\ have "x \ A\<^sub>P" by simp + with \x \ P\ \extractFrame P = \A\<^sub>P, \\<^sub>P\\ + have "x \ \\<^sub>P" + apply(drule_tac extractFrameFresh) + by(simp add: frameResChainFresh) (simp add: fresh_def name_list_supp) + from \x \ \\ \x \ \\<^sub>P\ have "x \ (\ \ \\<^sub>P)" by force + from \\ \ \\<^sub>P \ Q \ \M\\*xvec\\N\ \ Q'\ \x \ supp M\ \x \ (\ \ \\<^sub>P)\ + have QTrans: "\ \ \\<^sub>P \ \\x\Q \ \ \ \\x\(\\*xvec\Q')" + by(rule BrClose) + with \extractFrame P = \A\<^sub>P, \\<^sub>P\\ \A\<^sub>P \* \\ \x \ A\<^sub>P\ \A\<^sub>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 \\<^sub>Q P' A\<^sub>P \\<^sub>P Q' A\<^sub>Q) + from \\ \ \\<^sub>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 \\<^sub>Q P' A\<^sub>P \\<^sub>P Q' A\<^sub>Q) + from \\ \ \\<^sub>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\<^sub>P \\<^sub>P) + from \\ \ \\<^sub>P \ Q \\\ Q'\ \extractFrame P = \A\<^sub>P, \\<^sub>P\\ \bn \ \* P\ \A\<^sub>P \* \\ \A\<^sub>P \* Q\ \A\<^sub>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\<^sub>Q \\<^sub>Q) + from \\ \ \\<^sub>Q \ P \\ \ P'\ \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ \bn \ \* Q\ \A\<^sub>Q \* \\ \A\<^sub>Q \* P\ \A\<^sub>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 \\<^sub>P M N Q' A\<^sub>Q \\<^sub>Q K xvec P' A\<^sub>P) + note \\ \ \\<^sub>Q \ P \K\\*xvec\\N\ \ P'\ \extractFrame P = \A\<^sub>P, \\<^sub>P\\ + \\ \ \\<^sub>P \ Q \M\N\ \ Q'\ \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ + moreover from \\ \ \\<^sub>Q \ \\<^sub>P \ M \ K\ + have "\ \ \\<^sub>Q \ \\<^sub>P \ K \ M" + by(rule chanEqSym) + then have "\ \ \\<^sub>P \ \\<^sub>Q \ K \ M" + by(blast intro: statEqEnt compositionSym Commutativity) + ultimately have "\ \ P \ Q \\ \ \\*xvec\(P' \ Q')" + using \A\<^sub>P \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* Q\ \A\<^sub>Q \* A\<^sub>P\ \A\<^sub>Q \* \\ \A\<^sub>Q \* P\ \A\<^sub>Q \* Q\ \xvec \* Q\ \A\<^sub>P \* K\ \A\<^sub>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 \\<^sub>P M xvec N Q' A\<^sub>Q \\<^sub>Q K P' A\<^sub>P) + note \\ \ \\<^sub>Q \ P \K\N\ \ P'\ \extractFrame P = \A\<^sub>P, \\<^sub>P\\ + \\ \ \\<^sub>P \ Q \M\\*xvec\\N\ \ Q'\ \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ + moreover from \\ \ \\<^sub>Q \ \\<^sub>P \ M \ K\ + have "\ \ \\<^sub>Q \ \\<^sub>P \ K \ M" + by(rule chanEqSym) + then have "\ \ \\<^sub>P \ \\<^sub>Q \ K \ M" + by(blast intro: statEqEnt compositionSym Commutativity) + ultimately have "\ \ P \ Q \\ \ \\*xvec\(P' \ Q')" + using \A\<^sub>P \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* Q\ \A\<^sub>Q \* A\<^sub>P\ \A\<^sub>Q \* \\ \A\<^sub>Q \* P\ \A\<^sub>Q \* Q\ \xvec \* P\ \A\<^sub>P \* K\ \A\<^sub>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 \\<^sub>P M N Q' A\<^sub>Q \\<^sub>Q P' A\<^sub>P) + then have "\ \ P \ Q \ \M\N\ \ P' \ Q'" + by(intro BrMerge) (assumption|simp)+ + then show ?case by(blast intro: C1) + next + case(cBrComm1 \\<^sub>P M N Q' A\<^sub>Q \\<^sub>Q xvec P' A\<^sub>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 \\<^sub>P M xvec N Q' A\<^sub>Q \\<^sub>Q P' A\<^sub>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 diff --git a/thys/Broadcast_Psi/Simulation.thy b/thys/Broadcast_Psi/Simulation.thy new file mode 100644 --- /dev/null +++ b/thys/Broadcast_Psi/Simulation.thy @@ -0,0 +1,485 @@ +theory Simulation + imports Semantics +begin + +text \This file is a (heavily modified) variant of the theory {\it Psi\_Calculi.Simulation} +from~\cite{DBLP:journals/afp/Bengtson12}.\ + +context env begin + +definition + "simulation" :: "'b \ ('a, 'b, 'c) psi \ + ('b \ ('a, 'b, 'c) psi \ ('a, 'b, 'c) psi) set \ + ('a, 'b, 'c) psi \ bool" ("_ \ _ \[_] _" [80, 80, 80, 80] 80) + where + "\ \ P \[Rel] Q \ \\ Q'. \ \ Q \\ \ Q' \ bn \ \* \ \ bn \ \* P \ (\P'. \ \ P \\ \ P' \ (\, P', Q') \ Rel)" + +abbreviation + simulationNilJudge ("_ \[_] _" [80, 80, 80] 80) where "P \[Rel] Q \ SBottom' \ P \[Rel] Q" + +lemma simI[consumes 1, case_names cSim]: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and Rel :: "('b \ ('a, 'b, 'c) psi \ ('a, 'b, 'c) psi) set" + and Q :: "('a, 'b, 'c) psi" + and C :: "'d::fs_name" + +assumes Eqvt: "eqvt Rel" + and Sim: "\\ Q'. \\ \ Q \\ \ Q'; bn \ \* P; bn \ \* Q; bn \ \* \; distinct(bn \); + bn \ \* (subject \); bn \ \* C\ \ \P'. \ \ P \\ \ P' \ (\, P', Q') \ Rel" + +shows "\ \ P \[Rel] Q" +proof - + { + fix \ Q' + assume "\ \ Q \\ \ Q'" and "bn \ \* \" and "bn \ \* P" + then have "\P'. \ \ P \\ \ P' \ (\, P', Q') \ Rel" + proof(nominal_induct \ rule: action.strong_induct) + case(In M N) + then show ?case by(auto simp add: Sim) + next + case(BrIn M N) + then show ?case by(auto simp add: Sim) + next + case(Out M xvec N) + moreover { + fix M xvec N Q' + assume "(xvec::name list) \* \" and "xvec \* P" + obtain p where xvecFreshPsi: "((p::name prm) \ (xvec::name list)) \* \" + and xvecFreshM: "(p \ xvec) \* (M::'a)" + and xvecFreshN: "(p \ xvec) \* (N::'a)" + and xvecFreshP: "(p \ xvec) \* P" + and xvecFreshQ: "(p \ xvec) \* Q" + and xvecFreshQ': "(p \ xvec) \* (Q'::('a, 'b, 'c) psi)" + and xvecFreshC: "(p \ xvec) \* C" + and xvecFreshxvec: "(p \ xvec) \* xvec" + and S: "(set p) \ (set xvec) \ (set(p \ xvec))" + and dpr: "distinctPerm p" + by(rule name_list_avoiding[where xvec="xvec" and c="(\, M, Q, N, P, Q', xvec, C)"]) auto + from \(p \ xvec) \* M\ \distinctPerm p\ have "xvec \* (p \ M)" + by(subst pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst, where pi=p, symmetric]) simp + + assume Trans: "\ \ Q \M\\*xvec\\N\ \ Q'" + with xvecFreshN xvecFreshQ' S + have "\ \ Q \M\\*(p \ xvec)\\(p \ N)\ \ (p \ Q')" + by(simp add: boundOutputChainAlpha'' residualInject) + moreover then have "distinct(p \ xvec)" by(auto dest: boundOutputDistinct) + + moreover note xvecFreshPsi xvecFreshP xvecFreshQ xvecFreshM xvecFreshC + ultimately obtain P' where PTrans: "\ \ P \M\\*(p \ xvec)\\(p \ N)\ \ P'" + and P'RelQ': "(\, P', p \ Q') \ Rel" + using Sim by (metis bn.simps(3) optionFreshChain(1) subject.simps(3)) + then have "(p \ \) \ (p \ P) \(p \ (M\\*(p \ xvec)\\(p \ N)\ \ P'))" + by(simp add: semantics.eqvt) + with \xvec \* \\ xvecFreshPsi \xvec \* P\ xvecFreshP S dpr + have "\ \ P \(p \ M)\\*xvec\\N\ \ (p \ P')" + by(simp add: eqvts name_set_fresh_fresh) + with \xvec \* \\ xvecFreshPsi \xvec \* P\ xvecFreshP S \xvec \* (p \ M)\ + have "\ \ P \(p \ p \ M)\\*xvec\\N\ \ (p \ P')" + by(simp add: outputPermSubject) + + with dpr have "\ \ P \M\\*xvec\\N\ \ (p \ P')" + by simp + + moreover from P'RelQ' Eqvt have "(p \ \, p \ P', p \ p \ Q') \ Rel" + apply(simp add: eqvt_def eqvts) + apply(erule ballE[where x="(\, P', p \ Q')"]) + apply(erule allE[where x="p"]) + by(auto simp add: eqvts) + + with \xvec \* \\ xvecFreshPsi S dpr have "(\, p \ P', Q') \ Rel" + by simp + ultimately have "\P'. \ \ P \ M\\*xvec\\N\ \ P' \ (\, P', Q') \ Rel" + by blast + } + ultimately show ?case by force + next + case(BrOut M xvec N) + moreover { + fix M xvec N Q' + assume "(xvec::name list) \* \" and "xvec \* P" + obtain p where xvecFreshPsi: "((p::name prm) \ (xvec::name list)) \* \" + and xvecFreshM: "(p \ xvec) \* (M::'a)" + and xvecFreshN: "(p \ xvec) \* (N::'a)" + and xvecFreshP: "(p \ xvec) \* P" + and xvecFreshQ: "(p \ xvec) \* Q" + and xvecFreshQ': "(p \ xvec) \* (Q'::('a, 'b, 'c) psi)" + and xvecFreshC: "(p \ xvec) \* C" + and xvecFreshxvec: "(p \ xvec) \* xvec" + and S: "(set p) \ (set xvec) \ (set(p \ xvec))" + and dpr: "distinctPerm p" + by(rule name_list_avoiding[where xvec="xvec" and c="(\, M, Q, N, P, Q', xvec, C)"]) auto + + from \(p \ xvec) \* M\ \distinctPerm p\ have "xvec \* (p \ M)" + by(subst pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst, where pi=p, symmetric]) simp + + assume Trans: "\ \ Q \\M\\*xvec\\N\ \ Q'" + with xvecFreshN xvecFreshQ' S + have "\ \ Q \\M\\*(p \ xvec)\\(p \ N)\ \ (p \ Q')" + by(simp add: boundOutputChainAlpha'' residualInject) + moreover then have "distinct(p \ xvec)" + by(auto dest: boundOutputDistinct) + + moreover note xvecFreshPsi xvecFreshP xvecFreshQ xvecFreshM xvecFreshC + ultimately obtain P' where PTrans: "\ \ P \\M\\*(p \ xvec)\\(p \ N)\ \ P'" + and P'RelQ': "(\, P', p \ Q') \ Rel" + by(auto dest: Sim) + then have "(p \ \) \ (p \ P) \(p \ (\M\\*(p \ xvec)\\(p \ N)\ \ P'))" + by(metis semantics.eqvt) + with \xvec \* \\ xvecFreshPsi \xvec \* P\ xvecFreshP S dpr + have "\ \ P \\(p \ M)\\*xvec\\N\ \ (p \ P')" + by(simp add: eqvts name_set_fresh_fresh) + with \xvec \* \\ xvecFreshPsi \xvec \* P\ xvecFreshP S \xvec \* (p \ M)\ + have "\ \ P \\(p \ p \ M)\\*xvec\\N\ \ (p \ P')" + by(simp add: broutputPermSubject) + + with dpr have "\ \ P \\M\\*xvec\\N\ \ (p \ P')" + by simp + + moreover from P'RelQ' Eqvt have "(p \ \, p \ P', p \ p \ Q') \ Rel" + apply(simp add: eqvt_def eqvts) + apply(erule ballE[where x="(\, P', p \ Q')"]) + apply(erule allE[where x="p"]) + by(auto simp add: eqvts) + + with \xvec \* \\ xvecFreshPsi S dpr have "(\, p \ P', Q') \ Rel" + by simp + ultimately have "\P'. \ \ P \ \M\\*xvec\\N\ \ P' \ (\, P', Q') \ Rel" + by blast + } + ultimately show ?case by force + next + case Tau + then show ?case by (auto simp add: Sim) + qed + } + then show ?thesis + by (auto simp add: simulation_def) +qed + +lemma simI2[case_names cSim]: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and Rel :: "('b \ ('a, 'b, 'c) psi \ ('a, 'b, 'c) psi) set" + and Q :: "('a, 'b, 'c) psi" + and C :: "'d::fs_name" + +assumes Sim: "\\ Q'. \\ \ Q \\ \ Q'; bn \ \* P; bn \ \* \; distinct(bn \)\ \ \P'. \ \ P \\ \ P' \ (\, P', Q') \ Rel" + +shows "\ \ P \[Rel] Q" + using assms + by(auto simp add: simulation_def dest: boundOutputDistinct) + +lemma simIChainFresh[consumes 4, case_names cSim]: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and Rel :: "('b \ ('a, 'b, 'c) psi \ ('a, 'b, 'c) psi) set" + and Q :: "('a, 'b, 'c) psi" + and xvec :: "name list" + and C :: "'d::fs_name" + +assumes Eqvt: "eqvt Rel" + and "xvec \* \" + and "xvec \* P" + and "xvec \* Q" + and Sim: "\\ Q'. \\ \ Q \\ \ Q'; bn \ \* P; bn \ \* Q; bn \ \* \; + bn \ \* subject \; distinct(bn \); bn \ \* C; xvec \* \; xvec \* Q'\ \ + \P'. \ \ P \\ \ P' \ (\, P', Q') \ Rel" +shows "\ \ P \[Rel] Q" + using \eqvt Rel\ +proof(induct rule: simI[where C="(xvec, C)"]) + case(cSim \ Q') + from \bn \ \* (xvec, C)\ have "bn \ \* xvec" and "bn \ \* C" by simp+ + obtain p::"name prm" where "(p \ xvec) \* \" and "(p \ xvec) \* P" and "(p \ xvec) \* Q" + and "(p \ xvec) \* \" and S: "set p \ set xvec \ set(p \ xvec)" + and "distinctPerm p" + by(rule name_list_avoiding[where c="(\, P, Q, \)" and xvec=xvec]) auto + show ?case + proof(cases rule: actionCases[where \=\]) + case(cInput M N) + from \\ \ Q \\ \ Q'\ \\=M\N\\ have "(p \ \) \ (p \ Q) \(p \ (M\N\ \ Q'))" + by(auto intro: semantics.eqvt) + with \xvec \* \\ \(p \ xvec) \* \\ \xvec \* Q\ \(p \ xvec) \* Q\ S + have QTrans: "\ \ Q \(p \ M)\(p \ N)\ \ (p \ Q')" + by(simp add: eqvts) + moreover from \(p \ xvec) \* \\ have "(p \ (p \ xvec)) \* (p \ \)" + by(simp only: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \distinctPerm p\ \\=M\N\\ have "xvec \* (p \ M)" and "xvec \* (p \ N)" by simp+ + moreover with QTrans \xvec \* Q\ have "xvec \* (p \ Q')" + by(metis inputFreshChainDerivative) + ultimately have "\P'. \ \ P \ (p \ M)\(p \ N)\ \ P' \ (\, P', (p \ Q')) \ Rel" + by(simp add: Sim) + then obtain P' where PTrans: "\ \ P \ (p \ M)\(p \ N)\ \ P'" and P'RelQ': "(\, P', (p \ Q')) \ Rel" + by blast + from PTrans have "(p \ \) \ (p \ P) \ (p \ ((p \ M)\(p \ N)\ \ P'))" + by(rule semantics.eqvt) + with \xvec \* \\ \(p \ xvec) \* \\ \xvec \* P\ \(p \ xvec) \* P\ S \distinctPerm p\ + have "\ \ P \ M\N\ \ (p \ P')" by(simp add: eqvts) + moreover from P'RelQ' \eqvt Rel\ have "(p \ \, p \ P', p \ p \ Q') \ Rel" + by(auto simp add: eqvt_def) + with \xvec \* \\ \(p \ xvec) \* \\ S \distinctPerm p\ + have "(\, p \ P', Q') \ Rel" by simp + ultimately show ?thesis using \\=M\N\\ by blast + next + case(cBrInput M N) + from \\ \ Q \\ \ Q'\ \\=\M\N\\ have "(p \ \) \ (p \ Q) \(p \ (\M\N\ \ Q'))" + by(auto intro: semantics.eqvt) + with \xvec \* \\ \(p \ xvec) \* \\ \xvec \* Q\ \(p \ xvec) \* Q\ S + have QTrans: "\ \ Q \\(p \ M)\(p \ N)\ \ (p \ Q')" + by(simp add: eqvts) + moreover from \(p \ xvec) \* \\ have "(p \ (p \ xvec)) \* (p \ \)" + by(simp only: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \distinctPerm p\ \\=\M\N\\ have "xvec \* (p \ M)" and "xvec \* (p \ N)" by simp+ + moreover with QTrans \xvec \* Q\ have "xvec \* (p \ Q')" by(metis brinputFreshChainDerivative) + ultimately have "\P'. \ \ P \ \(p \ M)\(p \ N)\ \ P' \ (\, P', (p \ Q')) \ Rel" + by(simp add: Sim) + then obtain P' where PTrans: "\ \ P \ \(p \ M)\(p \ N)\ \ P'" and P'RelQ': "(\, P', (p \ Q')) \ Rel" + by blast + from PTrans have "(p \ \) \ (p \ P) \ (p \ (\(p \ M)\(p \ N)\ \ P'))" + by(rule semantics.eqvt) + with \xvec \* \\ \(p \ xvec) \* \\ \xvec \* P\ \(p \ xvec) \* P\ S \distinctPerm p\ + have "\ \ P \ \M\N\ \ (p \ P')" by(simp add: eqvts) + moreover from P'RelQ' \eqvt Rel\ have "(p \ \, p \ P', p \ p \ Q') \ Rel" + by(auto simp add: eqvt_def) + with \xvec \* \\ \(p \ xvec) \* \\ S \distinctPerm p\ + have "(\, p \ P', Q') \ Rel" by simp + ultimately show ?thesis using \\=\M\N\\ by blast + next + case(cOutput M yvec N) + from \distinct(bn \)\ \bn \ \* subject \\ \\=M\\*yvec\\N\\ have "distinct yvec" and "yvec \* M" by simp+ + from \\ \ Q \\ \ Q'\ \\=M\\*yvec\\N\\ have "(p \ \) \ (p \ Q) \(p \ (M\\*yvec\\N\ \ Q'))" + by(auto intro: semantics.eqvt) + with \xvec \* \\ \(p \ xvec) \* \\ \xvec \* Q\ \(p \ xvec) \* Q\ S + have QTrans: "\ \ Q \(p \ M)\\*(p \ yvec)\\(p \ N)\ \ (p \ Q')" + by(simp add: eqvts) + with S \bn \ \* xvec\ \(p \ xvec) \* \\ \\=M\\*yvec\\N\\ have "\ \ Q \(p \ M)\\*yvec\\(p \ N)\ \ (p \ Q')" + by simp + moreover from \(p \ xvec) \* \\ have "(p \ (p \ xvec)) \* (p \ \)" + by(simp only: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \distinctPerm p\ \\=M\\*yvec\\N\\ have "xvec \* (p \ M)" and "xvec \* (p \ N)" and "xvec \* (p \ yvec)" by simp+ + moreover with QTrans \xvec \* Q\ \distinct yvec\ \yvec \* M\ have "xvec \* (p \ Q')" + apply(drule_tac outputFreshChainDerivative(2)) + by(auto simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + moreover from \yvec \* M\ S \bn \ \* xvec\ \(p \ xvec) \* \\ \\=M\\*yvec\\N\\ \distinctPerm p\ + have "yvec \* (p \ M)" by(subst pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst, symmetric, where pi=p]) simp + ultimately have "\P'. \ \ P \ (p \ M)\\*yvec\\(p \ N)\ \ P' \ (\, P', (p \ Q')) \ Rel" + using \bn \ \* \\ \bn \ \* P\ \bn \ \* Q\\bn \ \* xvec\ \bn \ \* C\ \yvec \* M\ \\=M\\*yvec\\N\\ \distinct yvec\ + by(simp add: Sim) + then obtain P' where PTrans: "\ \ P \ (p \ M)\\*yvec\\(p \ N)\ \ P'" and P'RelQ': "(\, P', (p \ Q')) \ Rel" + by blast + from PTrans have "(p \ \) \ (p \ P) \ (p \ ((p \ M)\\*yvec\\(p \ N)\ \ P'))" + by(rule semantics.eqvt) + with \xvec \* \\ \(p \ xvec) \* \\ \xvec \* P\ \(p \ xvec) \* P\ S \distinctPerm p\ \bn \ \* xvec\ \(p \ xvec) \* \\ \\=M\\*yvec\\N\\ + have "\ \ P \ M\\*yvec\\N\ \ (p \ P')" by(simp add: eqvts) + moreover from P'RelQ' \eqvt Rel\ have "(p \ \, p \ P', p \ p \ Q') \ Rel" + by(auto simp add: eqvt_def) + with \xvec \* \\ \(p \ xvec) \* \\ S \distinctPerm p\ + have "(\, p \ P', Q') \ Rel" by simp + ultimately show ?thesis using \\=M\\*yvec\\N\\ by blast + next + case(cBrOutput M yvec N) + from \distinct(bn \)\ \bn \ \* subject \\ \\=\M\\*yvec\\N\\ have "distinct yvec" and "yvec \* M" by simp+ + from \\ \ Q \\ \ Q'\ \\=\M\\*yvec\\N\\ have "(p \ \) \ (p \ Q) \(p \ (\M\\*yvec\\N\ \ Q'))" + by(auto intro: semantics.eqvt) + with \xvec \* \\ \(p \ xvec) \* \\ \xvec \* Q\ \(p \ xvec) \* Q\ S + have QTrans: "\ \ Q \\(p \ M)\\*(p \ yvec)\\(p \ N)\ \ (p \ Q')" + by(simp add: eqvts) + with S \bn \ \* xvec\ \(p \ xvec) \* \\ \\=\M\\*yvec\\N\\ have "\ \ Q \\(p \ M)\\*yvec\\(p \ N)\ \ (p \ Q')" + by simp + moreover from \(p \ xvec) \* \\ have "(p \ (p \ xvec)) \* (p \ \)" + by(simp only: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + with \distinctPerm p\ \\=\M\\*yvec\\N\\ have "xvec \* (p \ M)" and "xvec \* (p \ N)" and "xvec \* (p \ yvec)" by simp+ + moreover with QTrans \xvec \* Q\ \distinct yvec\ \yvec \* M\ have "xvec \* (p \ Q')" + apply(drule_tac broutputFreshChainDerivative(2)) + by (auto simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) + moreover from \yvec \* M\ S \bn \ \* xvec\ \(p \ xvec) \* \\ \\=\M\\*yvec\\N\\ \distinctPerm p\ + have "yvec \* (p \ M)" by(subst pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst, symmetric, where pi=p]) simp + ultimately have "\P'. \ \ P \ \(p \ M)\\*yvec\\(p \ N)\ \ P' \ (\, P', (p \ Q')) \ Rel" + using \bn \ \* \\ \bn \ \* P\ \bn \ \* Q\\bn \ \* xvec\ \bn \ \* C\ \yvec \* M\ \\=\M\\*yvec\\N\\ \distinct yvec\ + by(simp add: Sim) + then obtain P' where PTrans: "\ \ P \ \(p \ M)\\*yvec\\(p \ N)\ \ P'" and P'RelQ': "(\, P', (p \ Q')) \ Rel" + by blast + from PTrans have "(p \ \) \ (p \ P) \ (p \ (\(p \ M)\\*yvec\\(p \ N)\ \ P'))" + by(rule semantics.eqvt) + with \xvec \* \\ \(p \ xvec) \* \\ \xvec \* P\ \(p \ xvec) \* P\ S \distinctPerm p\ \bn \ \* xvec\ \(p \ xvec) \* \\ \\=\M\\*yvec\\N\\ + have "\ \ P \ \M\\*yvec\\N\ \ (p \ P')" by(simp add: eqvts) + moreover from P'RelQ' \eqvt Rel\ have "(p \ \, p \ P', p \ p \ Q') \ Rel" + by(auto simp add: eqvt_def) + with \xvec \* \\ \(p \ xvec) \* \\ S \distinctPerm p\ + have "(\, p \ P', Q') \ Rel" by simp + ultimately show ?thesis using \\=\M\\*yvec\\N\\ by blast + next + case cTau + from \\ \ Q \\ \ Q'\ \\ = \\ \xvec \* Q\ have "xvec \* Q'" + by(blast dest: tauFreshChainDerivative) + with \\ \ Q \\ \ Q'\ \\ = \\ + show ?thesis + using Sim by simp + qed +qed + +lemma simIFresh[consumes 4, case_names cSim]: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and Rel :: "('b \ ('a, 'b, 'c) psi \ ('a, 'b, 'c) psi) set" + and Q :: "('a, 'b, 'c) psi" + and x :: name + and C :: "'d::fs_name" + +assumes Eqvt: "eqvt Rel" + and "x \ \" + and "x \ P" + and "x \ Q" + and "\\ Q'. \\ \ Q \\ \ Q'; bn \ \* P; bn \ \* Q; bn \ \* \; + bn \ \* subject \; distinct(bn \); bn \ \* C; x \ \; x \ Q'\ \ + \P'. \ \ P \\ \ P' \ (\, P', Q') \ Rel" + +shows "\ \ P \[Rel] Q" + using assms + by (auto intro: simIChainFresh[where xvec="[x]" and C=C]) + +lemma simE: + fixes F :: 'b + and P :: "('a, 'b, 'c) psi" + and Rel :: "('b \ ('a, 'b, 'c) psi \ ('a, 'b, 'c) psi) set" + and Q :: "('a, 'b, 'c) psi" + +assumes "\ \ P \[Rel] Q" + +shows "\\ Q'. \\ \ Q \\ \ Q'; bn \ \* \; bn \ \* P\ \ \P'. \ \ P \\ \ P' \ (\, P', Q') \ Rel" + using assms + by(auto simp add: simulation_def) + +lemma simClosedAux: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and Rel :: "('b \ ('a, 'b, 'c) psi \ ('a, 'b, 'c) psi) set" + and Q :: "('a, 'b, 'c) psi" + and p :: "name prm" + +assumes EqvtRel: "eqvt Rel" + and PSimQ: "\ \ P \[Rel] Q" + +shows "(p \ \) \ (p \ P) \[Rel] (p \ Q)" + using EqvtRel +proof(induct rule: simI[of _ _ _ _ "(\, P, p)"]) + case(cSim \ Q') + from \p \ \ \ p \ Q \\ \ Q'\ + have "(rev p \ p \ \) \ (rev p \ p \ Q) \(rev p \ (\ \ Q'))" + by(blast dest: semantics.eqvt) + then have "\ \ Q \(rev p \ \) \ (rev p \ Q')" + by(simp add: eqvts) + moreover with \bn \ \* (\, P, p)\ have "bn \ \* \" and "bn \ \* P" and "bn \ \* p" by simp+ + ultimately obtain P' where PTrans: "\ \ P \(rev p \ \) \ P'" + and P'RelQ': "(\, P', rev p \ Q') \ Rel" + using PSimQ + by(force dest: simE freshChainPermSimp simp add: eqvts) + from PTrans have "(p \ \) \ (p \ P) \(p \ ((rev p \ \) \ P'))" + by(rule semantics.eqvt) + with \bn \ \* p\ have "(p \ \) \ (p \ P) \\ \ (p \ P')" + by(simp add: eqvts freshChainPermSimp) + moreover from P'RelQ' EqvtRel have "(p \ (\, P', rev p \ Q')) \ Rel" + by(simp only: eqvt_def) + then have "(p \ \, p \ P', Q') \ Rel" by simp + ultimately show ?case by blast +qed + +lemma simClosed: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and Rel :: "('b \ ('a, 'b, 'c) psi \ ('a, 'b, 'c) psi) set" + and Q :: "('a, 'b, 'c) psi" + and p :: "name prm" + +assumes EqvtRel: "eqvt Rel" + +shows "\ \ P \[Rel] Q \ (p \ \) \ (p \ P) \[Rel] (p \ Q)" + and "P \[Rel] Q \ (p \ P) \[Rel] (p \ Q)" + using EqvtRel + by(force dest: simClosedAux simp add: permBottom)+ + +lemma reflexive: + fixes Rel :: "('b \ ('a, 'b, 'c) psi \ ('a, 'b, 'c) psi) set" + and \ :: 'b + and P :: "('a, 'b, 'c) psi" + +assumes "{(\, P, P) | \ P. True} \ Rel" + +shows "\ \ P \[Rel] P" + using assms + by(auto simp add: simulation_def) + +lemma transitive: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and Rel :: "('b \ ('a, 'b, 'c) psi \ ('a, 'b, 'c) psi) set" + and Q :: "('a, 'b, 'c) psi" + and Rel' :: "('b \ ('a, 'b, 'c) psi \ ('a, 'b, 'c) psi) set" + and R :: "('a, 'b, 'c) psi" + and Rel'' :: "('b \ ('a, 'b, 'c) psi \ ('a, 'b, 'c) psi) set" + +assumes PSimQ: "\ \ P \[Rel] Q" + and QSimR: "\ \ Q \[Rel'] R" + and Eqvt: "eqvt Rel''" + and Set: "{(\, P, R) | \ P R. \Q. (\, P, Q) \ Rel \ (\, Q, R) \ Rel'} \ Rel''" + +shows "\ \ P \[Rel''] R" + using \eqvt Rel''\ +proof(induct rule: simI[where C=Q]) + case(cSim \ R') + from QSimR \\ \ R \\ \ R'\ \(bn \) \* \\ \(bn \) \* Q\ + obtain Q' where QTrans: "\ \ Q \\ \ Q'" and Q'Rel'R': "(\, Q', R') \ Rel'" + by(blast dest: simE) + from PSimQ QTrans \bn \ \* \\ \bn \ \* P\ + obtain P' where PTrans: "\ \ P \\ \ P'" and P'RelQ': "(\, P', Q') \ Rel" + by(blast dest: simE) + with PTrans Q'Rel'R' P'RelQ' Set + show ?case by blast +qed + +lemma statEqSim: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and Rel :: "('b \ ('a, 'b, 'c) psi \ ('a, 'b, 'c) psi) set" + and Q :: "('a, 'b, 'c) psi" + and \' :: 'b + +assumes PSimQ: "\ \ P \[Rel] Q" + and "eqvt Rel'" + and "\ \ \'" + and C1: "\\'' R S \'''. \(\'', R, S) \ Rel; \'' \ \'''\ \ (\''', R, S) \ Rel'" + +shows "\' \ P \[Rel'] Q" + using \eqvt Rel'\ +proof(induct rule: simI[of _ _ _ _ \]) + case(cSim \ Q') + from \\' \ Q \\ \ Q'\ \\ \ \'\ + have "\ \ Q \\ \ Q'" by(metis statEqTransition AssertionStatEqSym) + with PSimQ \bn \ \* \\ \bn \ \* P\ + obtain P' where "\ \ P \\ \ P'" and "(\, P', Q') \ Rel" + by(blast dest: simE) + + from \\ \ P \\ \ P'\ \\ \ \'\ have "\' \ P \\ \ P'" + by(rule statEqTransition) + moreover from \(\, P', Q') \ Rel\ \\ \ \'\ have "(\', P', Q') \ Rel'" + by(rule C1) + ultimately show ?case by blast +qed + +lemma monotonic: + fixes \ :: 'b + and P :: "('a, 'b, 'c) psi" + and A :: "('b \ ('a, 'b, 'c) psi \ ('a, 'b, 'c) psi) set" + and Q :: "('a, 'b, 'c) psi" + and B :: "('b \ ('a, 'b, 'c) psi \ ('a, 'b, 'c) psi) set" + +assumes "\ \ P \[A] Q" + and "A \ B" + +shows "\ \ P \[B] Q" + using assms + by(simp (no_asm) add: simulation_def) (auto dest: simE) + +end + +end diff --git a/thys/Broadcast_Psi/document/root.bib b/thys/Broadcast_Psi/document/root.bib new file mode 100644 --- /dev/null +++ b/thys/Broadcast_Psi/document/root.bib @@ -0,0 +1,88 @@ +@inproceedings{DBLP:conf/sefm/BorgstromHJRVPP11, + author = {Johannes Borgstr{\"{o}}m and + Shuqin Huang and + Magnus Johansson and + Palle Raabjerg and + Bj{\"{o}}rn Victor and + Johannes {\AA}man Pohjola and + Joachim Parrow}, + editor = {Gilles Barthe and + Alberto Pardo and + Gerardo Schneider}, + title = {Broadcast {Psi}-calculi with an Application to Wireless Protocols}, + booktitle = {Software Engineering and Formal Methods - 9th International Conference, + {SEFM} 2011, Montevideo, Uruguay, November 14-18, 2011. Proceedings}, + series = {Lecture Notes in Computer Science}, + volume = {7041}, + pages = {74--89}, + publisher = {Springer}, + year = {2011}, + url = {https://doi.org/10.1007/978-3-642-24690-6\_7}, + doi = {10.1007/978-3-642-24690-6\_7}, + biburl = {https://dblp.org/rec/conf/sefm/BorgstromHJRVPP11.bib}, + bibsource = {dblp computer science bibliography, https://dblp.org} +} + +@article{DBLP:journals/sosym/BorgstromHJRVPP15, + author = {Johannes Borgstr{\"{o}}m and + Shuqin Huang and + Magnus Johansson and + Palle Raabjerg and + Bj{\"{o}}rn Victor and + Johannes {\AA}man Pohjola and + Joachim Parrow}, + title = {Broadcast psi-calculi with an application to wireless protocols}, + journal = {Softw. Syst. Model.}, + volume = {14}, + number = {1}, + pages = {201--216}, + year = {2015}, + url = {https://doi.org/10.1007/s10270-013-0375-z}, + doi = {10.1007/s10270-013-0375-z}, + biburl = {https://dblp.org/rec/journals/sosym/BorgstromHJRVPP15.bib}, + bibsource = {dblp computer science bibliography, https://dblp.org} +} + +@article{DBLP:journals/corr/abs-1101-3262, + author = {Jesper Bengtson and + Magnus Johansson and + Joachim Parrow and + Bj{\"{o}}rn Victor}, + title = {{Psi}-calculi: a framework for mobile processes with nominal data and + logic}, + journal = {Log. Methods Comput. Sci.}, + volume = {7}, + number = {1}, + year = {2011}, + url = {https://doi.org/10.2168/LMCS-7(1:11)2011}, + doi = {10.2168/LMCS-7(1:11)2011}, + biburl = {https://dblp.org/rec/journals/corr/abs-1101-3262.bib}, + bibsource = {dblp computer science bibliography, https://dblp.org} +} + +@article{DBLP:journals/jar/BengtsonPW16, + author = {Jesper Bengtson and + Joachim Parrow and + Tjark Weber}, + title = {{Psi}-Calculi in {Isabelle}}, + journal = {J. Autom. Reason.}, + volume = {56}, + number = {1}, + pages = {1--47}, + year = {2016}, + url = {https://doi.org/10.1007/s10817-015-9336-2}, + doi = {10.1007/s10817-015-9336-2}, + biburl = {https://dblp.org/rec/journals/jar/BengtsonPW16.bib}, + bibsource = {dblp computer science bibliography, https://dblp.org} +} + +@article{DBLP:journals/afp/Bengtson12, + author = {Jesper Bengtson}, + title = {{Psi}-calculi in {Isabelle}}, + journal = {Arch. Formal Proofs}, + year = {2012}, + note = {\url{https://www.isa-afp.org/entries/Psi_Calculi.html}, Formal proof development}, + url = {https://www.isa-afp.org/entries/Psi\_Calculi.shtml}, + biburl = {https://dblp.org/rec/journals/afp/Bengtson12.bib}, + bibsource = {dblp computer science bibliography, https://dblp.org} +} diff --git a/thys/Broadcast_Psi/document/root.tex b/thys/Broadcast_Psi/document/root.tex new file mode 100644 --- /dev/null +++ b/thys/Broadcast_Psi/document/root.tex @@ -0,0 +1,53 @@ +\documentclass[11pt,a4paper]{article} +\usepackage[T1]{fontenc} +\usepackage{isabelle,isabellesym} +\usepackage{amssymb} +\usepackage[english]{babel} + +% this should be the last package used +\usepackage{pdfsetup} + +% urls in roman style, theory text in math-similar italics +\urlstyle{rm} +\isabellestyle{it} + +\begin{document} + +\title{Broadcast Psi-calculi} +\author{Palle Raabjerg \and Johannes {\AA}man Pohjola \and Tjark Weber} +\maketitle + +\begin{abstract} + We provide an Isabelle/HOL-Nominal formalisation of the definitions, + theorems and proofs in the paper \emph{Broadcast Psi-calculi with an + Application to Wireless Protocols} by Borgstr\"om et al., which + extends the Psi-calculi framework with primitives for broadcast + communication in order to model wireless protocols. +\end{abstract} + +%\tableofcontents + +\section{Introduction} + +We provide an Isabelle/HOL-Nominal formalisation of the definitions, +theorems and proofs in the paper \emph{Broadcast Psi-calculi with an +Application to Wireless +Protocols}~\cite{DBLP:conf/sefm/BorgstromHJRVPP11,DBLP:journals/sosym/BorgstromHJRVPP15}, +which extends the Psi-calculi +framework~\cite{DBLP:journals/corr/abs-1101-3262,DBLP:journals/jar/BengtsonPW16,DBLP:journals/afp/Bengtson12} +with primitives for broadcast communication in order to model wireless +protocols. + +The file \texttt{Broadcast\_Thms.thy} contains a collection of the +relevant definitions and theorems, with comments relating them +directly to the paper. + +% include generated text of all theories +\section{Formalisation} + +\input{session} + +\bibliographystyle{plain} +\bibliography{root} + +\end{document} diff --git a/thys/ROOTS b/thys/ROOTS --- a/thys/ROOTS +++ b/thys/ROOTS @@ -1,810 +1,811 @@ ABY3_Protocols ADS_Functor AI_Planning_Languages_Semantics AODV AOT AVL-Trees AWN Abortable_Linearizable_Modules Abs_Int_ITP2012 Abstract-Hoare-Logics Abstract-Rewriting Abstract_Completeness Abstract_Soundness Ackermanns_not_PR Actuarial_Mathematics Adaptive_State_Counting Affine_Arithmetic Aggregation_Algebras Akra_Bazzi Algebraic_Numbers Algebraic_VCs Allen_Calculus Amicable_Numbers Amortized_Complexity AnselmGod Applicative_Lifting Approximate_Model_Counting Approximation_Algorithms Architectural_Design_Patterns Aristotles_Assertoric_Syllogistic Arith_Prog_Rel_Primes ArrowImpossibilityGS Attack_Trees Auto2_HOL Auto2_Imperative_HOL AutoFocus-Stream Automated_Stateful_Protocol_Verification Automatic_Refinement AxiomaticCategoryTheory BDD BD_Security_Compositional BNF_CC BNF_Operations BTree Balog_Szemeredi_Gowers Banach_Steinhaus Belief_Revision Bell_Numbers_Spivey BenOr_Kozen_Reif Berlekamp_Zassenhaus Bernoulli Bertrands_Postulate Bicategory BinarySearchTree Binary_Code_Imprimitive Binding_Syntax_Theory Binomial-Heaps Binomial-Queues BirdKMP Birkhoff_Finite_Distributive_Lattices Blue_Eyes Bondy Boolean_Expression_Checkers Boolos_Curious_Inference Boolos_Curious_Inference_Automated Bounded_Deducibility_Security +Broadcast_Psi Buchi_Complementation Budan_Fourier Buffons_Needle Buildings BytecodeLogicJmlTypes C2KA_DistributedSystems CAVA_Automata CAVA_LTL_Modelchecker CCS CHERI-C_Memory_Model CISC-Kernel Concentration_Inequalities CRDT CRYSTALS-Kyber CRYSTALS-Kyber_Security CSP_RefTK CVP_Hardness CYK CZH_Elementary_Categories CZH_Foundations CZH_Universal_Constructions CakeML CakeML_Codegen Call_Arity Card_Equiv_Relations Card_Multisets Card_Number_Partitions Card_Partitions Cardinality_Continuum Cartan_FP Case_Labeling Catalan_Numbers Category Category2 Category3 Catoids Cauchy Cayley_Hamilton Certification_Monads Ceva Chandy_Lamport Chebyshev_Polynomials Chord_Segments Circus Clean Clique_and_Monotone_Circuits ClockSynchInst Closest_Pair_Points CoCon CoSMeDis CoSMed CofGroups Coinductive Coinductive_Languages Collections Combinable_Wands Combinatorial_Enumeration_Algorithms Combinatorics_Words Combinatorics_Words_Graph_Lemma Combinatorics_Words_Lyndon CommCSL Commuting_Hermitian Comparison_Sort_Lower_Bound Compiling-Exceptions-Correctly Complete_Non_Orders Completeness Complex_Bounded_Operators Complex_Geometry Complx ComponentDependencies ConcurrentGC ConcurrentIMP Concurrent_Ref_Alg Concurrent_Revisions Conditional_Simplification Conditional_Transfer_Rule CondNormReasHOL Consensus_Refined Constructive_Cryptography Constructive_Cryptography_CM Constructor_Funs Containers Continued_Fractions Cook_Levin CoreC++ Core_DOM Core_SC_DOM Correctness_Algebras Cotangent_PFD_Formula Count_Complex_Roots Coupledsim_Contrasim CryptHOL CryptoBasedCompositionalProperties Crypto_Standards CubicalCategories Cubic_Quartic_Equations DCR-ExecutionEquivalence DFS_Framework DOM_Components DPRM_Theorem DPT-SAT-Solver DataRefinementIBP Datatype_Order_Generator Decl_Sem_Fun_PL Decreasing-Diagrams Decreasing-Diagrams-II Dedekind_Real Deep_Learning Delta_System_Lemma Density_Compiler Dependent_SIFUM_Refinement Dependent_SIFUM_Type_Systems Depth-First-Search Derangements Deriving Descartes_Sign_Rule Design_Theory Dict_Construction Differential_Dynamic_Logic Differential_Game_Logic DigitsInBase Digit_Expansions Dijkstra_Shortest_Path Diophantine_Eqns_Lin_Hom Directed_Sets Dirichlet_L Dirichlet_Series DiscretePricing Discrete_Summation Disintegration DiskPaxos Distributed_Distinct_Elements Dominance_CHK Doob_Convergence DynamicArchitectures Dynamic_Tables E_Transcendental Earley_Parser Echelon_Form EdmondsKarp_Maxflow Edwards_Elliptic_Curves_Group Efficient-Mergesort Efficient_Weighted_Path_Order Elimination_Of_Repeated_Factors Elliptic_Curves_Group_Law Encodability_Process_Calculi Epistemic_Logic Equivalence_Relation_Enumeration Ergodic_Theory Error_Function Euler_MacLaurin Euler_Partition Euler_Polyhedron_Formula Eudoxus_Reals Eval_FO Example-Submission Executable_Randomized_Algorithms Expander_Graphs Extended_Finite_State_Machine_Inference Extended_Finite_State_Machines FFT FLP FOL-Fitting FOL_Axiomatic FOL_Harrison FOL_Seq_Calc1 FOL_Seq_Calc2 FOL_Seq_Calc3 FO_Theory_Rewriting FSM_Tests Factor_Algebraic_Polynomial Factored_Transition_System_Bounding Falling_Factorial_Sum Farkas FeatherweightJava Featherweight_OCL Fermat3_4 FileRefinement FinFun Finger-Trees Finite-Map-Extras Finite_Automata_HF Finite_Fields Finitely_Generated_Abelian_Groups First_Order_Terms First_Welfare_Theorem Fishburn_Impossibility Fisher_Yates Fishers_Inequality Fixed_Length_Vector Flow_Networks Floyd_Warshall Flyspeck-Tame FocusStreamsCaseStudies Forcing Formal_Puiseux_Series Formal_SSA Formula_Derivatives Foundation_of_geometry Fourier Free-Boolean-Algebra Free-Groups Frequency_Moments Fresh_Identifiers FunWithFunctions FunWithTilings Functional-Automata Functional_Ordered_Resolution_Prover Furstenberg_Topology GPU_Kernel_PL Gabow_SCC GaleStewart_Games Gale_Shapley Game_Based_Crypto Gauss-Jordan-Elim-Fun Gauss_Jordan Gauss_Sums Gaussian_Integers GenClock General-Triangle Generalized_Counting_Sort Generic_Deriving Generic_Join GewirthPGCProof Girth_Chromatic Given_Clause_Loops Go GoedelGod Goedel_HFSet_Semantic Goedel_HFSet_Semanticless Goedel_Incompleteness Goodstein_Lambda GraphMarkingIBP Graph_Saturation Graph_Theory Gray_Codes Green Groebner_Bases Groebner_Macaulay Gromov_Hyperbolicity Grothendieck_Schemes Group-Ring-Module HOL-CSP HOL-CSPM HOL-CSP_OpSem HOLCF-Prelude HRB-Slicing Hahn_Jordan_Decomposition Hales_Jewett Heard_Of Hello_World HereditarilyFinite Hermite Hermite_Lindemann Hidden_Markov_Models Higher_Order_Terms HoareForDivergence Hoare_Time Hood_Melville_Queue HotelKeyCards Huffman Hybrid_Logic Hybrid_Multi_Lane_Spatial_Logic Hybrid_Systems_VCs HyperCTL HyperHoareLogic Hyperdual Hypergraph_Basics Hypergraph_Colourings IEEE_Floating_Point IFC_Tracking IMAP-CRDT IMO2019 IMP2 IMP2_Binary_Heap IMP_Compiler IMP_Compiler_Reuse Interval_Analysis IO_Language_Conformance IP_Addresses Imperative_Insertion_Sort Implicational_Logic Impossible_Geometry IMP_Noninterference Incompleteness Incredible_Proof_Machine Independence_CH Inductive_Confidentiality Inductive_Inference InfPathElimination InformationFlowSlicing InformationFlowSlicing_Inter Integration Interpolation_Polynomials_HOL_Algebra Interpreter_Optimizations Interval_Arithmetic_Word32 Intro_Dest_Elim Involutions2Squares Iptables_Semantics Irrational_Series_Erdos_Straus Irrationality_J_Hancl Irrationals_From_THEBOOK IsaGeoCoq IsaNet Isabelle_C Isabelle_hoops Isabelle_Marries_Dirac Isabelle_Meta_Model Jacobson_Basic_Algebra Jinja JinjaDCI JinjaThreads JiveDataStoreModel Jordan_Hoelder Jordan_Normal_Form KAD Karatsuba KAT_and_DRA KBPs KD_Tree Key_Agreement_Strong_Adversaries Khovanskii_Theorem Kleene_Algebra Kneser_Cauchy_Davenport Knights_Tour Knot_Theory Knuth_Bendix_Order Knuth_Morris_Pratt KnuthMorrisPratt Koenigsberg_Friendship Kruskal Kummer_Congruence Kuratowski_Closure_Complement LLL_Basis_Reduction LLL_Factorization LOFT LP_Duality LTL LTL_Master_Theorem LTL_Normal_Form LTL_to_DRA LTL_to_GBA Lam-ml-Normalization LambdaAuth LambdaMu Lambda_Free_EPO Lambda_Free_KBOs Lambda_Free_RPOs Lambert_Series Lambert_W Landau_Symbols Laplace_Transform Latin_Square LatticeProperties Launchbury Laws_of_Large_Numbers Lazy-Lists-II Lazy_Case Lehmer Lifting_Definition_Option Lifting_the_Exponent LightweightJava LinearQuantifierElim Linear_Inequalities Linear_Programming Linear_Recurrences Liouville_Numbers List-Index List-Infinite List_Interleaving List_Inversions List_Update LocalLexing Localization_Ring Locally-Nameless-Sigma Logging_Independent_Anonymity Lovasz_Local Lowe_Ontological_Argument Lower_Semicontinuous Lp Lucas_Theorem MDP-Algorithms MDP-Rewards MFMC_Countable MFODL_Monitor_Optimized MFOTL_Monitor MSO_Regex_Equivalence Markov_Models Marriage Martingales Mason_Stothers Matrices_for_ODEs Matrix Matrix_Tensor Matroids Max-Card-Matching Maximum_Segment_Sum Median_Method Median_Of_Medians_Selection Menger Mereology Mersenne_Primes Metalogic_ProofChecker MHComputation MiniML MiniSail Minimal_SSA Minkowskis_Theorem Minsky_Machines MLSS_Decision_Proc ML_Unification Modal_Logics_for_NTS Modular_Assembly_Kit_Security Modular_arithmetic_LLL_and_HNF_algorithms Monad_Memo_DP Monad_Normalisation MonoBoolTranAlgebra MonoidalCategory Monomorphic_Monad MuchAdoAboutTwo Multi_Party_Computation Multirelations Multirelations_Heterogeneous Multiset_Ordering_NPC Multitape_To_Singletape_TM Myhill-Nerode Name_Carrying_Type_Inference Nano_JSON Nash_Williams Nat-Interval-Logic Native_Word Nested_Multisets_Ordinals Network_Security_Policy_Verification Neumann_Morgenstern_Utility No_FTL_observers No_FTL_observers_Gen_Rel Nominal2 Nominal_Myhill_Nerode Noninterference_CSP Noninterference_Concurrent_Composition Noninterference_Generic_Unwinding Noninterference_Inductive_Unwinding Noninterference_Ipurge_Unwinding Noninterference_Sequential_Composition NormByEval Nullstellensatz Number_Theoretic_Transform Octonions OmegaCatoidsQuantales OpSets Open_Induction Optics Optimal_BST Orbit_Stabiliser Order_Lattice_Props Ordered_Resolution_Prover Ordinal Ordinal_Partitions Ordinals_and_Cardinals Ordinary_Differential_Equations PAC_Checker PAL PAPP_Impossibility PCF PLM POPLmark-deBruijn PSemigroupsConvolution Package_logic Padic_Field Padic_Ints Pairing_Heap Paraconsistency Parity_Game Partial_Function_MR Partial_Order_Reduction Password_Authentication_Protocol Pell Perfect_Fields Perfect-Number-Thm Perron_Frobenius Physical_Quantities Pi_Calculus Pi_Transcendental Planarity_Certificates Pluennecke_Ruzsa_Inequality Poincare_Bendixson Poincare_Disc Polygonal_Number_Theorem Polylog Polynomial_Crit_Geometry Polynomial_Factorization Polynomial_Interpolation Polynomials Pop_Refinement Posix-Lexing Possibilistic_Noninterference Power_Sum_Polynomials Pratt_Certificate Prefix_Free_Code_Combinators Presburger-Automata Prim_Dijkstra_Simple Prime_Distribution_Elementary Prime_Harmonic_Series Prime_Number_Theorem Priority_Queue_Braun Priority_Search_Trees Probabilistic_Noninterference Probabilistic_Prime_Tests Probabilistic_System_Zoo Probabilistic_Timed_Automata Probabilistic_While Probability_Inequality_Completeness Program-Conflict-Analysis Progress_Tracking Projective_Geometry Projective_Measurements Promela Proof_Strategy_Language PropResPI Propositional_Logic_Class Propositional_Proof_Systems Prpu_Maxflow PseudoHoops Psi_Calculi Ptolemys_Theorem Public_Announcement_Logic Q0_Metatheory Q0_Soundness QBF_Solver_Verification QHLProver QR_Decomposition Quantales Quantales_Converse Quantifier_Elimination_Hybrid Quasi_Borel_Spaces Quaternions Query_Optimization Quick_Sort_Cost RIPEMD-160-SPARK ROBDD RSAPSS Ramsey-Infinite Random_BSTs Random_Graph_Subgraph_Threshold Randomised_BSTs Randomised_Social_Choice Rank_Nullity_Theorem Real_Impl Real_Power Real_Time_Deque Recursion-Addition Recursion-Theory-I Refine_Imperative_HOL Refine_Monadic RefinementReactive Regex_Equivalence Region_Quadtrees Registers Regression_Test_Selection Regular-Sets Regular_Algebras Regular_Tree_Relations Relation_Algebra Relational-Incorrectness-Logic Relational_Cardinality Relational_Disjoint_Set_Forests Relational_Forests Relational_Method Relational_Minimum_Spanning_Trees Relational_Paths Rensets Rep_Fin_Groups ResiduatedTransitionSystem Residuated_Lattices Resolution_FOL Rewrite_Properties_Reduction Rewriting_Z Ribbon_Proofs Risk_Free_Lending Robbins-Conjecture Robinson_Arithmetic Root_Balanced_Tree Roth_Arithmetic_Progressions Routing Roy_Floyd_Warshall SATSolverVerification SCC_Bloemen_Sequential SC_DOM_Components SDS_Impossibility SIFPL SIFUM_Type_Systems SPARCv8 S_Finite_Measure_Monad Safe_Distance Safe_OCL Safe_Range_RC Saturation_Framework Saturation_Framework_Extensions Sauer_Shelah_Lemma Schutz_Spacetime Schwartz_Zippel Secondary_Sylow Security_Protocol_Refinement Selection_Heap_Sort SenSocialChoice Separata Separation_Algebra Separation_Logic_Imperative_HOL Separation_Logic_Unbounded SequentInvertibility Shadow_DOM Shadow_SC_DOM Shivers-CFA ShortestPath Show Sigma_Commit_Crypto Signature_Groebner Simpl Simple_Clause_Learning Simple_Firewall Simplex Simplicial_complexes_and_boolean_functions SimplifiedOntologicalArgument Skew_Heap Skip_Lists Slicing Sliding_Window_Algorithm Smith_Normal_Form Smooth_Manifolds Solidity Sophomores_Dream Sort_Encodings Source_Coding_Theorem SpecCheck Special_Function_Bounds Splay_Tree Sqrt_Babylonian Stable_Matching Stalnaker_Logic Standard_Borel_Spaces Statecharts Stateful_Protocol_Composition_and_Typing Stellar_Quorums Stern_Brocot Stewart_Apollonius Stirling_Formula Stochastic_Matrices Stone_Algebras Stone_Kleene_Relation_Algebras Stone_Relation_Algebras Store_Buffer_Reduction Stream-Fusion Stream_Fusion_Code StrictOmegaCategories Strong_Security Sturm_Sequences Sturm_Tarski Stuttering_Equivalence Subresultants Subset_Boolean_Algebras Sumcheck_Protocol SumSquares Sunflowers SuperCalc Suppes_Theorem Surprise_Paradox Symmetric_Polynomials Syntax_Independent_Logic Synthetic_Completeness Szemeredi_Regularity Szpilrajn TESL_Language TLA Tail_Recursive_Functions Tarskis_Geometry Taylor_Models Three_Circles Three_Squares Timed_Automata Topological_Semantics Topology TortoiseHare TsirelsonBound Transcendence_Series_Hancl_Rucki Transformer_Semantics Transition_Systems_and_Automata Transitive-Closure Transitive-Closure-II Transitive_Models Transport Treaps Tree-Automata Tree_Decomposition Tree_Enumeration Triangle Trie Turans_Graph_Theorem Twelvefold_Way Two_Generated_Word_Monoids_Intersection Tycon Types_Tableaus_and_Goedels_God Types_To_Sets_Extension UPF UPF_Firewall UTP Undirected_Graph_Theory Universal_Hash_Families Universal_Turing_Machine UpDown_Scheme VYDRA_MDL Valuation Van_Emde_Boas_Trees Van_der_Waerden VectorSpace VeriComp Verified-Prover Verified_SAT_Based_AI_Planning VerifyThis2018 VerifyThis2019 Vickrey_Clarke_Groves Virtual_Substitution VolpanoSmith WHATandWHERE_Security WOOT_Strong_Eventual_Consistency WebAssembly Weight_Balanced_Trees Weighted_Arithmetic_Geometric_Mean Weighted_Path_Order Well_Quasi_Orders Wetzels_Problem Wieferich_Kempner Winding_Number_Eval Word_Lib WorkerWrapper X86_Semantics XML Youngs_Inequality ZFC_in_HOL Zeckendorf Zeta_3_Irrational Zeta_Function pGCL Labeled_Transition_Systems Pushdown_Systems