Theory Semantics

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 "⦇ν*xvecB  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  (⦇ν*xvecB) = ⦇ν*(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 "(⦇ν*xvecN ≺' P = N' ≺' P') = (xvec = []  N = N'  P = P')"
  and "(N' ≺' P' = ⦇ν*xvecN ≺' P) = (xvec = []  N = N'  P = P')"
  and "(N' ≺' P' = N ≺' P) = (N = N'  P = P')"
  and "(⦇ν*xvecB = ⦇ν*xvecB') = (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  (⦇ν*xvecB)) = (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 ♯* (⦇ν*xvecB) = (xXs. x  set xvec  x  B)"
  and "yvec ♯* (⦇ν*xvecB) = (x(set yvec). x  set xvec  x  B)"
  and "Xs ♯* (⦇νxB) = Xs ♯* [x].B"
  and "xvec ♯* (⦇νxB) = 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(⦇ν*xvecB)::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 ♯* (⦇ν*xvecB)) = (Xs ♯* B)"
  and "yvec ♯* xvec  yvec ♯* (⦇ν*xvecB) = yvec ♯* B"
  and "xvec ♯* (⦇ν*xvecB)"
  and "x  xvec  x  ⦇ν*xvecB = 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 "(⦇ν*yvecB) = (⦇ν*(p  yvec)(p  B))"
proof -
  note pt_name_inst at_name_inst S
  moreover from (set xvec)  (set yvec) have "set xvec ♯* (⦇ν*yvecB)"
    by(force simp add: boundOutputFreshSet)
  moreover from xvecFreshB (set xvec)  (set yvec) have "set (p  xvec) ♯* (⦇ν*yvecB)"
    by (simp add: boundOutputFreshSet) (simp add: fresh_star_def)
  ultimately have "(⦇ν*yvecB) = p  (⦇ν*yvecB)"
    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 ♯* (⦇ν*zvecB)"

shows "(⦇ν*zvecB) = (⦇ν*(p  zvec)(p  B))"
proof -
  note pt_name_inst at_name_inst S yvec ♯* (⦇ν*zvecB)
  moreover from xvecFreshB have "set (xvec) ♯* (⦇ν*zvecB)"
    by (simp add: boundOutputFreshSet) (simp add: fresh_star_def)
  ultimately have "(⦇ν*zvecB) = p  (⦇ν*zvecB)"
    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 "(⦇ν*yvecM ≺' 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 "⦇ν*xvecN ≺' 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 "⦇νxB = ⦇ν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 "⦇νxB = ⦇νyC"
  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 "⦇νxB = ⦇νyC"
  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 "⦇ν*xvecB = ⦇ν*yvecB'"
  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: "⦇ν*xvecB = ⦇ν*yvecB'" 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 ⦇ν*xvecB = ⦇ν*yvecB' 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) = ⦇ν*yvecB'; 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 "⦇ν*xvecM ≺' P = ⦇ν*yvecN ≺' 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 ⦇ν*xvecM ≺' P = ⦇ν*yvecN ≺' 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 ⦇ν*xvecM ≺' P = ⦇ν*yvecN ≺' 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. ⦇ν*xvecM ≺' P = ⦇ν*yvecN ≺' (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 "⦇ν*xvecM ≺' P = ⦇ν*yvecN ≺' 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 "⦇ν*xvecM ≺' P = ⦇ν*yvecN ≺' 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: "⦇ν*xvecM ≺' P = ⦇ν*yvecN ≺' 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 ⦇ν*xvecM ≺' P = ⦇ν*yvecN ≺' 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) = ⦇ν*yvecN ≺' 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(⦇ν*xvecM ≺' P) = ⦇νy(⦇ν*yvecN ≺' 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 "⦇ν*xvecM ≺' 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(⦇νxB) = ⦇ν*yvec(⦇νyB')"
  and   L: "length xvec = length yvec"
  and   xFreshB': "x  B'"
  and   xFreshxvec: "x  xvec"
  and   xFreshyvec: "x  yvec"

shows "⦇ν*xvecB = ⦇ν*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⦇νxB = ⦇ν*yvec⦇νyB'" 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⦇νxB = ⦇ν*yvec⦇νyB'" by fact
  with xEq yEq have Eq: "⦇νx'(⦇ν*xvec'⦇νxB) = ⦇νy'(⦇ν*yvec'⦇νyB')" by simp
  have IH: "xvec yvec y B'.
            n = length xvec; ⦇ν*xvec⦇νxB = ⦇ν*yvec⦇νyB'; length xvec = length yvec; x  B'; x  xvec; x  yvec
             ⦇ν*xvecB = ⦇ν*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'⦇νxB = ⦇ν*yvec'⦇νyB'" 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'⦇νxB = ⦇ν*([(x', y')]  yvec')⦇ν([(x', y')]  y)([(x', y')]  B')"
      and x'FreshB': "x'  ⦇ν*yvec'⦇νyB'"
      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'  ⦇νyB'"
        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 "⦇ν*xvecM ≺' P = ⦇ν*yvecN ≺' (Q  R)"
  and   "xvec ♯* R"
  and   "yvec ♯* R"

obtains T where "P = T  R" and "⦇ν*xvecM ≺' T = ⦇ν*yvecN ≺' Q"
proof -
  assume "T. P = T  R; ⦇ν*xvecM ≺' T = ⦇ν*yvecN ≺' Q  thesis"
  moreover obtain n where "n = length xvec" by auto
  with assms have "T. P = T  R  ⦇ν*xvecM ≺' T = ⦇ν*yvecN ≺' Q"
  proof(induct n arbitrary: xvec yvec M N P Q R)
    case(0 xvec yvec M N P Q R)
    have Eq: "⦇ν*xvecM ≺' P = ⦇ν*yvecN ≺' (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 ⦇ν*xvecM ≺' P = ⦇ν*yvecN ≺' (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. ⦇ν*xvecM ≺' (P::('a, 'b, 'c) psi) = ⦇ν*yvecN ≺' (Q  R); xvec ♯* R; yvec ♯* R; n = length xvec  T. P = T  R  ⦇ν*xvecM ≺' T = ⦇ν*yvecN ≺' 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 "⦇ν*xvecM ≺' P = ⦇ν*yvecN ≺' (Q  R)"
  and   "xvec ♯* yvec"

obtains T p where "set p  set xvec × set yvec" and "P = T  (p  R)" and "⦇ν*xvecM ≺' T = ⦇ν*yvecN ≺' Q"
proof -
  assume "p T. set p  set xvec × set yvec; P = T  (p  R); ⦇ν*xvecM ≺' T = ⦇ν*yvecN ≺' 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)  ⦇ν*xvecM ≺' T = ⦇ν*yvecN ≺' Q"
  proof(induct n arbitrary: xvec yvec M N P Q R)
    case(0 xvec yvec M N P Q R)
    have Eq: "⦇ν*xvecM ≺' P = ⦇ν*yvecN ≺' (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 ⦇ν*xvecM ≺' P = ⦇ν*yvecN ≺' (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. ⦇ν*xvecM ≺' (P::('a, 'b, 'c) psi) = ⦇ν*yvecN ≺' (Q  R);  xvec ♯* yvec; n = length xvec  p T. set p  set xvec × set yvec  P = T  (p  R)  ⦇ν*xvecM ≺' T = ⦇ν*yvecN ≺' 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 "⦇ν*xvecM ≺' P = ⦇ν*yvecN ≺' (Q  R)"
  and   "xvec ♯* Q"
  and   "yvec ♯* Q"

obtains T where "P = Q  T" and "⦇ν*xvecM ≺' T = ⦇ν*yvecN ≺' R"
proof -
  assume "T. P = Q  T; ⦇ν*xvecM ≺' T = ⦇ν*yvecN ≺' R  thesis"
  moreover obtain n where "n = length xvec" by auto
  with assms have "T. P = Q  T  ⦇ν*xvecM ≺' T = ⦇ν*yvecN ≺' R"
  proof(induct n arbitrary: xvec yvec M N P Q R)
    case(0 xvec yvec M N P Q R)
    have Eq: "⦇ν*xvecM ≺' P = ⦇ν*yvecN ≺' (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 ⦇ν*xvecM ≺' P = ⦇ν*yvecN ≺' (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. ⦇ν*xvecM ≺' (P::('a, 'b, 'c) psi) = ⦇ν*yvecN ≺' (Q  R); xvec ♯* Q; yvec ♯* Q; n = length xvec  T. P = Q  T  ⦇ν*xvecM ≺' T = ⦇ν*yvecN ≺' 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 "⦇ν*xvecM ≺' P = ⦇ν*yvecN ≺' (Q  R)"
  and   "xvec ♯* yvec"

obtains T p where "set p  set xvec × set yvec" and "P = (p  Q)  T" and "⦇ν*xvecM ≺' T = ⦇ν*yvecN ≺' R"
proof -
  assume "p T. set p  set xvec × set yvec; P = (p  Q)  T; ⦇ν*xvecM ≺' T = ⦇ν*yvecN ≺' 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  ⦇ν*xvecM ≺' T = ⦇ν*yvecN ≺' R"
  proof(induct n arbitrary: xvec yvec M N P Q R)
    case(0 xvec yvec M N P Q R)
    have Eq: "⦇ν*xvecM ≺' P = ⦇ν*yvecN ≺' (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 ⦇ν*xvecM ≺' P = ⦇ν*yvecN ≺' (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. ⦇ν*xvecM ≺' (P::('a, 'b, 'c) psi) = ⦇ν*yvecN ≺' (Q  R);  xvec ♯* yvec; n = length xvec  p T. set p  set xvec × set yvec  P = (p  Q)  T  ⦇ν*xvecM ≺' T = ⦇ν*yvecN ≺' 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(⦇ν*yvecB)"
  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 = ⦇ν*yvecN ≺' 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 = ⦇ν*yvecN ≺' 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 "⦇ν*xvecM ≺' P = ⦇ν*yvecN ≺' ⦇νzQ"
  and   "z  xvec"
  and   "z  yvec"

obtains R where "P = ⦇νzR" and "⦇ν*xvecM ≺' R = ⦇ν*yvecN ≺' Q"
proof -
  assume "R. P = ⦇νzR; ⦇ν*xvecM ≺' R = ⦇ν*yvecN ≺' Q  thesis"
  moreover obtain n where "n = length xvec" by auto
  with assms have "R. P = ⦇νzR  ⦇ν*xvecM ≺' R = ⦇ν*yvecN ≺' Q"
  proof(induct n arbitrary: xvec yvec M N P Q z)
    case(0 xvec yvec M N P Q z)
    have Eq: "⦇ν*xvecM ≺' P = ⦇ν*yvecN ≺' ⦇νzQ" 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 ⦇ν*xvecM ≺' P = ⦇ν*yvecN ≺' (⦇νzQ) xvec = x # xvec'
    obtain y yvec' where "⦇ν*(x#xvec')M ≺' P = ⦇ν*(y#yvec')N ≺' ⦇νzQ"
      and "yvec = y#yvec'"
      by(cases yvec) auto
    then have EQ: "⦇νx(⦇ν*xvec'M ≺' P) = ⦇νy(⦇ν*yvec'N ≺' ⦇νzQ)"
      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. ⦇ν*xvecM ≺' (P::('a, 'b, 'c) psi) = ⦇ν*yvecN ≺' ⦇νzQ; z  xvec; z  yvec; n = length xvec  R. P = ⦇νzR  ⦇ν*xvecM ≺' R = ⦇ν*yvecN ≺' Q"
      by fact
    show ?case
    proof(cases "x = y")
      assume "x = y"
      with EQ have "⦇ν*xvec'M ≺' P = ⦇ν*yvec'N ≺' ⦇νzQ"
        by(simp add: boundOutput.inject alpha)
      with z  xvec' z  yvec' length xvec' = n
      obtain R where "P = ⦇νzR" 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 ≺' ⦇νzQ"
        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 = ⦇νzR" 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 = ⦇νzR 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 (MN) = []"
  | "bn (¿MN) = []"
  | "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
    "(MN)  P = RIn M N P"
  | "(¿MN)  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 (MN) = Some M"
  | "subject (¿MN) = 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 (MN) = Some N"
  | "object (¿MN) = 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  (MN  P) = (x  M  x  N  x  P)"
  and "x  (¿ MN  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  α = MN)"
  and "(α  P = RBrIn M N Q) = (P = Q  α = ¿MN)"
  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  α = MN)"
  and "(RBrIn M N Q = α  P) = (P = Q  α = ¿MN)"
  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  α = MN)"
    by(nominal_induct α rule: action.strong_induct)
      (auto simp add: residual.inject action.inject)
next
  show "(α  P = RBrIn M N Q) = (P = Q  α = ¿MN)"
    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  α = MN)"
    by(nominal_induct α rule: action.strong_induct)
      (auto simp add: residual.inject action.inject)
next
  show "(RBrIn M N Q = α  P) = (P = Q  α = ¿MN)"
    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 ♯* (MN  P) = (xvec ♯* M  xvec ♯* N  xvec ♯* P)"
  and "xvec ♯* (¿MN  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 ♯* (MN  P) = (X ♯* M  X ♯* N  X ♯* P)"
  and "X ♯* (¿MN  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. α = MN  Prop"
  and   "M N. α = ¿MN  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 = β  ⦇νxQ"
  and   "x  bn α"
  and   "x  bn β"

obtains R where "P = ⦇νxR" 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 "MN  M⦇ν*([])⦈⟨N"

abbreviation
  brOutputJudge ("¡__" [110, 110] 110) where "¡MN  ¡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 ("_ F _" [85, 85] 85)
abbreviation
  FBottomJudge ("F" 90) where "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 ("_ 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  Ψ  MN⟩.P  KN  P"
  | BrOutput: "Ψ  M  K  Ψ  MN⟩.P  ¡KN  P"
  | Case:   "Ψ  P  Rs; (φ, P)  set Cs; Ψ  φ; guarded P  Ψ  Cases Cs  Rs"
  | cPar1:   "(Ψ  ΨQ)  P α  P'; extractFrame Q = AQ, ΨQ; distinct AQ;
             AQ ♯* P; AQ ♯* Q; AQ ♯* Ψ; AQ ♯* α; AQ ♯* P'; distinct(bn α);
             bn α ♯* Ψ; bn α ♯* ΨQ; bn α ♯* Q; bn α ♯* P; bn α ♯* (subject α) 
             Ψ  P  Q α  (P'  Q)"
  | cPar2:   "(Ψ  ΨP)  Q α  Q'; extractFrame P = AP, ΨP; distinct AP;
             AP ♯* P; AP ♯* Q; AP ♯* Ψ; AP ♯* α; AP ♯* Q'; distinct(bn α);
             bn α ♯* Ψ; bn α ♯* ΨP; bn α ♯* P; bn α ♯* Q; bn α ♯* (subject α) 
             Ψ  P  Q α  (P  Q')"
  | cComm1:   "Ψ  ΨQ  P  MN  P'; extractFrame P = AP, ΨP; distinct AP;
             Ψ  ΨP  Q  K⦇ν*xvec⦈⟨N  Q'; extractFrame Q = AQ, ΨQ; distinct AQ;
             Ψ  ΨP  ΨQ  M  K;
             AP ♯* Ψ; AP ♯* ΨQ; AP ♯* P; AP ♯* M; AP ♯* N; AP ♯* P';
             AP ♯* Q; AP ♯* Q'; AP ♯* AQ; AP ♯* xvec;
             AQ ♯* Ψ; AQ ♯* ΨP; AQ ♯* P; AQ ♯* N; AQ ♯* P';
             AQ ♯* Q; AQ ♯* K; AQ ♯* Q'; AQ ♯* xvec; distinct xvec;
             xvec ♯* Ψ; xvec ♯* ΨP; xvec ♯* ΨQ; xvec ♯* P; xvec ♯* M;
             xvec ♯* Q; xvec ♯* K 
             Ψ  P  Q  τ  ⦇ν*xvec(P'  Q')"
  | cComm2:   "Ψ  ΨQ  P  M⦇ν*xvec⦈⟨N  P'; extractFrame P = AP, ΨP; distinct AP;
             Ψ  ΨP  Q  KN  Q'; extractFrame Q = AQ, ΨQ; distinct AQ;
             Ψ  ΨP  ΨQ  M  K;
             AP ♯* Ψ; AP ♯* ΨQ; AP ♯* P; AP ♯* M; AP ♯* N; AP ♯* P';
             AP ♯* Q; AP ♯* Q'; AP ♯* AQ; AP ♯* xvec;
             AQ ♯* Ψ; AQ ♯* ΨP; AQ ♯* P; AQ ♯* N; AQ ♯* P';
             AQ ♯* Q; AQ ♯* K; AQ ♯* Q'; AQ ♯* xvec; distinct xvec;
             xvec ♯* Ψ; xvec ♯* ΨP; xvec ♯* ΨQ; xvec ♯* P; xvec ♯* M;
             xvec ♯* Q; xvec ♯* K 
             Ψ  P  Q  τ  ⦇ν*xvec(P'  Q')"
  | cBrMerge: "Ψ  ΨQ  P  ¿MN  P'; extractFrame P = AP, ΨP; distinct AP;
             Ψ  ΨP  Q  ¿MN  Q'; extractFrame Q = AQ, ΨQ; distinct AQ;
             AP ♯* Ψ; AP ♯* ΨQ; AP ♯* P; AP ♯* N; AP ♯* P';
             AP ♯* Q; AP ♯* Q'; AP ♯* AQ;
             AP ♯* M; AQ ♯* M;
             AQ ♯* Ψ; AQ ♯* ΨP; AQ ♯* P; AQ ♯* N; AQ ♯* P';
             AQ ♯* Q; AQ ♯* Q' 
             Ψ  P  Q  ¿MN  (P'  Q')"
    (* Removed: AP ♯* M; AQ ♯* K; *)
  | cBrComm1:  "Ψ  ΨQ  P  ¿MN  P'; extractFrame P = AP, ΨP; distinct AP;
             Ψ  ΨP  Q  ¡M⦇ν*xvec⦈⟨N  Q'; extractFrame Q = AQ, ΨQ; distinct AQ;
             AP ♯* Ψ; AP ♯* ΨQ; AP ♯* P; AP ♯* N; AP ♯* P';
             AP ♯* Q; AP ♯* Q'; AP ♯* AQ; AP ♯* xvec;
             AP ♯* M; AQ ♯* M;
             AQ ♯* Ψ; AQ ♯* ΨP; AQ ♯* P; AQ ♯* N; AQ ♯* P';
             AQ ♯* Q; AQ ♯* Q'; AQ ♯* xvec; distinct xvec;
             xvec ♯* Ψ; xvec ♯* ΨP; xvec ♯* ΨQ; xvec ♯* P;
             xvec ♯* Q; xvec ♯* M 
             Ψ  P  Q  ¡M⦇ν*xvec⦈⟨N  (P'  Q')"
  | cBrComm2:  "Ψ  ΨQ  P  ¡M⦇ν*xvec⦈⟨N  P'; extractFrame P = AP, ΨP; distinct AP;
             Ψ  ΨP  Q  ¿MN  Q'; extractFrame Q = AQ, ΨQ; distinct AQ;
             AP ♯* Ψ; AP ♯* ΨQ; AP ♯* P; AP ♯* N; AP ♯* P';
             AP ♯* Q; AP ♯* Q'; AP ♯* AQ; AP ♯* xvec;
             AP ♯* M; AQ ♯* M;
             AQ ♯* Ψ; AQ ♯* ΨP; AQ ♯* P; AQ ♯* N; AQ ♯* P';
             AQ ♯* Q; AQ ♯* Q'; AQ ♯* xvec; distinct xvec;
             xvec ♯* Ψ; xvec ♯* ΨP; xvec ♯* Ψ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 
            Ψ  ⦇νxP  τ  ⦇νx(⦇ν*xvecP')"
  | 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 
              Ψ  ⦇νxP  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 
              Ψ  ⦇νxP  ¡M⦇ν*(xvec@x#yvec)⦈⟨N  P'"
  | cScope:  "Ψ  P α  P'; x  Ψ; x  α; bn α ♯* Ψ; bn α ♯* P; bn α ♯* (subject α); distinct(bn α)  Ψ  ⦇νxP α  (⦇νxP')"
  | 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 AQ  set(bn α)"
  | cPar2: "set AP  set(bn α)"
  | cComm1: "set AP  set AQ  set xvec"
  | cComm2: "set AP  set AQ  set xvec"
  | cBrMerge: "set AP  set AQ"
  | cBrComm1: "set AP  set AQ  set xvec"
  | cBrComm2: "set AP  set AQ  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 "Ψ  MN⟩.P KN'  P'  False"
  and "Ψ  MN⟩.P ¿KN'  P'  False"
  and "Ψ  MN⟩.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 Ψ (MN⟩.P) (KN) P"
  and   rBrOutput:"Ψ M K N P C. Ψ  M  K  Prop C Ψ (MN⟩.P) (¡KN) 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: "Ψ ΨQ P α P' AQ Q C.
                   Ψ  ΨQ  P α  P'; extractFrame Q = AQ, ΨQ; distinct AQ;
                    C. Prop C (Ψ  ΨQ) P α P';
                    AQ ♯* P; AQ ♯* Q; AQ ♯* Ψ; AQ ♯* α; AQ ♯* P'; AQ ♯* C; distinct(bn α); bn α ♯* Q;
                    bn α ♯* Ψ; bn α ♯* ΨQ; bn α ♯* P; bn α ♯* subject α; bn α ♯* C 
                    Prop C Ψ (P  Q) α (P'  Q)"
  and   rPar2: "Ψ ΨP Q α Q' AP P C.
                   Ψ  ΨP  Q α  Q'; extractFrame P = AP, ΨP; distinct AP;
                    C. Prop C (Ψ  ΨP) Q α Q';
                    AP ♯* P; AP ♯* Q; AP ♯* Ψ; AP ♯* α; AP ♯* Q'; AP ♯* C; distinct(bn α); bn α ♯* Q;
                    bn α ♯* Ψ; bn α ♯* ΨP; bn α ♯* P; bn α ♯* subject α; bn α ♯* C 
                    Prop C Ψ (P  Q) α (P  Q')"
  and   rComm1: "Ψ ΨQ P M N P' AP ΨP Q K xvec Q' AQ C.
                   Ψ  ΨQ  P MN  P'; C. Prop C (Ψ  ΨQ) P (MN) P';
                    extractFrame P = AP, ΨP; distinct AP;
                    Ψ  ΨP  Q K⦇ν*xvec⦈⟨N  Q'; C. Prop C (Ψ  ΨP) Q (K⦇ν*xvec⦈⟨N) Q';
                    extractFrame Q = AQ, ΨQ; distinct AQ;
                    Ψ  ΨP  ΨQ  M  K;
                    AP ♯* Ψ; AP ♯* ΨQ; AP ♯* P; AP ♯* M; AP ♯* N; AP ♯* P';
                    AP ♯* Q; AP ♯* Q'; AP ♯* AQ; AP ♯* xvec; AQ ♯* Ψ; AQ ♯* ΨP;
                    AQ ♯* P; AQ ♯* N; AQ ♯* P'; AQ ♯* Q; AQ ♯* K; AQ ♯* Q'; distinct xvec;
                    AQ ♯* xvec; xvec ♯* Ψ; xvec ♯* ΨP; xvec ♯* ΨQ; xvec ♯* P; xvec ♯* M;
                    xvec ♯* Q; xvec ♯* K; AP ♯* C; AQ ♯* C; xvec ♯* C 
                    Prop C Ψ (P  Q) (τ) (⦇ν*xvec(P'  Q'))"
  and   rComm2: "Ψ ΨQ P M xvec N P' AP ΨP Q K Q' AQ C.
                   Ψ  ΨQ  P M⦇ν*xvec⦈⟨N  P'; C. Prop C (Ψ  ΨQ) P (M⦇ν*xvec⦈⟨N) P';
                    extractFrame P = AP, ΨP; distinct AP;
                    Ψ  ΨP  Q KN  Q'; C. Prop C (Ψ  ΨP) Q (KN) Q';
                    extractFrame Q = AQ, ΨQ; distinct AQ;
                    Ψ  ΨP  ΨQ  M  K;
                    AP ♯* Ψ; AP ♯* ΨQ; AP ♯* P; AP ♯* M; AP ♯* N; AP ♯* P';
                    AP ♯* Q; AP ♯* Q'; AP ♯* AQ; AP ♯* xvec; AQ ♯* Ψ; AQ ♯* ΨP;
                    AQ ♯* P; AQ ♯* N; AQ ♯* P'; AQ ♯* Q; AQ ♯* K; AQ ♯* Q'; distinct xvec;
                    AQ ♯* xvec; xvec ♯* Ψ; xvec ♯* ΨP; xvec ♯* ΨQ; xvec ♯* P; xvec ♯* M;
                    xvec ♯* Q; xvec ♯* K; AP ♯* C; AQ ♯* C; xvec ♯* C 
                    Prop C Ψ (P  Q) (τ) (⦇ν*xvec(P'  Q'))"
  and   rBrMerge: "Ψ ΨQ P M N P' AP ΨP Q Q' AQ C.
                    Ψ  ΨQ  P  ¿MN  P'; C. Prop C (Ψ  ΨQ) P (¿MN) P';
                    extractFrame P = AP, ΨP; distinct AP;
                    Ψ  ΨP  Q  ¿MN  Q'; C. Prop C (Ψ  ΨP) Q (¿MN) Q';
                    extractFrame Q = AQ, ΨQ; distinct AQ;
                    AP ♯* Ψ; AP ♯* ΨQ; AP ♯* P; AP ♯* N; AP ♯* P';
                    AP ♯* Q; AP ♯* Q'; AP ♯* AQ; AP ♯* M; AQ ♯* M;
                    AQ ♯* Ψ; AQ ♯* ΨP; AQ ♯* P; AQ ♯* N; AQ ♯* P';
                    AQ ♯* Q; AQ ♯* Q'; AP ♯* C; AQ ♯* C 
                    Prop C Ψ (P  Q) (¿MN) (P'  Q')"
  and   rBrComm1: "Ψ ΨQ P M N P' AP ΨP Q xvec Q' AQ C.
                   Ψ  ΨQ  P ¿MN  P'; C. Prop C (Ψ  ΨQ) P (¿MN) P';
                    extractFrame P = AP, ΨP; distinct AP;
                    Ψ  ΨP  Q ¡M⦇ν*xvec⦈⟨N  Q'; C. Prop C (Ψ  ΨP) Q (¡M⦇ν*xvec⦈⟨N) Q';
                    extractFrame Q = AQ, ΨQ; distinct AQ;
                    AP ♯* Ψ; AP ♯* ΨQ; AP ♯* P; AP ♯* N; AP ♯* P';
                    AP ♯* Q; AP ♯* Q'; AP ♯* AQ; AP ♯* xvec; AQ ♯* Ψ; AQ ♯* ΨP;
                    AQ ♯* P; AQ ♯* N; AQ ♯* P'; AQ ♯* Q; AQ ♯* Q'; distinct xvec;
                    AP ♯* M; AQ ♯* M; xvec ♯* M;
                    AQ ♯* xvec; xvec ♯* Ψ; xvec ♯* ΨP; xvec ♯* ΨQ; xvec ♯* P;
                    xvec ♯* Q; AP ♯* C; AQ ♯* C; xvec ♯* C 
                    Prop C Ψ (P  Q) (¡M⦇ν*xvec⦈⟨N) (P'  Q')"
  and   rBrComm2: "Ψ ΨQ P M xvec N P' AP ΨP Q Q' AQ C.
                   Ψ  ΨQ  P ¡M⦇ν*xvec⦈⟨N  P'; C. Prop C (Ψ  ΨQ) P (¡M⦇ν*xvec⦈⟨N) P';
                    extractFrame P = AP, ΨP; distinct AP;
                    Ψ  ΨP  Q ¿MN  Q'; C. Prop C (Ψ  ΨP) Q (¿MN) Q';
                    extractFrame Q = AQ, ΨQ; distinct AQ;
                    AP ♯* Ψ; AP ♯* ΨQ; AP ♯* P; AP ♯* N; AP ♯* P';
                    AP ♯* Q; AP ♯* Q'; AP ♯* AQ; AP ♯* xvec; AQ ♯* Ψ; AQ ♯* ΨP;
                    AQ ♯* P; AQ ♯* N; AQ ♯* P'; AQ ♯* Q; AQ ♯* Q'; distinct xvec;
                    AP ♯* M; AQ ♯* M; xvec ♯* M;
                    AQ ♯* xvec; xvec ♯* Ψ; xvec ♯* ΨP; xvec ♯* ΨQ; xvec ♯* P;
                    xvec ♯* Q; AP ♯* C; AQ ♯* 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 Ψ (⦇νxP) (τ) (⦇νx(⦇ν*xvecP'))"
  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 Ψ (⦇νxP) (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 Ψ (⦇νxP) (¡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 Ψ (⦇νxP) α (⦇νxP')"
  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 Ψ ΨQ P α P' Q AQ α' 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 Ψ  ΨQ  P α  P' extractFrame Q = AQ, ΨQ distinct AQ
  moreover from bn α ♯* subject α distinct(bn α)
  have "C. Prop C (Ψ  ΨQ) P α P'" by(metis cPar1)
  moreover note AQ ♯* P AQ ♯* Q AQ ♯* Ψ AQ ♯* α AQ ♯* P' AQ ♯* C
    bn α ♯* Q distinct(bn α) bn α ♯* Ψ bn α ♯* Ψ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) AQ ♯* 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 Ψ ΨP Q α Q' P AP α' 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 Ψ  ΨP  Q α  Q' extractFrame P = AP, ΨP distinct AP
  moreover from bn α ♯* subject α distinct(bn α)
  have "C. Prop C (Ψ  ΨP) Q α Q'" by(auto intro: cPar2)

  moreover note AP ♯* P AP ♯* Q AP ♯* Ψ AP ♯* α AP ♯* Q' AP ♯* C
    bn α ♯* Q distinct(bn α) bn α ♯* Ψ bn α ♯* Ψ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 Ψ ΨQ P M N P' AP ΨP Q K xvec Q' AQ α 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 Ψ ΨQ P M xvec N P' AP ΨP Q K Q' AQ α 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 Ψ ΨQ P M N P' AP ΨP Q Q' AQ α C P'')
  then show ?case by(simp add: rBrMerge residualInject)
next
  case(cBrComm1 Ψ ΨQ P M N P' AP ΨP Q xvec Q' AQ α C P'')
  from cBrComm1 AP ♯* M AQ ♯* 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 Ψ ΨQ P M xvec N P' AP ΨP Q Q' AQ α C P'')
  from cBrComm2 AP ♯* M AQ ♯* 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 Ψ (⦇νxP) (τ) (⦇νx(⦇ν*xvecP'))"
    by(fastforce intro: rBrClose)
  then show ?case using τ  ⦇νx(⦇ν*xvecP') = α  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 Ψ (⦇νxP) (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 Ψ (⦇νxP) (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 Ψ (⦇νxP) (¡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 Ψ (⦇νxP) (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 α  (⦇νxP') = α'  P''
  moreover from bn α ♯* α' have "bn α ♯* (bn α')" by auto
  moreover note distinct (bn α) distinct(bn α')
  moreover from bn α ♯* subject α bn α' ♯* subject α'
  have "bn α ♯* (α  ⦇νxP')" 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  (⦇νxP')" and "(bn(p  α)) ♯* α"
    and "(bn(p  α)) ♯* (⦇νxP')"
    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 Ψ (⦇νxP) α (⦇νxP')"
    by(rule rScope)
  with bn α ♯* Ψ bn α ♯* P x  α bn α ♯* subject α bn α ♯* C bn α ♯* (bn α') S distinctPerm p bn(p  α) ♯* α bn(p  α) ♯* (⦇νxP')
  have "Prop C Ψ (⦇νxP) (p  α) (p  (⦇νxP'))"
    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 Ψ (MN⟩.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: "Ψ ΨQ P M xvec N  P' AQ Q C.
                   Ψ  ΨQ  P M⦇ν*xvec⦈⟨N  P'; extractFrame Q = AQ, ΨQ; distinct AQ;
                    C. Prop C (Ψ  ΨQ) P M (⦇ν*xvecN ≺' P');
                    AQ ♯* P; AQ ♯* Q; AQ ♯* Ψ; AQ ♯* M;
                    AQ ♯* xvec; AQ ♯* N; AQ ♯* P'; AQ ♯* C; xvec ♯* Q;
                    xvec ♯* Ψ; xvec ♯* ΨQ; xvec ♯* P; xvec ♯* M; xvec ♯* C 
                    Prop C Ψ (P  Q) M (⦇ν*xvecN ≺' (P'  Q))"
  and   rPar2: "Ψ ΨP Q M xvec N  Q' AP P C.
                   Ψ  ΨP  Q M⦇ν*xvec⦈⟨N  Q'; extractFrame P = AP, ΨP; distinct AP;
                    C. Prop C (Ψ  ΨP) Q M (⦇ν*xvecN ≺' Q');
                    AP ♯* P; AP ♯* Q; AP ♯* Ψ; AP ♯* M;
                    AP ♯* xvec; AP ♯* N; AP ♯* Q'; AP ♯* C; xvec ♯* P;
                    xvec ♯* Ψ; xvec ♯* ΨP; xvec ♯* Q; xvec ♯* M; xvec ♯* C 
                    Prop C Ψ (P  Q) M (⦇ν*xvecN ≺' (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 Ψ (⦇νxP) 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 (⦇ν*xvecN ≺' P');
                    x  Ψ; x  M; x  xvec; x  N; xvec ♯* Ψ; xvec ♯* P; xvec ♯* M;
                    x  C; xvec ♯* C 
                    Prop C Ψ (⦇νxP) M (⦇ν*xvecN ≺' ⦇νxP')"
  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 Ψ ΨQ P α P' Q AQ C)
  then show ?case by(force intro: rPar1 simp add: residualInject)
next
  case(cPar2 Ψ ΨP Q α Q' P AP 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 Ψ ΨQ P M N P' AP ΨP Q xvec Q' AQ C B)
  then show ?case by(simp add: residualInject)
next
  case (cBrComm2 Ψ ΨQ P M xvec N P' AP ΨP Q Q' AQ 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 Ψ (MN⟩.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: "Ψ ΨQ P M xvec N  P' AQ Q C.
                   Ψ  ΨQ  P ¡M⦇ν*xvec⦈⟨N  P'; extractFrame Q = AQ, ΨQ; distinct AQ;
                    C. Prop C (Ψ  ΨQ) P M (⦇ν*xvecN ≺' P');
                    AQ ♯* P; AQ ♯* Q; AQ ♯* Ψ; AQ ♯* M;
                    AQ ♯* xvec; AQ ♯* N; AQ ♯* P'; AQ ♯* C; xvec ♯* Q;
                    xvec ♯* Ψ; xvec ♯* ΨQ; xvec ♯* P; xvec ♯* M; xvec ♯* C 
                    Prop C Ψ (P  Q) M (⦇ν*xvecN ≺' (P'  Q))"
  and   rPar2: "Ψ ΨP Q M xvec N  Q' AP P C.
                   Ψ  ΨP  Q ¡M⦇ν*xvec⦈⟨N  Q'; extractFrame P = AP, ΨP; distinct AP;
                    C. Prop C (Ψ  ΨP) Q M (⦇ν*xvecN ≺' Q');
                    AP ♯* P; AP ♯* Q; AP ♯* Ψ; AP ♯* M;
                    AP ♯* xvec; AP ♯* N; AP ♯* Q'; AP ♯* C; xvec ♯* P;
                    xvec ♯* Ψ; xvec ♯* ΨP; xvec ♯* Q; xvec ♯* M; xvec ♯* C 
                    Prop C Ψ (P  Q) M (⦇ν*xvecN ≺' (P  Q'))"
  and   rBrComm1: "Ψ ΨQ P M N P' AP ΨP Q xvec Q' AQ C.
                   Ψ  ΨQ  P ¿MN  P';
                    extractFrame P = AP, ΨP; distinct AP;
                    Ψ  ΨP  Q ¡M⦇ν*xvec⦈⟨N  Q'; C. Prop C (Ψ  ΨP) Q M (⦇ν*xvecN ≺' Q');
                    extractFrame Q = AQ, ΨQ; distinct AQ;
                    AP ♯* Ψ; AP ♯* ΨQ; AP ♯* P; AP ♯* N; AP ♯* P';
                    AP ♯* Q; AP ♯* Q'; AP ♯* AQ; AP ♯* xvec; AQ ♯* Ψ; AQ ♯* ΨP;
                    AQ ♯* P; AQ ♯* N; AQ ♯* P'; AQ ♯* Q; AQ ♯* Q'; distinct xvec;
                    AP ♯* M; AQ ♯* M; xvec ♯* M;
                    AQ ♯* xvec; xvec ♯* Ψ; xvec ♯* ΨP; xvec ♯* ΨQ; xvec ♯* P;
                    xvec ♯* Q; AP ♯* C; AQ ♯* C; xvec ♯* C 
                    Prop C Ψ (P  Q) M (⦇ν*xvecN ≺' (P'  Q'))" (* Removed: "⋀C. Prop C (Ψ ⊗ ΨQ) P (M⦇N⦈ ≺ P');" *)
  and   rBrComm2: "Ψ ΨQ P M xvec N P' AP ΨP Q Q' AQ C.
                   Ψ  ΨQ  P ¡M⦇ν*xvec⦈⟨N  P'; C. Prop C (Ψ  ΨQ) P M (⦇ν*xvecN ≺' P');
                    extractFrame P = AP, ΨP; distinct AP;
                    Ψ  ΨP  Q ¿MN  Q';
                    extractFrame Q = AQ, ΨQ; distinct AQ;
                    AP ♯* Ψ; AP ♯* ΨQ; AP ♯* P; AP ♯* N; AP ♯* P';
                    AP ♯* Q; AP ♯* Q'; AP ♯* AQ; AP ♯* xvec; AQ ♯* Ψ; AQ ♯* ΨP;
                    AQ ♯* P; AQ ♯* N; AQ ♯* P'; AQ ♯* Q; AQ ♯* Q'; distinct xvec;
                    AP ♯* M; AQ ♯* M; xvec ♯* M;
                    AQ ♯* xvec; xvec ♯* Ψ; xvec ♯* ΨP; xvec ♯* ΨQ; xvec ♯* P;
                    xvec ♯* Q; AP ♯* C; AQ ♯* C; xvec ♯* C 
                    Prop C Ψ (P  Q) M (⦇ν*xvecN ≺' (P'  Q'))" (* Removed: "⋀C. Prop C (Ψ ⊗ Ψ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 Ψ (⦇νxP) 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 (⦇ν*xvecN ≺' P');
                    x  Ψ; x  M; x  xvec; x  N; xvec ♯* Ψ; xvec ♯* P; xvec ♯* M;
                    x  C; xvec ♯* C 
                    Prop C Ψ (⦇νxP) M (⦇ν*xvecN ≺' ⦇νxP')"
  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 Ψ ΨQ P α P' Q AQ C)
  then show ?case by(force intro: rPar1 simp add: residualInject)
next
  case(cPar2 Ψ ΨP Q α Q' P AP 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 Ψ ΨQ P M N P' AP ΨP Q xvec Q' AQ C B)
  then show ?case by(force intro: rBrComm1 simp add: residualInject)
next
  case (cBrComm2 Ψ ΨQ P M xvec N P' AP ΨP Q Q' AQ 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 "⦇ν*xvecB = ⦇ν*yvec([xvec yvec] 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] v N) ≺' ([xvec yvec] v P))"
proof -
  from assms have "⦇ν*xvec(N ≺' P) = ⦇ν*yvec([xvec yvec] 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: "⦇ν*xvecN ≺' P = ⦇ν*yvecM ≺' 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: "⦇ν*xvecN ≺' P = ⦇ν*yvecM ≺' 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 "⦇ν*yvecM ≺' Q = ⦇ν*xvec([yvec xvec] v M) ≺' ([yvec xvec] v Q)"
          by(simp add: alphaBoundOutputChain'')
        with Eq have "N = [yvec xvec] v M" by simp
        with xEq yEq have "N = [(y, x)]  [yvec' xvec'] v M"
          by simp
        with xSuppN have ySuppM: "y  supp([yvec' xvec'] 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'] 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 Ψ ΨQ P M N P' AP ΨP Q xvec Q' AQ α P'')
  note Ψ  Ψ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 Ψ ΨQ P M xvec N P' AP ΨP Q Q' AQ α P'')
  note Ψ  Ψ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 "Ψ  ⦇νxP 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 "Ψ  ⦇νxP ¡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 Ψ (MN⟩.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: "Ψ ΨQ P M xvec N  P' AQ Q C.
                   Ψ  ΨQ  P M⦇ν*xvec⦈⟨N  P'; extractFrame Q = AQ, ΨQ; distinct AQ;
                    C. Prop C (Ψ  ΨQ) P M xvec N P';
                    AQ ♯* P; AQ ♯* Q; AQ ♯* Ψ; AQ ♯* M;
                    AQ ♯* xvec; AQ ♯* N; AQ ♯* P'; AQ ♯* C; xvec ♯* Q;
                    xvec ♯* Ψ; xvec ♯* ΨQ; xvec ♯* P; xvec ♯* M; xvec ♯* C 
                    Prop C Ψ (P  Q) M xvec N (P'  Q)"
  and   rPar2: "Ψ ΨP Q M xvec N  Q' AP P C.
                   Ψ  ΨP  Q M⦇ν*xvec⦈⟨N  Q'; extractFrame P = AP, ΨP;  distinct AP;
                    C. Prop C (Ψ  ΨP) Q M xvec N Q';
                    AP ♯* P; AP ♯* Q; AP ♯* Ψ; AP ♯* M;
                    AP ♯* xvec; AP ♯* N; AP ♯* Q'; AP ♯* C; xvec ♯* Q;
                    xvec ♯* Ψ; xvec ♯* Ψ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 Ψ (⦇νxP) 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 Ψ (⦇νxP) M xvec N (⦇νxP')"
  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 Ψ (MN⟩.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: "Ψ ΨQ P M xvec N  P' AQ Q C.
                   Ψ  ΨQ  P ¡M⦇ν*xvec⦈⟨N  P'; extractFrame Q = AQ, ΨQ; distinct AQ;
                    C. Prop C (Ψ  ΨQ) P M xvec N P';
                    AQ ♯* P; AQ ♯* Q; AQ ♯* Ψ; AQ ♯* M;
                    AQ ♯* xvec; AQ ♯* N; AQ ♯* P'; AQ ♯* C; xvec ♯* Q;
                    xvec ♯* Ψ; xvec ♯* ΨQ; xvec ♯* P; xvec ♯* M; xvec ♯* C 
                    Prop C Ψ (P  Q) M xvec N (P'  Q)"
  and   rPar2: "Ψ ΨP Q M xvec N  Q' AP P C.
                   Ψ  ΨP  Q ¡M⦇ν*xvec⦈⟨N  Q'; extractFrame P = AP, ΨP;  distinct AP;
                    C. Prop C (Ψ  ΨP) Q M xvec N Q';
                    AP ♯* P; AP ♯* Q; AP ♯* Ψ; AP ♯* M;
                    AP ♯* xvec; AP ♯* N; AP ♯* Q'; AP ♯* C; xvec ♯* Q;
                    xvec ♯* Ψ; xvec ♯* ΨP; xvec ♯* P; xvec ♯* M; xvec ♯* C 
                    Prop C Ψ (P  Q) M xvec N (P  Q')"
  and   rBrComm1: "Ψ ΨQ P M N P' AP ΨP Q xvec Q' AQ C.
                   Ψ  ΨQ  P ¿MN  P';
                    extractFrame P = AP, ΨP; distinct AP;
                    Ψ  ΨP  Q ¡M⦇ν*xvec⦈⟨N  Q'; C. Prop C (Ψ  ΨP) Q M xvec N Q';
                    extractFrame Q = AQ, ΨQ; distinct AQ;
                    AP ♯* Ψ; AP ♯* ΨQ; AP ♯* P; AP ♯* N; AP ♯* P';
                    AP ♯* Q; AP ♯* Q'; AP ♯* AQ; AP ♯* xvec; AQ ♯* Ψ; AQ ♯* ΨP;
                    AQ ♯* P; AQ ♯* N; AQ ♯* P'; AQ ♯* Q; AQ ♯* Q'; distinct xvec;
                    AP ♯* M; AQ ♯* M; xvec ♯* M;
                    AQ ♯* xvec; xvec ♯* Ψ; xvec ♯* ΨP; xvec ♯* ΨQ; xvec ♯* P;
                    xvec ♯* Q; AP ♯* C; AQ ♯* C; xvec ♯* C 
                    Prop C Ψ (P  Q) M xvec N (P'  Q')" (* Removed: "⋀C. Prop C (Ψ ⊗ ΨQ) P (M⦇N⦈ ≺ P');" *)
  and   rBrComm2: "Ψ ΨQ P M xvec N P' AP ΨP Q Q' AQ C.
                   Ψ  ΨQ  P ¡M⦇ν*xvec⦈⟨N  P'; C. Prop C (Ψ  ΨQ) P M xvec N P';
                    extractFrame P = AP, ΨP; distinct AP;
                    Ψ  ΨP  Q ¿MN  Q';
                    extractFrame Q = AQ, ΨQ; distinct AQ;
                    AP ♯* Ψ; AP ♯* ΨQ; AP ♯* P; AP ♯* N; AP ♯* P';
                    AP ♯* Q; AP ♯* Q'; AP ♯* AQ; AP ♯* xvec; AQ ♯* Ψ; AQ ♯* ΨP;
                    AQ ♯* P; AQ ♯* N; AQ ♯* P'; AQ ♯* Q; AQ ♯* Q'; distinct xvec;
                    AP ♯* M; AQ ♯* M; xvec ♯* M;
                    AQ ♯* xvec; xvec ♯* Ψ; xvec ♯* ΨP; xvec ♯* ΨQ; xvec ♯* P;
                    xvec ♯* Q; AP ♯* C; AQ ♯* C; xvec ♯* C 
                    Prop C Ψ (P  Q) M xvec N (P'  Q')" (* Removed: "⋀C. Prop C (Ψ ⊗ Ψ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 Ψ (⦇νxP) 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 Ψ (⦇νxP) M xvec N (⦇νxP')"
  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 MN  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 MN  P'; C. Prop C Ψ P M N P'; (φ, P)  set Cs; Ψ  φ; guarded P 
                                        Prop C Ψ (Cases Cs) M N P'"
  and   rPar1: "Ψ ΨQ P M N P' AQ Q C.
                   Ψ  ΨQ  P MN  P'; extractFrame Q = AQ, ΨQ; distinct AQ;
                   C. Prop C (Ψ  ΨQ) P M N P'; distinct AQ;
                   AQ ♯* P; AQ ♯* Q; AQ ♯* Ψ; AQ ♯* M; AQ ♯* N;
                   AQ ♯* P'; AQ ♯* C 
                   Prop C Ψ (P  Q) M N (P'  Q)"
  and   rPar2: "Ψ ΨP Q M N Q' AP P C.
                   Ψ  ΨP  Q MN  Q'; extractFrame P = AP, ΨP; distinct AP;
                   C. Prop C (Ψ  ΨP) Q M N Q'; distinct AP;
                   AP ♯* P; AP ♯* Q; AP ♯* Ψ; AP ♯* M; AP ♯* N;
                   AP ♯* Q'; AP ♯* C 
                   Prop C Ψ (P  Q) M N (P  Q')"
  and   rScope: "Ψ P M N P' x C.
                    Ψ  P MN  P'; C. Prop C Ψ P M N P'; x  Ψ; x  M; x  N; x  C 
                     Prop C Ψ (⦇νxP) M N (⦇νxP')"
  and   rBang:    "Ψ P M N P' C.
                     Ψ  P  !P MN  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=="MN  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 Ψ ΨQ P α P' Q AQ 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 ¿MN  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 ¿MN  P'; C. Prop C Ψ P M N P'; (φ, P)  set Cs; Ψ  φ; guarded P 
                                        Prop C Ψ (Cases Cs) M N P'"
  and   rPar1: "Ψ ΨQ P M N P' AQ Q C.
                   Ψ  ΨQ  P ¿MN  P'; extractFrame Q = AQ, ΨQ; distinct AQ;
                   C. Prop C (Ψ  ΨQ) P M N P'; distinct AQ;
                   AQ ♯* P; AQ ♯* Q; AQ ♯* Ψ; AQ ♯* M; AQ ♯* N;
                   AQ ♯* P'; AQ ♯* C 
                   Prop C Ψ (P  Q) M N (P'  Q)"
  and   rPar2: "Ψ ΨP Q M N Q' AP P C.
                   Ψ  ΨP  Q ¿MN  Q'; extractFrame P = AP, ΨP; distinct AP;
                   C. Prop C (Ψ  ΨP) Q M N Q'; distinct AP;
                   AP ♯* P; AP ♯* Q; AP ♯* Ψ; AP ♯* M; AP ♯* N;
                   AP ♯* Q'; AP ♯* C 
                   Prop C Ψ (P  Q) M N (P  Q')"
  and   rBrMerge: "Ψ ΨQ P M N P' AP ΨP Q Q' AQ C.
                    Ψ  ΨQ  P  ¿MN  P'; C. Prop C (Ψ  ΨQ) P M N P';
                    extractFrame P = AP, ΨP; distinct AP;
                    Ψ  ΨP  Q  ¿MN  Q'; C. Prop C (Ψ  ΨP) Q M N Q';
                    extractFrame Q = AQ, ΨQ; distinct AQ;
                    AP ♯* Ψ; AP ♯* ΨQ; AP ♯* P; AP ♯* N; AP ♯* P';
                    AP ♯* Q; AP ♯* Q'; AP ♯* AQ; AP ♯* M; AQ ♯* M;
                    AQ ♯* Ψ; AQ ♯* ΨP; AQ ♯* P; AQ ♯* N; AQ ♯* P';
                    AQ ♯* Q; AQ ♯* Q'; AP ♯* C; AQ ♯* C;
                    AP ♯* M; AQ ♯* M 
                    Prop C Ψ (P  Q) M N (P'  Q')"
  and   rScope: "Ψ P M N P' x C.
                    Ψ  P ¿MN  P'; C. Prop C Ψ P M N P'; x  Ψ; x  M; x  N; x  C 
                     Prop C Ψ (⦇νxP) M N (⦇νxP')"
  and   rBang:    "Ψ P M N P' C.
                     Ψ  P  !P ¿MN  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=="¿MN  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 Ψ ΨQ P α P' Q AQ 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: "Ψ ΨQ P P' AQ Q C.
                   Ψ  ΨQ  P τ  P'; extractFrame Q = AQ, ΨQ; distinct AQ;
                   C. Prop C (Ψ  ΨQ) P P';
                   AQ ♯* P; AQ ♯* Q; AQ ♯* Ψ;
                   AQ ♯* P'; AQ ♯* C 
                   Prop C Ψ (P  Q) (P'  Q)"
  and   rPar2: "Ψ ΨP Q Q' AP P C.
                   Ψ  ΨP  Q τ  Q'; extractFrame P = AP, ΨP; distinct AP;
                   C. Prop C (Ψ  ΨP) Q Q';
                   AP ♯* P; AP ♯* Q; AP ♯* Ψ;
                   AP ♯* Q'; AP ♯* C 
                   Prop C Ψ (P  Q) (P  Q')"
  and   rComm1: "Ψ ΨQ P M N P' AP ΨP Q K xvec Q' AQ C.
                   Ψ  ΨQ  P MN  P'; extractFrame P = AP, ΨP; distinct AP;
                    Ψ  ΨP  Q K⦇ν*xvec⦈⟨N  Q'; extractFrame Q = AQ, ΨQ; distinct AQ;
                    Ψ  ΨP  ΨQ  M  K;
                    AP ♯* Ψ; AP ♯* ΨQ; AP ♯* P; AP ♯* M; AP ♯* N; AP ♯* P';
                    AP ♯* Q; AP ♯* Q'; AP ♯* AQ; AP ♯* xvec; AQ ♯* Ψ; AQ ♯* ΨP;
                    AQ ♯* P; AQ ♯* N; AQ ♯* P'; AQ ♯* Q; AQ ♯* K; AQ ♯* Q';
                    AQ ♯* xvec; xvec ♯* Ψ; xvec ♯* ΨP; xvec ♯* ΨQ; xvec ♯* P; xvec ♯* M;
                    xvec ♯* Q; xvec ♯* K; AP ♯* C; AQ ♯* C; xvec ♯* C 
                    Prop C Ψ (P  Q) (⦇ν*xvec(P'  Q'))"
  and   rComm2: "Ψ ΨQ P M xvec N P' AP ΨP Q K Q' AQ C.
                   Ψ  ΨQ  P M⦇ν*xvec⦈⟨N  P';  extractFrame P = AP, ΨP; distinct AP;
                    Ψ  ΨP  Q KN  Q'; extractFrame Q = AQ, ΨQ; distinct AQ;
                    Ψ  ΨP  ΨQ  M  K;
                    AP ♯* Ψ; AP ♯* ΨQ; AP ♯* P; AP ♯* M; AP ♯* N; AP ♯* P';
                    AP ♯* Q; AP ♯* Q'; AP ♯* AQ; AP ♯* xvec; AQ ♯* Ψ; AQ ♯* ΨP;
                    AQ ♯* P; AQ ♯* N; AQ ♯* P'; AQ ♯* Q; AQ ♯* K; AQ ♯* Q';
                    AQ ♯* xvec; xvec ♯* Ψ; xvec ♯* ΨP; xvec ♯* ΨQ; xvec ♯* P; xvec ♯* M;
                    xvec ♯* Q; xvec ♯* K; AP ♯* C; AQ ♯* C; xvec ♯* C 
                    Prop C Ψ (P  Q) (⦇ν*xvec(P'  Q'))"
  and   rBrClose: "Ψ P M xvec N P' AP ΨP x C.
                    Ψ  P  ¡M⦇ν*xvec⦈⟨N  P';
                     x  supp M;
                     distinct xvec; xvec ♯* Ψ; xvec ♯* P;
                     xvec ♯* M;
                     x  Ψ; x  xvec 
                     Prop C Ψ (⦇νxP) (⦇νx(⦇ν*xvecP'))"
  and   rScope: "Ψ P P' x C.
                    Ψ  P τ  P'; C. Prop C Ψ P P'; x  Ψ; x  C 
                     Prop C Ψ (⦇νxP) (⦇νxP')"
  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 Ψ ΨQ P α P' AQ Q C P'')
  then show ?case by(force intro: rPar1 simp add: residualInject)
next
  case(cPar2 Ψ ΨP Q α Q' AP P C Q'')
  then show ?case by(force intro: rPar2 simp add: residualInject)
next
  case(cComm1 Ψ ΨQ P M N P' AP ΨP Q K xvec Q' AQ C PQ)
  then show ?case by(force intro: rComm1 simp add: residualInject)
next
  case(cComm2 Ψ ΨQ P M xvec N P' AP ΨP Q' AQ 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 AP   :: "name list"
    and Ψ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 = AP, ΨP"
  and   "distinct AP"
  and   rAlpha: "Ψ P AP ΨP p Rs C. AP ♯* Ψ; AP ♯* P; AP ♯* (p  AP); AP ♯* Rs; AP ♯* C;
                                         set p  set AP × set(p  AP); distinctPerm p;
                                          Prop C Ψ P Rs AP ΨP  Prop C Ψ P Rs (p  AP) (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: "Ψ 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 Ψ (MN⟩.P) (KN  P) ([]) (𝟭)"
  and   rBrOutput: "Ψ M K N P C. Ψ  M  K  Prop C Ψ (MN⟩.P) (¡KN  P) ([]) (𝟭)"
  and   rCase: "Ψ P Rs φ Cs AP ΨP C. Ψ  P  Rs; extractFrame P = AP, ΨP; distinct AP; C. Prop C Ψ P Rs AP ΨP;
                                            (φ, P)  set Cs; Ψ  φ; guarded P;  ΨP  𝟭; (supp ΨP) = ({}::name set);
                                            AP ♯* Ψ; AP ♯* P; AP ♯* Rs; AP ♯* C  Prop C Ψ (Cases Cs) Rs ([]) (𝟭)"
  and   rPar1: "Ψ ΨQ P α P' AQ Q AP ΨP C.
                   Ψ  ΨQ  P α  P';
                   extractFrame P = AP, ΨP; distinct AP;
                   extractFrame Q = AQ, ΨQ; distinct AQ;
                   C. Prop C (Ψ  ΨQ) P (α  P') AP ΨP; distinct(bn α);
                   AP ♯* P; AP ♯* Q; AP ♯* Ψ; AP ♯* α; AP ♯* P'; AP ♯* AQ; AP ♯* ΨQ;
                   AQ ♯* P; AQ ♯* Q; AQ ♯* Ψ; AQ ♯* α; AQ ♯* P'; AQ ♯* ΨP;
                   bn α ♯* Ψ; bn α ♯* P; bn α ♯* Q; bn α ♯* subject α; bn α ♯* ΨP; bn α ♯* ΨQ;
                   AP ♯* C; AQ ♯* C; bn α ♯* C 
                   Prop C Ψ (P  Q) (α  (P'  Q)) (AP@AQ) (ΨP  ΨQ)"
  and   rPar2: "Ψ ΨP Q α Q' AP P AQ ΨQ C.
                   Ψ  ΨP  Q α  Q';
                   extractFrame P = AP, ΨP; distinct AP;
                   extractFrame Q = AQ, ΨQ; distinct AQ;
                   C. Prop C (Ψ  ΨP) Q (α  Q') AQ ΨQ; distinct(bn α);
                   AP ♯* P; AP ♯* Q; AP ♯* Ψ; AP ♯* α; AP ♯* Q'; AP ♯* AQ; AP ♯* ΨQ;
                   AQ ♯* P; AQ ♯* Q; AQ ♯* Ψ; AQ ♯* α; AQ ♯* Q'; AQ ♯* ΨP;
                   bn α ♯* Ψ; bn α ♯* P; bn α ♯* Q; bn α ♯* subject α; bn α ♯* ΨP; bn α ♯* ΨQ;
                   AP ♯* C; AQ ♯* C; bn α ♯* C 
                   Prop C Ψ (P  Q) (α  (P  Q')) (AP@AQ) (ΨP  ΨQ)"
  and   rComm1: "Ψ ΨQ P M N P' AP ΨP Q K xvec Q' AQ C.
                   Ψ  ΨQ  P MN  P'; extractFrame P = AP, ΨP; distinct AP;
                   C. Prop C (Ψ  ΨQ) P ((MN)  P') AP ΨP;
                    Ψ  ΨP  Q K⦇ν*xvec⦈⟨N  Q'; extractFrame Q = AQ, ΨQ; distinct AQ;
                    Ψ  ΨP  ΨQ  M  K;
                   C. Prop C (Ψ  ΨP) Q (K⦇ν*xvec⦈⟨N  Q') AQ ΨQ; distinct xvec;
                    AP ♯* Ψ; AP ♯* ΨQ; AP ♯* P; AP ♯* M; AP ♯* N; AP ♯* P';
                    AP ♯* Q; AP ♯* Q'; AP ♯* AQ; AP ♯* xvec; AQ ♯* Ψ; AQ ♯* ΨP;
                    AQ ♯* P; AQ ♯* N; AQ ♯* P'; AQ ♯* Q; AQ ♯* K; AQ ♯* Q';
                    AQ ♯* xvec; xvec ♯* Ψ; xvec ♯* ΨP; xvec ♯* ΨQ; xvec ♯* P; xvec ♯* M;
                    xvec ♯* Q; xvec ♯* K; AP ♯* C; AQ ♯* C; xvec ♯* C 
                    Prop C Ψ (P  Q) (τ  ⦇ν*xvec(P'  Q')) (AP@AQ) (ΨP  ΨQ)"
  and   rComm2: "Ψ ΨQ P M xvec N P' AP ΨP Q K Q' AQ C.
                   Ψ  ΨQ  P M⦇ν*xvec⦈⟨N  P'; extractFrame P = AP, ΨP; distinct AP;
                   C. Prop C (Ψ  ΨQ) P (M⦇ν*xvec⦈⟨N  P') AP ΨP;
                    Ψ  ΨP  Q KN  Q'; extractFrame Q = AQ, ΨQ; distinct AQ;
                   C. Prop C (Ψ  ΨP) Q (KN  Q') AQ ΨQ;
                    Ψ  ΨP  ΨQ  M  K; distinct xvec;
                    AP ♯* Ψ; AP ♯* ΨQ; AP ♯* P; AP ♯* M; AP ♯* N; AP ♯* P';
                    AP ♯* Q; AP ♯* Q'; AP ♯* AQ; AP ♯* xvec; AQ ♯* Ψ; AQ ♯* ΨP;
                    AQ ♯* P; AQ ♯* N; AQ ♯* P'; AQ ♯* Q; AQ ♯* K; AQ ♯* Q';
                    AQ ♯* xvec; xvec ♯* Ψ; xvec ♯* ΨP; xvec ♯* ΨQ; xvec ♯* P; xvec ♯* M;
                    xvec ♯* Q; xvec ♯* K; AP ♯* C; AQ ♯* C; xvec ♯* C 
                    Prop C Ψ (P  Q) (τ  ⦇ν*xvec(P'  Q')) (AP@AQ) (ΨP  ΨQ)"
  and   rBrMerge: "Ψ ΨQ P M N P' AP ΨP Q Q' AQ C.
                    Ψ  ΨQ  P  ¿MN  P'; C. Prop C (Ψ  ΨQ) P (¿MN  P') AP ΨP;
                    extractFrame P = AP, ΨP; distinct AP;
                    Ψ  ΨP  Q  ¿MN  Q'; C. Prop C (Ψ  ΨP) Q (¿MN  Q') AQ ΨQ;
                    extractFrame Q = AQ, ΨQ; distinct AQ;
                    AP ♯* Ψ; AP ♯* ΨQ; AP ♯* P; AP ♯* N; AP ♯* P';
                    AP ♯* Q; AP ♯* Q'; AP ♯* AQ; AP ♯* M; AQ ♯* M;
                    AQ ♯* Ψ; AQ ♯* ΨP; AQ ♯* P; AQ ♯* N; AQ ♯* P';
                    AQ ♯* Q; AQ ♯* Q'; AP ♯* C; AQ ♯* C;
                    AP ♯* M; AQ ♯* M 
                    Prop C Ψ (P  Q) (¿MN  (P'  Q')) (AP@AQ) (ΨP  ΨQ)"
  and   rBrComm1:"Ψ ΨQ P M N P' AP ΨP Q xvec Q' AQ C.
                   Ψ  ΨQ  P ¿MN  P'; extractFrame P = AP, ΨP; distinct AP;
                   C. Prop C (Ψ  ΨQ) P ((¿MN)  P') AP ΨP;
                    Ψ  ΨP  Q ¡M⦇ν*xvec⦈⟨N  Q'; extractFrame Q = AQ, ΨQ; distinct AQ;
                   C. Prop C (Ψ  ΨP) Q (¡M⦇ν*xvec⦈⟨N  Q') AQ ΨQ; distinct xvec;
                    AP ♯* Ψ; AP ♯* ΨQ; AP ♯* P; AP ♯* N; AP ♯* P';
                    AP ♯* Q; AP ♯* Q'; AP ♯* AQ; AP ♯* xvec; AQ ♯* Ψ; AQ ♯* ΨP;
                    AQ ♯* P; AQ ♯* N; AQ ♯* P'; AQ ♯* Q; AQ ♯* Q';
                    AQ ♯* xvec; xvec ♯* Ψ; xvec ♯* ΨP; xvec ♯* ΨQ; xvec ♯* P;
                    xvec ♯* Q; AP ♯* C; AQ ♯* C; xvec ♯* C;
                    AP ♯* M; AQ ♯* M; xvec ♯* M 
                    Prop C Ψ (P  Q) (¡M⦇ν*xvec⦈⟨N  (P'  Q')) (AP@AQ) (ΨP  ΨQ)"
  and   rBrComm2:"Ψ ΨQ P M xvec N P' AP ΨP Q Q' AQ C.
                   Ψ  ΨQ  P ¡M⦇ν*xvec⦈⟨N  P'; extractFrame P = AP, ΨP; distinct AP;
                   C. Prop C (Ψ  ΨQ) P (¡M⦇ν*xvec⦈⟨N  P') AP ΨP;
                    Ψ  ΨP  Q ¿MN  Q'; extractFrame Q = AQ, ΨQ; distinct AQ;
                   C. Prop C (Ψ  ΨP) Q (¿MN  Q') AQ ΨQ; distinct xvec;
                    AP ♯* Ψ; AP ♯* ΨQ; AP ♯* P; AP ♯* N; AP ♯* P';
                    AP ♯* Q; AP ♯* Q'; AP ♯* AQ; AP ♯* xvec; AQ ♯* Ψ; AQ ♯* ΨP;
                    AQ ♯* P; AQ ♯* N; AQ ♯* P'; AQ ♯* Q; AQ ♯* Q';
                    AQ ♯* xvec; xvec ♯* Ψ; xvec ♯* ΨP; xvec ♯* ΨQ; xvec ♯* P;
                    xvec ♯* Q; AP ♯* C; AQ ♯* C; xvec ♯* C;
                    AP ♯* M; AQ ♯* M; xvec ♯* M 
                    Prop C Ψ (P  Q) (¡M⦇ν*xvec⦈⟨N  (P'  Q')) (AP@AQ) (ΨP  ΨQ)"
  and   rBrClose: "Ψ P M xvec N P' AP ΨP x C.
                    Ψ  P  ¡M⦇ν*xvec⦈⟨N  P';
                     x  supp M;
                     C. Prop C Ψ P (¡M⦇ν*xvec⦈⟨N  P') AP ΨP;
                     extractFrame P = AP, ΨP; distinct AP;
                     AP ♯* Ψ; AP ♯* P; AP ♯* M; AP ♯* N; AP ♯* P'; AP ♯* xvec;
                     distinct xvec; xvec ♯* Ψ; xvec ♯* ΨP; xvec ♯* P;
                     xvec ♯* M;
                     x  Ψ; x  xvec; x  AP;
                     AP ♯* C; xvec ♯* C; x  C 
                     Prop C Ψ (⦇νxP) (τ  (⦇νx(⦇ν*xvecP'))) (x#AP) ΨP"
  and   rOpen: "Ψ P M xvec yvec N P' x AP ΨP C.
                    Ψ  P M⦇ν*(xvec@yvec)⦈⟨N  P'; extractFrame P = AP, ΨP; distinct AP;
                    C. Prop C Ψ P (M⦇ν*(xvec@yvec)⦈⟨N  P') AP ΨP; x  supp N; x  Ψ; x  M;
                     x  AP; x  xvec; x  yvec; AP ♯* Ψ; AP ♯* P; AP ♯* M; AP ♯* N; AP ♯* P';
                     AP ♯* xvec; AP ♯* yvec; xvec ♯* yvec; distinct xvec; distinct yvec;
                     xvec ♯* Ψ; xvec ♯* P; xvec ♯* M; xvec ♯* ΨP; yvec ♯* ΨP;
                     yvec ♯* Ψ; yvec ♯* P; yvec ♯* M; AP ♯* C; x  C; xvec ♯* C; yvec ♯* C 
                     Prop C Ψ (⦇νxP) (M⦇ν*(xvec@x#yvec)⦈⟨N  P') (x#AP) ΨP"
  and   rBrOpen: "Ψ P M xvec yvec N P' x AP ΨP C.
                    Ψ  P ¡M⦇ν*(xvec@yvec)⦈⟨N  P'; extractFrame P = AP, ΨP; distinct AP;
                    C. Prop C Ψ P (¡M⦇ν*(xvec@yvec)⦈⟨N  P') AP ΨP; x  supp N; x  Ψ; x  M;
                     x  AP; x  xvec; x  yvec; AP ♯* Ψ; AP ♯* P; AP ♯* M; AP ♯* N; AP ♯* P';
                     AP ♯* xvec; AP ♯* yvec; xvec ♯* yvec; distinct xvec; distinct yvec;
                     xvec ♯* Ψ; xvec ♯* P; xvec ♯* M; xvec ♯* ΨP; yvec ♯* ΨP;
                     yvec ♯* Ψ; yvec ♯* P; yvec ♯* M; AP ♯* C; x  C; xvec ♯* C; yvec ♯* C 
                     Prop C Ψ (⦇νxP) (¡M⦇ν*(xvec@x#yvec)⦈⟨N  P') (x#AP) ΨP"
  and   rScope: "Ψ P α P' x AP ΨP C.
                    Ψ  P α  P'; extractFrame P = AP, ΨP; distinct AP;
                    C. Prop C Ψ P (α  P') AP ΨP;
                     x  Ψ; x  α; x  AP; AP ♯* Ψ; AP ♯* P;
                     AP ♯* α; AP ♯* P'; distinct(bn α);
                     bn α ♯* Ψ; bn α ♯* P; bn α ♯* subject α; bn α ♯* ΨP;
                     AP ♯* C; x  C; bn α ♯* C 
                     Prop C Ψ (⦇νxP) (α  (⦇νxP')) (x#AP) ΨP"
  and   rBang:    "Ψ P Rs AP ΨP C.
                     Ψ  P  !P  Rs; guarded P; extractFrame P = AP, ΨP; distinct AP;
                      C. Prop C Ψ (P  !P) Rs AP (ΨP  𝟭); ΨP  𝟭; supp ΨP = ({}::name set);
                      AP ♯* Ψ; AP ♯* P; AP ♯* Rs; AP ♯* C  Prop C Ψ (!P) Rs ([]) (𝟭)"
shows "Prop C Ψ P Rs AP ΨP"
  using Trans FrP distinct AP
proof(nominal_induct  avoiding: AP ΨP C rule: semantics.strong_induct)
  case(cInput Ψ M K xvec N Tvec P AP ΨP C)
  from extractFrame (M⦇λ*xvec N⦈.P) = AP, ΨP
  have "AP = []" and "Ψ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 AP ΨP C)
  from extractFrame (M⦇λ*xvec N⦈.P) = AP, ΨP
  have "AP = []" and "Ψ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 AP ΨP)
  from extractFrame (MN⟩.P) = AP, ΨP
  have "AP = []" and "ΨP = 𝟭"
    by auto
  with Ψ  M  K show ?case
    by(blast intro: rOutput)
next
  case(BrOutput Ψ M K N P AP ΨP)
  from extractFrame (MN⟩.P) = AP, ΨP
  have "AP = []" and "ΨP = 𝟭"
    by auto
  with Ψ  M  K show ?case
    by(blast intro: rBrOutput)
next
  case(Case Ψ P Rs φ Cs AcP ΨcP C)
  obtain AP ΨP where FrP: "extractFrame P = AP, ΨP" and "distinct AP"
    and "AP ♯* (Ψ, P, Rs, C)"
    by(rule freshFrame)
  then have "AP ♯* Ψ" and "AP ♯* P" and "AP ♯* Rs" and "AP ♯* C"
    by simp+
  note Ψ  P  Rs FrP distinct AP
  moreover from FrP distinct AP AP ΨP C. extractFrame P = AP, ΨP; distinct AP  Prop C Ψ P Rs AP ΨP
  have "C. Prop C Ψ P Rs AP ΨP" by simp
  moreover note (φ, P)  set Cs Ψ  φ guarded P
  moreover from guarded P FrP have "ΨP  𝟭" and "supp ΨP = ({}::name set)" by(metis guardedStatEq)+
  moreover note AP ♯* Ψ AP ♯* P AP ♯* Rs AP ♯* C
  ultimately have "Prop C Ψ (Cases Cs) Rs ([]) (𝟭)"
    by(rule rCase)
  then show ?case using extractFrame(Cases Cs) = AcP, ΨcP by simp
next
  case(cPar1 Ψ ΨQ P α P' Q AQ APQ ΨPQ C)
  obtain AP ΨP where FrP: "extractFrame P = AP, ΨP" and "distinct AP"
    "AP ♯* (P, Q, Ψ, α, P', AQ, APQ, C, ΨQ)"
    by(rule freshFrame)
  then have "AP ♯* P" and "AP ♯* Q" and "AP ♯* Ψ" and "AP ♯* α" and "AP ♯* P'"
    and "AP ♯* AQ" and "AP ♯* APQ" and "AP ♯* C" and "AP ♯* ΨQ"
    by simp+

  have FrQ: "extractFrame Q = AQ, ΨQ" by fact

  from AQ ♯* P AP ♯* AQ FrP have "AQ ♯* ΨP"
    by(force dest: extractFrameFreshChain)

  from bn α ♯* P AP ♯* α FrP have "bn α ♯* ΨP"
    by(force dest: extractFrameFreshChain)

  from extractFrame(P  Q) = APQ, ΨPQ FrP FrQ AP ♯* AQ AP ♯* ΨQ AQ ♯* ΨP
  have "(AP@AQ), ΨP  ΨQ = APQ, ΨPQ"
    by simp
  moreover from distinct AP distinct AQ AP ♯* AQ have "distinct(AP@AQ)"
    by(auto simp add: fresh_star_def fresh_def name_list_supp)
  ultimately obtain p where S: "set p  set(AP@AQ) × set((p  AP)@(p  AQ))"  and "distinctPerm p"
    and Ψeq: "ΨPQ = p  (ΨP  ΨQ)" and Aeq: "APQ = (p  AP)@(p  AQ)"
    using AP ♯* APQ AQ ♯* APQ distinct APQ
    apply -
    apply(rule frameChainEq')
    by (assumption | simp add: eqvts)+
  note Ψ  ΨQ  P α  P' FrP distinct AP FrQ distinct AQ

  moreover from FrP distinct AP AP ΨP C. extractFrame P = AP, ΨP; distinct AP  Prop C (Ψ  ΨQ) P (α  P') AP ΨP
  have "C. Prop C (Ψ  ΨQ) P (α  P') AP ΨP" by simp

  moreover note AP ♯* P AP ♯* Q AP ♯* Ψ AP ♯* α AP ♯* P' AP ♯* AQ AP ♯* ΨQ
    AQ ♯* P AQ ♯* Q AQ ♯* Ψ AQ ♯* α AQ ♯* P' AQ ♯* ΨP distinct(bn α)
    bn α ♯* Ψ bn α ♯* P  bn α ♯* Q  bn α ♯* subject α  bn α ♯* ΨP  bn α ♯* ΨQ
    AP ♯* C AQ ♯* C bn α ♯* C
  ultimately have "Prop C Ψ (P  Q) (α  (P'  Q)) (AP@AQ) (ΨP  ΨQ)"
    by(metis rPar1)
  with AP ♯* Ψ AP ♯* P AP ♯* Q AP ♯* α AP ♯* P' AP ♯* APQ AP ♯* C
    AQ ♯* Ψ AQ ♯* P AQ ♯* Q AQ ♯* α AQ ♯* P' AQ ♯* APQ AQ ♯* C
    S distinctPerm p Aeq
  have "Prop C Ψ (P  Q) (α  (P'  Q)) (p  (AP@AQ)) (p  (ΨP  ΨQ))"
    apply -
    apply(rule rAlpha)
    by(assumption | simp add: eqvts)+
  with Ψeq Aeq show ?case by(simp add: eqvts)
next
  case(cPar2 Ψ ΨP Q α Q' P AP APQ ΨPQ C)
  obtain AQ ΨQ where FrQ: "extractFrame Q = AQ, ΨQ" and "distinct AQ"
    "AQ ♯* (P, Q, Ψ, α, Q', AP, APQ, C, ΨP)"
    by(rule freshFrame)
  then have "AQ ♯* P" and "AQ ♯* Q" and "AQ ♯* Ψ" and "AQ ♯* α" and "AQ ♯* Q'"
    and "AQ ♯* AP" and "AQ ♯* APQ" and "AQ ♯* C" and "AQ ♯* ΨP"
    by simp+

  from AQ ♯* AP have "AP ♯* AQ" by simp
  have FrP: "extractFrame P = AP, ΨP" by fact

  from AP ♯* Q AQ ♯* AP FrQ have "AP ♯* ΨQ"
    by(force dest: extractFrameFreshChain)
  from bn α ♯* Q AQ ♯* α FrQ have "bn α ♯* ΨQ"
    by(force dest: extractFrameFreshChain)

  from extractFrame(P  Q) = APQ, ΨPQ FrP FrQ AP ♯* AQ AP ♯* ΨQ AQ ♯* ΨP
  have "(AP@AQ), ΨP  ΨQ = APQ, ΨPQ"
    by simp
  moreover from distinct AP distinct AQ AP ♯* AQ have "distinct(AP@AQ)"
    by(auto simp add: fresh_star_def fresh_def name_list_supp)
  ultimately obtain p where S: "(set p  (set(AP@AQ)) × (set APQ))"  and "distinctPerm p"
    and Ψeq: "ΨPQ = p  (ΨP  ΨQ)" and Aeq: "APQ = ((p  AP)@(p  AQ))"
    using AP ♯* APQ AQ ♯* APQ distinct APQ
    apply -
    apply(rule frameChainEq')
    by(assumption | simp add: eqvts)+

  note Ψ  ΨP  Q α  Q' FrP distinct AP FrQ distinct AQ

  moreover from FrQ distinct AQ AQ ΨQ C. extractFrame Q = AQ, ΨQ; distinct AQ  Prop C (Ψ  ΨP) Q (α  Q') AQ ΨQ
  have "C. Prop C (Ψ  ΨP) Q (α  Q') AQ ΨQ" by simp

  moreover note AP ♯* P AP ♯* Q AP ♯* Ψ AP ♯* α AP ♯* Q' AP ♯* AQ AP ♯* ΨQ
    AQ ♯* P AQ ♯* Q AQ ♯* Ψ AQ ♯* α AQ ♯* Q' AQ ♯* ΨP distinct(bn α)
    bn α ♯* Ψ bn α ♯* P  bn α ♯* Q  bn α ♯* subject α  bn α ♯* ΨP  bn α ♯* ΨQ
    AP ♯* C AQ ♯* C bn α ♯* C
  ultimately have "Prop C Ψ (P  Q) (α  (P  Q')) (AP@AQ) (ΨP  ΨQ)"
    by(metis rPar2)

  with AP ♯* Ψ AP ♯* P AP ♯* Q AP ♯* α AP ♯* Q' AP ♯* APQ AP ♯* C
    AQ ♯* Ψ AQ ♯* P AQ ♯* Q AQ ♯* α AQ ♯* Q' AQ ♯* APQ AQ ♯* C
    S distinctPerm p Aeq
  have "Prop C Ψ (P  Q) (α  (P  Q')) (p  (AP@AQ)) (p  (ΨP  ΨQ))"
    apply -
    apply(rule rAlpha)
    by(assumption | simp add: eqvts)+
  with Ψeq Aeq show ?case by(simp add: eqvts)
next
  case(cComm1 Ψ ΨQ P M N P' AP ΨP Q K xvec Q' AQ APQ ΨPQ C)
  from distinct AP distinct AQ AP ♯* AQ have "distinct(AP@AQ)"
    by(auto simp add: fresh_star_def fresh_def name_list_supp)
  from cComm1 have  "Prop C Ψ (P  Q) (τ  ⦇ν*xvec(P'  Q')) (AP@AQ) (ΨP  ΨQ)"
    by(metis rComm1)
  moreover from extractFrame(P  Q) = APQ, ΨPQ extractFrame P = AP, ΨP extractFrame Q = AQ, ΨQ
    AP ♯* AQ AP ♯* ΨQ AQ ♯* ΨP
  have "(AP@AQ), (ΨP  ΨQ) = APQ, ΨPQ"
    by simp
  with AP ♯* APQ AQ ♯* APQ distinct(AP@AQ) distinct APQ
  obtain p where S: "(set p  (set(AP@AQ)) × (set APQ))"  and "distinctPerm p"
    and Ψeq: "ΨPQ = p  (ΨP  ΨQ)" and Aeq: "APQ = p  (AP@AQ)"
    apply -
    apply(rule frameChainEq')
    by(assumption | simp)+
  moreover note AP ♯* Ψ AQ ♯* Ψ AP ♯* P AQ ♯* P AP ♯* Q AQ ♯* Q AP ♯* xvec
    AQ ♯* xvec AP ♯* P' AQ ♯* P' AP ♯* Q' AQ ♯* Q' AP ♯* APQ AQ ♯* APQ
    AP ♯* C AQ ♯* C
  ultimately have "Prop C Ψ (P  Q) (τ  ⦇ν*xvec(P'  Q')) (p  (AP@AQ)) (p  (ΨP  ΨQ))"
    by(fastforce simp add: rAlpha)
  with Ψeq Aeq show ?case by simp
next
  case(cComm2 Ψ ΨQ P M xvec N P' AP ΨP Q K Q' AQ APQ ΨPQ C)
  from distinct AP distinct AQ AP ♯* AQ have "distinct(AP@AQ)"
    by(auto simp add: fresh_star_def fresh_def name_list_supp)
  from cComm2 have  "Prop C Ψ (P  Q) (τ  ⦇ν*xvec(P'  Q')) (AP@AQ) (ΨP  ΨQ)"
    by(metis rComm2)
  moreover from extractFrame(P  Q) = APQ, ΨPQ extractFrame P = AP, ΨP extractFrame Q = AQ, ΨQ
    AP ♯* AQ AP ♯* ΨQ AQ ♯* ΨP
  have "(AP@AQ), (ΨP  ΨQ) = APQ, ΨPQ"
    by simp
  with AP ♯* APQ AQ ♯* APQ distinct(AP@AQ) distinct APQ
  obtain p where S: "(set p  (set(AP@AQ)) × (set APQ))"  and "distinctPerm p"
    and Ψeq: "ΨPQ = p  (ΨP  ΨQ)" and Aeq: "APQ = p  (AP@AQ)"
    apply -
    apply(rule frameChainEq')
    by(assumption | simp)+
  moreover note AP ♯* Ψ AQ ♯* Ψ AP ♯* P AQ ♯* P AP ♯* Q AQ ♯* Q AP ♯* xvec
    AQ ♯* xvec AP ♯* P' AQ ♯* P' AP ♯* Q' AQ ♯* Q' AP ♯* APQ AQ ♯* APQ
    AP ♯* C AQ ♯* C
  ultimately have "Prop C Ψ (P  Q) (τ  ⦇ν*xvec(P'  Q')) (p  (AP@AQ)) (p  (ΨP  ΨQ))"
    by(fastforce intro: rAlpha)
  with Ψeq Aeq show ?case by simp
next
  case(cBrMerge Ψ ΨQ P M N P' AP ΨP Q Q' AQ APQ ΨPQ C)
  from distinct AP distinct AQ AP ♯* AQ have "distinct(AP@AQ)"
    by(auto simp add: fresh_star_def fresh_def name_list_supp)
  from cBrMerge have "Prop C Ψ (P  Q) (¿MN  (P'  Q')) (AP@AQ) (ΨP  ΨQ)"
    by(fastforce intro!: rBrMerge)
  moreover from extractFrame(P  Q) = APQ, ΨPQ extractFrame P = AP, ΨP extractFrame Q = AQ, ΨQ
    AP ♯* AQ AP ♯* ΨQ AQ ♯* ΨP
  have "(AP@AQ), (ΨP  ΨQ) = APQ, ΨPQ"
    by simp
  with AP ♯* APQ AQ ♯* APQ distinct(AP@AQ) distinct APQ
  obtain p where S: "(set p  (set(AP@AQ)) × (set APQ))"  and "distinctPerm p"
    and Ψeq: "ΨPQ = p  (ΨP  ΨQ)" and Aeq: "APQ = p  (AP@AQ)"
    apply -
    apply(rule frameChainEq')
    by(assumption | simp)+
  moreover note AP ♯* Ψ AQ ♯* Ψ AP ♯* P AQ ♯* P AP ♯* Q AQ ♯* Q
    AP ♯* P' AQ ♯* P' AP ♯* Q' AQ ♯* Q' AP ♯* APQ AQ ♯* APQ
    AP ♯* C AQ ♯* C AP ♯* M AQ ♯* M AP ♯* N AQ ♯* N
  ultimately have "Prop C Ψ (P  Q) (¿MN  (P'  Q')) (p  (AP@AQ)) (p  (ΨP  ΨQ))"
    by(fastforce intro: rAlpha)
  with Ψeq Aeq show ?case by simp
next
  case(cBrComm1 Ψ ΨQ P M N P' AP ΨP Q xvec Q' AQ APQ ΨPQ C)
  from distinct AP distinct AQ AP ♯* AQ have "distinct(AP@AQ)"
    by(auto simp add: fresh_star_def fresh_def name_list_supp)
  from cBrComm1 have  "Prop C Ψ (P  Q) (¡M⦇ν*xvec⦈⟨N  (P'  Q')) (AP@AQ) (ΨP  ΨQ)"
    by(metis rBrComm1)
  moreover from extractFrame(P  Q) = APQ, ΨPQ extractFrame P = AP, ΨP extractFrame Q = AQ, ΨQ
    AP ♯* AQ AP ♯* ΨQ AQ ♯* ΨP
  have "(AP@AQ), (ΨP  ΨQ) = APQ, ΨPQ"
    by simp
  with AP ♯* APQ AQ ♯* APQ distinct(AP@AQ) distinct APQ
  obtain p where S: "(set p  (set(AP@AQ)) × (set APQ))"  and "distinctPerm p"
    and Ψeq: "ΨPQ = p  (ΨP  ΨQ)" and Aeq: "APQ = p  (AP@AQ)"
    apply -
    apply(rule frameChainEq')
    by(assumption | simp)+
  moreover note AP ♯* Ψ AQ ♯* Ψ AP ♯* P AQ ♯* P AP ♯* Q AQ ♯* Q AP ♯* xvec
    AQ ♯* xvec AP ♯* P' AQ ♯* P' AP ♯* Q' AQ ♯* Q' AP ♯* APQ AQ ♯* APQ
    AP ♯* C AQ ♯* C AP ♯* M AQ ♯* M AP ♯* N AQ ♯* N
  ultimately have "Prop C Ψ (P  Q) (¡M⦇ν*xvec⦈⟨N  (P'  Q')) (p  (AP@AQ)) (p  (ΨP  ΨQ))"
    by(fastforce intro: rAlpha)
  with Ψeq Aeq show ?case by simp
next
  case(cBrComm2 Ψ ΨQ P M xvec N P' AP ΨP Q Q' AQ APQ ΨPQ C)
  from distinct AP distinct AQ AP ♯* AQ have "distinct(AP@AQ)"
    by(auto simp add: fresh_star_def fresh_def name_list_supp)
  from cBrComm2 have "Prop C Ψ (P  Q) (¡M⦇ν*xvec⦈⟨N  (P'  Q')) (AP@AQ) (ΨP  ΨQ)"
    by(metis rBrComm2)
  moreover from extractFrame(P  Q) = APQ, ΨPQ extractFrame P = AP, ΨP extractFrame Q = AQ, ΨQ
    AP ♯* AQ AP ♯* ΨQ AQ ♯* ΨP
  have "(AP@AQ), (ΨP  ΨQ) = APQ, ΨPQ"
    by simp
  with AP ♯* APQ AQ ♯* APQ distinct(AP@AQ) distinct APQ
  obtain p where S: "(set p  (set(AP@AQ)) × (set APQ))"  and "distinctPerm p"
    and Ψeq: "ΨPQ = p  (ΨP  ΨQ)" and Aeq: "APQ = p  (AP@AQ)"
    apply -
    apply(rule frameChainEq')
    by(assumption | simp)+
  moreover note AP ♯* Ψ AQ ♯* Ψ AP ♯* P AQ ♯* P AP ♯* Q AQ ♯* Q AP ♯* xvec
    AQ ♯* xvec AP ♯* P' AQ ♯* P' AP ♯* Q' AQ ♯* Q' AP ♯* APQ AQ ♯* APQ
    AP ♯* C AQ ♯* C AP ♯* M AQ ♯* M AP ♯* N AQ ♯* N
  ultimately have "Prop C Ψ (P  Q) (¡M⦇ν*xvec⦈⟨N  (P'  Q')) (p  (AP@AQ)) (p  (ΨP  ΨQ))"
    by(fastforce intro: rAlpha)
  with Ψeq Aeq show ?case by simp
next
  case(cBrClose Ψ P M xvec N P' x AP' ΨP' C)
  obtain AP ΨP where FrP: "extractFrame P = AP, ΨP" and "distinct AP"
    and "AP ♯* (Ψ, P, M, xvec, N, P', AP', ΨP', C, x)"
    by(rule freshFrame)
  then have "AP ♯* Ψ" and "AP ♯* P" and "AP ♯* M" and "AP ♯* xvec" and "AP ♯* N" and "AP ♯* P'"
    and "AP ♯* AP'" and "AP ♯* ΨP'" and "AP ♯* C" and "x  AP"
    by simp+
  from FrP AP ♯* xvec xvec ♯* P have "xvec ♯* ΨP"
    by(force dest: extractFrameFreshChain)
  from AP ♯* xvec AP ♯* P' x  AP
  have "AP ♯* (τ  ⦇νx(⦇ν*xvecP'))" by simp
  from extractFrame P = AP, ΨP extractFrame (⦇νxP) = AP', ΨP'
  have "(x#AP), ΨP = AP', ΨP'" by simp
  with AP ♯* AP' x  AP' x  AP distinct AP distinct AP'
  obtain p where S: "(set p  (set (x#AP)) × (set AP'))"  and "distinctPerm p"
    and Ψeq: "ΨP' = p  ΨP" and Aeq: "AP' = p  (x#AP)"
    apply -
    apply(rule frameChainEq')
    by(assumption | simp)+

  from x  AP' Aeq have "x  (p  (x#AP))" by simp
  moreover from S Aeq distinct AP' x  AP' AP ♯* AP' have "(p  x)  AP"
    by simp
  from cBrClose FrP distinct AP AP ♯* Ψ AP ♯* P AP ♯* M AP ♯* xvec AP ♯* N AP ♯* P'
    AP ♯* AP' AP ♯* ΨP' AP ♯* C x  AP xvec ♯* ΨP AP ♯* C xvec ♯* C x  C
  have "Prop C Ψ (⦇νxP) (τ  ⦇νx(⦇ν*xvecP')) (x#AP) ΨP"
    by(force intro: rBrClose)

  moreover from Aeq AP ♯* AP' have "AP ♯* (p  AP)" by simp
  moreover from Aeq (set p  (set (x#AP)) × (set AP'))
  have "(set p  (set (x#AP)) × (set (p  (x#AP))))" by simp
  moreover from AP ♯* P x  AP have "AP ♯* (⦇νxP)" by simp
  moreover from S x  AP' have "p  x  x"
    using Aeq by fastforce
  moreover from x  AP' Aeq have "x  p  AP" by simp
  moreover note AP ♯* Ψ AP ♯* P AP ♯* (τ  ⦇νx(⦇ν*xvecP'))
    AP ♯* C distinctPerm p
    x  Ψ x  C (p  x)  AP
  ultimately
  have "Prop C Ψ (⦇νxP) (τ  ⦇νx(⦇ν*xvecP')) (p  (x#AP)) (p  Ψ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 AxP ΨxP C)
  obtain AP ΨP where FrP: "extractFrame P = AP, ΨP" and "distinct AP"
    and "AP ♯* (Ψ, P, M, xvec, yvec, N, P', AxP, ΨxP, C, x)"
    by(rule freshFrame)
  then have "AP ♯* Ψ" and "AP ♯* P" and "AP ♯* M" and "AP ♯* xvec"and "AP ♯* yvec" and "AP ♯* N" and "AP ♯* P'"
    and "AP ♯* AxP" and "AP ♯* ΨxP" and "AP ♯* C" and "x  AP"
    by simp+

  from xvec ♯* P AP ♯* xvec FrP have "xvec ♯* ΨP"
    by(force dest: extractFrameFreshChain)
  from yvec ♯* P AP ♯* yvec FrP have "yvec ♯* ΨP"
    by(force dest: extractFrameFreshChain)

  from extractFrame(⦇νxP) = AxP, ΨxP FrP
  have "(x#AP), ΨP = AxP, ΨxP"
    by simp
  moreover from x  AP distinct AP have "distinct(x#AP)" by simp
  ultimately obtain p where S: "set p  set (x#AP) × set (p  (x#AP))" and "distinctPerm p"
    and Ψeq: "ΨxP = p  ΨP" and Aeq: "AxP = (p  x)#(p  AP)"
    using AP ♯* AxPx  AxP distinct AxP
    apply -
    apply(rule frameChainEq')
    by(assumption | simp)+

  note Ψ  P M⦇ν*(xvec@yvec)⦈⟨N  P' FrP distinct AP
  moreover from FrP distinct AP AP ΨP C. extractFrame P = AP, ΨP; distinct AP  Prop C Ψ P (M⦇ν*(xvec@yvec)⦈⟨N  P') AP ΨP
  have "C. Prop C Ψ P (M⦇ν*(xvec@yvec)⦈⟨N  P') AP ΨP" by simp
  moreover note x  Ψ x  M x  xvec x  yvec x  supp N x  AP AP ♯* Ψ AP ♯* Ψ AP ♯* P AP ♯* M AP ♯* xvec AP ♯* yvec AP ♯* N AP ♯* P'
    xvec ♯* Ψ xvec ♯* P  xvec ♯* M  xvec ♯* ΨP yvec ♯* Ψ yvec ♯* P  yvec ♯* M  yvec ♯* ΨP
    AP ♯* C x  C xvec ♯* C yvec ♯* C xvec ♯* yvec distinct xvec distinct yvec
  ultimately have "Prop C Ψ (⦇νxP) (M⦇ν*(xvec@x#yvec)⦈⟨N  P') (x#AP) ΨP"
    by(metis rOpen)

  with AP ♯* Ψ AP ♯* P AP ♯* M AP ♯* xvec AP ♯* yvec AP ♯* N AP ♯* P' AP ♯* AxP AP ♯* C x  AxP AP ♯* AxP x  AP
    x  Ψ x  M x  C x  xvec x  yvec Aeq
    S distinctPerm p
  have "Prop C Ψ (⦇νxP) (M⦇ν*(xvec@x#yvec)⦈⟨N  P') (p  (x#AP)) (p  ΨP)"
    apply -
    apply(rule rAlpha[where AP="x#AP"])
    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 AxP ΨxP C)
  obtain AP ΨP where FrP: "extractFrame P = AP, ΨP" and "distinct AP"
    and "AP ♯* (Ψ, P, M, xvec, yvec, N, P', AxP, ΨxP, C, x)"
    by(rule freshFrame)
  then have "AP ♯* Ψ" and "AP ♯* P" and "AP ♯* M" and "AP ♯* xvec"and "AP ♯* yvec" and "AP ♯* N" and "AP ♯* P'"
    and "AP ♯* AxP" and "AP ♯* ΨxP" and "AP ♯* C" and "x  AP"
    by simp+

  from xvec ♯* P AP ♯* xvec FrP have "xvec ♯* ΨP"
    by(force dest: extractFrameFreshChain)
  from yvec ♯* P AP ♯* yvec FrP have "yvec ♯* ΨP"
    by(force dest: extractFrameFreshChain)

  from extractFrame(⦇νxP) = AxP, ΨxP FrP
  have "(x#AP), ΨP = AxP, ΨxP"
    by simp
  moreover from x  AP distinct AP have "distinct(x#AP)" by simp
  ultimately obtain p where S: "set p  set (x#AP) × set (p  (x#AP))" and "distinctPerm p"
    and Ψeq: "ΨxP = p  ΨP" and Aeq: "AxP = (p  x)#(p  AP)"
    using AP ♯* AxPx  AxP distinct AxP
    apply -
    apply(rule frameChainEq')
    by(assumption | simp)+

  note Ψ  P ¡M⦇ν*(xvec@yvec)⦈⟨N  P' FrP distinct AP
  moreover from FrP distinct AP AP ΨP C. extractFrame P = AP, ΨP; distinct AP  Prop C Ψ P (¡M⦇ν*(xvec@yvec)⦈⟨N  P') AP ΨP
  have "C. Prop C Ψ P (¡M⦇ν*(xvec@yvec)⦈⟨N  P') AP ΨP" by simp
  moreover note x  Ψ x  M x  xvec x  yvec x  supp N x  AP AP ♯* Ψ AP ♯* Ψ AP ♯* P AP ♯* M AP ♯* xvec AP ♯* yvec AP ♯* N AP ♯* P'
    xvec ♯* Ψ xvec ♯* P  xvec ♯* M  xvec ♯* ΨP yvec ♯* Ψ yvec ♯* P  yvec ♯* M  yvec ♯* ΨP
    AP ♯* C x  C xvec ♯* C yvec ♯* C xvec ♯* yvec distinct xvec distinct yvec
  ultimately have "Prop C Ψ (⦇νxP) (¡M⦇ν*(xvec@x#yvec)⦈⟨N  P') (x#AP) ΨP"
    by(metis rBrOpen)

  with AP ♯* Ψ AP ♯* P AP ♯* M AP ♯* xvec AP ♯* yvec AP ♯* N AP ♯* P' AP ♯* AxP AP ♯* C x  AxP AP ♯* AxP x  AP
    x  Ψ x  M x  C x  xvec x  yvec Aeq
    S distinctPerm p
  have "Prop C Ψ (⦇νxP) (¡M⦇ν*(xvec@x#yvec)⦈⟨N  P') (p  (x#AP)) (p  ΨP)"
    apply -
    apply(rule rAlpha[where AP="x#AP"])
    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 AxP ΨxP C)
  obtain AP ΨP where FrP: "extractFrame P = AP, ΨP" and "distinct AP"
    and "AP ♯* (Ψ, P, α, P', AxP, ΨxP, C, x)"
    by(rule freshFrame)
  then have "AP ♯* Ψ" and "AP ♯* P" and "AP ♯* α" and "AP ♯* P'"
    and "AP ♯* AxP" and "AP ♯* ΨxP" and "AP ♯* C" and "x  AP"
    by simp+

  from bn α ♯* P AP ♯* α FrP have "bn α ♯* ΨP"
    by(force dest: extractFrameFreshChain)

  from extractFrame(⦇νxP) = AxP, ΨxP FrP
  have "(x#AP), ΨP = AxP, ΨxP"
    by simp
  moreover from x  AP distinct AP have "distinct(x#AP)" by simp
  ultimately obtain p where S: "set p  set (x#AP) × set (p  (x#AP))" and "distinctPerm p"
    and Ψeq: "ΨxP = p  ΨP" and Aeq: "AxP = (p  x)#(p  AP)"
    using AP ♯* AxPx  AxP distinct AxP
    apply -
    apply(rule frameChainEq')
    by(assumption | simp)+

  note Ψ  P α  P' FrP distinct AP
  moreover from FrP distinct AP AP ΨP C. extractFrame P = AP, ΨP; distinct AP  Prop C Ψ P (α  P') AP ΨP
  have "C. Prop C Ψ P (α  P') AP ΨP" by simp
  moreover note x  Ψ x  α x  AP AP ♯* Ψ AP ♯* Ψ AP ♯* P AP ♯* α AP ♯* P' distinct(bn α)
    bn α ♯* Ψ bn α ♯* P  bn α ♯* subject α  bn α ♯* ΨP AP ♯* C x  C bn α ♯* C
  ultimately have "Prop C Ψ (⦇νxP) (α  (⦇νxP')) (x#AP) ΨP"
    by(metis rScope)

  with AP ♯* Ψ AP ♯* P AP ♯* α AP ♯* P' AP ♯* AxP AP ♯* C x  AxP AP ♯* AxP x  AP
    x  Ψ x  α x  C Aeq
    S distinctPerm p
  have "Prop C Ψ (⦇νxP) (α  (⦇νxP')) (p  (x#AP)) (p  ΨP)"
    apply -
    apply(rule rAlpha[where AP="x#AP"])
    by(assumption | simp add: abs_fresh fresh_star_def)+
  with Ψeq Aeq show ?case by(simp add: eqvts)
next
  case(Bang Ψ P Rs AbP ΨbP C)

  obtain AP ΨP where FrP: "extractFrame P = AP, ΨP" and "distinct AP"
    and "AP ♯* (Ψ, P, Rs, C)"
    by(rule freshFrame)
  then have "AP ♯* Ψ" and "AP ♯* P" and "AP ♯* Rs" and "AP ♯* C"
    by simp+

  note Ψ  P  !P  Rs guarded P FrP distinct AP
  moreover from FrP have "extractFrame (P  !P) = AP, ΨP  𝟭"
    by simp
  with distinct AP AP ΨP C. extractFrame (P  !P) = AP, ΨP; distinct AP  Prop C Ψ (P  !P) Rs AP ΨP
  have "C. Prop C Ψ (P  !P) Rs AP (ΨP  𝟭)" by simp
  moreover from guarded P FrP have "ΨP  𝟭" and "supp ΨP = ({}::name set)" by(metis guardedStatEq)+
  moreover note AP ♯* Ψ AP ♯* P AP ♯* Rs AP ♯* C
  ultimately have "Prop C Ψ (!P) Rs ([]) (𝟭)"
    by(rule rBang)
  then show ?case using extractFrame(!P) = AbP, ΨbP 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 AP   :: "name list"
    and Ψ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 = AP, ΨP"
  and   "distinct AP"
  and   "bn α ♯* subject α"
  and   "distinct(bn α)"
  and   rAlpha: "Ψ P α P' p AP ΨP C. bn α ♯* Ψ; bn α ♯* P; bn α ♯* subject α; bn α ♯* ΨP;
                                           bn α ♯* C; bn α ♯* (p  α); AP ♯* Ψ; AP ♯* P; AP ♯* α; AP ♯* P'; AP ♯* C;
                                           set p  set(bn α) × set(bn(p  α)); distinctPerm p;
                                           bn(p  α) ♯* α; (bn(p  α)) ♯* P'; Prop C Ψ P α P' AP ΨP 
                                           Prop C Ψ P (p  α) (p  P') AP ΨP"
  and   rFrameAlpha: "Ψ P AP ΨP p α P' C. AP ♯* Ψ; AP ♯* P; AP ♯* (p  AP); AP ♯* α; AP ♯* P'; AP ♯* C;
                                                set p  set AP × set(p  AP); distinctPerm p; AP ♯* subject α;
                                                Prop C Ψ P α P' AP ΨP  Prop C Ψ P α P' (p  AP) (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: "Ψ 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 Ψ (MN⟩.P) (KN) P ([]) (𝟭)"
  and   rBrOutput: "Ψ M K N P C. Ψ  M  K  Prop C Ψ (MN⟩.P) (¡KN) P ([]) (𝟭)"
  and   rCase: "Ψ P α P' φ Cs AP ΨP C. Ψ  P α  P'; extractFrame P = AP, ΨP; distinct AP; C. Prop C Ψ P α P' AP ΨP;
                                            (φ, P)  set Cs; Ψ  φ; guarded P;  ΨP  𝟭; (supp ΨP) = ({}::name set);
                                            AP ♯* Ψ; AP ♯* P; AP ♯* α; AP ♯* P'; AP ♯* C  Prop C Ψ (Cases Cs) α P' ([]) (𝟭)"
  and   rPar1: "Ψ ΨQ P α P' AQ Q AP ΨP C.
                   Ψ  ΨQ  P α  P';
                   extractFrame P = AP, ΨP; distinct AP;
                   extractFrame Q = AQ, ΨQ; distinct AQ;
                   C. Prop C (Ψ  ΨQ) P α P' AP ΨP;
                   AP ♯* P; AP ♯* Q; AP ♯* Ψ; AP ♯* α; AP ♯* P'; AP ♯* AQ; AP ♯* ΨQ;
                   AQ ♯* P; AQ ♯* Q; AQ ♯* Ψ; AQ ♯* α; AQ ♯* P'; AQ ♯* ΨP;
                   bn α ♯* Ψ; bn α ♯* P; bn α ♯* Q; bn α ♯* subject α; bn α ♯* ΨP; bn α ♯* ΨQ;
                   AP ♯* C; AQ ♯* C; bn α ♯* C 
                   Prop C Ψ (P  Q) α (P'  Q) (AP@AQ) (ΨP  ΨQ)"
  and   rPar2: "Ψ ΨP Q α Q' AP P AQ ΨQ C.
                   Ψ  ΨP  Q α  Q';
                   extractFrame P = AP, ΨP; distinct AP;
                   extractFrame Q = AQ, ΨQ; distinct AQ;
                   C. Prop C (Ψ  ΨP) Q α Q' AQ ΨQ;
                   AP ♯* P; AP ♯* Q; AP ♯* Ψ; AP ♯* α; AP ♯* Q'; AP ♯* AQ; AP ♯* ΨQ;
                   AQ ♯* P; AQ ♯* Q; AQ ♯* Ψ; AQ ♯* α; AQ ♯* Q'; AQ ♯* ΨP;
                   bn α ♯* Ψ; bn α ♯* P; bn α ♯* Q; bn α ♯* subject α; bn α ♯* ΨP; bn α ♯* ΨQ;
                   AP ♯* C; AQ ♯* C; bn α ♯* C 
                   Prop C Ψ (P  Q) α (P  Q') (AP@AQ) (ΨP  ΨQ)"
  and   rComm1: "Ψ ΨQ P M N P' AP ΨP Q K xvec Q' AQ C.
                   Ψ  ΨQ  P MN  P'; extractFrame P = AP, ΨP; distinct AP;
                   C. Prop C (Ψ  ΨQ) P (MN) P' AP ΨP;
                    Ψ  ΨP  Q K⦇ν*xvec⦈⟨N  Q'; extractFrame Q = AQ, ΨQ; distinct AQ;
                    Ψ  ΨP  ΨQ  M  K; distinct xvec;
                   C. Prop C (Ψ  ΨP) Q (K⦇ν*xvec⦈⟨N) Q' AQ ΨQ;
                    AP ♯* Ψ; AP ♯* ΨQ; AP ♯* P; AP ♯* M; AP ♯* N; AP ♯* P';
                    AP ♯* Q; AP ♯* Q'; AP ♯* AQ; AP ♯* xvec; AQ ♯* Ψ; AQ ♯* ΨP;
                    AQ ♯* P; AQ ♯* N; AQ ♯* P'; AQ ♯* Q; AQ ♯* K; AQ ♯* Q';
                    AQ ♯* xvec; xvec ♯* Ψ; xvec ♯* ΨP; xvec ♯* ΨQ; xvec ♯* P; xvec ♯* M;
                    xvec ♯* Q; xvec ♯* K; AP ♯* C; AQ ♯* C; xvec ♯* C 
                    Prop C Ψ (P  Q) (τ) (⦇ν*xvec(P'  Q')) (AP@AQ) (ΨP  ΨQ)"
  and   rComm2: "Ψ ΨQ P M xvec N P' AP ΨP Q K Q' AQ C.
                   Ψ  ΨQ  P M⦇ν*xvec⦈⟨N  P'; extractFrame P = AP, ΨP; distinct AP;
                   C. Prop C (Ψ  ΨQ) P (M⦇ν*xvec⦈⟨N) P' AP ΨP;
                    Ψ  ΨP  Q KN  Q'; extractFrame Q = AQ, ΨQ; distinct AQ;
                   C. Prop C (Ψ  ΨP) Q (KN) Q' AQ ΨQ;
                    Ψ  ΨP  ΨQ  M  K; distinct xvec;
                    AP ♯* Ψ; AP ♯* ΨQ; AP ♯* P; AP ♯* M; AP ♯* N; AP ♯* P';
                    AP ♯* Q; AP ♯* Q'; AP ♯* AQ; AP ♯* xvec; AQ ♯* Ψ; AQ ♯* ΨP;
                    AQ ♯* P; AQ ♯* N; AQ ♯* P'; AQ ♯* Q; AQ ♯* K; AQ ♯* Q';
                    AQ ♯* xvec; xvec ♯* Ψ; xvec ♯* ΨP; xvec ♯* ΨQ; xvec ♯* P; xvec ♯* M;
                    xvec ♯* Q; xvec ♯* K; AP ♯* C; AQ ♯* C; xvec ♯* C 
                    Prop C Ψ (P  Q) (τ) (⦇ν*xvec(P'  Q')) (AP@AQ) (ΨP  ΨQ)"
  and   rBrMerge: "Ψ ΨQ P M N P' AP ΨP Q Q' AQ C.
                    Ψ  ΨQ  P  ¿MN  P'; C. Prop C (Ψ  ΨQ) P (¿MN) P' AP ΨP;
                    extractFrame P = AP, ΨP; distinct AP;
                    Ψ  ΨP  Q  ¿MN  Q'; C. Prop C (Ψ  ΨP) Q (¿MN) Q' AQ ΨQ;
                    extractFrame Q = AQ, ΨQ; distinct AQ;
                    AP ♯* Ψ; AP ♯* ΨQ; AP ♯* P; AP ♯* N; AP ♯* P';
                    AP ♯* Q; AP ♯* Q'; AP ♯* AQ; AP ♯* M; AQ ♯* M;
                    AQ ♯* Ψ; AQ ♯* ΨP; AQ ♯* P; AQ ♯* N; AQ ♯* P';
                    AQ ♯* Q; AQ ♯* Q'; AP ♯* C; AQ ♯* C 
                    Prop C Ψ (P  Q) (¿MN) (P'  Q') (AP@AQ) (ΨP  ΨQ)"
  and   rBrComm1:"Ψ ΨQ P M N P' AP ΨP Q xvec Q' AQ C.
                   Ψ  ΨQ  P ¿MN  P'; extractFrame P = AP, ΨP; distinct AP;
                   C. Prop C (Ψ  ΨQ) P (¿MN) P' AP ΨP;
                    Ψ  ΨP  Q ¡M⦇ν*xvec⦈⟨N  Q'; extractFrame Q = AQ, ΨQ; distinct AQ;
                    distinct xvec;
                   C. Prop C (Ψ  ΨP) Q (¡M⦇ν*xvec⦈⟨N) Q' AQ ΨQ;
                    AP ♯* Ψ; AP ♯* ΨQ; AP ♯* P; AP ♯* N; AP ♯* P';
                    AP ♯* Q; AP ♯* Q'; AP ♯* AQ; AP ♯* xvec; AQ ♯* Ψ; AQ ♯* ΨP;
                    AQ ♯* P; AQ ♯* N; AQ ♯* P'; AQ ♯* Q; AQ ♯* Q';
                    AQ ♯* xvec; xvec ♯* Ψ; xvec ♯* ΨP; xvec ♯* ΨQ; xvec ♯* P;
                    xvec ♯* Q; AP ♯* C; AQ ♯* C; xvec ♯* C;
                    AP ♯* M; AQ ♯* M; xvec ♯* M 
                    Prop C Ψ (P  Q) (¡M⦇ν*xvec⦈⟨N) (P'  Q') (AP@AQ) (ΨP  ΨQ)"
  and   rBrComm2:"Ψ ΨQ P M xvec N P' AP ΨP Q Q' AQ C.
                   Ψ  ΨQ  P ¡M⦇ν*xvec⦈⟨N  P'; extractFrame P = AP, ΨP; distinct AP;
                   C. Prop C (Ψ  ΨQ) P (¡M⦇ν*xvec⦈⟨N) P' AP ΨP;
                    Ψ  ΨP  Q ¿MN  Q'; extractFrame Q = AQ, ΨQ; distinct AQ;
                   C. Prop C (Ψ  ΨP) Q (¿MN) Q' AQ ΨQ;
                    distinct xvec;
                    AP ♯* Ψ; AP ♯* ΨQ; AP ♯* P; AP ♯* N; AP ♯* P';
                    AP ♯* Q; AP ♯* Q'; AP ♯* AQ; AP ♯* xvec; AQ ♯* Ψ; AQ ♯* ΨP;
                    AQ ♯* P; AQ ♯* N; AQ ♯* P'; AQ ♯* Q; AQ ♯* Q';
                    AQ ♯* xvec; xvec ♯* Ψ; xvec ♯* ΨP; xvec ♯* ΨQ; xvec ♯* P;
                    xvec ♯* Q; AP ♯* C; AQ ♯* C; xvec ♯* C;
                    AP ♯* M; AQ ♯* M; xvec ♯* M 
                    Prop C Ψ (P  Q) (¡M⦇ν*xvec⦈⟨N) (P'  Q') (AP@AQ) (ΨP  ΨQ)"
  and   rBrClose: "Ψ P M xvec N P' AP ΨP x C.
                    Ψ  P  ¡M⦇ν*xvec⦈⟨N  P';
                     x  supp M;
                     C. Prop C Ψ P (¡M⦇ν*xvec⦈⟨N) P' AP ΨP;
                     extractFrame P = AP, ΨP; distinct AP;
                     AP ♯* Ψ; AP ♯* P; AP ♯* M; AP ♯* N; AP ♯* P'; AP ♯* xvec;
                     distinct xvec; xvec ♯* Ψ; xvec ♯* ΨP; xvec ♯* P;
                     xvec ♯* M;
                     x  Ψ; x  xvec; x  AP;
                     AP ♯* C; xvec ♯* C; x  C 
                     Prop C Ψ (⦇νxP) (τ) (⦇νx(⦇ν*xvecP')) (x#AP) ΨP"
  and   rOpen: "Ψ P M xvec yvec N P' x AP ΨP y C.
                    Ψ  P M⦇ν*(xvec@yvec)⦈⟨N  P'; extractFrame P = AP, ΨP; distinct AP;
                    C. Prop C Ψ P (M⦇ν*(xvec@yvec)⦈⟨N) P' AP ΨP; x  supp N; x  Ψ; x  M;
                     x  AP; x  xvec; x  yvec; AP ♯* Ψ; AP ♯* P; AP ♯* M; AP ♯* N; AP ♯* P';
                     AP ♯* xvec; AP ♯* yvec; xvec ♯* yvec; distinct xvec; distinct yvec;
                     xvec ♯* Ψ; xvec ♯* P; xvec ♯* M; xvec ♯* ΨP;
                     yvec ♯* Ψ; yvec ♯* P; yvec ♯* M; AP ♯* C; x  C; xvec ♯* C; yvec ♯* C;
                     y  x; y  Ψ; y  P; y  M; y  xvec; y  yvec; y  N; y  P'; y  AP; y  ΨP; y  C 
                     Prop C Ψ (⦇νxP) (M⦇ν*(xvec@y#yvec)⦈⟨([(x, y)]  N)) ([(x, y)]  P') (x#AP) ΨP"
  and   rBrOpen: "Ψ P M xvec yvec N P' x AP ΨP y C.
                    Ψ  P ¡M⦇ν*(xvec@yvec)⦈⟨N  P'; extractFrame P = AP, ΨP; distinct AP;
                    C. Prop C Ψ P (¡M⦇ν*(xvec@yvec)⦈⟨N) P' AP ΨP; x  supp N; x  Ψ; x  M;
                     x  AP; x  xvec; x  yvec; AP ♯* Ψ; AP ♯* P; AP ♯* M; AP ♯* N; AP ♯* P';
                     AP ♯* xvec; AP ♯* yvec; xvec ♯* yvec; distinct xvec; distinct yvec;
                     xvec ♯* Ψ; xvec ♯* P; xvec ♯* M; xvec ♯* ΨP;
                     yvec ♯* Ψ; yvec ♯* P; yvec ♯* M; AP ♯* C; x  C; xvec ♯* C; yvec ♯* C;
                     y  x; y  Ψ; y  P; y  M; y  xvec; y  yvec; y  N; y  P'; y  AP; y  ΨP; y  C 
                     Prop C Ψ (⦇νxP) (¡M⦇ν*(xvec@y#yvec)⦈⟨([(x, y)]  N)) ([(x, y)]  P') (x#AP) ΨP"
  and   rScope: "Ψ P α P' x AP ΨP C.
                    Ψ  P α  P'; extractFrame P = AP, ΨP; distinct AP;
                    C. Prop C Ψ P α P' AP ΨP;
                     x  Ψ; x  α; x  AP; AP ♯* Ψ; AP ♯* P;
                     AP ♯* α; AP ♯* P';
                     bn α ♯* Ψ; bn α ♯* P; bn α ♯* subject α; bn α ♯* ΨP;
                     AP ♯* C; x  C; bn α ♯* C 
                     Prop C Ψ (⦇νxP) α (⦇νxP') (x#AP) ΨP"
  and   rBang:    "Ψ P α P' AP ΨP C.
                     Ψ  P  !P α  P'; guarded P; extractFrame P = AP, ΨP; distinct AP;
                      C. Prop C Ψ (P  !P) α P' AP (ΨP  𝟭); ΨP  𝟭; supp ΨP = ({}::name set);
                      AP ♯* Ψ; AP ♯* P; AP ♯* α; AP ♯* P'; AP ♯* C  Prop C Ψ (!P) α P' ([]) (𝟭)"
shows "Prop C Ψ P α P' AP ΨP"
  using Trans FrP distinct AP bn α ♯* subject α distinct(bn α)
proof(nominal_induct Ψ P Rs=="α  P'" AP Ψ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 Ψ ΨQ P α P' AQ Q AP Ψ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 Ψ  ΨQ  P α  P' extractFrame Q = AQ, ΨQ distinct AQ
  moreover from bn α ♯* subject α distinct(bn α) AP ♯* α
  have "C. Prop C (Ψ  ΨQ) P α P' AP ΨP" by(fastforce intro: cPar1)

  moreover note AQ ♯* P AQ ♯* Q AQ ♯* Ψ AQ ♯* α AQ ♯* P' AQ ♯* C AP ♯* P AP ♯* Q AP ♯* Ψ AP ♯* α AP ♯* P' AP ♯* C
    bn α ♯* Q distinct(bn α) bn α ♯* Ψ bn α ♯* ΨQ bn α ♯* P bn α ♯* subject α bn α ♯* C
    extractFrame P = AP, ΨP distinct AP AP ♯* AQ AP ♯* ΨQ AQ ♯* ΨP bn α ♯* ΨP
  ultimately have "Prop C Ψ (P  Q) α (P'  Q) (AP@AQ) (Ψ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) bn α ♯* ΨP bn α ♯* ΨQ AP ♯* α AQ ♯* α AP ♯* α' AQ ♯* α' αEq bn α ♯* ΨP bn α ♯* α' AP ♯* Ψ AQ ♯* Ψ AP ♯* P AQ ♯* P AP ♯* Q AQ ♯* Q AP ♯* P' AQ ♯* P' AP ♯* C AQ ♯* C
  have "Prop C Ψ (P  Q) (p  α) (p  (P'  Q)) (AP@AQ) (ΨP  ΨQ)"
    by(fastforce intro!: rAlpha)
  with αEq P'eq distinctPerm p show ?case by simp
next
  case(cPar2 Ψ ΨP Q α Q' AP P AQ Ψ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 Ψ  ΨP  Q α  Q' extractFrame P = AP, ΨP distinct AP
  moreover from bn α ♯* subject α distinct(bn α) AQ ♯* α
  have "C. Prop C (Ψ  ΨP) Q α Q' AQ ΨQ" by(fastforce intro!: cPar2)

  moreover note AQ ♯* P AQ ♯* Q AQ ♯* Ψ AQ ♯* α AQ ♯* Q' AQ ♯* C AP ♯* P AP ♯* Q AP ♯* Ψ AP ♯* α AP ♯* Q' AP ♯* C
    bn α ♯* Q distinct(bn α) bn α ♯* Ψ bn α ♯* ΨQ bn α ♯* P bn α ♯* subject α bn α ♯* C
    extractFrame Q = AQ, ΨQ distinct AQ AP ♯* AQ AP ♯* ΨQ AQ ♯* ΨP bn α ♯* ΨP
  ultimately have "Prop C Ψ (P  Q) α (P  Q') (AP@AQ) (ΨP  Ψ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 α ♯* ΨP bn α ♯* ΨQ AP ♯* α AQ ♯* α AP ♯* α' AQ ♯* α' αEq bn α ♯* α' AP ♯* Ψ AQ ♯* Ψ AP ♯* P AQ ♯* P AP ♯* Q AQ ♯* Q AP ♯* Q' AQ ♯* Q' AP ♯* C AQ ♯* C
  have "Prop C Ψ (P  Q) (p  α) (p  (P  Q')) (AP@AQ) (ΨP  ΨQ)"
    by(fastforce intro!: rAlpha)
  with αEq Q'eq distinctPerm p show ?case by simp
next
  case(cComm1 Ψ ΨQ P M N P' AP ΨP Q K xvec Q' AQ C α P'')
  then show ?case using rComm1
    apply -
    apply(drule meta_spec[where x="MN"])
    apply(drule meta_spec[where x="K⦇ν*xvec⦈⟨N"])
    by(auto simp add: residualInject)
next
  case(cComm2 Ψ ΨQ P M xvec N P' AP ΨP Q K Q' AQ C α Q'')
  then show ?case using rComm2
    apply -
    apply(drule meta_spec[where x="M⦇ν*xvec⦈⟨N"])
    apply(drule meta_spec[where x="KN"])
    by(auto simp add: residualInject)
next
  case(cBrMerge Ψ ΨQ P M N P' AP ΨP Q Q' AQ C α P'')
  then show ?case using rBrMerge
    apply -
    apply(drule meta_spec[where x="¿MN"])
    apply(drule meta_spec[where x="¿MN"])
    by(auto simp add: residualInject)
next
  case(cBrComm1 Ψ ΨQ P M N P' AP ΨP Q xvec Q' AQ C α P'')
  have "bn (¿MN) ♯* subject (¿MN)" by simp
  moreover have "distinct (bn (¿MN))" by simp
  moreover have "¿MN  P' = ¿MN  P'" by simp
  moreover note cBrComm1
  ultimately have inProp: "C. Prop C (Ψ  ΨQ) P (¿MN) P' AP ΨP" by simp

  note xvec ♯* M distinct xvec cBrComm1
  then have outProp: "C. Prop C (Ψ  ΨP) Q (¡M⦇ν*xvec⦈⟨N) Q' AQ ΨQ" by simp

  note inProp outProp cBrComm1
  then have bigProp: "Prop C Ψ (P  Q) (¡M⦇ν*xvec⦈⟨N) (P'  Q') (AP @ AQ) (ΨP  Ψ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 ♯* ΨP xvec ♯* ΨQ have "bn (¡M⦇ν*xvec⦈⟨N) ♯* (ΨP  Ψ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 AP ♯* Ψ AQ ♯* Ψ have "(AP @ AQ) ♯* Ψ" by simp
  moreover from AP ♯* P AQ ♯* Q AP ♯* Q AQ ♯* P have "(AP @ AQ) ♯* (P  Q)" by simp
  moreover from AP ♯* α AQ ♯* α have "(AP @ AQ) ♯* α" by simp
  moreover from AP ♯* P' AQ ♯* Q' AP ♯* Q' AQ ♯* P' have "(AP @ AQ) ♯* (P'  Q')" by simp
  moreover from AP ♯* C AQ ♯* C have "(AP @ AQ) ♯* C" by simp
  moreover note S distinctPerm p (bn(p  (¡M⦇ν*xvec⦈⟨N))) ♯* (¡M⦇ν*xvec⦈⟨N)
    (bn(p  (¡M⦇ν*xvec⦈⟨N))) ♯* (P'  Q') bigProp
    AP ♯* M AQ ♯* M AP ♯* xvec AQ ♯* xvec AP ♯* N AQ ♯* N

  ultimately have "Prop C Ψ (P  Q) (p  (¡M⦇ν*xvec⦈⟨N)) (p  (P'  Q')) (AP @ AQ) (ΨP  ΨQ)"
    by(fastforce intro!: rAlpha)
  then show ?case using αEq P'eq by simp
next
  case(cBrComm2 Ψ ΨQ P M xvec N P' AP ΨP Q Q' AQ C α Q'')
  have "bn (¿MN) ♯* subject (¿MN)" by simp
  moreover have "distinct (bn (¿MN))" by simp
  moreover have "¿MN  Q' = ¿MN  Q'" by simp
  moreover note cBrComm2
  ultimately have inProp: "C. Prop C (Ψ  ΨP) Q (¿MN) Q' AQ Ψ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 (Ψ  ΨQ) P (¡M⦇ν*xvec⦈⟨N) P' AP ΨP" by simp

  note inProp outProp cBrComm2
  then have bigProp: "Prop C Ψ (P  Q) (¡M⦇ν*xvec⦈⟨N) (P'  Q') (AP @ AQ) (ΨP  Ψ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 ♯* ΨP xvec ♯* ΨQ have "bn (¡M⦇ν*xvec⦈⟨N) ♯* (ΨP  Ψ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 AP ♯* Ψ AQ ♯* Ψ have "(AP @ AQ) ♯* Ψ" by simp
  moreover from AP ♯* P AQ ♯* Q AP ♯* Q AQ ♯* P have "(AP @ AQ) ♯* (P  Q)" by simp
  moreover from AP ♯* α AQ ♯* α have "(AP @ AQ) ♯* α" by simp
  moreover from AP ♯* P' AQ ♯* Q' AP ♯* Q' AQ ♯* P' have "(AP @ AQ) ♯* (P'  Q')" by simp
  moreover from AP ♯* C AQ ♯* C have "(AP @ AQ) ♯* C" by simp
  moreover note S distinctPerm p (bn(p  (¡M⦇ν*xvec⦈⟨N))) ♯* (¡M⦇ν*xvec⦈⟨N)
    (bn(p  (¡M⦇ν*xvec⦈⟨N))) ♯* (P'  Q') bigProp
    AP ♯* M AQ ♯* M AP ♯* xvec AQ ♯* xvec AP ♯* N AQ ♯* N

  ultimately have "Prop C Ψ (P  Q) (p  (¡M⦇ν*xvec⦈⟨N)) (p  (P'  Q')) (AP @ AQ) (ΨP  ΨQ)"
    by(fastforce intro!: rAlpha)
  then show ?case using αEq P'eq by simp
next
  case(cBrClose Ψ P M xvec N P' AP ΨP x C α P'')
  note τ  ⦇νx(⦇ν*xvecP') = α  P''
  moreover have "bn (τ) ♯* (bn α)" by simp
  moreover have "distinct (bn (τ))" by simp
  moreover note distinct (bn α)
  moreover have "(bn (τ) ♯* (τ  ⦇νx(⦇ν*xvecP')))" 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(⦇ν*xvecP'))"
    and "bn (τ) ♯* α" and "bn (τ) ♯* P''"
    and "(bn(p  (τ))) ♯* (τ)" and "(bn(p  (τ))) ♯* (⦇νx(⦇ν*xvecP'))"
    by(rule residualEq) simp
  moreover from cBrClose have "C. Prop C Ψ P (¡M⦇ν*xvec⦈⟨N) P' AP ΨP" by simp
  moreover with cBrClose have "Prop C Ψ (⦇νxP) (τ) (⦇νx(⦇ν*xvecP')) (x#AP) ΨP"
    by(simp add: rBrClose)
  with S have "Prop C Ψ (⦇νxP) (p  τ) (p  ⦇νx(⦇ν*xvecP')) (x#AP) ΨP" by simp
  then show ?case using αEq P'eq
    by simp
next
  case(cOpen Ψ P M xvec yvec N P' x AP Ψ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' AP ΨP" using AP ♯* xvec AP ♯* yvec AP ♯* M AP ♯* 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  AP" and "y  Ψ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 = AP, ΨP distinct AP x  AP xvec ♯* yvec xvec ♯* ΨP
    AP ♯* Ψ AP ♯* P AP ♯* M AP ♯* xvec AP ♯* yvec AP ♯* N AP ♯* P' AP ♯* C
  ultimately have "Prop C Ψ (⦇νxP) (M⦇ν*(xvec@y#yvec)⦈⟨([(x, y)]  N)) ([(x, y)]  P') (x#AP) Ψ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 ♯* ΨP yvec ♯* ΨP x  Ψ
    AP ♯* xvec x  AP AP ♯* yvec AP ♯* M x  xvec x  yvec x  M x  AP AP ♯* N
    A B C  αeq AP ♯* α y  Ψ y  x y  P y  M y  ΨP y  C xvec ♯* α x  α yvec ♯* α y  α AP ♯* P AP ♯* Ψ y  AP y  N AP ♯* P' y  P' AP ♯* C P'eq
  ultimately have "Prop C Ψ (⦇νxP) (([(x, y)]  p)  (M⦇ν*(xvec@y#yvec)⦈⟨([(x, y)]  N))) (([(x, y)]  p)  [(x, y)]  P') (x#AP) Ψ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 AP Ψ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' AP ΨP" using AP ♯* xvec AP ♯* yvec AP ♯* M AP ♯* 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  AP" and "y  Ψ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 = AP, ΨP distinct AP x  AP xvec ♯* yvec xvec ♯* ΨP
    AP ♯* Ψ AP ♯* P AP ♯* M AP ♯* xvec AP ♯* yvec AP ♯* N AP ♯* P' AP ♯* C
  ultimately have "Prop C Ψ (⦇νxP) (¡M⦇ν*(xvec@y#yvec)⦈⟨([(x, y)]  N)) ([(x, y)]  P') (x#AP) Ψ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 ♯* ΨP yvec ♯* ΨP x  Ψ
    AP ♯* xvec x  AP AP ♯* yvec AP ♯* M x  xvec x  yvec x  M x  AP AP ♯* N
    A B C  αeq AP ♯* α y  Ψ y  x y  P y  M y  ΨP y  C xvec ♯* α x  α yvec ♯* α y  α AP ♯* P AP ♯* Ψ y  AP y  N AP ♯* P' y  P' AP ♯* C P'eq
  ultimately have "Prop C Ψ (⦇νxP) (([(x, y)]  p)  (¡M⦇ν*(xvec@y#yvec)⦈⟨([(x, y)]  N))) (([(x, y)]  p)  [(x, y)]  P') (x#AP) Ψ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 AP ΨP C α' P'')
  note α  (⦇νxP') = α'  P''
  moreover from bn α ♯* α' have "bn α ♯* (bn α')" by auto
  moreover note distinct (bn α) distinct(bn α')
  moreover from bn α ♯* subject α bn α' ♯* subject α'
  have "bn α ♯* (α  ⦇νxP')" 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  (⦇νxP')" and "(bn(p  α)) ♯* α"
    and "(bn(p  α)) ♯* (⦇νxP')"
    by(rule residualEq)

  note Ψ  P α  P'
  moreover from bn α ♯* subject α distinct(bn α)
  have "C. Prop C Ψ P α P' AP ΨP" by(fastforce intro!: cScope)

  moreover note x  Ψ x  α bn α ♯* Ψ bn α ♯* P bn α ♯* subject α bn α ♯* ΨP
    x  C bn α ♯* C distinct(bn α) extractFrame P = AP, ΨP
    distinct AP x  AP AP ♯* Ψ AP ♯* P AP ♯* α AP ♯* P' AP ♯* C
  ultimately have "Prop C Ψ (⦇νxP) α (⦇νxP') (x#AP) ΨP"
    by(metis rScope)
  with bn α ♯* Ψ bn α ♯* P x  α bn α ♯* subject α bn α ♯* C bn α ♯* (bn α') S distinctPerm p bn(p  α) ♯* α bn(p  α) ♯* (⦇νxP') AP ♯* α AP ♯* α' αEq x  α' bn α ♯* ΨP bn α ♯* α' x  Ψ AP ♯* Ψ x  AP AP ♯* P AP ♯* P' x  C AP ♯* C
  have "Prop C Ψ (⦇νxP) (p  α) (p  (⦇νxP'))  (x#AP) ΨP"
    by(fastforce intro!: rAlpha simp add: abs_fresh)
  with αEq P'eq distinctPerm p show ?case by simp
next
  case(cBang Ψ P AP Ψ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 MN  P'"
  and   FrP: "extractFrame P = AP, ΨP"
  and   "distinct AP"
  and   rAlpha: "Ψ P M N P' AP ΨP p C. AP ♯* Ψ; AP ♯* P; AP ♯* M; AP ♯* N; AP ♯* P'; AP ♯* (p  AP); AP ♯* C;
                                            set p  set AP × set(p  AP); distinctPerm p;
                                             Prop C Ψ P M N P' AP ΨP  Prop C Ψ P M N P' (p  AP) (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   rCase: "Ψ P M N P' φ Cs AP ΨP C. Ψ  P MN  P'; extractFrame P = AP, ΨP; distinct AP; C. Prop C Ψ P M N P' AP ΨP;
                                              (φ, P)  set Cs; Ψ  φ; guarded P;  ΨP  𝟭; (supp ΨP) = ({}::name set);
                                              AP ♯* Ψ; AP ♯* P; AP ♯* M; AP ♯* N; AP ♯* P'; AP ♯* C  Prop C Ψ (Cases Cs) M N P' ([]) (𝟭)"
  and   rPar1: "Ψ ΨQ P M N P' AQ Q AP ΨP C.
                   Ψ  ΨQ  P MN  P';
                   extractFrame P = AP, ΨP; distinct AP;
                   extractFrame Q = AQ, ΨQ; distinct AQ;
                   C. Prop C (Ψ  ΨQ) P M N P' AP ΨP;
                   AP ♯* P; AP ♯* Q; AP ♯* Ψ; AP ♯* M; AP ♯* N; AP ♯* P'; AP ♯* AQ; AP ♯* ΨQ;
                   AQ ♯* P; AQ ♯* Q; AQ ♯* Ψ; AQ ♯* M; AQ ♯* N; AQ ♯* P'; AQ ♯* ΨP;
                   AP ♯* C; AQ ♯* C 
                   Prop C Ψ (P  Q) M N (P'  Q) (AP@AQ) (ΨP  ΨQ)"
  and   rPar2: "Ψ ΨP Q M N Q' AP P AQ ΨQ C.
                   Ψ  ΨP  Q MN  Q';
                   extractFrame P = AP, ΨP; distinct AP;
                   extractFrame Q = AQ, ΨQ; distinct AQ;
                   C. Prop C (Ψ  ΨP) Q M N Q' AQ ΨQ;
                   AP ♯* P; AP ♯* Q; AP ♯* Ψ; AP ♯* M; AP ♯* N; AP ♯* Q'; AP ♯* AQ; AP ♯* ΨQ;
                   AQ ♯* P; AQ ♯* Q; AQ ♯* Ψ; AQ ♯* M; AQ ♯* N; AQ ♯* Q'; AQ ♯* ΨP;
                   AP ♯* C; AQ ♯* C 
                   Prop C Ψ (P  Q) M N (P  Q') (AP@AQ) (ΨP  ΨQ)"
  and   rScope: "Ψ P M N P' x AP ΨP C.
                    Ψ  P MN  P'; extractFrame P = AP, ΨP; distinct AP;
                    C. Prop C Ψ P M N P' AP ΨP; x  Ψ; x  M; x  N;
                     x  AP; AP ♯* Ψ; AP ♯* P; AP ♯* M; AP ♯* N; AP ♯* P';
                     AP ♯* C; x  C 
                     Prop C Ψ (⦇νxP) M N (⦇νxP') (x#AP) ΨP"
  and   rBang:    "Ψ P M N P' AP ΨP C.
                     Ψ  P  !P MN  P'; guarded P; extractFrame P = AP, ΨP;  distinct AP;
                      C. Prop C Ψ (P  !P) M N P' AP (ΨP  𝟭); ΨP  𝟭; (supp ΨP) = ({}::name set);
                      AP ♯* Ψ; AP ♯* P; AP ♯* M; AP ♯* N; AP ♯* P'; AP ♯* C  Prop C Ψ (!P) M N P' ([]) (𝟭)"
shows "Prop C Ψ P M N P' AP ΨP"
  using Trans FrP distinct AP
proof(nominal_induct Ψ P Rs=="MN  P'" AP Ψ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 ¿MN  P'"
  and   FrP: "extractFrame P = AP, ΨP"
  and   "distinct AP"
  and   rAlpha: "Ψ P M N P' AP ΨP p C. AP ♯* Ψ; AP ♯* P; AP ♯* M; AP ♯* N; AP ♯* P'; AP ♯* (p  AP); AP ♯* C;
                                            set p  set AP × set(p  AP); distinctPerm p;
                                             Prop C Ψ P M N P' AP ΨP  Prop C Ψ P M N P' (p  AP) (p  Ψ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 AP ΨP C. Ψ  P ¿MN  P'; extractFrame P = AP, ΨP; distinct AP; C. Prop C Ψ P M N P' AP ΨP;
                                              (φ, P)  set Cs; Ψ  φ; guarded P;  ΨP  𝟭; (supp ΨP) = ({}::name set);
                                              AP ♯* Ψ; AP ♯* P; AP ♯* M; AP ♯* N; AP ♯* P'; AP ♯* C  Prop C Ψ (Cases Cs) M N P' ([]) (𝟭)"
  and   rPar1: "Ψ ΨQ P M N P' AQ Q AP ΨP C.
                   Ψ  ΨQ  P ¿MN  P';
                   extractFrame P = AP, ΨP; distinct AP;
                   extractFrame Q = AQ, ΨQ; distinct AQ;
                   C. Prop C (Ψ  ΨQ) P M N P' AP ΨP;
                   AP ♯* P; AP ♯* Q; AP ♯* Ψ; AP ♯* M; AP ♯* N; AP ♯* P'; AP ♯* AQ; AP ♯* ΨQ;
                   AQ ♯* P; AQ ♯* Q; AQ ♯* Ψ; AQ ♯* M; AQ ♯* N; AQ ♯* P'; AQ ♯* ΨP;
                   AP ♯* C; AQ ♯* C 
                   Prop C Ψ (P  Q) M N (P'  Q) (AP@AQ) (ΨP  ΨQ)"
  and   rPar2: "Ψ ΨP Q M N Q' AP P AQ ΨQ C.
                   Ψ  ΨP  Q ¿MN  Q';
                   extractFrame P = AP, ΨP; distinct AP;
                   extractFrame Q = AQ, ΨQ; distinct AQ;
                   C. Prop C (Ψ  ΨP) Q M N Q' AQ ΨQ;
                   AP ♯* P; AP ♯* Q; AP ♯* Ψ; AP ♯* M; AP ♯* N; AP ♯* Q'; AP ♯* AQ; AP ♯* ΨQ;
                   AQ ♯* P; AQ ♯* Q; AQ ♯* Ψ; AQ ♯* M; AQ ♯* N; AQ ♯* Q'; AQ ♯* ΨP;
                   AP ♯* C; AQ ♯* C 
                   Prop C Ψ (P  Q) M N (P  Q') (AP@AQ) (ΨP  ΨQ)"
  and   rBrMerge: "Ψ ΨQ P M N P' AP ΨP Q Q' AQ C.
                    Ψ  ΨQ  P  ¿MN  P'; C. Prop C (Ψ  ΨQ) P M N P' AP ΨP;
                    extractFrame P = AP, ΨP; distinct AP;
                    Ψ  ΨP  Q  ¿MN  Q'; C. Prop C (Ψ  ΨP) Q M N Q' AQ ΨQ;
                    extractFrame Q = AQ, ΨQ; distinct AQ;
                    AP ♯* Ψ; AP ♯* ΨQ; AP ♯* P; AP ♯* N; AP ♯* P';
                    AP ♯* Q; AP ♯* Q'; AP ♯* AQ; AP ♯* M; AQ ♯* M;
                    AQ ♯* Ψ; AQ ♯* ΨP; AQ ♯* P; AQ ♯* N; AQ ♯* P';
                    AQ ♯* Q; AQ ♯* Q'; AP ♯* C; AQ ♯* C 
                    Prop C Ψ (P  Q) M N (P'  Q') (AP@AQ) (ΨP  ΨQ)"
  and   rScope: "Ψ P M N P' x AP ΨP C.
                    Ψ  P ¿MN  P'; extractFrame P = AP, ΨP; distinct AP;
                    C. Prop C Ψ P M N P' AP ΨP; x  Ψ; x  M; x  N;
                     x  AP; AP ♯* Ψ; AP ♯* P; AP ♯* M; AP ♯* N; AP ♯* P';
                     AP ♯* C; x  C 
                     Prop C Ψ (⦇νxP) M N (⦇νxP') (x#AP) ΨP"
  and   rBang:    "Ψ P M N P' AP ΨP C.
                     Ψ  P  !P ¿MN  P'; guarded P; extractFrame P = AP, ΨP;  distinct AP;
                      C. Prop C Ψ (P  !P) M N P' AP (ΨP  𝟭); ΨP  𝟭; (supp ΨP) = ({}::name set);
                      AP ♯* Ψ; AP ♯* P; AP ♯* M; AP ♯* N; AP ♯* P'; AP ♯* C  Prop C Ψ (!P) M N P' ([]) (𝟭)"
shows "Prop C Ψ P M N P' AP ΨP"
  using Trans FrP distinct AP
proof(nominal_induct Ψ P Rs=="¿MN  P'" AP Ψ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 AP   :: "name list"
    and Ψ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 = AP, ΨP"
  and   "distinct AP"
  and   rAlpha: "Ψ P M AP ΨP p B C. AP ♯* Ψ; AP ♯* P; AP ♯* M; AP ♯* (p  AP); AP ♯* B; AP ♯* C;
                                         set p  set AP × set(p  AP); distinctPerm p;
                                          Prop C Ψ P M B AP ΨP  Prop C Ψ P M B (p  AP) (p  ΨP)"
  and   rOutput: "Ψ M K N P C. Ψ  M  K  Prop C Ψ (MN⟩.P) K (N ≺' P) ([]) (𝟭)"
  and   rCase: "Ψ P M B φ Cs AP ΨP C. Ψ  P (ROut M B); extractFrame P = AP, ΨP; distinct AP; C. Prop C Ψ P M B AP ΨP;
                                            (φ, P)  set Cs; Ψ  φ; guarded P;  ΨP  𝟭; (supp ΨP) = ({}::name set);
                                            AP ♯* Ψ; AP ♯* P; AP ♯* M; AP ♯* B; AP ♯* C  Prop C Ψ (Cases Cs) M B ([]) (𝟭)"
  and   rPar1: "Ψ ΨQ P M xvec N P' AQ Q AP ΨP C.
                   Ψ  ΨQ  P M⦇ν*xvec⦈⟨N  P';
                   extractFrame P = AP, ΨP; distinct AP;
                   extractFrame Q = AQ, ΨQ; distinct AQ;
                   C. Prop C (Ψ  ΨQ) P M (⦇ν*xvecN ≺' P') AP ΨP;
                   AP ♯* P; AP ♯* Q; AP ♯* Ψ; AP ♯* M;  AP ♯* xvec; AP ♯* N; AP ♯* P'; AP ♯* AQ; AP ♯* ΨQ;
                   AQ ♯* P; AQ ♯* Q; AQ ♯* Ψ; AQ ♯* M; AQ ♯* xvec; AQ ♯* N; AQ ♯* P'; AQ ♯* ΨP;
                   xvec ♯* Ψ; xvec ♯* P; xvec ♯* Q; xvec ♯* M; xvec ♯* ΨP; xvec ♯* ΨQ;
                   AP ♯* C; AQ ♯* C; xvec ♯* C 
                   Prop C Ψ (P  Q) M (⦇ν*xvecN ≺' (P'  Q)) (AP@AQ) (ΨP  ΨQ)"
  and   rPar2: "Ψ ΨP Q M xvec N Q' AP P AQ ΨQ C.
                   Ψ  ΨP  Q M⦇ν*xvec⦈⟨N  Q';
                   extractFrame P = AP, ΨP; distinct AP;
                   extractFrame Q = AQ, ΨQ; distinct AQ;
                   C. Prop C (Ψ  ΨP) Q M (⦇ν*xvecN ≺' Q') AQ ΨQ;
                   AP ♯* P; AP ♯* Q; AP ♯* Ψ; AP ♯* M; AP ♯* xvec; AP ♯* N; AP ♯* Q'; AP ♯* AQ; AP ♯* ΨQ;
                   AQ ♯* P; AQ ♯* Q; AQ ♯* Ψ; AQ ♯* M; AQ ♯* xvec; AQ ♯* N; AQ ♯* Q'; AQ ♯* ΨP;
                   xvec ♯* Ψ; xvec ♯* P; xvec ♯* Q; xvec ♯* M; xvec ♯* ΨP; xvec ♯* ΨQ;
                   AP ♯* C; AQ ♯* C; xvec ♯* C 
                   Prop C Ψ (P  Q) M (⦇ν*xvecN ≺' (P  Q')) (AP@AQ) (ΨP  ΨQ)"
  and   rOpen: "Ψ P M xvec yvec N P' x AP ΨP C.
                    Ψ  P M⦇ν*(xvec@yvec)⦈⟨N  P'; extractFrame P = AP, ΨP; distinct AP;
                    C. Prop C Ψ P M (⦇ν*(xvec@yvec)N ≺' P') AP ΨP; x  supp N; x  Ψ; x  M;
                     x  AP; x  xvec; x  yvec; AP ♯* Ψ; AP ♯* P; AP ♯* M; AP ♯* N; AP ♯* P';
                     AP ♯* xvec; AP ♯* yvec;
                     xvec ♯* Ψ; xvec ♯* P; xvec ♯* M; xvec ♯* ΨP;
                     yvec ♯* Ψ; yvec ♯* P; yvec ♯* M; AP ♯* C; x  C; xvec ♯* C; yvec ♯* C 
                     Prop C Ψ (⦇νxP) M (⦇ν*(xvec@x#yvec)N ≺' P') (x#AP) ΨP"
  and   rScope: "Ψ P M xvec N P' x AP ΨP C.
                    Ψ  P M⦇ν*xvec⦈⟨N  P'; extractFrame P = AP, ΨP; distinct AP;
                    C. Prop C Ψ P M (⦇ν*xvecN ≺' P') AP ΨP;
                     x  Ψ; x  M; x  xvec; x  N; x  AP; AP ♯* Ψ; AP ♯* P;
                     AP ♯* M; AP ♯* N; AP ♯* P'; AP ♯* xvec;
                     xvec ♯* Ψ; xvec ♯* P; xvec ♯* M; xvec ♯* ΨP;
                     AP ♯* C; x  C; xvec ♯* C 
                     Prop C Ψ (⦇νxP) M (⦇ν*xvecN ≺' (⦇νxP')) (x#AP) ΨP"
  and   rBang:    "Ψ P M B AP ΨP C.
                     Ψ  P  !P ROut M B; guarded P; extractFrame P = AP, ΨP; distinct AP;
                      C. Prop C Ψ (P  !P) M B AP (ΨP  𝟭); ΨP  𝟭; supp ΨP = ({}::name set);
                      AP ♯* Ψ; AP ♯* P; AP ♯* M; AP ♯* C  Prop C Ψ (!P) M B ([]) (𝟭)"
shows "Prop C Ψ P M B AP ΨP"
proof -
  {
    fix B
    assume "Ψ  P ROut M B"
    then have "Prop C Ψ P M B AP ΨP" using FrP distinct AP
    proof(nominal_induct Ψ P Rs=="ROut M B" AP Ψ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 AP   :: "name list"
    and Ψ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 = AP, ΨP"
  and   "distinct AP"
  and   rAlpha: "Ψ P M AP ΨP p B C. AP ♯* Ψ; AP ♯* P; AP ♯* M; AP ♯* (p  AP); AP ♯* B; AP ♯* C;
                                         set p  set AP × set(p  AP); distinctPerm p;
                                          Prop C Ψ P M B AP ΨP  Prop C Ψ P M B (p  AP) (p  ΨP)"
  and   rBrOutput: "Ψ M K N P C. Ψ  M  K  Prop C Ψ (MN⟩.P) K (N ≺' P) ([]) (𝟭)"
  and   rCase: "Ψ P M B φ Cs AP ΨP C. Ψ  P (RBrOut M B); extractFrame P = AP, ΨP; distinct AP; C. Prop C Ψ P M B AP ΨP;
                                            (φ, P)  set Cs; Ψ  φ; guarded P;  ΨP  𝟭; (supp ΨP) = ({}::name set);
                                            AP ♯* Ψ; AP ♯* P; AP ♯* M; AP ♯* B; AP ♯* C  Prop C Ψ (Cases Cs) M B ([]) (𝟭)"
  and   rPar1: "Ψ ΨQ P M xvec N P' AQ Q AP ΨP C.
                   Ψ  ΨQ  P ¡M⦇ν*xvec⦈⟨N  P';
                   extractFrame P = AP, ΨP; distinct AP;
                   extractFrame Q = AQ, ΨQ; distinct AQ;
                   C. Prop C (Ψ  ΨQ) P M (⦇ν*xvecN ≺' P') AP ΨP;
                   AP ♯* P; AP ♯* Q; AP ♯* Ψ; AP ♯* M;  AP ♯* xvec; AP ♯* N; AP ♯* P'; AP ♯* AQ; AP ♯* ΨQ;
                   AQ ♯* P; AQ ♯* Q; AQ ♯* Ψ; AQ ♯* M; AQ ♯* xvec; AQ ♯* N; AQ ♯* P'; AQ ♯* ΨP;
                   xvec ♯* Ψ; xvec ♯* P; xvec ♯* Q; xvec ♯* M; xvec ♯* ΨP; xvec ♯* ΨQ;
                   AP ♯* C; AQ ♯* C; xvec ♯* C 
                   Prop C Ψ (P  Q) M (⦇ν*xvecN ≺' (P'  Q)) (AP@AQ) (ΨP  ΨQ)"
  and   rPar2: "Ψ ΨP Q M xvec N Q' AP P AQ ΨQ C.
                   Ψ  ΨP  Q ¡M⦇ν*xvec⦈⟨N  Q';
                   extractFrame P = AP, ΨP; distinct AP;
                   extractFrame Q = AQ, ΨQ; distinct AQ;
                   C. Prop C (Ψ  ΨP) Q M (⦇ν*xvecN ≺' Q') AQ ΨQ;
                   AP ♯* P; AP ♯* Q; AP ♯* Ψ; AP ♯* M; AP ♯* xvec; AP ♯* N; AP ♯* Q'; AP ♯* AQ; AP ♯* ΨQ;
                   AQ ♯* P; AQ ♯* Q; AQ ♯* Ψ; AQ ♯* M; AQ ♯* xvec; AQ ♯* N; AQ ♯* Q'; AQ ♯* ΨP;
                   xvec ♯* Ψ; xvec ♯* P; xvec ♯* Q; xvec ♯* M; xvec ♯* ΨP; xvec ♯* ΨQ;
                   AP ♯* C; AQ ♯* C; xvec ♯* C 
                   Prop C Ψ (P  Q) M (⦇ν*xvecN ≺' (P  Q')) (AP@AQ) (ΨP  ΨQ)"
  and   rBrComm1:"Ψ ΨQ P M N P' AP ΨP Q xvec Q' AQ C.
                   Ψ  ΨQ  P ¿MN  P'; extractFrame P = AP, ΨP; distinct AP;
                    Ψ  ΨP  Q ¡M⦇ν*xvec⦈⟨N  Q'; extractFrame Q = AQ, ΨQ; distinct AQ;
                    distinct xvec;
                   C. Prop C (Ψ  ΨP) Q M (⦇ν*xvecN ≺' Q') AQ ΨQ;
                    AP ♯* Ψ; AP ♯* ΨQ; AP ♯* P; AP ♯* N; AP ♯* P';
                    AP ♯* Q; AP ♯* Q'; AP ♯* AQ; AP ♯* xvec; AQ ♯* Ψ; AQ ♯* ΨP;
                    AQ ♯* P; AQ ♯* N; AQ ♯* P'; AQ ♯* Q; AQ ♯* Q';
                    AQ ♯* xvec; xvec ♯* Ψ; xvec ♯* ΨP; xvec ♯* ΨQ; xvec ♯* P;
                    xvec ♯* Q; AP ♯* C; AQ ♯* C; xvec ♯* C;
                    AP ♯* M; AQ ♯* M; xvec ♯* M 
                    Prop C Ψ (P  Q) M (⦇ν*xvecN ≺' (P'  Q')) (AP@AQ) (ΨP  ΨQ)" (* Not applicable: ⋀C. Prop C (Ψ ⊗ ΨQ) P (M⦇N⦈) P' AP ΨP; *)
  and   rBrComm2:"Ψ ΨQ P M xvec N P' AP ΨP Q Q' AQ C.
                   Ψ  ΨQ  P ¡M⦇ν*xvec⦈⟨N  P'; extractFrame P = AP, ΨP; distinct AP;
                   C. Prop C (Ψ  ΨQ) P M (⦇ν*xvecN ≺' P') AP ΨP;
                    Ψ  ΨP  Q ¿MN  Q'; extractFrame Q = AQ, ΨQ; distinct AQ;
                    distinct xvec;
                    AP ♯* Ψ; AP ♯* ΨQ; AP ♯* P; AP ♯* N; AP ♯* P';
                    AP ♯* Q; AP ♯* Q'; AP ♯* AQ; AP ♯* xvec; AQ ♯* Ψ; AQ ♯* ΨP;
                    AQ ♯* P; AQ ♯* N; AQ ♯* P'; AQ ♯* Q; AQ ♯* Q';
                    AQ ♯* xvec; xvec ♯* Ψ; xvec ♯* ΨP; xvec ♯* ΨQ; xvec ♯* P;
                    xvec ♯* Q; AP ♯* C; AQ ♯* C; xvec ♯* C;
                    AP ♯* M; AQ ♯* M; xvec ♯* M 
                    Prop C Ψ (P  Q) M (⦇ν*xvecN ≺' (P'  Q')) (AP@AQ) (ΨP  ΨQ)" (* Not applicable: ⋀C. Prop C (Ψ ⊗ ΨP) Q (K⦇N⦈) Q' AQ ΨQ; *)
  and   rBrOpen: "Ψ P M xvec yvec N P' x AP ΨP C.
                    Ψ  P ¡M⦇ν*(xvec@yvec)⦈⟨N  P'; extractFrame P = AP, ΨP; distinct AP;
                    C. Prop C Ψ P M (⦇ν*(xvec@yvec)N ≺' P') AP ΨP; x  supp N; x  Ψ; x  M;
                     x  AP; x  xvec; x  yvec; AP ♯* Ψ; AP ♯* P; AP ♯* M; AP ♯* N; AP ♯* P';
                     AP ♯* xvec; AP ♯* yvec;
                     xvec ♯* Ψ; xvec ♯* P; xvec ♯* M; xvec ♯* ΨP;
                     yvec ♯* Ψ; yvec ♯* P; yvec ♯* M; AP ♯* C; x  C; xvec ♯* C; yvec ♯* C 
                     Prop C Ψ (⦇νxP) M (⦇ν*(xvec@x#yvec)N ≺' P') (x#AP) ΨP"
  and   rScope: "Ψ P M xvec N P' x AP ΨP C.
                    Ψ  P ¡M⦇ν*xvec⦈⟨N  P'; extractFrame P = AP, ΨP; distinct AP;
                    C. Prop C Ψ P M (⦇ν*xvecN ≺' P') AP ΨP;
                     x  Ψ; x  M; x  xvec; x  N; x  AP; AP ♯* Ψ; AP ♯* P;
                     AP ♯* M; AP ♯* N; AP ♯* P'; AP ♯* xvec;
                     xvec ♯* Ψ; xvec ♯* P; xvec ♯* M; xvec ♯* ΨP;
                     AP ♯* C; x  C; xvec ♯* C 
                     Prop C Ψ (⦇νxP) M (⦇ν*xvecN ≺' (⦇νxP')) (x#AP) ΨP"
  and   rBang:    "Ψ P M B AP ΨP C.
                     Ψ  P  !P RBrOut M B; guarded P; extractFrame P = AP, ΨP; distinct AP;
                      C. Prop C Ψ (P  !P) M B AP (ΨP  𝟭); ΨP  𝟭; supp ΨP = ({}::name set);
                      AP ♯* Ψ; AP ♯* P; AP ♯* M; AP ♯* C  Prop C Ψ (!P) M B ([]) (𝟭)"
shows "Prop C Ψ P M B AP ΨP"
proof -
  {
    fix B
    assume "Ψ  P RBrOut M B"
    then have "Prop C Ψ P M B AP ΨP" using FrP distinct AP
    proof(nominal_induct Ψ P Rs=="RBrOut M B" AP Ψ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 = AP, ΨP"
  and   "distinct AP"
  and   rAlpha: "Ψ P P' AP ΨP p C. AP ♯* Ψ; AP ♯* P; AP ♯* P'; AP ♯* (p  AP); AP ♯* C;
                                        set p  set AP × set (p  AP); distinctPerm p;
                                         Prop C Ψ P P' AP ΨP  Prop C Ψ P P' (p  AP) (p  ΨP)"
  and   rCase: "Ψ P P' φ Cs AP ΨP C. Ψ  P τ  P'; extractFrame P = AP, ΨP; distinct AP; C. Prop C Ψ P P' AP ΨP;
                                          (φ, P)  set Cs; Ψ  φ; guarded P;  ΨP  𝟭; (supp ΨP) = ({}::name set);
                                          AP ♯* Ψ; AP ♯* P; AP ♯* P'; AP ♯* C  Prop C Ψ (Cases Cs) P' ([]) (𝟭)"
  and   rPar1: "Ψ ΨQ P P' AQ Q AP ΨP C.
                   Ψ  ΨQ  P τ  P';
                   extractFrame P = AP, ΨP; distinct AP;
                   extractFrame Q = AQ, ΨQ; distinct AQ;
                   C. Prop C (Ψ  ΨQ) P P' AP ΨP;
                   AP ♯* P; AP ♯* Q; AP ♯* Ψ; AP ♯* P'; AP ♯* AQ; AP ♯* ΨQ;
                   AQ ♯* P; AQ ♯* Q; AQ ♯* Ψ; AQ ♯* P'; AQ ♯* ΨP;
                   AP ♯* C; AQ ♯* C 
                   Prop C Ψ (P  Q) (P'  Q) (AP@AQ) (ΨP  ΨQ)"
  and   rPar2: "Ψ ΨP Q Q' AP P AQ ΨQ C.
                   Ψ  ΨP  Q τ  Q';
                   extractFrame P = AP, ΨP; distinct AP;
                   extractFrame Q = AQ, ΨQ; distinct AQ;
                   C. Prop C (Ψ  ΨP) Q Q' AQ ΨQ;
                   AP ♯* P; AP ♯* Q; AP ♯* Ψ; AP ♯* Q'; AP ♯* AQ; AP ♯* ΨQ;
                   AQ ♯* P; AQ ♯* Q; AQ ♯* Ψ; AQ ♯* Q'; AQ ♯* ΨP;
                   AP ♯* C; AQ ♯* C 
                   Prop C Ψ (P  Q) (P  Q') (AP@AQ) (ΨP  ΨQ)"
  and   rComm1: "Ψ ΨQ P M N P' AP ΨP Q K xvec Q' AQ C.
                   Ψ  ΨQ  P MN  P'; extractFrame P = AP, ΨP; distinct AP;
                    Ψ  ΨP  Q K⦇ν*xvec⦈⟨N  Q'; extractFrame Q = AQ, ΨQ; distinct AQ;
                    Ψ  ΨP  ΨQ  M  K; distinct xvec;
                    AP ♯* Ψ; AP ♯* ΨQ; AP ♯* P; AP ♯* M; AP ♯* N; AP ♯* P';
                    AP ♯* Q; AP ♯* Q'; AP ♯* AQ; AP ♯* xvec; AQ ♯* Ψ; AQ ♯* ΨP;
                    AQ ♯* P; AQ ♯* N; AQ ♯* P'; AQ ♯* Q; AQ ♯* K; AQ ♯* Q';
                    AQ ♯* xvec; xvec ♯* Ψ; xvec ♯* ΨP; xvec ♯* ΨQ; xvec ♯* P; xvec ♯* M;
                    xvec ♯* Q; xvec ♯* K; AP ♯* C; AQ ♯* C; xvec ♯* C 
                    Prop C Ψ (P  Q) (⦇ν*xvec(P'  Q')) (AP@AQ) (ΨP  ΨQ)"
  and   rComm2: "Ψ ΨQ P M xvec N P' AP ΨP Q K Q' AQ C.
                   Ψ  ΨQ  P M⦇ν*xvec⦈⟨N  P'; extractFrame P = AP, ΨP; distinct AP;
                    Ψ  ΨP  Q KN  Q'; extractFrame Q = AQ, ΨQ; distinct AQ;
                    Ψ  ΨP  ΨQ  M  K; distinct xvec;
                    AP ♯* Ψ; AP ♯* ΨQ; AP ♯* P; AP ♯* M; AP ♯* N; AP ♯* P';
                    AP ♯* Q; AP ♯* Q'; AP ♯* AQ; AP ♯* xvec; AQ ♯* Ψ; AQ ♯* ΨP;
                    AQ ♯* P; AQ ♯* N; AQ ♯* P'; AQ ♯* Q; AQ ♯* K; AQ ♯* Q';
                    AQ ♯* xvec; xvec ♯* Ψ; xvec ♯* ΨP; xvec ♯* ΨQ; xvec ♯* P; xvec ♯* M;
                    xvec ♯* Q; xvec ♯* K; AP ♯* C; AQ ♯* C; xvec ♯* C 
                    Prop C Ψ (P  Q) (⦇ν*xvec(P'  Q')) (AP@AQ) (ΨP  ΨQ)"
  and   rBrClose: "Ψ P M xvec N P' AP ΨP x C.
                    Ψ  P  ¡M⦇ν*xvec⦈⟨N  P';
                     x  supp M;
                     extractFrame P = AP, ΨP; distinct AP;
                     AP ♯* Ψ; AP ♯* P; AP ♯* M; AP ♯* N; AP ♯* P'; AP ♯* xvec;
                     distinct xvec; xvec ♯* Ψ; xvec ♯* ΨP; xvec ♯* P;
                     xvec ♯* M;
                     x  Ψ; x  xvec; x  AP;
                     AP ♯* C; xvec ♯* C; x  C 
                     Prop C Ψ (⦇νxP) (⦇νx(⦇ν*xvecP')) (x#AP) ΨP"
  and   rScope: "Ψ P P' x AP ΨP C.
                    Ψ  P τ  P'; extractFrame P = AP, ΨP; distinct AP;
                    C. Prop C Ψ P P' AP ΨP; x  Ψ;
                     x  AP; AP ♯* Ψ; AP ♯* P; AP ♯* P';
                     AP ♯* C; x  C 
                     Prop C Ψ (⦇νxP) (⦇νxP') (x#AP) ΨP"
  and   rBang:    "Ψ P P' AP ΨP C.
                     Ψ  P  !P τ  P'; guarded P; extractFrame P = AP, ΨP;  distinct AP;
                      C. Prop C Ψ (P  !P) P' AP (ΨP  𝟭); ΨP  𝟭; supp ΨP = ({}::name set);
                      AP ♯* Ψ; AP ♯* P; AP ♯* P'; AP ♯* C  Prop C Ψ (!P) P' ([]) (𝟭)"
shows "Prop C Ψ P P' AP ΨP"
  using Trans FrP distinct AP
proof(nominal_induct Ψ P Rs=="τ  P'" AP Ψ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 MN  P'"
  and   "x  P"
  and   "x  N"

shows "x  P'"
proof -
  have "bn(MN) ♯* subject(MN)" and "distinct(bn(MN))" by simp+
  with Ψ  P MN  P' show ?thesis using x  P x  N
  proof(nominal_induct Ψ P α=="MN" 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]) = MN 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 Ψ ΨQ P P' xvec Q x)
    then show ?case by simp
  next
    case(cPar2 Ψ ΨP Q Q' xvec P x)
    then show ?case by simp
  next
    case(cComm1 Ψ ΨQ P M N P' AP ΨP Q K xvec Q' AQ x)
    then show ?case by simp
  next
    case(cComm2 Ψ ΨQ P M xwec N P' AP ΨP Q K Q' AQ 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 ¿MN  P'"
  and   "x  P"
  and   "x  N"

shows "x  P'"
proof -
  have "bn(¿MN) ♯* subject(¿MN)" and "distinct(bn(¿MN))" by simp+
  with Ψ  P ¿MN  P' show ?thesis using x  P x  N
  proof(nominal_induct Ψ P α=="¿MN" 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]) = ¿MN 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 Ψ ΨQ P P' xvec Q x)
    then show ?case by simp
  next
    case(cPar2 Ψ ΨP Q Q' xvec P x)
    then show ?case by simp
  next
    case(cComm1 Ψ ΨQ P M N P' AP ΨP Q K xvec Q' AQ x)
    then show ?case by simp
  next
    case(cComm2 Ψ ΨQ P M xwec N P' AP ΨP Q K Q' AQ 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 MN  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 ¿MN  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  ⦇νxP x  y  have "y  P" by(simp add: abs_fresh)
    moreover from y  zvec zvec = xvec@x#yvechave "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  ⦇νxP x  y  have "y  P" by(simp add: abs_fresh)
    moreover from y  zvec zvec = xvec@x#yvechave "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  ⦇νxP x  y  have "y  P" by(simp add: abs_fresh)
    moreover from y  zvec zvec = xvec@x#yvechave "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 Ψ ΨQ P M N P' AP ΨP Q xvec Q' AQ x M' zvec N')
    from x  (P  Q) have "x  P" and "x  Q" by simp+

    from Ψ  ΨP  Q  ¡M⦇ν*xvec⦈⟨N  Q' xvec ♯* M distinct xvec x  Q xvec ♯* x
    have "x  N" by(simp add: broutputFreshDerivativeN)

    with Ψ  ΨQ  P  ¿MN  P' x  P have "x  P'" by(simp add: brinputFreshDerivative)

    then show ?case using cBrComm1 by simp
  next
    case (cBrComm2 Ψ ΨQ P M xvec N P' AP ΨP Q Q' AQ x M' zvec N')
    from x  (P  Q) have "x  P" and "x  Q" by simp+

    from Ψ  ΨQ  P  ¡M⦇ν*xvec⦈⟨N  P' xvec ♯* M distinct xvec x  P xvec ♯* x
    have "x  N" by(simp add: broutputFreshDerivativeN)

    with Ψ  ΨP  Q  ¿MN  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  ⦇νxP x  y  have "y  P" by(simp add: abs_fresh)
    moreover from y  zvec zvec = xvec@x#yvechave "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 KN[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 ¿KN[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 ΨQ   :: 'b
    and P    :: "('a, 'b, 'c) psi"
    and α    :: "'a action"
    and P'   :: "('a, 'b, 'c) psi"
    and AQ   :: "name list"
    and Q    :: "('a, 'b, 'c) psi"

assumes Trans: "Ψ  ΨQ  P α  P'"
  and   "extractFrame Q = AQ, ΨQ"
  and   "bn α ♯* Q"
  and   "AQ ♯* Ψ"
  and   "AQ ♯* P"
  and   "AQ ♯* α"

shows "Ψ  P  Q α  (P'  Q)"
proof -
  {
    fix Ψ    :: 'b
      and ΨQ   :: 'b
      and P    :: "('a, 'b, 'c) psi"
      and α    :: "'a action"
      and P'   :: "('a, 'b, 'c) psi"
      and AQ   :: "name list"
      and Q    :: "('a, 'b, 'c) psi"

    assume "Ψ  ΨQ  P α  P'"
      and   "extractFrame Q = AQ, ΨQ"
      and   "bn α ♯* Q"
      and   "bn α ♯* subject α"
      and   "AQ ♯* Ψ"
      and   "AQ ♯* P"
      and   "AQ ♯* α"
      and   "distinct AQ"

    have  "Ψ  P  Q α  (P'  Q)"
    proof -
      from Ψ  Ψ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  α) ♯* AQ" and "bn(q  α) ♯* P'" and "bn(q  α) ♯* ΨQ"
        and Sq: "(set q)  (set (bn α)) × (set(bn(q  α)))"
        apply -
        by(rule name_list_avoiding[where xvec="bn α" and c="(Ψ, P, Q, α, AQ, ΨQ, P')"]) (auto simp add: eqvts)
      obtain p::"name prm" where "(p  AQ) ♯* Ψ" and "(p  AQ) ♯* P" and "(p  AQ) ♯* Q" and "(p  AQ) ♯* α"
        and "(p  AQ) ♯* α" and "(p  AQ) ♯* (q  α)" and "(p  AQ) ♯* P'"
        and "(p  AQ) ♯* (q  P')" and "(p  AQ) ♯* ΨQ" and Sp: "(set p)  (set AQ) × (set(p  AQ))"
        apply -
        by(rule name_list_avoiding[where xvec=AQ and c="(Ψ, P, Q, α, bn α, q  α, P', (q  P'), ΨQ)"]) auto
      from distinct(bn α) have "distinct(bn(q  α))"
        by - (rule actionCases[where α=α], auto simp add: eqvts)
      from AQ ♯* α bn(q  α) ♯* AQ Sq have "AQ ♯* (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 Ψ  ΨQ  P α  P' bn(q  α) ♯* α bn(q  α) ♯* P' bn α ♯* (subject α) Sq
      have Trans: "Ψ  ΨQ  P (q  α)  (q  P')"
        by(force simp add: residualAlpha)
      then have "AQ ♯* (q  P')" using  bn(q  α) ♯* subject(q  α) distinct(bn(q  α)) AQ ♯* P AQ ♯* (q  α)
        by(auto intro: freeFreshChainDerivative)
      from Trans have "(p  (Ψ  ΨQ))  (p  P) p  ((q  α)  (q  P'))"
        by(rule semantics.eqvt)
      with AQ ♯* Ψ AQ ♯* P AQ ♯* (q  α) (p  AQ) ♯* (q  α) AQ ♯* (q  P')
        (p  AQ) ♯* Ψ (p  AQ) ♯* P (p  AQ) ♯* (q  P') Sp
      have "Ψ  (p  ΨQ)  P (q  α)  (q  P')" by(simp add: eqvts)
      moreover from extractFrame Q = AQ, ΨQ (p  AQ) ♯* ΨQ Sp have  "extractFrame Q = (p  AQ), (p  ΨQ)"
        by(simp add: frameChainAlpha' eqvts)
      moreover from (bn(q  α)) ♯* ΨQ (bn(q  α)) ♯* AQ (p  AQ) ♯* (q  α) Sp
      have "(bn(q  α)) ♯* (p  ΨQ)"
        by(simp add: freshAlphaPerm)
      moreover from distinct AQ have "distinct(p  AQ)" by simp
      ultimately have "Ψ  P  Q (q  α)  ((q  P')  Q)"
        using (p  AQ) ♯* P (p  AQ) ♯* Q (p  AQ) ♯* Ψ (p  AQ) ♯* (q  α)
          (p  AQ) ♯* (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 = AQ, ΨQ AQ ♯* Ψ AQ ♯* P AQ ♯* α
  obtain AQ' where FrQ: "extractFrame Q = AQ', ΨQ" and "distinct AQ'" and "AQ' ♯* Ψ" and "AQ' ♯* P" and "AQ' ♯* α"
    apply -
    by(rule distinctFrame[where C="(Ψ, P, α)"]) auto
  show ?thesis
  proof(induct rule: actionCases[where α=α])
    case(cInput M N)
    from Trans FrQ AQ' ♯* Ψ AQ' ♯* P AQ' ♯* α distinct AQ' bn α ♯* Q
    show ?case using α = MN by(force intro: Goal)
  next
    case(cBrInput M N)
    from Trans FrQ AQ' ♯* Ψ AQ' ♯* P AQ' ♯* α distinct AQ' bn α ♯* Q
    show ?case using α = ¿MN by(force intro: Goal)
  next
    case cTau
    from Trans FrQ AQ' ♯* Ψ AQ' ♯* P AQ' ♯* α distinct AQ' bn α ♯* Q
    show ?case using α = τ by(force intro: Goal)
  next
    case(cOutput M xvec N)
    from α = M⦇ν*xvec⦈⟨N AQ' ♯* α bn α ♯* Q have "xvec ♯* AQ'" 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) ♯* AQ'"
      and S: "set p  set xvec × set(p  xvec)"
      apply -
      by(rule name_list_avoiding[where xvec=xvec and c="(N, P', Q, M, AQ')"]) auto
    from Trans α=M⦇ν*xvec⦈⟨N have "Ψ  ΨQ  P M⦇ν*xvec⦈⟨N  P'" by simp
    with (p  xvec) ♯* N (p  xvec) ♯* P' S
    have "Ψ  ΨQ  P M⦇ν*(p  xvec)⦈⟨(p  N)  (p  P')"
      by(simp add: boundOutputChainAlpha'' create_residual.simps)
    moreover from xvec ♯* AQ' (p  xvec) ♯* AQ' AQ' ♯* α S
    have "AQ' ♯* (p  α)" by(simp add: freshChainSimps del: actionFreshChain)
    ultimately have "Ψ  P  Q M⦇ν*(p  xvec)⦈⟨(p  N)  (p  P')  Q"
      using FrQ AQ' ♯* Ψ AQ' ♯* P distinct AQ' (p  xvec) ♯* Q AQ' ♯* α
        (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 AQ' ♯* α bn α ♯* Q have "xvec ♯* AQ'" 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) ♯* AQ'"
      and S: "set p  set xvec × set(p  xvec)"
      by(rule name_list_avoiding[where xvec=xvec and c="(N, P', Q, M, AQ')"]) auto
    from Trans α=¡M⦇ν*xvec⦈⟨N have "Ψ  ΨQ  P ¡M⦇ν*xvec⦈⟨N  P'" by simp
    with (p  xvec) ♯* N (p  xvec) ♯* P' S
    have "Ψ  ΨQ  P ¡M⦇ν*(p  xvec)⦈⟨(p  N)  (p  P')"
      by(simp add: boundOutputChainAlpha'' create_residual.simps)
    moreover from xvec ♯* AQ' (p  xvec) ♯* AQ' AQ' ♯* α S
    have "AQ' ♯* (p  α)" by(simp add: freshChainSimps del: actionFreshChain)
    ultimately have "Ψ  P  Q ¡M⦇ν*(p  xvec)⦈⟨(p  N)  (p  P')  Q"
      using FrQ AQ' ♯* Ψ AQ' ♯* P distinct AQ' (p  xvec) ♯* Q AQ' ♯* α
        (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 ΨP   :: 'b
    and Q    :: "('a, 'b, 'c) psi"
    and α    :: "'a action"
    and Q'   :: "('a, 'b, 'c) psi"
    and AP   :: "name list"
    and P    :: "('a, 'b, 'c) psi"

assumes Trans: "Ψ  ΨP  Q α  Q'"
  and   "extractFrame P = AP, ΨP"
  and   "bn α ♯* P"
  and   "AP ♯* Ψ"
  and   "AP ♯* Q"
  and   "AP ♯* α"

shows "Ψ  P  Q α  (P  Q')"
proof -
  {
    fix Ψ    :: 'b
      and ΨP   :: 'b
      and Q    :: "('a, 'b, 'c) psi"
      and α    :: "'a action"
      and Q'   :: "('a, 'b, 'c) psi"
      and AP   :: "name list"
      and P    :: "('a, 'b, 'c) psi"

    assume "Ψ  ΨP  Q α  Q'"
      and   "extractFrame P = AP, ΨP"
      and   "bn α ♯* P"
      and   "bn α ♯* subject α"
      and   "AP ♯* Ψ"
      and   "AP ♯* Q"
      and   "AP ♯* α"
      and   "distinct AP"

    have  "Ψ  P  Q α  (P  Q')"
    proof -
      from Ψ  Ψ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  α) ♯* AP" and "bn(q  α) ♯* Q'" and "bn(q  α) ♯* ΨP"
        and Sq: "(set q)  (set (bn α)) × (set(bn(q  α)))"
        by(rule name_list_avoiding[where xvec="bn α" and c="(Ψ, P, Q, α, AP, ΨP, Q')"]) (auto simp add: eqvts)
      obtain p::"name prm" where "(p  AP) ♯* Ψ" and "(p  AP) ♯* P" and "(p  AP) ♯* Q" and "(p  AP) ♯* α"
        and "(p  AP) ♯* α" and "(p  AP) ♯* (q  α)" and "(p  AP) ♯* Q'"
        and "(p  AP) ♯* (q  Q')" and "(p  AP) ♯* ΨP"
        and Sp: "(set p)  (set AP) × (set(p  AP))"
        by(rule name_list_avoiding[where xvec=AP and c="(Ψ, P, Q, α, q  α, Q', (q  Q'), ΨP)"]) auto
      from distinct(bn α) have "distinct(bn(q  α))"
        apply -
        by(rule actionCases[where α=α]) (auto simp add: eqvts)
      from AP ♯* α bn(q  α) ♯* AP Sq have "AP ♯* (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 Ψ  ΨP  Q α  Q' bn(q  α) ♯* α bn(q  α) ♯* Q' bn α ♯* (subject α) Sq
      have Trans: "Ψ  ΨP  Q (q  α)  (q  Q')"
        by(force simp add: residualAlpha)
      then have "AP ♯* (q  Q')" using  bn(q  α) ♯* subject(q  α) distinct(bn(q  α)) AP ♯* Q AP ♯* (q  α)
        by(auto intro: freeFreshChainDerivative)
      from Trans have "(p  (Ψ  ΨP))  (p  Q) p  ((q  α)  (q  Q'))"
        by(rule semantics.eqvt)
      with AP ♯* Ψ AP ♯* Q AP ♯* (q  α) (p  AP) ♯* (q  α) AP ♯* (q  Q')
        (p  AP) ♯* Ψ (p  AP) ♯* Q (p  AP) ♯* (q  Q') Sp
      have "Ψ  (p  ΨP)  Q (q  α)  (q  Q')" by(simp add: eqvts)
      moreover from extractFrame P = AP, ΨP (p  AP) ♯* ΨP Sp have  "extractFrame P = (p  AP), (p  ΨP)"
        by(simp add: frameChainAlpha' eqvts)
      moreover from (bn(q  α)) ♯* ΨP (bn(q  α)) ♯* AP (p  AP) ♯* (q  α) Sp
      have "(bn(q  α)) ♯* (p  ΨP)"
        by(simp add: freshAlphaPerm)
      moreover from distinct AP have "distinct(p  AP)" by simp
      ultimately have "Ψ  P  Q (q  α)  (P  (q  Q'))"
        using (p  AP) ♯* P (p  AP) ♯* Q (p  AP) ♯* Ψ (p  AP) ♯* (q  α)
          (p  AP) ♯* (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 = AP, ΨP AP ♯* Ψ AP ♯* Q AP ♯* α
  obtain AP' where FrP: "extractFrame P = AP', ΨP" and "distinct AP'" and "AP' ♯* Ψ" and "AP' ♯* Q" and "AP' ♯* α"
    apply -
    by(rule distinctFrame[where C="(Ψ, Q, α)"]) auto
  show ?thesis
  proof(induct rule: actionCases[where α=α])
    case(cInput M N)
    from Trans FrP AP' ♯* Ψ AP' ♯* Q AP' ♯* α distinct AP' bn α ♯* P
    show ?case using α = MN by(force intro: Goal)
  next
    case(cBrInput M N)
    from Trans FrP AP' ♯* Ψ AP' ♯* Q AP' ♯* α distinct AP' bn α ♯* P
    show ?case using α = ¿MN by(force intro: Goal)
  next
    case cTau
    from Trans FrP AP' ♯* Ψ AP' ♯* Q AP' ♯* α distinct AP' bn α ♯* P
    show ?case using α = τ by(force intro: Goal)
  next
    case(cOutput M xvec N)
    from α = M⦇ν*xvec⦈⟨N AP' ♯* α bn α ♯* P have "xvec ♯* AP'" 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) ♯* AP'"
      and S: "set p  set xvec × set(p  xvec)"
      by(rule name_list_avoiding[where xvec=xvec and c="(N, Q', P, M, AP')"]) auto
    from Trans α=M⦇ν*xvec⦈⟨N have "Ψ  ΨP  Q M⦇ν*xvec⦈⟨N  Q'" by simp
    with (p  xvec) ♯* N (p  xvec) ♯* Q' S
    have "Ψ  ΨP  Q M⦇ν*(p  xvec)⦈⟨(p  N)  (p  Q')"
      by(simp add: boundOutputChainAlpha'' create_residual.simps)
    moreover from xvec ♯* AP' (p  xvec) ♯* AP' AP' ♯* α S
    have "AP' ♯* (p  α)" by(simp add: freshChainSimps del: actionFreshChain)
    ultimately have "Ψ  P  Q M⦇ν*(p  xvec)⦈⟨(p  N)  P  (p  Q')"
      using FrP AP' ♯* Ψ AP' ♯* Q distinct AP' (p  xvec) ♯* P AP' ♯* α
        (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 AP' ♯* α bn α ♯* P have "xvec ♯* AP'" 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) ♯* AP'"
      and S: "set p  set xvec × set(p  xvec)"
      by(rule name_list_avoiding[where xvec=xvec and c="(N, Q', P, M, AP')"]) auto
    from Trans α=¡M⦇ν*xvec⦈⟨N have "Ψ  ΨP  Q ¡M⦇ν*xvec⦈⟨N  Q'" by simp
    with (p  xvec) ♯* N (p  xvec) ♯* Q' S
    have "Ψ  ΨP  Q ¡M⦇ν*(p  xvec)⦈⟨(p  N)  (p  Q')"
      by(simp add: boundOutputChainAlpha'' create_residual.simps)
    moreover from xvec ♯* AP' (p  xvec) ♯* AP' AP' ♯* α S
    have "AP' ♯* (p  α)" by(simp add: freshChainSimps del: actionFreshChain)
    ultimately have "Ψ  P  Q ¡M⦇ν*(p  xvec)⦈⟨(p  N)  P  (p  Q')"
      using FrP AP' ♯* Ψ AP' ♯* Q distinct AP' (p  xvec) ♯* P AP' ♯* α
        (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 "Ψ  ⦇νxP 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 "Ψ  ⦇νxP 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 "Ψ  ⦇νxP 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 "Ψ  ⦇νxP ¡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 "Ψ  ⦇νxP ¡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 "Ψ  ⦇νxP ¡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 "Ψ  ⦇νxP α  ⦇νxP'"
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 "Ψ  ⦇νxP 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 "Ψ  ⦇νxP M⦇ν*(p  xvec)⦈⟨(p  N)  (p  (⦇νxP'))" by simp
    moreover from (p  xvec) ♯* P' x  xvec x  (p  xvec) have "(p  xvec) ♯* ⦇νxP'"
      by(simp add: abs_fresh_star)
    ultimately have "Ψ  ⦇νxP M⦇ν*xvec⦈⟨N  ⦇νxP'" 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 
    Ψ  ⦇νxP M⦇ν*xvec⦈⟨N  ⦇νxP'" 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 "Ψ  ⦇νxP ¡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 "Ψ  ⦇νxP ¡M⦇ν*(p  xvec)⦈⟨(p  N)  (p  (⦇νxP'))" by simp
    moreover from (p  xvec) ♯* P' x  xvec x  (p  xvec) have "(p  xvec) ♯* ⦇νxP'"
      by(simp add: abs_fresh_star)
    ultimately have "Ψ  ⦇νxP ¡M⦇ν*xvec⦈⟨N  ⦇νxP'" 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 
    Ψ  ⦇νxP ¡M⦇ν*xvec⦈⟨N  ⦇νxP'" 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 MN  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 Ψ ΨQ P M N P' AQ 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)]  (Ψ  ΨQ))  P ([(x, y)]  M)N  P'
  have "([(x, y)]  Ψ)  ([(x, y)]  ΨQ)  P ([(x, y)]  M)N  P'"
    by(simp add: eqvts)

  moreover from extractFrame Q = AQ, ΨQ have "([(x, y)]  (extractFrame Q)) = ([(x, y)]  AQ, ΨQ)"
    by simp
  with AQ ♯* x x  Q AQ ♯* y y  Q have "AQ, ([(x, y)]  ΨQ) = extractFrame Q"
    by(simp add: eqvts)
  moreover from AQ ♯* Ψ have "([(x, y)]  AQ) ♯* ([(x, y)]  Ψ)"
    by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
  with AQ ♯* x AQ ♯* y have "AQ ♯* ([(x, y)]  Ψ)" by simp
  moreover from AQ ♯* M have "([(x, y)]  AQ) ♯* ([(x, y)]  M)"
    by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
  with AQ ♯* x AQ ♯* y have "AQ ♯* ([(x, y)]  M)" by simp
  ultimately show ?case using AQ ♯* P AQ ♯* N
    by(force intro!: Par1)
next
  case(cPar2 Ψ ΨP Q M N Q' AP 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)]  (Ψ  ΨP))  Q ([(x, y)]  M)N  Q'
  have "([(x, y)]  Ψ)  ([(x, y)]  ΨP)  Q ([(x, y)]  M)N  Q'"
    by(simp add: eqvts)

  moreover from extractFrame P = AP, ΨP have "([(x, y)]  (extractFrame P)) = ([(x, y)]  AP, ΨP)"
    by simp
  with AP ♯* x x  P AP ♯* y y  P have "AP, ([(x, y)]  ΨP) = extractFrame P"
    by(simp add: eqvts)
  moreover from AP ♯* Ψ have "([(x, y)]  AP) ♯* ([(x, y)]  Ψ)"
    by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
  with AP ♯* x AP ♯* y have "AP ♯* ([(x, y)]  Ψ)" by simp
  moreover from AP ♯* M have "([(x, y)]  AP) ♯* ([(x, y)]  M)"
    by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
  with AP ♯* x AP ♯* y have "AP ♯* ([(x, y)]  M)" by simp
  ultimately show ?case using AP ♯* Q AP ♯* N
    by(force intro: Par2)
next
  case(cScope Ψ P M N P' z x y)
  from x  ⦇νzP z  x have "x  P" by(simp add: abs_fresh)
  from y  ⦇νzP 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 ¿MN  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 Ψ ΨQ P M N P' AQ 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)]  (Ψ  ΨQ))  P ¿([(x, y)]  M)N  P'
  have "([(x, y)]  Ψ)  ([(x, y)]  ΨQ)  P ¿([(x, y)]  M)N  P'"
    by(simp add: eqvts)

  moreover from extractFrame Q = AQ, ΨQ have "([(x, y)]  (extractFrame Q)) = ([(x, y)]  AQ, ΨQ)"
    by simp
  with AQ ♯* x x  Q AQ ♯* y y  Q have "AQ, ([(x, y)]  ΨQ) = extractFrame Q"
    by(simp add: eqvts)
  moreover from AQ ♯* Ψ have "([(x, y)]  AQ) ♯* ([(x, y)]  Ψ)"
    by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
  with AQ ♯* x AQ ♯* y have "AQ ♯* ([(x, y)]  Ψ)" by simp
  moreover from AQ ♯* M have "([(x, y)]  AQ) ♯* ([(x, y)]  M)"
    by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
  with AQ ♯* x AQ ♯* y have "AQ ♯* ([(x, y)]  M)" by simp
  ultimately show ?case using AQ ♯* P AQ ♯* N
    by(force intro!: Par1)
next
  case(cPar2 Ψ ΨP Q M N Q' AP 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)]  (Ψ  ΨP))  Q ¿([(x, y)]  M)N  Q'
  have "([(x, y)]  Ψ)  ([(x, y)]  ΨP)  Q ¿([(x, y)]  M)N  Q'"
    by(simp add: eqvts)

  moreover from extractFrame P = AP, ΨP have "([(x, y)]  (extractFrame P)) = ([(x, y)]  AP, ΨP)"
    by simp
  with AP ♯* x x  P AP ♯* y y  P have "AP, ([(x, y)]  ΨP) = extractFrame P"
    by(simp add: eqvts)
  moreover from AP ♯* Ψ have "([(x, y)]  AP) ♯* ([(x, y)]  Ψ)"
    by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
  with AP ♯* x AP ♯* y have "AP ♯* ([(x, y)]  Ψ)" by simp
  moreover from AP ♯* M have "([(x, y)]  AP) ♯* ([(x, y)]  M)"
    by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
  with AP ♯* x AP ♯* y have "AP ♯* ([(x, y)]  M)" by simp
  ultimately show ?case using AP ♯* Q AP ♯* N
    by(force intro!: Par2)
next
  case (cBrMerge Ψ ΨQ P M N P' AP ΨP Q Q' AQ 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 = AP, ΨP have "([(x, y)]  (extractFrame P)) = ([(x, y)]  AP, ΨP)"
    by simp
  with AP ♯* x x  P AP ♯* y y  P have "extractFrame P = AP, ([(x, y)]  ΨP)"
    by(simp add: eqvts)

  from extractFrame Q = AQ, ΨQ have "([(x, y)]  (extractFrame Q)) = ([(x, y)]  AQ, ΨQ)"
    by simp
  with AQ ♯* x x  Q AQ ♯* y y  Q have "extractFrame Q = AQ, ([(x, y)]  ΨQ)"
    by(simp add: eqvts)

  from x  P y  P x y. x  P; y  P  ([(x, y)]  (Ψ  ΨQ))  P ¿([(x, y)]  M)N  P'
  have "([(x, y)]  Ψ)  ([(x, y)]  Ψ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)]  (Ψ  ΨP))  Q ¿([(x, y)]  M)N  Q'
  have "([(x, y)]  Ψ)  ([(x, y)]  ΨP)  Q ¿([(x, y)]  M)N  Q'"
    by(simp add: eqvts)
  moreover note extractFrame P = AP, ([(x, y)]  ΨP) extractFrame Q = AQ, ([(x, y)]  ΨQ)

  moreover from AP ♯* Ψ have "([(x, y)]  AP) ♯* ([(x, y)]  Ψ)"
    by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
  with AP ♯* x AP ♯* y have "AP ♯* ([(x, y)]  Ψ)" by simp
  moreover from AP ♯* ΨQ have "([(x, y)]  AP) ♯* ([(x, y)]  ΨQ)"
    by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
  with AP ♯* x AP ♯* y have "AP ♯* ([(x, y)]  ΨQ)" by simp
  moreover from AP ♯* M have "([(x, y)]  AP) ♯* ([(x, y)]  M)"
    by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
  with AP ♯* x AP ♯* y have "AP ♯* ([(x, y)]  M)" by simp

  moreover from AQ ♯* Ψ have "([(x, y)]  AQ) ♯* ([(x, y)]  Ψ)"
    by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
  with AQ ♯* x AQ ♯* y have "AQ ♯* ([(x, y)]  Ψ)" by simp
  moreover from AQ ♯* ΨP have "([(x, y)]  AQ) ♯* ([(x, y)]  ΨP)"
    by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
  with AQ ♯* x AQ ♯* y have "AQ ♯* ([(x, y)]  ΨP)" by simp
  moreover from AQ ♯* M have "([(x, y)]  AQ) ♯* ([(x, y)]  M)"
    by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
  with AQ ♯* x AQ ♯* y have "AQ ♯* ([(x, y)]  M)" by simp

  moreover note distinct AP distinct AQ
    AP ♯* P AP ♯* N AP ♯* P' AP ♯* Q AP ♯* Q' AP ♯* AQ
    AQ ♯* P AQ ♯* N AQ ♯* P' AQ ♯* Q AQ ♯* Q'
  ultimately show ?case
    by(force intro!: semantics.cBrMerge)
next
  case(cScope Ψ P M N P' z x y)
  from x  ⦇νzP z  x have "x  P" by(simp add: abs_fresh)
  from y  ⦇νzP 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 MN  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 MN  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 ¿MN  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 ¿MN  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 MN  P'"
  and   "x  P"
  and   "y  P"
  and   "x  Ψ"
  and   "y  Ψ"

shows "Ψ  P  ([(x, y)]  M)N  P'"
proof -
  from Ψ  P MN  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 ¿MN  P'"
  and   "x  P"
  and   "y  P"
  and   "x  Ψ"
  and   "y  Ψ"

shows "Ψ  P  ¿([(x, y)]  M)N  P'"
proof -
  from Ψ  P ¿MN  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 MN  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 MN  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 ¿MN  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 ¿MN  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 MN  P'"
  and   "x  P"
  and   "y  P"
  and   "x  M"
  and   "y  M"

shows "([(x, y)]  Ψ)  P  MN  P'"
proof -
  from Ψ  P MN  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 ¿MN  P'"
  and   "x  P"
  and   "y  P"
  and   "x  M"
  and   "y  M"

shows "([(x, y)]  Ψ)  P  ¿MN  P'"
proof -
  from Ψ  P ¿MN  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 MN  P'"
  and   S: "set p  Xs × Ys"
  and   "Xs ♯* P"
  and   "Ys ♯* P"
  and   "Xs ♯* M"
  and   "Ys ♯* M"

shows "(p  Ψ)  P  MN  P'"
proof -
  from Ψ  P MN  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 ¿MN  P'"
  and   S: "set p  Xs × Ys"
  and   "Xs ♯* P"
  and   "Ys ♯* P"
  and   "Xs ♯* M"
  and   "Ys ♯* M"

shows "(p  Ψ)  P  ¿MN  P'"
proof -
  from Ψ  P ¿MN  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 MN  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 MN  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 ¿MN  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 ¿MN  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 AF :: "name list"
    and ΨF :: 'b

assumes "x  AF"
  and   "x  AF, ΨF"

shows "x  Ψ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  MN⟩.P have "x  M" by simp
  from y  MN⟩.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 Ψ ΨQ P M xvec N P' AQ 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)]  (Ψ  ΨQ))  P ([(x, y)]  M)⦇ν*xvec⦈⟨N  P'
  have "([(x, y)]  Ψ)  ([(x, y)]  ΨQ)  P ([(x, y)]  M)⦇ν*xvec⦈⟨N  P'"
    by(simp add: eqvts)

  moreover from extractFrame Q = AQ, ΨQ have "([(x, y)]  AQ, ΨQ) = ([(x, y)]  (extractFrame Q))"
    by simp
  with AQ ♯* x x  Q AQ ♯* y y  Q have "AQ, ([(x, y)]  ΨQ) = extractFrame Q"
    by(simp add: eqvts)
  moreover from AQ ♯* Ψ have "([(x, y)]  AQ) ♯* ([(x, y)]  Ψ)"
    by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
  with AQ ♯* x AQ ♯* y have "AQ ♯* ([(x, y)]  Ψ)" by simp
  moreover from AQ ♯* M have "([(x, y)]  AQ) ♯* ([(x, y)]  M)"
    by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
  with AQ ♯* x AQ ♯* y have "AQ ♯* ([(x, y)]  M)" by simp
  ultimately show ?case using AQ ♯* P AQ ♯* N xvec ♯* Q AQ ♯* xvec
    by(force intro: Par1)
next
  case(cPar2 Ψ ΨP Q M xvec N Q' AP 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)]  (Ψ  ΨP))  Q ([(x, y)]  M)⦇ν*xvec⦈⟨N  Q'
  have "([(x, y)]  Ψ)  ([(x, y)]  ΨP)  Q ([(x, y)]  M)⦇ν*xvec⦈⟨N  Q'"
    by(simp add: eqvts)

  moreover from extractFrame P = AP, ΨP have "([(x, y)]  AP, ΨP) = ([(x, y)]  (extractFrame P))"
    by simp
  with AP ♯* x x  P AP ♯* y y  P have "AP, ([(x, y)]  ΨP) = extractFrame P"
    by(simp add: eqvts)
  moreover from AP ♯* Ψ have "([(x, y)]  AP) ♯* ([(x, y)]  Ψ)"
    by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
  with AP ♯* x AP ♯* y have "AP ♯* ([(x, y)]  Ψ)" by simp
  moreover from AP ♯* M have "([(x, y)]  AP) ♯* ([(x, y)]  M)"
    by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
  with AP ♯* x AP ♯* y have "AP ♯* ([(x, y)]  M)" by simp
  ultimately show ?case using AP ♯* Q AP ♯* N xvec ♯* P AP ♯* xvec
    by(force intro: Par2)
next
  case(cOpen Ψ P M xvec yvec N P' z x y)
  from x  ⦇νzP z  x have "x  P" by(simp add: abs_fresh)
  from y  ⦇νzP 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  ⦇νzP z  x have "x  P" by(simp add: abs_fresh)
  from y  ⦇νzP 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  MN⟩.P have "x  M" by simp
  from y  MN⟩.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 Ψ ΨQ P M xvec N P' AQ 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)]  (Ψ  ΨQ))  P ¡([(x, y)]  M)⦇ν*xvec⦈⟨N  P'
  have "([(x, y)]  Ψ)  ([(x, y)]  ΨQ)  P ¡([(x, y)]  M)⦇ν*xvec⦈⟨N  P'"
    by(simp add: eqvts)

  moreover from extractFrame Q = AQ, ΨQ have "([(x, y)]  AQ, ΨQ) = ([(x, y)]  (extractFrame Q))"
    by simp
  with AQ ♯* x x  Q AQ ♯* y y  Q have "AQ, ([(x, y)]  ΨQ) = extractFrame Q"
    by(simp add: eqvts)
  moreover from AQ ♯* Ψ have "([(x, y)]  AQ) ♯* ([(x, y)]  Ψ)"
    by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
  with AQ ♯* x AQ ♯* y have "AQ ♯* ([(x, y)]  Ψ)" by simp
  moreover from AQ ♯* M have "([(x, y)]  AQ) ♯* ([(x, y)]  M)"
    by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
  with AQ ♯* x AQ ♯* y have "AQ ♯* ([(x, y)]  M)" by simp
  ultimately show ?case using AQ ♯* P AQ ♯* N xvec ♯* Q AQ ♯* xvec
    by(force intro: Par1)
next
  case(cPar2 Ψ ΨP Q M xvec N Q' AP 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)]  (Ψ  ΨP))  Q ¡([(x, y)]  M)⦇ν*xvec⦈⟨N  Q'
  have "([(x, y)]  Ψ)  ([(x, y)]  ΨP)  Q ¡([(x, y)]  M)⦇ν*xvec⦈⟨N  Q'"
    by(simp add: eqvts)

  moreover from extractFrame P = AP, ΨP have "([(x, y)]  AP, ΨP) = ([(x, y)]  (extractFrame P))"
    by simp
  with AP ♯* x x  P AP ♯* y y  P have "AP, ([(x, y)]  ΨP) = extractFrame P"
    by(simp add: eqvts)
  moreover from AP ♯* Ψ have "([(x, y)]  AP) ♯* ([(x, y)]  Ψ)"
    by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
  with AP ♯* x AP ♯* y have "AP ♯* ([(x, y)]  Ψ)" by simp
  moreover from AP ♯* M have "([(x, y)]  AP) ♯* ([(x, y)]  M)"
    by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
  with AP ♯* x AP ♯* y have "AP ♯* ([(x, y)]  M)" by simp
  ultimately show ?case using AP ♯* Q AP ♯* N xvec ♯* P AP ♯* xvec
    by(force intro: Par2)
next
  case(cBrComm1 Ψ ΨQ P M N P' AP ΨP Q xvec Q' AQ 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 Ψ  ΨQ  P  ¿MN  P' x  P y  P
  have "([(x, y)]  (Ψ  ΨQ))  P ¿([(x, y)]  M)N  P'"
    by(rule brinputSwapFrameSubject)
  then have permIn: "(([(x, y)]  Ψ)  ([(x, y)]  Ψ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)]  (Ψ  ΨP))  Q ¡([(x, y)]  M)⦇ν*xvec⦈⟨N  Q'
  have permOut: "([(x, y)]  Ψ)  ([(x, y)]  ΨP)  Q ¡([(x, y)]  M)⦇ν*xvec⦈⟨N  Q'"
    by(simp add: eqvts)

  moreover from extractFrame P = AP, ΨP have "([(x, y)]  AP, ΨP) = ([(x, y)]  (extractFrame P))"
    by simp
  with AP ♯* x x  P AP ♯* y y  P have "extractFrame P = AP, ([(x, y)]  ΨP)"
    by(simp add: eqvts)
  moreover from extractFrame Q = AQ, ΨQ have "([(x, y)]  AQ, ΨQ) = ([(x, y)]  (extractFrame Q))"
    by simp
  with AQ ♯* x x  Q AQ ♯* y y  Q have "extractFrame Q = AQ, ([(x, y)]  ΨQ)"
    by(simp add: eqvts)
  moreover from AP ♯* Ψ have "([(x, y)]  AP) ♯* ([(x, y)]  Ψ)"
    by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
  with AP ♯* x AP ♯* y have "AP ♯* ([(x, y)]  Ψ)" by simp
  moreover from AQ ♯* Ψ have "([(x, y)]  AQ) ♯* ([(x, y)]  Ψ)"
    by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
  with AQ ♯* x AQ ♯* y have "AQ ♯* ([(x, y)]  Ψ)" by simp
  moreover from AP ♯* ΨQ have "([(x, y)]  AP) ♯* ([(x, y)]  ΨQ)"
    by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
  with AP ♯* x AP ♯* y have "AP ♯* ([(x, y)]  ΨQ)" by simp
  moreover from AQ ♯* ΨP have "([(x, y)]  AQ) ♯* ([(x, y)]  ΨP)"
    by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
  with AQ ♯* x AQ ♯* y have "AQ ♯* ([(x, y)]  ΨP)" by simp
  moreover from AP ♯* M have "([(x, y)]  AP) ♯* ([(x, y)]  M)"
    by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
  with AP ♯* x AP ♯* y have "AP ♯* ([(x, y)]  M)" by simp
  moreover from AQ ♯* M have "([(x, y)]  AQ) ♯* ([(x, y)]  M)"
    by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
  with AQ ♯* x AQ ♯* y have "AQ ♯* ([(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 ♯* Ψ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 xvec ♯* y have "xvec ♯* ([(x, y)]  ΨQ)" by simp
  moreover from xvec ♯* ΨP have "([(x, y)]  xvec) ♯* ([(x, y)]  Ψ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)]  Ψ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 AP distinct AQ AP ♯* P AP ♯* N AP ♯* P' AP ♯* Q AP ♯* Q' AP ♯* AQ
    AP ♯* xvec AQ ♯* P AQ ♯* N AQ ♯* P' AQ ♯* Q AQ ♯* Q' AQ ♯* xvec
    distinct xvec xvec ♯* P xvec ♯* Q
  ultimately show ?case
    by(simp add: semantics.cBrComm1)
next
  case(cBrComm2 Ψ ΨQ P M xvec N P' AP ΨP Q Q' AQ 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 Ψ  ΨP  Q  ¿MN  Q' x  Q y  Q
  have "([(x, y)]  (Ψ  ΨP))  Q ¿([(x, y)]  M)N  Q'"
    by(rule brinputSwapFrameSubject)
  then have permIn: "(([(x, y)]  Ψ)  ([(x, y)]  Ψ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)]  (Ψ  ΨQ))  P ¡([(x, y)]  M)⦇ν*xvec⦈⟨N  P'
  have permOut: "([(x, y)]  Ψ)  ([(x, y)]  ΨQ)  P ¡([(x, y)]  M)⦇ν*xvec⦈⟨N  P'"
    by(simp add: eqvts)

  moreover from extractFrame P = AP, ΨP have "([(x, y)]  AP, ΨP) = ([(x, y)]  (extractFrame P))"
    by simp
  with AP ♯* x x  P AP ♯* y y  P have "extractFrame P = AP, ([(x, y)]  ΨP)"
    by(simp add: eqvts)
  moreover from extractFrame Q = AQ, ΨQ have "([(x, y)]  AQ, ΨQ) = ([(x, y)]  (extractFrame Q))"
    by simp
  with AQ ♯* x x  Q AQ ♯* y y  Q have "extractFrame Q = AQ, ([(x, y)]  ΨQ)"
    by(simp add: eqvts)
  moreover from AP ♯* Ψ have "([(x, y)]  AP) ♯* ([(x, y)]  Ψ)"
    by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
  with AP ♯* x AP ♯* y have "AP ♯* ([(x, y)]  Ψ)" by simp
  moreover from AQ ♯* Ψ have "([(x, y)]  AQ) ♯* ([(x, y)]  Ψ)"
    by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
  with AQ ♯* x AQ ♯* y have "AQ ♯* ([(x, y)]  Ψ)" by simp
  moreover from AP ♯* ΨQ have "([(x, y)]  AP) ♯* ([(x, y)]  ΨQ)"
    by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
  with AP ♯* x AP ♯* y have "AP ♯* ([(x, y)]  ΨQ)" by simp
  moreover from AQ ♯* ΨP have "([(x, y)]  AQ) ♯* ([(x, y)]  ΨP)"
    by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
  with AQ ♯* x AQ ♯* y have "AQ ♯* ([(x, y)]  ΨP)" by simp
  moreover from AP ♯* M have "([(x, y)]  AP) ♯* ([(x, y)]  M)"
    by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
  with AP ♯* x AP ♯* y have "AP ♯* ([(x, y)]  M)" by simp
  moreover from AQ ♯* M have "([(x, y)]  AQ) ♯* ([(x, y)]  M)"
    by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
  with AQ ♯* x AQ ♯* y have "AQ ♯* ([(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 ♯* Ψ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 xvec ♯* y have "xvec ♯* ([(x, y)]  ΨQ)" by simp
  moreover from xvec ♯* ΨP have "([(x, y)]  xvec) ♯* ([(x, y)]  Ψ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)]  Ψ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 AP distinct AQ AP ♯* P AP ♯* N AP ♯* P' AP ♯* Q AP ♯* Q' AP ♯* AQ
    AP ♯* xvec AQ ♯* P AQ ♯* N AQ ♯* P' AQ ♯* Q AQ ♯* Q' AQ ♯* 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  ⦇νzP z  x have "x  P" by(simp add: abs_fresh)
  from y  ⦇νzP 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  ⦇νzP z  x have "x  P" by(simp add: abs_fresh)
  from y  ⦇νzP 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 ΨQ  :: 'b
    and P    :: "('a, 'b, 'c) psi"
    and M    :: 'a
    and N    :: 'a
    and P'   :: "('a, 'b, 'c) psi"
    and AP   :: "name list"
    and ΨP  :: 'b
    and Q    :: "('a, 'b, 'c) psi"
    and K    :: 'a
    and xvec :: "name list"
    and Q'   :: "('a, 'b, 'c) psi"
    and AQ   :: "name list"

assumes "Ψ  ΨQ  P MN  P'"
  and   "extractFrame P = AP, ΨP"
  and   "Ψ  ΨP  Q K⦇ν*xvec⦈⟨N  Q'"
  and   "extractFrame Q = AQ, ΨQ"
  and   "Ψ  ΨP  ΨQ  M  K"
  and   "AP ♯* Ψ"
  and   "AP ♯* P"
  and   "AP ♯* Q"
  and   "AP ♯* M"
  and   "AP ♯* AQ"
  and   "AQ ♯* Ψ"
  and   "AQ ♯* P"
  and   "AQ ♯* Q"
  and   "AQ ♯* K"
  and   "xvec ♯* P"

shows "Ψ  P  Q τ  ⦇ν*xvec(P'  Q')"
proof -
  {
    fix Ψ    :: 'b
      and ΨQ  :: 'b
      and P    :: "('a, 'b, 'c) psi"
      and M    :: 'a
      and N    :: 'a
      and P'   :: "('a, 'b, 'c) psi"
      and AP   :: "name list"
      and ΨP  :: 'b
      and Q    :: "('a, 'b, 'c) psi"
      and K    :: 'a
      and xvec :: "name list"
      and Q'   :: "('a, 'b, 'c) psi"
      and AQ   :: "name list"

    assume "Ψ  ΨQ  P MN  P'"
      and  "extractFrame P = AP, ΨP"
      and  "distinct AP"
      and  "Ψ  ΨP  Q K⦇ν*xvec⦈⟨N  Q'"
      and  "extractFrame Q = AQ, ΨQ"
      and  "distinct AQ"
      and  "Ψ  ΨP  ΨQ  M  K"
      and  "AP ♯* Ψ"
      and  "AP ♯* P"
      and  "AP ♯* Q"
      and  "AP ♯* M"
      and  "AP ♯* AQ"
      and  "AQ ♯* Ψ"
      and  "AQ ♯* P"
      and  "AQ ♯* Q"
      and  "AQ ♯* 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) ♯* AP" and "(r  xvec) ♯* AQ"
        and "(r  xvec) ♯* P'" and "(r  xvec) ♯* Q'" and "(r  xvec) ♯* ΨP" and "(r  xvec) ♯* Ψ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, AP, AQ, ΨP, ΨQ, P', Q')"])
          (auto simp add: eqvts fresh_star_prod)
      obtain q::"name prm" where "(q  AQ) ♯* Ψ" and "(q  AQ) ♯* P" and "(q  AQ) ♯* Q" and "(q  AQ) ♯* K"
        and "(q  AQ) ♯* (r  N)" and "(q  AQ) ♯* (r  xvec)" and "(q  AQ) ♯* (r  Q')"
        and "(q  AQ) ♯* (r  P')" and "(q  AQ) ♯* ΨP" and "(q  AQ) ♯* AP" and "(q  AQ) ♯* ΨQ"
        and Sq: "set q  set AQ × set(q  AQ)"
        by(rule name_list_avoiding[where xvec=AQ and c="(Ψ, P, Q, K, r  N, r  xvec, ΨQ, AP, ΨP, r  Q', r  P')"])
          (auto simp add: eqvts fresh_star_prod)
      obtain p::"name prm" where "(p  AP) ♯* Ψ" and "(p  AP) ♯* P" and "(p  AP) ♯* Q" and "(p  AP) ♯* M"
        and "(p  AP) ♯* (r  N)" and "(p  AP) ♯* (r  xvec)" and "(p  AP) ♯* (r  Q')"
        and "(p  AP) ♯* (r  P')" and "(p  AP) ♯* ΨP" and "(p  AP) ♯* ΨQ" and "(p  AP) ♯* AQ"
        and "(p  AP) ♯* (q  AQ)" and Sp: "(set p)  (set AP) × (set(p  AP))"
        by(rule name_list_avoiding[where xvec=AP and c="(Ψ, P, Q, M, r  N, r  xvec, AQ, q  AQ, ΨQ, ΨP, r  Q', r  P')"])
          (auto simp add: eqvts fresh_star_prod)
      have FrP: "extractFrame P = AP, ΨP" by fact
      have FrQ: "extractFrame Q = AQ, ΨQ" by fact

      from AP ♯* Q FrQ AP ♯* AQ have "AP ♯* ΨQ"
        by(force dest: extractFrameFreshChain)
      from AQ ♯* P FrP AP ♯* AQ have "AQ ♯* ΨP"
        by(force dest: extractFrameFreshChain)
      from (r  xvec) ♯* AP (p  AP) ♯* (r  xvec) (r  xvec) ♯* AP Sp have "(r  xvec) ♯* (p  AP)"
        by(simp add: freshChainSimps)

      from Ψ  ΨQ  P MN  P' Sr distinctPerm r xvec ♯* P (r  xvec) ♯* P
      have "Ψ  ΨQ  P M(r  N)  (r  P')"
        by(rule inputAlpha)
      then have "(q  (Ψ  ΨQ))  P (q  M)(r  N)  (r  P')" using Sq AQ ♯* P (q  AQ) ♯* P
        by - (rule inputPermFrameSubject, (assumption | simp)+)
      then have PTrans: "Ψ  (q  ΨQ)  P (q  M)(r  N)  (r  P')" using Sq AQ ♯* Ψ (q  AQ) ♯* Ψ
        by(simp add: eqvts)

      moreover from extractFrame P = AP, ΨP  Sp (p  AP) ♯* ΨP
      have FrP: "extractFrame P = (p  AP), (p  ΨP)"
        by(simp add: frameChainAlpha)
      moreover from distinct AP have "distinct(p  AP)"  by simp

      moreover from Ψ  ΨP  Q K⦇ν*xvec⦈⟨N  Q' Sr (r  xvec) ♯* N (r  xvec) ♯* Q'
      have "Ψ  ΨP  Q K⦇ν*(r  xvec)⦈⟨(r  N)  (r  Q')"
        by(simp add: boundOutputChainAlpha'' create_residual.simps)
      then have "(p  (Ψ  ΨP))  Q (p  K)⦇ν*(r  xvec)⦈⟨(r  N)  (r  Q')" using Sp AP ♯* Q (p  AP) ♯* Q (r  xvec) ♯* K (r  xvec) ♯* AP (r  xvec) ♯* (p  AP)
        by(fastforce intro: outputPermFrameSubject)
      then have QTrans: "Ψ  (p  ΨP)  Q (p  K)⦇ν*(r  xvec)⦈⟨(r  N)  (r  Q')" using Sp AP ♯* Ψ (p  AP) ♯* Ψ
        by(simp add: eqvts)
      moreover then have "distinct(r  xvec)" by(force dest: boundOutputDistinct)
      moreover from extractFrame Q = AQ, ΨQ  Sq (q  AQ) ♯* ΨQ
      have FrQ: "extractFrame Q = (q  AQ), (q  ΨQ)"
        by(simp add: frameChainAlpha)
      moreover from distinct AQ have "distinct(q  AQ)"  by simp

      moreover from Ψ  ΨP  ΨQ  M  K have "(p  q  (Ψ  ΨP  ΨQ))  (p  q  M)  (p  q  K)"
        by(metis chanEqClosed)
      with AP ♯* Ψ (p  AP) ♯* Ψ AQ ♯* Ψ (q  AQ) ♯* Ψ AQ ♯* ΨP (q  AQ) ♯* ΨP
        AP ♯* ΨQ (p  AP) ♯* ΨQ AP ♯* M (p  AP) ♯* M (q  AQ) ♯* AP (p  AP) ♯* (q  AQ)
        AQ ♯* K (q  AQ) ♯* K AP ♯* AQ (p  AP) ♯* AQ  Sp Sq
      have "Ψ  (p  ΨP)  (q  ΨQ)  (q  M)  (p  K)" by(simp add: eqvts freshChainSimps)
      moreover note (p  AP) ♯* Ψ
      moreover from (p  AP) ♯* AQ (p  AP) ♯* (q  AQ) (p  AP) ♯* ΨQ Sq have "(p  AP) ♯* (q  ΨQ)"
        by(simp add: freshChainSimps)
      moreover note (p  AP) ♯* P
      moreover from (p  AP) ♯* AQ (p  AP) ♯* (q  AQ) (p  AP) ♯* M Sq have "(p  AP) ♯* (q  M)"
        by(simp add: freshChainSimps)
      moreover note  (p  AP) ♯* (r  N) (p  AP) ♯* (r  P') (p  AP) ♯* Q (p  AP) ♯* (r  Q') (p  AP) ♯* (q  AQ)
        (p  AP) ♯* (r  xvec) (q  AQ) ♯* Ψ
      moreover from (q  AQ) ♯* AP (p  AP) ♯* AQ (q  AQ) ♯* ΨP (p  AP) ♯* (q  AQ) Sp Sq have "(q  AQ) ♯* (p  ΨP)"
        by(simp add: freshChainSimps)
      moreover note (q  AQ) ♯* P (q  AQ) ♯* (r  N)(q  AQ) ♯* (r  P') (q  AQ) ♯* Q
      moreover from (q  AQ) ♯* AP (p  AP) ♯* AQ (q  AQ) ♯* K (p  AP) ♯* (q  AQ) Sp Sq have "(q  AQ) ♯* (p  K)"
        by(simp add: freshChainSimps)
      moreover note  (q  AQ) ♯* (r  Q') (q  AQ) ♯* (r  xvec) (r  xvec) ♯* Ψ
      moreover from (r  xvec) ♯* AP (p  AP) ♯* (r  xvec) (r  xvec) ♯* ΨP Sp have "(r  xvec) ♯* (p  ΨP)"
        by(simp add: freshChainSimps)
      moreover from (r  xvec) ♯* AQ (q  AQ) ♯* (r  xvec) (r  xvec) ♯* ΨQ Sq have "(r  xvec) ♯* (q  ΨQ)"
        by(simp add: freshChainSimps)
      moreover note (r  xvec) ♯* P
      moreover from (r  xvec) ♯* AQ (q  AQ) ♯* (r  xvec) (r  xvec) ♯* M Sq have "(r  xvec) ♯* (q  M)"
        by(simp add: freshChainSimps)
      moreover note (r  xvec) ♯* Q
      moreover from (r  xvec) ♯* AP (p  AP) ♯* (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 Ψ  ΨQ  P MN  P' Ψ  ΨP  Q K⦇ν*xvec⦈⟨N  Q' Ψ  ΨP  ΨQ  M  K
  moreover from extractFrame P = AP, ΨP AP ♯* Ψ AP ♯* P AP ♯* Q AP ♯* M AP ♯* AQ
  obtain AP' where "extractFrame P = AP', ΨP" and "distinct AP'" and "AP' ♯* Ψ" and "AP' ♯* P" and "AP' ♯* Q" and "AP' ♯* M" and "AP' ♯* AQ"
    by - (rule distinctFrame[where C="(Ψ, P, Q, M, AQ)"], auto)
  moreover from extractFrame Q = AQ, ΨQ AQ ♯* Ψ AQ ♯* P AQ ♯* Q AQ ♯* K AP' ♯* AQ
  obtain AQ' where "extractFrame Q = AQ', ΨQ" and "distinct AQ'" and "AQ' ♯* Ψ" and "AQ' ♯* P" and "AQ' ♯* Q" and "AQ' ♯* K" and "AP' ♯* AQ'"
    by - (rule distinctFrame[where C="(Ψ, P, Q, K, AP')"], auto)
  ultimately show ?thesis using xvec ♯* P
    by(metis Goal)
qed

lemma Comm2:
  fixes Ψ    :: 'b
    and Ψ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 AP   :: "name list"
    and ΨP  :: 'b
    and Q    :: "('a, 'b, 'c) psi"
    and K    :: 'a
    and Q'   :: "('a, 'b, 'c) psi"
    and AQ   :: "name list"

assumes "Ψ  ΨQ  P M⦇ν*xvec⦈⟨N  P'"
  and   "extractFrame P = AP, ΨP"
  and   "Ψ  ΨP  Q KN  Q'"
  and   "extractFrame Q = AQ, ΨQ"
  and   "Ψ  ΨP  ΨQ  M  K"
  and   "AP ♯* Ψ"
  and   "AP ♯* P"
  and   "AP ♯* Q"
  and   "AP ♯* M"
  and   "AP ♯* AQ"
  and   "AQ ♯* Ψ"
  and   "AQ ♯* P"
  and   "AQ ♯* Q"
  and   "AQ ♯* K"
  and   "xvec ♯* Q"

shows "Ψ  P  Q τ  ⦇ν*xvec(P'  Q')"
proof -
  {
    fix Ψ    :: 'b
      and Ψ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 AP   :: "name list"
      and ΨP  :: 'b
      and Q    :: "('a, 'b, 'c) psi"
      and K    :: 'a
      and Q'   :: "('a, 'b, 'c) psi"
      and AQ   :: "name list"

    assume "Ψ  ΨQ  P M⦇ν*xvec⦈⟨N  P'"
      and  "extractFrame P = AP, ΨP"
      and  "distinct AP"
      and  "Ψ  ΨP  Q KN  Q'"
      and  "extractFrame Q = AQ, ΨQ"
      and  "distinct AQ"
      and  "Ψ  ΨP  ΨQ  M  K"
      and  "AP ♯* Ψ"
      and  "AP ♯* P"
      and  "AP ♯* Q"
      and  "AP ♯* M"
      and  "AP ♯* AQ"
      and  "AQ ♯* Ψ"
      and  "AQ ♯* P"
      and  "AQ ♯* Q"
      and  "AQ ♯* 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) ♯* AP" and "(r  xvec) ♯* AQ"
        and "(r  xvec) ♯* P'" and "(r  xvec) ♯* Q'" and "(r  xvec) ♯* ΨP" and "(r  xvec) ♯* Ψ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, AP, AQ, ΨP, ΨQ, P', Q')"])
          (auto simp add: eqvts fresh_star_prod)
      obtain q::"name prm" where "(q  AQ) ♯* Ψ" and "(q  AQ) ♯* P" and "(q  AQ) ♯* Q" and "(q  AQ) ♯* K"
        and "(q  AQ) ♯* (r  N)" and "(q  AQ) ♯* (r  xvec)" and "(q  AQ) ♯* (r  Q')"
        and "(q  AQ) ♯* (r  P')" and "(q  AQ) ♯* ΨP" and "(q  AQ) ♯* AP" and "(q  AQ) ♯* ΨQ"
        and Sq: "set q  set AQ × set(q  AQ)"
        by(rule name_list_avoiding[where xvec=AQ and c="(Ψ, P, Q, K, r  N, r  xvec, ΨQ, AP, ΨP, r  Q', r  P')"])
          (auto simp add: eqvts fresh_star_prod)
      obtain p::"name prm" where "(p  AP) ♯* Ψ" and "(p  AP) ♯* P" and "(p  AP) ♯* Q" and "(p  AP) ♯* M"
        and "(p  AP) ♯* (r  N)" and "(p  AP) ♯* (r  xvec)" and "(p  AP) ♯* (r  Q')"
        and "(p  AP) ♯* (r  P')" and "(p  AP) ♯* ΨP" and "(p  AP) ♯* ΨQ" and "(p  AP) ♯* AQ"
        and "(p  AP) ♯* (q  AQ)" and Sp: "(set p)  (set AP) × (set(p  AP))"
        by(rule name_list_avoiding[where xvec=AP and c="(Ψ, P, Q, M, r  N, r  xvec, AQ, q  AQ, ΨQ, ΨP, r  Q', r  P')"])
          (auto simp add: eqvts fresh_star_prod)

      have FrP: "extractFrame P = AP, ΨP" by fact
      have FrQ: "extractFrame Q = AQ, ΨQ" by fact

      from AP ♯* Q FrQ AP ♯* AQ have "AP ♯* ΨQ"
        by(auto dest: extractFrameFreshChain)
      from AQ ♯* P FrP AP ♯* AQ have "AQ ♯* ΨP"
        by(auto dest: extractFrameFreshChain)

      from (r  xvec) ♯* AQ (q  AQ) ♯* (r  xvec) (r  xvec) ♯* AQ Sq have "(r  xvec) ♯* (q  AQ)"
        by(simp add: freshChainSimps)

      from Ψ  ΨQ  P M⦇ν*xvec⦈⟨N  P' Sr (r  xvec) ♯* N (r  xvec) ♯* P'
      have "Ψ  ΨQ  P M⦇ν*(r  xvec)⦈⟨(r  N)  (r  P')"
        by(simp add: boundOutputChainAlpha'' create_residual.simps)
      then have "(q  (Ψ  ΨQ))  P (q  M)⦇ν*(r  xvec)⦈⟨(r  N)  (r  P')" using Sq AQ ♯* P (q  AQ) ♯* P (r  xvec) ♯* M (r  xvec) ♯* AQ (r  xvec) ♯* (q  AQ)
        by(fastforce intro: outputPermFrameSubject)
      then have PTrans: "Ψ  (q  ΨQ)  P (q  M)⦇ν*(r  xvec)⦈⟨(r  N)  (r  P')" using Sq AQ ♯* Ψ (q  AQ) ♯* Ψ
        by(simp add: eqvts)
      moreover then have "distinct(r  xvec)" by(force dest: boundOutputDistinct)

      moreover from extractFrame P = AP, ΨP  Sp (p  AP) ♯* ΨP
      have FrP: "extractFrame P = (p  AP), (p  ΨP)"
        by(simp add: frameChainAlpha)
      moreover from distinct AP have "distinct(p  AP)"  by simp

      moreover from Ψ  ΨP  Q KN  Q' Sr distinctPerm r xvec ♯* Q (r  xvec) ♯* Q
      have "Ψ  ΨP  Q K(r  N)  (r  Q')"
        by(rule inputAlpha)
      then have "(p  (Ψ  ΨP))  Q (p  K)(r  N)  (r  Q')" using Sp AP ♯* Q (p  AP) ♯* Q
        by - (rule inputPermFrameSubject, (assumption | simp)+)
      then have QTrans: "Ψ  (p  ΨP)  Q (p  K)(r  N)  (r  Q')" using Sp AP ♯* Ψ (p  AP) ♯* Ψ
        by(simp add: eqvts)

      moreover from extractFrame Q = AQ, ΨQ  Sq (q  AQ) ♯* ΨQ
      have FrQ: "extractFrame Q = (q  AQ), (q  ΨQ)"
        by(simp add: frameChainAlpha)
      moreover from distinct AQ have "distinct(q  AQ)"  by simp

      moreover from Ψ  ΨP  ΨQ  M  K have "(p  q  (Ψ  ΨP  ΨQ))  (p  q  M)  (p  q  K)"
        by(metis chanEqClosed)
      with AP ♯* Ψ (p  AP) ♯* Ψ AQ ♯* Ψ (q  AQ) ♯* Ψ AQ ♯* ΨP (q  AQ) ♯* ΨP
        AP ♯* ΨQ (p  AP) ♯* ΨQ AP ♯* M (p  AP) ♯* M (q  AQ) ♯* AP (p  AP) ♯* (q  AQ)
        AQ ♯* K (q  AQ) ♯* K AP ♯* AQ (p  AP) ♯* AQ  Sp Sq
      have "Ψ  (p  ΨP)  (q  ΨQ)  (q  M)  (p  K)"
        by(simp add: eqvts freshChainSimps)
      moreover note (p  AP) ♯* Ψ
      moreover from (p  AP) ♯* AQ (p  AP) ♯* (q  AQ) (p  AP) ♯* ΨQ Sq have "(p  AP) ♯* (q  ΨQ)"
        by(simp add: freshChainSimps)
      moreover note (p  AP) ♯* P
      moreover from (p  AP) ♯* AQ (p  AP) ♯* (q  AQ) (p  AP) ♯* M Sq have "(p  AP) ♯* (q  M)"
        by(simp add: freshChainSimps)
      moreover note  (p  AP) ♯* (r  N) (p  AP) ♯* (r  P') (p  AP) ♯* Q (p  AP) ♯* (r  Q') (p  AP) ♯* (q  AQ)
        (p  AP) ♯* (r  xvec) (q  AQ) ♯* Ψ
      moreover from (q  AQ) ♯* AP (p  AP) ♯* AQ (q  AQ) ♯* ΨP (p  AP) ♯* (q  AQ) Sp Sq have "(q  AQ) ♯* (p  ΨP)"
        by(simp add: freshChainSimps)
      moreover note (q  AQ) ♯* P (q  AQ) ♯* (r  N)(q  AQ) ♯* (r  P') (q  AQ) ♯* Q
      moreover from (q  AQ) ♯* AP (p  AP) ♯* AQ (q  AQ) ♯* K (p  AP) ♯* (q  AQ) Sp Sq have "(q  AQ) ♯* (p  K)"
        by(simp add: freshChainSimps)
      moreover note  (q  AQ) ♯* (r  Q') (q  AQ) ♯* (r  xvec) (r  xvec) ♯* Ψ
      moreover from (r  xvec) ♯* AP (p  AP) ♯* (r  xvec) (r  xvec) ♯* ΨP Sp have "(r  xvec) ♯* (p  ΨP)"
        by(simp add: freshChainSimps)
      moreover from (r  xvec) ♯* AQ (q  AQ) ♯* (r  xvec) (r  xvec) ♯* ΨQ Sq have "(r  xvec) ♯* (q  ΨQ)"
        by(simp add: freshChainSimps)
      moreover note (r  xvec) ♯* P
      moreover from (r  xvec) ♯* AQ (q  AQ) ♯* (r  xvec) (r  xvec) ♯* M Sq have "(r  xvec) ♯* (q  M)"
        by(simp add: freshChainSimps)
      moreover note (r  xvec) ♯* Q
      moreover from (r  xvec) ♯* AP (p  AP) ♯* (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 Ψ  ΨQ  P M⦇ν*xvec⦈⟨N  P' Ψ  ΨP  Q KN  Q' Ψ  ΨP  ΨQ  M  K
  moreover from extractFrame P = AP, ΨP AP ♯* Ψ AP ♯* P AP ♯* Q AP ♯* M AP ♯* AQ
  obtain AP' where "extractFrame P = AP', ΨP" and "distinct AP'" and "AP' ♯* Ψ" and "AP' ♯* P" and "AP' ♯* Q" and "AP' ♯* M" and "AP' ♯* AQ"
    by - (rule distinctFrame[where C="(Ψ, P, Q, M, AQ)"], auto)
  moreover from extractFrame Q = AQ, ΨQ AQ ♯* Ψ AQ ♯* P AQ ♯* Q AQ ♯* K AP' ♯* AQ
  obtain AQ' where "extractFrame Q = AQ', ΨQ" and "distinct AQ'" and "AQ' ♯* Ψ" and "AQ' ♯* P" and "AQ' ♯* Q" and "AQ' ♯* K" and "AP' ♯* AQ'"
    by - (rule distinctFrame[where C="(Ψ, P, Q, K, AP')"], auto)
  ultimately show ?thesis using xvec ♯* Q
    by(metis Goal)
qed

lemma BrMerge:
  fixes Ψ    :: 'b
    and ΨQ  :: 'b
    and P    :: "('a, 'b, 'c) psi"
    and M    :: 'a
    and N    :: 'a
    and P'   :: "('a, 'b, 'c) psi"
    and AP   :: "name list"
    and ΨP  :: 'b
    and Q    :: "('a, 'b, 'c) psi"
    and Q'   :: "('a, 'b, 'c) psi"
    and AQ   :: "name list"

assumes "Ψ  ΨQ  P ¿MN  P'"
  and   "extractFrame P = AP, ΨP"
  and   "Ψ  ΨP  Q ¿MN  Q'"
  and   "extractFrame Q = AQ, ΨQ"
  and   "AP ♯* Ψ"
  and   "AP ♯* P"
  and   "AP ♯* Q"
  and   "AP ♯* M"
  and   "AP ♯* AQ"
  and   "AQ ♯* Ψ"
  and   "AQ ♯* P"
  and   "AQ ♯* Q"
  and   "AQ ♯* M"

shows "Ψ  P  Q ¿MN  (P'  Q')"
proof -
  {
    fix Ψ    :: 'b
      and ΨQ  :: 'b
      and P    :: "('a, 'b, 'c) psi"
      and M    :: 'a
      and N    :: 'a
      and P'   :: "('a, 'b, 'c) psi"
      and AP   :: "name list"
      and ΨP  :: 'b
      and Q    :: "('a, 'b, 'c) psi"
      and Q'   :: "('a, 'b, 'c) psi"
      and AQ   :: "name list"
      and svec :: "name list"

    assume "Ψ  ΨQ  P ¿MN  P'"
      and  "extractFrame P = AP, ΨP"
      and  "distinct AP"
      and  "Ψ  ΨP  Q ¿MN  Q'"
      and  "extractFrame Q = AQ, ΨQ"
      and  "distinct AQ"
      and  "AP ♯* Ψ"
      and  "AP ♯* P"
      and  "AP ♯* Q"
      and  "AP ♯* M"
      and  "AP ♯* AQ"
      and  "AQ ♯* Ψ"
      and  "AQ ♯* P"
      and  "AQ ♯* Q"
      and  "AQ ♯* M"

    have "Ψ  P  Q ¿MN  (P'  Q')"
    proof -
      obtain q::"name prm" where "(q  (AQ::name list)) ♯* Ψ" and "(q  AQ) ♯* P"
        and "(q  AQ) ♯* Q" and "(q  AQ) ♯* M"
        and "(q  AQ) ♯* ΨP" and "(q  AQ) ♯* AP" and "(q  AQ) ♯* ΨQ"
        and "(q  AQ) ♯* N" and "(q  AQ) ♯* P'" and "(q  AQ) ♯* Q'"
        and Sq: "set q  set AQ × set(q  AQ)"
        and "distinctPerm q"
        by(rule name_list_avoiding[where c="(Ψ, P, M, N, P', Q', Q, ΨQ, AP, ΨP)"])
          (auto simp add: eqvts fresh_star_prod)
      obtain p::"name prm" where "(p  AP) ♯* Ψ" and "(p  AP) ♯* P"
        and "(p  AP) ♯* Q" and "(p  AP) ♯* M"
        and "(p  AP) ♯* ΨP" and "(p  AP) ♯* ΨQ" and "(p  AP) ♯* AQ"
        and "(p  AP) ♯* N" and "(p  AP) ♯* P'" and "(p  AP) ♯* Q'"
        and Sp: "set p  set AP × set(p  AP)"
        and "(p  AP) ♯* (q  AQ)"
        by(rule name_list_avoiding[where c="(Ψ, P, N, P', Q', Q, M, AQ, q  AQ, ΨQ, ΨP)"])
          (auto simp add: eqvts fresh_star_prod)

      have FrP: "extractFrame P = AP, ΨP" by fact
      have FrQ: "extractFrame Q = AQ, ΨQ" by fact

      from AP ♯* Q FrQ AP ♯* AQ have "AP ♯* ΨQ"
        by(force dest: extractFrameFreshChain)
      from AQ ♯* P FrP AP ♯* AQ have "AQ ♯* ΨP"
        by(force dest: extractFrameFreshChain)

      from Sq AQ ♯* M (q  AQ) ♯* M
      have "(q  M) = M"
        by simp
      from Sp AP ♯* M (p  AP) ♯* M
      have "(p  M) = M"
        by simp

      from distinct AP have "distinct(p  AP)" by simp
      moreover from distinct AQ have "distinct(q  AQ)"  by simp
      moreover from extractFrame P = AP, ΨP Sp (p  AP) ♯* ΨP
      have FrP: "extractFrame P = (p  AP), (p  ΨP)"
        by(simp add: frameChainAlpha)
      moreover from extractFrame Q = AQ, ΨQ Sq (q  AQ) ♯* ΨQ
      have FrQ: "extractFrame Q = (q  AQ), (q  ΨQ)"
        by(simp add: frameChainAlpha)

      moreover have "(q  (Ψ  ΨQ))  P ¿(q  M)N  P'" using Sq AQ ♯* P (q  AQ) ♯* P Ψ  ΨQ  P ¿MN  P'
        by - (rule brinputPermFrameSubject, (assumption | simp)+)
      then have "(Ψ  (q  ΨQ))  P ¿(q  M)N  P'" using Sq AQ ♯* Ψ (q  AQ) ♯* Ψ
        by(simp add: eqvts)
      with (q  M) = M have PTrans: "(Ψ  (q  ΨQ))  P ¿MN  P'"
        by simp

      moreover have "(p  (Ψ  ΨP))  Q ¿(p  M)N  Q'" using Sp AP ♯* Q (p  AP) ♯* Q Ψ  ΨP  Q ¿MN  Q'
        by - (rule brinputPermFrameSubject, (assumption | simp)+)
      then have "(Ψ  (p  ΨP))  Q ¿(p  M)N  Q'" using Sp AP ♯* Ψ (p  AP) ♯* Ψ
        by(simp add: eqvts)
      with (p  M) = M have PTrans: "(Ψ  (p  ΨP))  Q ¿MN  Q'"
        by simp
      moreover from (p  AP) ♯* AQ (p  AP) ♯* (q  AQ) (p  AP) ♯* ΨQ Sq have "(p  AP) ♯* (q  ΨQ)"
        by(simp add: freshChainSimps)
      moreover from (q  AQ) ♯* AP (p  AP) ♯* AQ (q  AQ) ♯* ΨP (p  AP) ♯* (q  AQ) Sp Sq have "(q  AQ) ♯* (p  ΨP)"
        by(simp add: freshChainSimps)

      moreover note
        (p  AP) ♯* Ψ (p  AP) ♯* M
        (p  AP) ♯* P (p  AP) ♯* N (p  AP) ♯* P'
        (p  AP) ♯* Q (p  AP) ♯* Q' (p  AP) ♯* (q  AQ)
        (q  AQ) ♯* Ψ (q  AQ) ♯* M
        (q  AQ) ♯* P (q  AQ) ♯* N (q  AQ) ♯* P'
        (q  AQ) ♯* Q (q  AQ) ♯* Q'
      ultimately show ?thesis
        by(simp add: cBrMerge)
    qed
  }
  note Goal = this

  note Ψ  ΨQ  P ¿MN  P' Ψ  ΨP  Q ¿MN  Q'
  moreover from extractFrame P = AP, ΨP AP ♯* Ψ AP ♯* P AP ♯* Q AP ♯* M AP ♯* AQ
  obtain AP' where "extractFrame P = AP', ΨP" and "distinct AP'" and "AP' ♯* Ψ" and "AP' ♯* P" and "AP' ♯* Q" and "AP' ♯* M" and "AP' ♯* AQ"
    by - (rule distinctFrame[where C="(Ψ, P, Q, M, AQ)"], auto)
  moreover from extractFrame Q = AQ, ΨQ AQ ♯* Ψ AQ ♯* P AQ ♯* Q AQ ♯* M AP' ♯* AQ
  obtain AQ' where "extractFrame Q = AQ', ΨQ" and "distinct AQ'" and "AQ' ♯* Ψ" and "AQ' ♯* P" and "AQ' ♯* Q" and "AQ' ♯* M" and "AP' ♯* AQ'"
    by - (rule distinctFrame[where C="(Ψ, P, Q, M, AP')"], auto)
  ultimately show ?thesis
    by(metis Goal)
qed

lemma BrComm1:
  fixes Ψ    :: 'b
    and ΨQ  :: 'b
    and P    :: "('a, 'b, 'c) psi"
    and M    :: 'a
    and N    :: 'a
    and P'   :: "('a, 'b, 'c) psi"
    and AP   :: "name list"
    and ΨP  :: 'b
    and Q    :: "('a, 'b, 'c) psi"
    and xvec :: "name list"
    and Q'   :: "('a, 'b, 'c) psi"
    and AQ   :: "name list"

assumes "Ψ  ΨQ  P ¿MN  P'"
  and   "extractFrame P = AP, ΨP"
  and   "Ψ  ΨP  Q ¡M⦇ν*xvec⦈⟨N  Q'"
  and   "extractFrame Q = AQ, ΨQ"
  and   "AP ♯* Ψ"
  and   "AP ♯* P"
  and   "AP ♯* Q"
  and   "AP ♯* M"
  and   "AP ♯* AQ"
  and   "AQ ♯* Ψ"
  and   "AQ ♯* P"
  and   "AQ ♯* Q"
  and   "AQ ♯* M"
  and   "xvec ♯* P"

shows "Ψ  P  Q ¡M⦇ν*xvec⦈⟨N  (P'  Q')"
proof -
  {
    fix Ψ    :: 'b
      and ΨQ  :: 'b
      and P    :: "('a, 'b, 'c) psi"
      and M    :: 'a
      and N    :: 'a
      and P'   :: "('a, 'b, 'c) psi"
      and AP   :: "name list"
      and ΨP  :: 'b
      and Q    :: "('a, 'b, 'c) psi"
      and xvec :: "name list"
      and Q'   :: "('a, 'b, 'c) psi"
      and AQ   :: "name list"

    assume "Ψ  ΨQ  P ¿MN  P'"
      and  "extractFrame P = AP, ΨP"
      and  "distinct AP"
      and  "Ψ  ΨP  Q ¡M⦇ν*xvec⦈⟨N  Q'"
      and  "extractFrame Q = AQ, ΨQ"
      and  "distinct AQ"
      and  "AP ♯* Ψ"
      and  "AP ♯* P"
      and  "AP ♯* Q"
      and  "AP ♯* M"
      and  "AP ♯* AQ"
      and  "AQ ♯* Ψ"
      and  "AQ ♯* P"
      and  "AQ ♯* Q"
      and  "AQ ♯* 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) ♯* AP" and "(r  xvec) ♯* AQ"
        and "(r  xvec) ♯* P'" and "(r  xvec) ♯* Q'" and "(r  xvec) ♯* ΨP" and "(r  xvec) ♯* Ψ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, AP, AQ, ΨP, ΨQ, P', Q')"])
          (auto simp add: eqvts fresh_star_prod)
      obtain q::"name prm" where "(q  AQ) ♯* Ψ" and "(q  AQ) ♯* P" and "(q  AQ) ♯* Q" and "(q  AQ) ♯* M"
        and "(q  AQ) ♯* (r  N)" and "(q  AQ) ♯* (r  xvec)" and "(q  AQ) ♯* (r  Q')"
        and "(q  AQ) ♯* (r  P')" and "(q  AQ) ♯* ΨP" and "(q  AQ) ♯* AP" and "(q  AQ) ♯* ΨQ"
        and Sq: "set q  set AQ × set(q  AQ)"
        by(rule name_list_avoiding[where xvec=AQ and c="(Ψ, P, Q, M, r  N, r  xvec, ΨQ, AP, ΨP, r  Q', r  P')"])
          (auto simp add: eqvts fresh_star_prod)
      obtain p::"name prm" where "(p  AP) ♯* Ψ" and "(p  AP) ♯* P" and "(p  AP) ♯* Q" and "(p  AP) ♯* M"
        and "(p  AP) ♯* (r  N)" and "(p  AP) ♯* (r  xvec)" and "(p  AP) ♯* (r  Q')"
        and "(p  AP) ♯* (r  P')" and "(p  AP) ♯* ΨP" and "(p  AP) ♯* ΨQ" and "(p  AP) ♯* AQ"
        and "(p  AP) ♯* (q  AQ)" and Sp: "(set p)  (set AP) × (set(p  AP))"
        by(rule name_list_avoiding[where xvec=AP and c="(Ψ, P, Q, M, r  N, r  xvec, AQ, q  AQ, ΨQ, ΨP, r  Q', r  P')"])
          (auto simp add: eqvts fresh_star_prod)

      have FrP: "extractFrame P = AP, ΨP" by fact
      have FrQ: "extractFrame Q = AQ, ΨQ" by fact

      from Sp AP ♯* M (p  AP) ♯* M
      have "(p  M) = M"
        by simp
      from Sq AQ ♯* M (q  AQ) ♯* M
      have "(q  M) = M"
        by simp

      from AP ♯* Q FrQ AP ♯* AQ have "AP ♯* ΨQ"
        by(auto dest: extractFrameFreshChain)
      from AQ ♯* P FrP AP ♯* AQ have "AQ ♯* ΨP"
        by(auto dest: extractFrameFreshChain)
      from (r  xvec) ♯* AP (p  AP) ♯* (r  xvec) (r  xvec) ♯* AP Sp have "(r  xvec) ♯* (p  AP)"
        by(simp add: freshChainSimps)

      from Ψ  ΨQ  P ¿MN  P' Sr distinctPerm r xvec ♯* P (r  xvec) ♯* P
      have "Ψ  ΨQ  P ¿M(r  N)  (r  P')"
        by(rule brinputAlpha)
      then have "(q  (Ψ  ΨQ))  P ¿(q  M)(r  N)  (r  P')" using Sq AQ ♯* P (q  AQ) ♯* P
        by - (rule brinputPermFrameSubject, (assumption | simp)+)
      then have "Ψ  (q  ΨQ)  P ¿(q  M)(r  N)  (r  P')" using Sq AQ ♯* Ψ (q  AQ) ♯* Ψ
        by(simp add: eqvts)
      with (q  M) = M have PTrans: "Ψ  (q  ΨQ)  P ¿M(r  N)  (r  P')" by simp

      moreover from extractFrame P = AP, ΨP  Sp (p  AP) ♯* ΨP
      have FrP: "extractFrame P = (p  AP), (p  ΨP)"
        by(simp add: frameChainAlpha)
      moreover from distinct AP have "distinct(p  AP)" by simp

      moreover from Ψ  ΨP  Q ¡M⦇ν*xvec⦈⟨N  Q' Sr (r  xvec) ♯* N (r  xvec) ♯* Q'
      have "Ψ  ΨP  Q ¡M⦇ν*(r  xvec)⦈⟨(r  N)  (r  Q')"
        by(simp add: boundOutputChainAlpha'' create_residual.simps)
      then have "(p  (Ψ  ΨP))  Q ¡(p  M)⦇ν*(r  xvec)⦈⟨(r  N)  (r  Q')" using Sp AP ♯* Q (p  AP) ♯* Q (r  xvec) ♯* M (r  xvec) ♯* AP (r  xvec) ♯* (p  AP)
        by(fastforce intro: broutputPermFrameSubject)
      then have "Ψ  (p  ΨP)  Q ¡(p  M)⦇ν*(r  xvec)⦈⟨(r  N)  (r  Q')" using Sp AP ♯* Ψ (p  AP) ♯* Ψ
        by(simp add: eqvts)
      with (p  M) = M have QTrans: "Ψ  (p  Ψ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 = AQ, ΨQ  Sq (q  AQ) ♯* ΨQ
      have FrQ: "extractFrame Q = (q  AQ), (q  ΨQ)"
        by(simp add: frameChainAlpha)
      moreover from distinct AQ have "distinct(q  AQ)"  by simp

      moreover note (p  AP) ♯* Ψ
      moreover from (p  AP) ♯* AQ (p  AP) ♯* (q  AQ) (p  AP) ♯* ΨQ Sq have "(p  AP) ♯* (q  ΨQ)"
        by(simp add: freshChainSimps)
      moreover note (p  AP) ♯* P (p  AP) ♯* M
      moreover note  (p  AP) ♯* (r  N) (p  AP) ♯* (r  P') (p  AP) ♯* Q (p  AP) ♯* (r  Q') (p  AP) ♯* (q  AQ)
        (p  AP) ♯* (r  xvec) (q  AQ) ♯* Ψ
      moreover from (q  AQ) ♯* AP (p  AP) ♯* AQ (q  AQ) ♯* ΨP (p  AP) ♯* (q  AQ) Sp Sq have "(q  AQ) ♯* (p  ΨP)"
        by(simp add: freshChainSimps)
      moreover note (q  AQ) ♯* P (q  AQ) ♯* (r  N)(q  AQ) ♯* (r  P') (q  AQ) ♯* Q (q  AQ) ♯* M
      moreover note  (q  AQ) ♯* (r  Q') (q  AQ) ♯* (r  xvec) (r  xvec) ♯* Ψ
      moreover from (r  xvec) ♯* AP (p  AP) ♯* (r  xvec) (r  xvec) ♯* ΨP Sp have "(r  xvec) ♯* (p  ΨP)"
        by(simp add: freshChainSimps)
      moreover from (r  xvec) ♯* AQ (q  AQ) ♯* (r  xvec) (r  xvec) ♯* ΨQ Sq have "(r  xvec) ♯* (q  Ψ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 "⦇ν*xvecN ≺' (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 Ψ  ΨQ  P ¿MN  P' Ψ  ΨP  Q ¡M⦇ν*xvec⦈⟨N  Q'
  moreover from extractFrame P = AP, ΨP AP ♯* Ψ AP ♯* P AP ♯* Q AP ♯* M AP ♯* AQ
  obtain AP' where "extractFrame P = AP', ΨP" and "distinct AP'" and "AP' ♯* Ψ" and "AP' ♯* P" and "AP' ♯* Q" and "AP' ♯* M" and "AP' ♯* AQ"
    by - (rule distinctFrame[where C="(Ψ, P, Q, M, AQ)"], auto)
  moreover from extractFrame Q = AQ, ΨQ AQ ♯* Ψ AQ ♯* P AQ ♯* Q AQ ♯* M AP' ♯* AQ
  obtain AQ' where "extractFrame Q = AQ', ΨQ" and "distinct AQ'" and "AQ' ♯* Ψ" and "AQ' ♯* P" and "AQ' ♯* Q" and "AQ' ♯* M" and "AP' ♯* AQ'"
    by - (rule distinctFrame[where C="(Ψ, P, Q, M, AP')"], auto)
  ultimately show ?thesis using xvec ♯* P
    by(metis Goal)
qed

lemma BrComm2:
  fixes Ψ    :: 'b
    and Ψ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 AP   :: "name list"
    and ΨP  :: 'b
    and Q    :: "('a, 'b, 'c) psi"
    and Q'   :: "('a, 'b, 'c) psi"
    and AQ   :: "name list"

assumes "Ψ  ΨQ  P ¡M⦇ν*xvec⦈⟨N  P'"
  and   "extractFrame P = AP, ΨP"
  and   "Ψ  ΨP  Q ¿MN  Q'"
  and   "extractFrame Q = AQ, ΨQ"
  and   "AP ♯* Ψ"
  and   "AP ♯* P"
  and   "AP ♯* Q"
  and   "AP ♯* M"
  and   "AP ♯* AQ"
  and   "AQ ♯* Ψ"
  and   "AQ ♯* P"
  and   "AQ ♯* Q"
  and   "AQ ♯* M"
  and   "xvec ♯* Q"

shows "Ψ  P  Q ¡M⦇ν*xvec⦈⟨N  (P'  Q')"
proof -
  {
    fix Ψ    :: 'b
      and Ψ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 AP   :: "name list"
      and ΨP  :: 'b
      and Q    :: "('a, 'b, 'c) psi"
      and Q'   :: "('a, 'b, 'c) psi"
      and AQ   :: "name list"

    assume "Ψ  ΨQ  P ¡M⦇ν*xvec⦈⟨N  P'"
      and  "extractFrame P = AP, ΨP"
      and  "distinct AP"
      and  "Ψ  ΨP  Q ¿MN  Q'"
      and  "extractFrame Q = AQ, ΨQ"
      and  "distinct AQ"
      and  "AP ♯* Ψ"
      and  "AP ♯* P"
      and  "AP ♯* Q"
      and  "AP ♯* M"
      and  "AP ♯* AQ"
      and  "AQ ♯* Ψ"
      and  "AQ ♯* P"
      and  "AQ ♯* Q"
      and  "AQ ♯* 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) ♯* AP" and "(r  xvec) ♯* AQ"
        and "(r  xvec) ♯* P'" and "(r  xvec) ♯* Q'" and "(r  xvec) ♯* ΨP" and "(r  xvec) ♯* Ψ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, AP, AQ, ΨP, ΨQ, P', Q')"])
          (auto simp add: eqvts fresh_star_prod)
      obtain q::"name prm" where "(q  AQ) ♯* Ψ" and "(q  AQ) ♯* P" and "(q  AQ) ♯* Q" and "(q  AQ) ♯* M"
        and "(q  AQ) ♯* (r  N)" and "(q  AQ) ♯* (r  xvec)" and "(q  AQ) ♯* (r  Q')"
        and "(q  AQ) ♯* (r  P')" and "(q  AQ) ♯* ΨP" and "(q  AQ) ♯* AP" and "(q  AQ) ♯* ΨQ"
        and Sq: "set q  set AQ × set(q  AQ)"
        by(rule name_list_avoiding[where xvec=AQ and c="(Ψ, P, Q, M, r  N, r  xvec, ΨQ, AP, ΨP, r  Q', r  P')"])
          (auto simp add: eqvts fresh_star_prod)
      obtain p::"name prm" where "(p  AP) ♯* Ψ" and "(p  AP) ♯* P" and "(p  AP) ♯* Q" and "(p  AP) ♯* M"
        and "(p  AP) ♯* (r  N)" and "(p  AP) ♯* (r  xvec)" and "(p  AP) ♯* (r  Q')"
        and "(p  AP) ♯* (r  P')" and "(p  AP) ♯* ΨP" and "(p  AP) ♯* ΨQ" and "(p  AP) ♯* AQ"
        and "(p  AP) ♯* (q  AQ)" and Sp: "(set p)  (set AP) × (set(p  AP))"
        by(rule name_list_avoiding[where xvec=AP and c="(Ψ, P, Q, M, r  N, r  xvec, AQ, q  AQ, ΨQ, ΨP, r  Q', r  P')"])
          (auto simp add: eqvts fresh_star_prod)

      have FrP: "extractFrame P = AP, ΨP" by fact
      have FrQ: "extractFrame Q = AQ, ΨQ" by fact

      from Sp AP ♯* M (p  AP) ♯* M
      have "(p  M) = M"
        by simp
      from Sq AQ ♯* M (q  AQ) ♯* M
      have "(q  M) = M"
        by simp

      from AP ♯* Q FrQ AP ♯* AQ have "AP ♯* ΨQ"
        by(auto dest: extractFrameFreshChain)
      from AQ ♯* P FrP AP ♯* AQ have "AQ ♯* ΨP"
        by(auto dest: extractFrameFreshChain)
      from (r  xvec) ♯* AQ (q  AQ) ♯* (r  xvec) (r  xvec) ♯* AQ Sq have "(r  xvec) ♯* (q  AQ)"
        by(simp add: freshChainSimps)

      from Ψ  ΨQ  P ¡M⦇ν*xvec⦈⟨N  P' Sr (r  xvec) ♯* N (r  xvec) ♯* P'
      have "Ψ  ΨQ  P ¡M⦇ν*(r  xvec)⦈⟨(r  N)  (r  P')"
        by(simp add: boundOutputChainAlpha'' create_residual.simps)
      then have "(q  (Ψ  ΨQ))  P ¡(q  M)⦇ν*(r  xvec)⦈⟨(r  N)  (r  P')" using Sq AQ ♯* P (q  AQ) ♯* P (r  xvec) ♯* M (r  xvec) ♯* AQ (r  xvec) ♯* (q  AQ)
        by(fastforce intro: broutputPermFrameSubject)
      then have "Ψ  (q  ΨQ)  P ¡(q  M)⦇ν*(r  xvec)⦈⟨(r  N)  (r  P')" using Sq AQ ♯* Ψ (q  AQ) ♯* Ψ
        by(simp add: eqvts)
      with (q  M) = M have PTrans: "Ψ  (q  Ψ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 = AP, ΨP  Sp (p  AP) ♯* ΨP
      have FrP: "extractFrame P = (p  AP), (p  ΨP)"
        by(simp add: frameChainAlpha)
      moreover from distinct AP have "distinct(p  AP)"  by simp

      moreover from Ψ  ΨP  Q ¿MN  Q' Sr distinctPerm r xvec ♯* Q (r  xvec) ♯* Q
      have "Ψ  ΨP  Q ¿M(r  N)  (r  Q')"
        by(rule brinputAlpha)
      then have "(p  (Ψ  ΨP))  Q ¿(p  M)(r  N)  (r  Q')" using Sp AP ♯* Q (p  AP) ♯* Q
        by - (rule brinputPermFrameSubject, (assumption | simp)+)
      then have QTrans: "Ψ  (p  ΨP)  Q ¿(p  M)(r  N)  (r  Q')" using Sp AP ♯* Ψ (p  AP) ♯* Ψ
        by(simp add: eqvts)
      with (p  M) = M have "Ψ  (p  ΨP)  Q ¿M(r  N)  (r  Q')" by simp

      moreover from extractFrame Q = AQ, ΨQ  Sq (q  AQ) ♯* ΨQ
      have FrQ: "extractFrame Q = (q  AQ), (q  ΨQ)"
        by(simp add: frameChainAlpha)
      moreover from distinct AQ have "distinct(q  AQ)"  by simp

      moreover note (p  AP) ♯* Ψ
      moreover from (p  AP) ♯* AQ (p  AP) ♯* (q  AQ) (p  AP) ♯* ΨQ Sq have "(p  AP) ♯* (q  ΨQ)"
        by(simp add: freshChainSimps)
      moreover note (p  AP) ♯* P (p  AP) ♯* M
      moreover note  (p  AP) ♯* (r  N) (p  AP) ♯* (r  P') (p  AP) ♯* Q (p  AP) ♯* (r  Q') (p  AP) ♯* (q  AQ)
        (p  AP) ♯* (r  xvec) (q  AQ) ♯* Ψ
      moreover from (q  AQ) ♯* AP (p  AP) ♯* AQ (q  AQ) ♯* ΨP (p  AP) ♯* (q  AQ) Sp Sq have "(q  AQ) ♯* (p  ΨP)"
        by(simp add: freshChainSimps)
      moreover note (q  AQ) ♯* P (q  AQ) ♯* (r  N)(q  AQ) ♯* (r  P') (q  AQ) ♯* Q (q  AQ) ♯* M
      moreover note  (q  AQ) ♯* (r  Q') (q  AQ) ♯* (r  xvec) (r  xvec) ♯* Ψ
      moreover from (r  xvec) ♯* AP (p  AP) ♯* (r  xvec) (r  xvec) ♯* ΨP Sp have "(r  xvec) ♯* (p  ΨP)"
        by(simp add: freshChainSimps)
      moreover from (r  xvec) ♯* AQ (q  AQ) ♯* (r  xvec) (r  xvec) ♯* ΨQ Sq have "(r  xvec) ♯* (q  Ψ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 "⦇ν*xvecN ≺' (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 Ψ  ΨQ  P ¡M⦇ν*xvec⦈⟨N  P' Ψ  ΨP  Q ¿MN  Q'
  moreover from extractFrame P = AP, ΨP AP ♯* Ψ AP ♯* P AP ♯* Q AP ♯* M AP ♯* AQ
  obtain AP' where "extractFrame P = AP', ΨP" and "distinct AP'" and "AP' ♯* Ψ" and "AP' ♯* P" and "AP' ♯* Q" and "AP' ♯* M" and "AP' ♯* AQ"
    by - (rule distinctFrame[where C="(Ψ, P, Q, M, AQ)"], auto)
  moreover from extractFrame Q = AQ, ΨQ AQ ♯* Ψ AQ ♯* P AQ ♯* Q AQ ♯* M AP' ♯* AQ
  obtain AQ' where "extractFrame Q = AQ', ΨQ" and "distinct AQ'" and "AQ' ♯* Ψ" and "AQ' ♯* P" and "AQ' ♯* Q" and "AQ' ♯* M" and "AP' ♯* AQ'"
    by - (rule distinctFrame[where C="(Ψ, P, Q, M, AP')"], 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 "Ψ  ⦇νxP  τ  ⦇νx(⦇ν*xvecP')"
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 "Ψ  ⦇νxP  τ  ⦇ν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 = MN⟩.P; cRs = KN  P; Ψ  M  K  Prop"
  and   rBrOutput: "M K N P. cP = MN⟩.P; cRs = ¡KN  P; Ψ  M  K  Prop"
  and   rCase: "Cs P φ. cP = Cases Cs; Ψ  P  cRs; (φ, P)  set Cs; Ψ  φ; guarded P  Prop"

and  rPar1: "ΨQ P α P' Q AQ. cP = P  Q; cRs = α  (P'  Q);
               (Ψ  ΨQ)  P  (α  P'); extractFrame Q = AQ, ΨQ; distinct AQ;
               AQ ♯* P; AQ ♯* Q; AQ ♯* Ψ; AQ ♯* α; AQ ♯* C; AQ ♯* P'; bn α ♯* Ψ; bn α  ♯* ΨQ;
               bn α  ♯* Q; bn α  ♯* P; bn α ♯* subject α; bn α ♯* C; distinct(bn α) 
               Prop"
and   rPar2:   "ΨP Q α Q' P AP. cP = P  Q; cRs = α  (P  Q');
                      (Ψ  ΨP)  Q α  Q'; extractFrame P = AP, ΨP; distinct AP;
             AP ♯* P; AP ♯* Q; AP ♯* Ψ; AP ♯* α; AP ♯* C;
             AP ♯* Q'; bn α ♯* Ψ; bn α ♯* ΨP; bn α ♯* P; bn α ♯* Q; bn α ♯* subject α; bn α ♯* C; distinct(bn α)  Prop"
and   rComm1: "ΨQ P M N P' AP ΨP Q K xvec Q' AQ.
                   cP = P  Q; cRs = τ  ⦇ν*xvecP'  Q';
                    Ψ  ΨQ  P MN  P'; extractFrame P = AP, ΨP; distinct AP;
                    Ψ  ΨP  Q  K⦇ν*xvec⦈⟨N  Q'; extractFrame Q = AQ, ΨQ; distinct AQ;
                    Ψ  ΨP  ΨQ  M  K; AP ♯* Ψ; AP ♯* ΨQ; AP ♯* P; AP ♯* M; AP ♯* N;
                    AP ♯* P'; AP ♯* Q; AP ♯* Q'; AP ♯* AQ; AP ♯* xvec; AQ ♯* Ψ; AQ ♯* ΨP;
                    AQ ♯* P; AQ ♯* K; AQ ♯* N; AQ ♯* P'; AQ ♯* Q; AQ ♯* Q'; AQ ♯* xvec;
                    xvec ♯* Ψ; xvec ♯* ΨP; xvec ♯* ΨQ; xvec ♯* P; xvec ♯* M; xvec ♯* Q;
                    xvec ♯* K; AP ♯* C; AQ ♯* C; xvec ♯* C; distinct xvec  Prop"
and   rComm2: "ΨQ P M xvec N P' AP ΨP Q K Q' AQ.
                   cP = P  Q; cRs = τ  ⦇ν*xvecP'  Q';
                    Ψ  ΨQ  P M⦇ν*xvec⦈⟨N  P'; extractFrame P = AP, ΨP; distinct AP;
                    Ψ  ΨP  Q  KN  Q'; extractFrame Q = AQ, ΨQ; distinct AQ;
                    Ψ  ΨP  ΨQ  M  K; AP ♯* Ψ; AP ♯* ΨQ; AP ♯* P; AP ♯* M; AP ♯* N;
                    AP ♯* P'; AP ♯* Q; AP ♯* Q'; AP ♯* AQ; AP ♯* xvec; AQ ♯* Ψ; AQ ♯* ΨP;
                    AQ ♯* P; AQ ♯* K; AQ ♯* N; AQ ♯* P'; AQ ♯* Q; AQ ♯* Q'; AQ ♯* xvec;
                    xvec ♯* Ψ; xvec ♯* ΨP; xvec ♯* ΨQ; xvec ♯* P; xvec ♯* M; xvec ♯* Q;
                    xvec ♯* K; AP ♯* C; AQ ♯* C; xvec ♯* C; distinct xvec  Prop"
and   rBrMerge: "ΨQ P M N P' AP ΨP Q Q' AQ.
                    cP = (P  Q); cRs = ¿MN  (P'  Q');
                    Ψ  ΨQ  P  ¿MN  P'; extractFrame P = AP, ΨP; distinct AP;
                    Ψ  ΨP  Q  ¿MN  Q'; extractFrame Q = AQ, ΨQ; distinct AQ;
                    AP ♯* Ψ; AP ♯* ΨQ; AP ♯* P; AP ♯* N; AP ♯* P';
                    AP ♯* Q; AP ♯* Q'; AP ♯* M; AP ♯* AQ;
                    AQ ♯* Ψ; AQ ♯* ΨP; AQ ♯* P; AQ ♯* N; AQ ♯* P';
                    AQ ♯* Q; AQ ♯* Q'; AQ ♯* M; AP ♯* C; AQ ♯* C  Prop"
and   rBrComm1: "ΨQ P M N P' AP ΨP Q xvec Q' AQ.
                   cP = P  Q; cRs = ¡M⦇ν*xvec⦈⟨N  (P'  Q');
                    Ψ  ΨQ  P ¿MN  P'; extractFrame P = AP, ΨP; distinct AP;
                    Ψ  ΨP  Q  ¡M⦇ν*xvec⦈⟨N  Q'; extractFrame Q = AQ, ΨQ; distinct AQ;
                    AP ♯* Ψ; AP ♯* ΨQ; AP ♯* P; AP ♯* N;
                    AP ♯* P'; AP ♯* Q; AP ♯* Q'; AP ♯* M; AP ♯* AQ; AP ♯* xvec; AQ ♯* Ψ; AQ ♯* ΨP;
                    AQ ♯* P; AQ ♯* N; AQ ♯* P'; AQ ♯* M; AQ ♯* Q; AQ ♯* Q'; AQ ♯* xvec;
                    xvec ♯* Ψ; xvec ♯* ΨP; xvec ♯* ΨQ; xvec ♯* P; xvec ♯* Q; xvec ♯* M;
                    AP ♯* C; AQ ♯* C; xvec ♯* C; distinct xvec  Prop"
and   rBrComm2: "ΨQ P M xvec N P' AP ΨP Q Q' AQ.
                   cP = P  Q; cRs = ¡M⦇ν*xvec⦈⟨N  (P'  Q');
                    Ψ  ΨQ  P ¡M⦇ν*xvec⦈⟨N  P'; extractFrame P = AP, ΨP; distinct AP;
                    Ψ  ΨP  Q  ¿MN  Q'; extractFrame Q = AQ, ΨQ; distinct AQ;
                    AP ♯* Ψ; AP ♯* ΨQ; AP ♯* P; AP ♯* N;
                    AP ♯* P'; AP ♯* Q; AP ♯* Q'; AP ♯* M; AP ♯* AQ; AP ♯* xvec; AQ ♯* Ψ; AQ ♯* ΨP;
                    AQ ♯* P; AQ ♯* N; AQ ♯* P'; AQ ♯* M; AQ ♯* Q; AQ ♯* Q'; AQ ♯* xvec;
                    xvec ♯* Ψ; xvec ♯* ΨP; xvec ♯* ΨQ; xvec ♯* P; xvec ♯* Q; xvec ♯* M;
                    AP ♯* C; AQ ♯* C; xvec ♯* C; distinct xvec  Prop"
and   rBrClose: "P M xvec N P' x.
                   cP = (⦇νxP); cRs = τ  ⦇νx(⦇ν*xvecP');
                    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 = ⦇νxP; 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 = ⦇νxP; 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 = ⦇νxP; cRs = α  ⦇νxP';
                                 Ψ  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 ΨQ P α P' Q AQ)
  obtain q::"name prm" where "(bn(q  α)) ♯* Ψ" and "(bn(q  α)) ♯* P" and "(bn(q  α)) ♯* Q"
    and "(bn(q  α)) ♯* α" and "(bn(q  α)) ♯* AQ" and "(bn(q  α)) ♯* P'" and "(bn(q  α)) ♯* Ψ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, α, AQ, ΨQ, P', C)"]) (auto simp add: eqvts)
  obtain p::"name prm" where "(p  AQ) ♯* Ψ" and "(p  AQ) ♯* P" and "(p  AQ) ♯* Q"
    and "(p  AQ) ♯* α" and "(p  AQ) ♯* (q  α)" and "(p  AQ) ♯* P'"
    and "(p  AQ) ♯* (q  P')" and "(p  AQ) ♯* ΨQ" and "(p  AQ) ♯* C"
    and Sp: "(set p)  (set AQ) × (set(p  AQ))" and "distinctPerm p"
    by(rule name_list_avoiding[where xvec=AQ and c="(Ψ, P, Q, α, q  α, P', (q  P'), ΨQ, C)"]) auto
  from AQ ♯* α bn(q  α) ♯* AQ Sq distinctPerm q have "AQ ♯* (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 Ψ  ΨQ  P α  P' bn α ♯* subject α (bn(q  α)) ♯* α (bn(q  α)) ♯* P' Sq
  have Trans: "Ψ  ΨQ  P (q  α)  (q  P')"
    by(force simp add: residualAlpha)
  then have "AQ ♯* (q  P')" using bn(q  α) ♯* subject(q  α) distinct(bn(q  α)) AQ ♯* P AQ ♯* (q  α)
    by(force dest: freeFreshChainDerivative)

  from Trans have "(p  (Ψ  ΨQ))  (p  P) p  ((q  α)  (q  P'))"
    by(rule semantics.eqvt)
  with AQ ♯* Ψ AQ ♯* P AQ ♯* (q  α) AQ ♯* (q  P') (bn(q  α)) ♯* AQ  (p  AQ) ♯* (q  α)
    (p  AQ) ♯* Ψ (p  AQ) ♯* P  (p  AQ) ♯* (q  P') Sp
  have "Ψ  (p  ΨQ)  P (q  α)  (q  P')" by(simp add: eqvts)
  moreover from extractFrame Q = AQ, ΨQ (p  AQ) ♯* ΨQ Sp have  "extractFrame Q = (p  AQ), (p  ΨQ)"
    by(simp add: frameChainAlpha' eqvts)
  moreover from (bn(q  α)) ♯* ΨQ (bn(q  α)) ♯* AQ (p  AQ) ♯* (q  α) Sp have "(bn(q  α)) ♯* (p  ΨQ)"
    by(simp add: freshAlphaPerm)
  moreover from distinct AQ have "distinct(p  AQ)" by simp
  ultimately show ?thesis
    using (p  AQ) ♯* P (p  AQ) ♯* Q (p  AQ) ♯* Ψ (p  AQ) ♯* (q  α)
      (p  AQ) ♯* (q  P') (bn(q  α)) ♯* Ψ (bn(q  α)) ♯* Q (bn(q  α)) ♯* P
      (bn(q  α)) ♯* C (p  AQ) ♯* C bn (q  α) ♯* subject (q  α) distinct(bn(q  α))
    by(metis rPar1)
next
  case(cPar2 ΨP Q α Q' P AP)
  obtain q::"name prm" where "(bn(q  α)) ♯* Ψ" and "(bn(q  α)) ♯* P" and "(bn(q  α)) ♯* Q"
    and "(bn(q  α)) ♯* α" and "(bn(q  α)) ♯* AP" and "(bn(q  α)) ♯* Q'" and "(bn(q  α)) ♯* Ψ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, α, AP, ΨP, Q', C)"]) (auto simp add: eqvts)
  obtain p::"name prm" where "(p  AP) ♯* Ψ" and "(p  AP) ♯* P" and "(p  AP) ♯* Q"
    and "(p  AP) ♯* α" and "(p  AP) ♯* (q  α)" and "(p  AP) ♯* Q'"
    and "(p  AP) ♯* (q  Q')" and "(p  AP) ♯* ΨP" and "(p  AP) ♯* C"
    and Sp: "(set p)  (set AP) × (set(p  AP))" and "distinctPerm p"
    by(rule name_list_avoiding[where xvec=AP and c="(Ψ, P, Q, α, q  α, Q', (q  Q'), ΨP, C)"]) auto
  from AP ♯* α bn(q  α) ♯* AP Sq distinctPerm q have "AP ♯* (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 Ψ  ΨP  Q α  Q' bn α ♯* subject α (bn(q  α)) ♯* α (bn(q  α)) ♯* Q' Sq
  have Trans: "Ψ  ΨP  Q (q  α)  (q  Q')"
    by(force simp add: residualAlpha)
  then have "AP ♯* (q  Q')" using bn(q  α) ♯* subject(q  α) distinct(bn(q  α)) AP ♯* Q AP ♯* (q  α)
    by(auto dest: freeFreshChainDerivative)

  from Trans have "(p  (Ψ  ΨP))  (p  Q) p  ((q  α)  (q  Q'))"
    by(rule semantics.eqvt)
  with AP ♯* Ψ AP ♯* Q AP ♯* (q  α) AP ♯* (q  Q') (bn(q  α)) ♯* AP  (p  AP) ♯* (q  α)
    (p  AP) ♯* Ψ (p  AP) ♯* Q  (p  AP) ♯* (q  Q') Sp
  have "Ψ  (p  ΨP)  Q (q  α)  (q  Q')" by(simp add: eqvts)
  moreover from extractFrame P = AP, ΨP (p  AP) ♯* ΨP Sp have  "extractFrame P = (p  AP), (p  ΨP)"
    by(simp add: frameChainAlpha' eqvts)
  moreover from (bn(q  α)) ♯* ΨP (bn(q  α)) ♯* AP (p  AP) ♯* (q  α) Sp have "(bn(q  α)) ♯* (p  ΨP)"
    by(simp add: freshAlphaPerm)
  moreover from distinct AP have "distinct(p  AP)" by simp
  ultimately show ?thesis
    using (p  AP) ♯* P (p  AP) ♯* Q (p  AP) ♯* Ψ (p  AP) ♯* (q  α)
      (p  AP) ♯* (q  Q') (bn(q  α)) ♯* Ψ (bn(q  α)) ♯* Q (bn(q  α)) ♯* P
      (bn(q  α)) ♯* C (p  AP) ♯* C bn (q  α) ♯* subject (q  α) distinct(bn(q  α))
    by(metis rPar2)
next
  case(cComm1 ΨQ P M N P' AP ΨP Q K xvec Q' AQ)
  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) ♯* AP" and "(r  xvec) ♯* AQ"
    and "(r  xvec) ♯* P'" and "(r  xvec) ♯* Q'" and "(r  xvec) ♯* ΨP" and "(r  xvec) ♯* Ψ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, AP, AQ, ΨP, ΨQ, P', Q', C)"])
      (auto simp add: eqvts)

  obtain q::"name prm" where "(q  AQ) ♯* Ψ" and "(q  AQ) ♯* P" and "(q  AQ) ♯* Q" and "(q  AQ) ♯* K"
    and "(q  AQ) ♯* N" and "(q  AQ) ♯* xvec" and "(q  AQ) ♯* Q'" and "(q  AQ) ♯* P'"
    and "(q  AQ) ♯* ΨP" and  "(q  AQ) ♯* AP" and "(q  AQ) ♯* ΨQ" and "(q  AQ) ♯* (r  xvec)"
    and "(q  AQ) ♯* C" and Sq: "(set q)  (set AQ) × (set(q  AQ))"
    by(rule name_list_avoiding[where xvec=AQ and c="(Ψ, P, Q, K, N, xvec, r  xvec, ΨQ, AP, ΨP, Q', P', C)"]) clarsimp

  obtain p::"name prm"  where "(p  AP) ♯* Ψ" and "(p  AP) ♯* P" and "(p  AP) ♯* Q" and "(p  AP) ♯* M"
    and "(p  AP) ♯* N" and "(p  AP) ♯* xvec" and "(p  AP) ♯* Q'" and "(p  AP) ♯* AQ"
    and "(p  AP) ♯* P'" and "(p  AP) ♯* ΨP" and "(p  AP) ♯* ΨQ" and "(p  AP) ♯* (q  AQ)"
    and "(p  AP) ♯* C" and "(p  AP) ♯* (r  xvec)" and Sp: "(set p)  (set AP) × (set(p  AP))"
    by(rule name_list_avoiding[where xvec=AP and c="(Ψ, P, Q, M, N, xvec, r  xvec, AQ, q  AQ, ΨQ, ΨP, Q', P', C)"])
      (auto simp add: eqvts fresh_star_prod)

  have FrP: "extractFrame P = AP, ΨP" by fact
  have FrQ: "extractFrame Q = AQ, ΨQ" by fact

  from AP ♯* Q FrQ AP ♯* AQ have "AP ♯* ΨQ"
    by(auto dest: extractFrameFreshChain)
  from AQ ♯* P FrP AP ♯* AQ have "AQ ♯* Ψ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 Ψ  ΨQ  P MN  P' Sr distinctPerm r xvec ♯* P (r  xvec) ♯* P
  have "Ψ  ΨQ  P M(r  N)  (r  P')"
    by(rule inputAlpha)
  then have "(q  (Ψ  ΨQ))  P (q  M)(r  N)  (r  P')" using Sq AQ ♯* P (q  AQ) ♯* P
    by - (rule inputPermFrameSubject, (assumption | simp)+)
  then have PTrans: "Ψ  (q  ΨQ)  P (q  M)(r  N)  (r  P')" using Sq AQ ♯* Ψ (q  AQ) ♯* Ψ
    by(simp add: eqvts)

  moreover from extractFrame P = AP, ΨP  Sp (p  AP) ♯* ΨP
  have FrP: "extractFrame P = (p  AP), (p  ΨP)"
    by(simp add: frameChainAlpha)
  moreover from distinct AP have "distinct(p  AP)"  by simp

  moreover from Ψ  ΨP  Q K⦇ν*xvec⦈⟨N  Q' Sr (r  xvec) ♯* N (r  xvec) ♯* Q'
  have "Ψ  ΨP  Q K⦇ν*(r  xvec)⦈⟨(r  N)  (r  Q')"
    by(simp add: boundOutputChainAlpha'' residualInject)
  then have "(p  (Ψ  ΨP))  Q (p  K)⦇ν*(r  xvec)⦈⟨(r  N)  (r  Q')" using Sp AP ♯* Q (p  AP) ♯* Q (r  xvec) ♯* K(p  AP) ♯* (r  xvec) (r  xvec) ♯* AP
    by(fastforce intro: outputPermFrameSubject)
  then have QTrans: "Ψ  (p  ΨP)  Q (p  K)⦇ν*(r  xvec)⦈⟨(r  N)  (r  Q')" using Sp AP ♯* Ψ (p  AP) ♯* Ψ
    by(simp add: eqvts)

  moreover from extractFrame Q = AQ, ΨQ  Sq (q  AQ) ♯* ΨQ
  have FrQ: "extractFrame Q = (q  AQ), (q  ΨQ)"
    by(simp add: frameChainAlpha)
  moreover from distinct AQ have "distinct(q  AQ)"  by simp

  moreover from Ψ  ΨP  ΨQ  M  K have "(p  q  (Ψ  ΨP  ΨQ))  (p  q  M)  (p  q  K)"
    by(metis chanEqClosed)
  with AP ♯* Ψ (p  AP) ♯* Ψ AQ ♯* Ψ (q  AQ) ♯* Ψ AQ ♯* ΨP (q  AQ) ♯* ΨP
    AP ♯* ΨQ (p  AP) ♯* ΨQ AP ♯* M (p  AP) ♯* M (q  AQ) ♯* AP (p  AP) ♯* (q  AQ)
    AQ ♯* K (q  AQ) ♯* K AP ♯* AQ (p  AP) ♯* AQ  Sp Sq
  have "Ψ  (p  ΨP)  (q  ΨQ)  (q  M)  (p  K)"
    by(simp add: eqvts freshChainSimps)
  moreover note (p  AP) ♯* Ψ
  moreover from (p  AP) ♯* AQ (p  AP) ♯* (q  AQ) (p  AP) ♯* ΨQ Sq have "(p  AP) ♯* (q  ΨQ)"
    by(simp add: freshChainSimps)
  moreover note (p  AP) ♯* P
  moreover from (p  AP) ♯* AQ (p  AP) ♯* (q  AQ) (p  AP) ♯* M Sq have "(p  AP) ♯* (q  M)"
    by(simp add: freshChainSimps)
  moreover from (p  AP) ♯* xvec (p  AP) ♯* (r  xvec) (p  AP) ♯* N Sr have "(p  AP) ♯* (r  N)"
    by(simp add: freshChainSimps)
  moreover from (p  AP) ♯* xvec (p  AP) ♯* (r  xvec) (p  AP) ♯* P' Sr have "(p  AP) ♯* (r  P')"
    by(simp add: freshChainSimps)
  moreover note  (p  AP) ♯* Q
  moreover from (p  AP) ♯* xvec (p  AP) ♯* (r  xvec) (p  AP) ♯* Q' Sr have "(p  AP) ♯* (r  Q')"
    by(simp add: freshChainSimps)
  moreover note (p  AP) ♯* (q  AQ) (p  AP) ♯* (r  xvec) (q  AQ) ♯* Ψ
  moreover from (q  AQ) ♯* AP (p  AP) ♯* AQ (q  AQ) ♯* ΨP (p  AP) ♯* (q  AQ) Sp Sq have "(q  AQ) ♯* (p  ΨP)"
    by(simp add: freshChainSimps)
  moreover note (q  AQ) ♯* P
  moreover from (q  AQ) ♯* AP (p  AP) ♯* AQ (q  AQ) ♯* K (p  AP) ♯* (q  AQ) Sp Sq have "(q  AQ) ♯* (p  K)"
    by(simp add: freshChainSimps)
  moreover from (q  AQ) ♯* xvec (q  AQ) ♯* (r  xvec) (q  AQ) ♯* N Sr have "(q  AQ) ♯* (r  N)"
    by(simp add: freshChainSimps)
  moreover from (q  AQ) ♯* xvec (q  AQ) ♯* (r  xvec) (q  AQ) ♯* P' Sr have "(q  AQ) ♯* (r  P')"
    by(simp add: freshChainSimps)
  moreover note (q  AQ) ♯* Q
  moreover from (q  AQ) ♯* xvec (q  AQ) ♯* (r  xvec) (q  AQ) ♯* Q' Sr have "(q  AQ) ♯* (r  Q')"
    by(simp add: freshChainSimps)
  moreover note (q  AQ) ♯* (r  xvec) (r  xvec) ♯* Ψ
  moreover from (r  xvec) ♯* AP (p  AP) ♯* (r  xvec) (r  xvec) ♯* ΨP Sp have "(r  xvec) ♯* (p  ΨP)"
    by(simp add: freshChainSimps)
  moreover from (r  xvec) ♯* AQ (q  AQ) ♯* (r  xvec) (r  xvec) ♯* ΨQ Sq have "(r  xvec) ♯* (q  ΨQ)"
    by(simp add: freshChainSimps)
  moreover note (r  xvec) ♯* P
  moreover from (r  xvec) ♯* AQ (q  AQ) ♯* (r  xvec) (r  xvec) ♯* M Sq have "(r  xvec) ♯* (q  M)"
    by(simp add: freshChainSimps)
  moreover note (r  xvec) ♯* Q
  moreover from (r  xvec) ♯* AP (p  AP) ♯* (r  xvec) (r  xvec) ♯* K Sp have "(r  xvec) ♯* (p  K)"
    by(simp add: freshChainSimps)
  moreover note (p  AP) ♯* C (q  AQ) ♯* C (r  xvec) ♯* C
  moreover from distinct xvec have "distinct(r  xvec)" by simp
  ultimately show ?thesis by(rule rComm1)
next
  case(cComm2 ΨQ P M xvec N P' AP ΨP Q K Q' AQ)
  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) ♯* AP" and "(r  xvec) ♯* AQ"
    and "(r  xvec) ♯* P'" and "(r  xvec) ♯* Q'" and "(r  xvec) ♯* ΨP" and "(r  xvec) ♯* Ψ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, AP, AQ, ΨP, ΨQ, P', Q', C)"])
      (auto simp add: eqvts)

  obtain q::"name prm" where "(q  AQ) ♯* Ψ" and "(q  AQ) ♯* P" and "(q  AQ) ♯* Q" and "(q  AQ) ♯* K"
    and "(q  AQ) ♯* N" and "(q  AQ) ♯* xvec" and "(q  AQ) ♯* Q'" and "(q  AQ) ♯* P'"
    and "(q  AQ) ♯* ΨP" and  "(q  AQ) ♯* AP" and "(q  AQ) ♯* ΨQ" and "(q  AQ) ♯* (r  xvec)"
    and "(q  AQ) ♯* C" and Sq: "(set q)  (set AQ) × (set(q  AQ))"
    by(rule name_list_avoiding[where xvec=AQ and c="(Ψ, P, Q, K, N, xvec, r  xvec, ΨQ, AP, ΨP, Q', P', C)"]) clarsimp

  obtain p::"name prm"  where "(p  AP) ♯* Ψ" and "(p  AP) ♯* P" and "(p  AP) ♯* Q" and "(p  AP) ♯* M"
    and "(p  AP) ♯* N" and "(p  AP) ♯* xvec" and "(p  AP) ♯* Q'" and "(p  AP) ♯* AQ"
    and "(p  AP) ♯* P'" and "(p  AP) ♯* ΨP" and "(p  AP) ♯* ΨQ" and "(p  AP) ♯* (q  AQ)"
    and "(p  AP) ♯* C" and "(p  AP) ♯* (r  xvec)" and Sp: "(set p)  (set AP) × (set(p  AP))"
    by(rule name_list_avoiding[where xvec=AP and c="(Ψ, P, Q, M, N, xvec, r  xvec, AQ, q  AQ, ΨQ, ΨP, Q', P', C)"])
      (auto simp add: eqvts fresh_star_prod)

  have FrP: "extractFrame P = AP, ΨP" by fact
  have FrQ: "extractFrame Q = AQ, ΨQ" by fact

  from AP ♯* Q FrQ AP ♯* AQ have "AP ♯* ΨQ"
    by(auto dest: extractFrameFreshChain)
  from AQ ♯* P FrP AP ♯* AQ have "AQ ♯* Ψ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 Ψ  ΨQ  P M⦇ν*xvec⦈⟨N  P' Sr (r  xvec) ♯* N (r  xvec) ♯* P'
  have "Ψ  ΨQ  P M⦇ν*(r  xvec)⦈⟨(r  N)  (r  P')" by(simp add: boundOutputChainAlpha'' residualInject)
  then have "(q  (Ψ  ΨQ))  P (q  M)⦇ν*(r  xvec)⦈⟨(r  N)  (r  P')" using Sq AQ ♯* P (q  AQ) ♯* P (r  xvec) ♯* M (r  xvec) ♯* AQ (q  AQ) ♯* (r  xvec)
    by(fastforce intro: outputPermFrameSubject)
  then have PTrans: "Ψ  (q  ΨQ)  P (q  M)⦇ν*(r  xvec)⦈⟨(r  N)  (r  P')" using Sq AQ ♯* Ψ (q  AQ) ♯* Ψ
    by(simp add: eqvts)

  moreover from extractFrame P = AP, ΨP  Sp (p  AP) ♯* ΨP
  have FrP: "extractFrame P = (p  AP), (p  ΨP)"
    by(simp add: frameChainAlpha)
  moreover from distinct AP have "distinct(p  AP)"  by simp

  moreover from Ψ  ΨP  Q KN  Q' Sr  distinctPerm r xvec ♯* Q (r  xvec) ♯* Q
  have "Ψ  ΨP  Q K(r  N)  (r  Q')" by(rule inputAlpha)
  then have "(p  (Ψ  ΨP))  Q (p  K)(r  N)  (r  Q')" using Sp AP ♯* Q (p  AP) ♯* Q
    by - (rule inputPermFrameSubject, (assumption | simp)+)
  then have QTrans: "Ψ  (p  ΨP)  Q (p  K)(r  N)  (r  Q')" using Sp AP ♯* Ψ (p  AP) ♯* Ψ
    by(simp add: eqvts)

  moreover from extractFrame Q = AQ, ΨQ  Sq (q  AQ) ♯* ΨQ
  have FrQ: "extractFrame Q = (q  AQ), (q  ΨQ)"
    by(simp add: frameChainAlpha)
  moreover from distinct AQ have "distinct(q  AQ)"  by simp

  moreover from Ψ  ΨP  ΨQ  M  K have "(p  q  (Ψ  ΨP  ΨQ))  (p  q  M)  (p  q  K)"
    by(metis chanEqClosed)
  with AP ♯* Ψ (p  AP) ♯* Ψ AQ ♯* Ψ (q  AQ) ♯* Ψ AQ ♯* ΨP (q  AQ) ♯* ΨP
    AP ♯* ΨQ (p  AP) ♯* ΨQ AP ♯* M (p  AP) ♯* M (q  AQ) ♯* AP (p  AP) ♯* (q  AQ)
    AQ ♯* K (q  AQ) ♯* K AP ♯* AQ (p  AP) ♯* AQ  Sp Sq
  have "Ψ  (p  ΨP)  (q  ΨQ)  (q  M)  (p  K)"
    by(simp add: eqvts freshChainSimps)
  moreover note (p  AP) ♯* Ψ
  moreover from (p  AP) ♯* AQ (p  AP) ♯* (q  AQ) (p  AP) ♯* ΨQ Sq have "(p  AP) ♯* (q  ΨQ)"
    by(simp add: freshChainSimps)
  moreover note (p  AP) ♯* P
  moreover from (p  AP) ♯* AQ (p  AP) ♯* (q  AQ) (p  AP) ♯* M Sq have "(p  AP) ♯* (q  M)"
    by(simp add: freshChainSimps)
  moreover from (p  AP) ♯* xvec (p  AP) ♯* (r  xvec) (p  AP) ♯* N Sr have "(p  AP) ♯* (r  N)"
    by(simp add: freshChainSimps)
  moreover from (p  AP) ♯* xvec (p  AP) ♯* (r  xvec) (p  AP) ♯* P' Sr have "(p  AP) ♯* (r  P')"
    by(simp add: freshChainSimps)
  moreover note  (p  AP) ♯* Q
  moreover from (p  AP) ♯* xvec (p  AP) ♯* (r  xvec) (p  AP) ♯* Q' Sr have "(p  AP) ♯* (r  Q')"
    by(simp add: freshChainSimps)
  moreover note (p  AP) ♯* (q  AQ) (p  AP) ♯* (r  xvec) (q  AQ) ♯* Ψ
  moreover from (q  AQ) ♯* AP (p  AP) ♯* AQ (q  AQ) ♯* ΨP (p  AP) ♯* (q  AQ) Sp Sq have "(q  AQ) ♯* (p  ΨP)"
    by(simp add: freshChainSimps)
  moreover note (q  AQ) ♯* P
  moreover from (q  AQ) ♯* AP (p  AP) ♯* AQ (q  AQ) ♯* K (p  AP) ♯* (q  AQ) Sp Sq have "(q  AQ) ♯* (p  K)"
    by(simp add: freshChainSimps)
  moreover from (q  AQ) ♯* xvec (q  AQ) ♯* (r  xvec) (q  AQ) ♯* N Sr have "(q  AQ) ♯* (r  N)"
    by(simp add: freshChainSimps)
  moreover from (q  AQ) ♯* xvec (q  AQ) ♯* (r  xvec) (q  AQ) ♯* P' Sr have "(q  AQ) ♯* (r  P')"
    by(simp add: freshChainSimps)
  moreover note (q  AQ) ♯* Q
  moreover from (q  AQ) ♯* xvec (q  AQ) ♯* (r  xvec) (q  AQ) ♯* Q' Sr have "(q  AQ) ♯* (r  Q')"
    by(simp add: freshChainSimps)
  moreover note (q  AQ) ♯* (r  xvec) (r  xvec) ♯* Ψ
  moreover from (r  xvec) ♯* AP (p  AP) ♯* (r  xvec) (r  xvec) ♯* ΨP Sp have "(r  xvec) ♯* (p  ΨP)"
    by(simp add: freshChainSimps)
  moreover from (r  xvec) ♯* AQ (q  AQ) ♯* (r  xvec) (r  xvec) ♯* ΨQ Sq have "(r  xvec) ♯* (q  ΨQ)"
    by(simp add: freshChainSimps)
  moreover note (r  xvec) ♯* P
  moreover from (r  xvec) ♯* AQ (q  AQ) ♯* (r  xvec) (r  xvec) ♯* M Sq have "(r  xvec) ♯* (q  M)"
    by(simp add: freshChainSimps)
  moreover note (r  xvec) ♯* Q
  moreover from (r  xvec) ♯* AP (p  AP) ♯* (r  xvec) (r  xvec) ♯* K Sp have "(r  xvec) ♯* (p  K)"
    by(simp add: freshChainSimps)
  moreover note (p  AP) ♯* C (q  AQ) ♯* C (r  xvec) ♯* C
  moreover from distinct xvec have "distinct(r  xvec)" by simp
  ultimately show ?thesis by(rule rComm2)
next
  case(cBrMerge ΨQ P M N P' AP ΨP Q Q' AQ)
  obtain q::"name prm" where "(q  AQ) ♯* Ψ" and "(q  AQ) ♯* P"
    and "(q  AQ) ♯* Q" and "(q  AQ) ♯* M"
    and "(q  AQ) ♯* ΨP" and "(q  AQ) ♯* AP" and "(q  AQ) ♯* ΨQ"
    and "(q  AQ) ♯* N" and "(q  AQ) ♯* P'" and "(q  AQ) ♯* Q'"
    and "(q  AQ) ♯* C"
    and Sq: "set q  set AQ × set(q  AQ)"
    by(rule name_list_avoiding[where c="(Ψ, P, N, M, P', Q', Q, ΨQ, AP, ΨP, C)"])
      (auto simp add: eqvts fresh_star_prod emptyFresh)
  obtain p::"name prm" where "(p  AP) ♯* Ψ" and "(p  AP) ♯* P"
    and "(p  AP) ♯* Q" and "(p  AP) ♯* M"
    and "(p  AP) ♯* ΨP" and "(p  AP) ♯* ΨQ" and "(p  AP) ♯* AQ"
    and "(p  AP) ♯* N" and "(p  AP) ♯* P'" and "(p  AP) ♯* Q'"
    and "(p  AP) ♯* C"
    and Sp: "set p  set AP × set(p  AP)"
    and "(p  AP) ♯* (q  AQ)"
    by(rule name_list_avoiding[where c="(Ψ, P, N, P', Q', Q, M, AQ, q  AQ, ΨQ, ΨP, C)"])
      (auto simp add: eqvts fresh_star_prod)

  have FrP: "extractFrame P = AP, ΨP" by fact
  have FrQ: "extractFrame Q = AQ, ΨQ" by fact

  from AP ♯* Q FrQ AP ♯* AQ have "AP ♯* ΨQ"
    by(auto dest: extractFrameFreshChain)
  from AQ ♯* P FrP AP ♯* AQ have "AQ ♯* ΨP"
    by(auto dest: extractFrameFreshChain)

  from Sp AP ♯* M (p  AP) ♯* M
  have "(p  M) = M"
    by simp
  from Sq AQ ♯* M (q  AQ) ♯* M
  have "(q  M) = M"
    by simp

  note cP = P  Q cRs = ¿MN  (P'  Q')
  moreover from distinct AP have "distinct(p  AP)" by simp
  moreover from distinct AQ have "distinct(q  AQ)"  by simp
  moreover from extractFrame P = AP, ΨP Sp (p  AP) ♯* ΨP
  have FrP: "extractFrame P = (p  AP), (p  ΨP)"
    by(simp add: frameChainAlpha)
  moreover from extractFrame Q = AQ, ΨQ Sq (q  AQ) ♯* ΨQ
  have FrQ: "extractFrame Q = (q  AQ), (q  ΨQ)"
    by(simp add: frameChainAlpha)

  moreover have "(q  (Ψ  ΨQ))  P ¿(q  M)N  P'" using Sq AQ ♯* P (q  AQ) ♯* P Ψ  ΨQ  P ¿MN  P'
    by - (rule brinputPermFrameSubject, (assumption | simp)+)
  then have "(Ψ  (q  ΨQ))  P ¿(q  M)N  P'" using Sq AQ ♯* Ψ (q  AQ) ♯* Ψ
    by(simp add: eqvts)
  with (q  M) = M have PTrans: "(Ψ  (q  ΨQ))  P ¿MN  P'"
    by simp

  moreover have "(p  (Ψ  ΨP))  Q ¿(p  M)N  Q'" using Sp AP ♯* Q (p  AP) ♯* Q Ψ  ΨP  Q ¿MN  Q'
    by - (rule brinputPermFrameSubject, (assumption | simp)+)
  then have "(Ψ  (p  ΨP))  Q ¿(p  M)N  Q'" using Sp AP ♯* Ψ (p  AP) ♯* Ψ
    by(simp add: eqvts)
  with (p  M) = M have PTrans: "(Ψ  (p  ΨP))  Q ¿MN  Q'"
    by simp
  moreover from (p  AP) ♯* AQ (p  AP) ♯* (q  AQ) (p  AP) ♯* ΨQ Sq have "(p  AP) ♯* (q  ΨQ)"
    by(simp add: freshChainSimps)
  moreover from (q  AQ) ♯* AP (p  AP) ♯* AQ (q  AQ) ♯* ΨP (p  AP) ♯* (q  AQ) Sp Sq have "(q  AQ) ♯* (p  ΨP)"
    by(simp add: freshChainSimps)

  moreover note
    (p  AP) ♯* Ψ
    (p  AP) ♯* P (p  AP) ♯* N (p  AP) ♯* P' (p  AP) ♯* M
    (p  AP) ♯* Q (p  AP) ♯* Q' (p  AP) ♯* (q  AQ) (q  AQ) ♯* Ψ
    (q  AQ) ♯* P (q  AQ) ♯* N (q  AQ) ♯* P' (q  AQ) ♯* M
    (q  AQ) ♯* Q (q  AQ) ♯* Q' (p  AP) ♯* C (q  AQ) ♯* C
  ultimately show ?thesis
    by(auto simp add: rBrMerge)
next
  case(cBrComm1 ΨQ P M N P' AP ΨP Q xvec Q' AQ)
  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) ♯* AP" and "(r  xvec) ♯* AQ"
    and "(r  xvec) ♯* P'" and "(r  xvec) ♯* Q'" and "(r  xvec) ♯* ΨP" and "(r  xvec) ♯* Ψ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, AP, AQ, ΨP, ΨQ, P', Q', C)"])
      (auto simp add: eqvts)

  obtain q::"name prm" where "(q  AQ) ♯* Ψ" and "(q  AQ) ♯* P" and "(q  AQ) ♯* Q" and "(q  AQ) ♯* M"
    and "(q  AQ) ♯* N" and "(q  AQ) ♯* xvec" and "(q  AQ) ♯* Q'" and "(q  AQ) ♯* P'"
    and "(q  AQ) ♯* ΨP" and  "(q  AQ) ♯* AP" and "(q  AQ) ♯* ΨQ" and "(q  AQ) ♯* (r  xvec)"
    and "(q  AQ) ♯* C" and Sq: "(set q)  (set AQ) × (set(q  AQ))"
    by(rule name_list_avoiding[where xvec=AQ and c="(Ψ, P, Q, M, N, xvec, r  xvec, ΨQ, AP, ΨP, Q', P', C)"]) clarsimp

  obtain p::"name prm"  where "(p  AP) ♯* Ψ" and "(p  AP) ♯* P" and "(p  AP) ♯* Q" and "(p  AP) ♯* M"
    and "(p  AP) ♯* N" and "(p  AP) ♯* xvec" and "(p  AP) ♯* Q'" and "(p  AP) ♯* AQ"
    and "(p  AP) ♯* P'" and "(p  AP) ♯* ΨP" and "(p  AP) ♯* ΨQ" and "(p  AP) ♯* (q  AQ)"
    and "(p  AP) ♯* C" and "(p  AP) ♯* (r  xvec)" and Sp: "(set p)  (set AP) × (set(p  AP))"
    by(rule name_list_avoiding[where xvec=AP and c="(Ψ, P, Q, M, N, xvec, r  xvec, AQ, q  AQ, ΨQ, ΨP, Q', P', C)"])
      (auto simp add: eqvts fresh_star_prod)

  from Sp AP ♯* M (p  AP) ♯* M
  have "(p  M) = M"
    by simp
  from Sq AQ ♯* M (q  AQ) ♯* M
  have "(q  M) = M"
    by simp
  from Sr xvec ♯* M (r  xvec) ♯* M
  have "(r  M) = M"
    by simp

  have FrP: "extractFrame P = AP, ΨP" by fact
  have FrQ: "extractFrame Q = AQ, ΨQ" by fact

  from AP ♯* Q FrQ AP ♯* AQ have "AP ♯* ΨQ"
    by(auto dest: extractFrameFreshChain)
  from AQ ♯* P FrP AP ♯* AQ have "AQ ♯* Ψ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 Ψ  ΨQ  P ¿MN  P' Sr distinctPerm r xvec ♯* P (r  xvec) ♯* P
  have "Ψ  ΨQ  P ¿M(r  N)  (r  P')"
    by(rule brinputAlpha)
  then have "(q  (Ψ  ΨQ))  P ¿(q  M)(r  N)  (r  P')" using Sq AQ ♯* P (q  AQ) ♯* P
    by - (rule brinputPermFrameSubject, (assumption | simp)+)
  then have PTrans: "Ψ  (q  ΨQ)  P ¿M(r  N)  (r  P')" using Sq AQ ♯* Ψ (q  AQ) ♯* Ψ (q  M) = M
    by(simp add: eqvts)

  moreover from extractFrame P = AP, ΨP  Sp (p  AP) ♯* ΨP
  have FrP: "extractFrame P = (p  AP), (p  ΨP)"
    by(simp add: frameChainAlpha)
  moreover from distinct AP have "distinct(p  AP)"  by simp

  moreover from Ψ  ΨP  Q ¡M⦇ν*xvec⦈⟨N  Q' Sr (r  xvec) ♯* N (r  xvec) ♯* Q'
  have "Ψ  ΨP  Q ¡M⦇ν*(r  xvec)⦈⟨(r  N)  (r  Q')"
    by(simp add: boundOutputChainAlpha'' residualInject)
  then have "(p  (Ψ  ΨP))  Q ¡(p  M)⦇ν*(r  xvec)⦈⟨(r  N)  (r  Q')" using Sp AP ♯* Q (p  AP) ♯* Q (r  xvec) ♯* M (p  AP) ♯* (r  xvec) (r  xvec) ♯* AP
    by(fastforce intro: broutputPermFrameSubject)
  then have QTrans: "Ψ  (p  ΨP)  Q ¡M⦇ν*(r  xvec)⦈⟨(r  N)  (r  Q')" using Sp AP ♯* Ψ (p  AP) ♯* Ψ (p  M) = M
    by(simp add: eqvts)

  moreover from extractFrame Q = AQ, ΨQ  Sq (q  AQ) ♯* ΨQ
  have FrQ: "extractFrame Q = (q  AQ), (q  ΨQ)"
    by(simp add: frameChainAlpha)
  moreover from distinct AQ have "distinct(q  AQ)"  by simp

  moreover note (p  AP) ♯* Ψ
  moreover from (p  AP) ♯* AQ (p  AP) ♯* (q  AQ) (p  AP) ♯* ΨQ Sq have "(p  AP) ♯* (q  ΨQ)"
    by(simp add: freshChainSimps)
  moreover note (p  AP) ♯* P
  moreover from (p  AP) ♯* xvec (p  AP) ♯* (r  xvec) (p  AP) ♯* N Sr have "(p  AP) ♯* (r  N)"
    by(simp add: freshChainSimps)
  moreover from (p  AP) ♯* xvec (p  AP) ♯* (r  xvec) (p  AP) ♯* P' Sr have "(p  AP) ♯* (r  P')"
    by(simp add: freshChainSimps)
  moreover note  (p  AP) ♯* Q
  moreover from (p  AP) ♯* xvec (p  AP) ♯* (r  xvec) (p  AP) ♯* Q' Sr have "(p  AP) ♯* (r  Q')"
    by(simp add: freshChainSimps)
  moreover note (p  AP) ♯* (q  AQ) (p  AP) ♯* (r  xvec) (q  AQ) ♯* Ψ
  moreover from (q  AQ) ♯* AP (p  AP) ♯* AQ (q  AQ) ♯* ΨP (p  AP) ♯* (q  AQ) Sp Sq have "(q  AQ) ♯* (p  ΨP)"
    by(simp add: freshChainSimps)
  moreover note (q  AQ) ♯* P
  moreover from (q  AQ) ♯* xvec (q  AQ) ♯* (r  xvec) (q  AQ) ♯* N Sr have "(q  AQ) ♯* (r  N)"
    by(simp add: freshChainSimps)
  moreover from (q  AQ) ♯* xvec (q  AQ) ♯* (r  xvec) (q  AQ) ♯* P' Sr have "(q  AQ) ♯* (r  P')"
    by(simp add: freshChainSimps)
  moreover note (q  AQ) ♯* Q
  moreover from (q  AQ) ♯* xvec (q  AQ) ♯* (r  xvec) (q  AQ) ♯* Q' Sr have "(q  AQ) ♯* (r  Q')"
    by(simp add: freshChainSimps)
  moreover note (q  AQ) ♯* (r  xvec) (r  xvec) ♯* Ψ
  moreover from (r  xvec) ♯* AP (p  AP) ♯* (r  xvec) (r  xvec) ♯* ΨP Sp have "(r  xvec) ♯* (p  ΨP)"
    by(simp add: freshChainSimps)
  moreover from (r  xvec) ♯* AQ (q  AQ) ♯* (r  xvec) (r  xvec) ♯* ΨQ Sq have "(r  xvec) ♯* (q  ΨQ)"
    by(simp add: freshChainSimps)
  moreover note (p  AP) ♯* M (q  AQ) ♯* M
  moreover note (r  xvec) ♯* P
  moreover note (r  xvec) ♯* Q (r  xvec) ♯* M
  moreover note (p  AP) ♯* C (q  AQ) ♯* C (r  xvec) ♯* C
  moreover from distinct xvec have "distinct(r  xvec)" by simp
  ultimately show ?thesis by(simp add: rBrComm1)
next
  case(cBrComm2 ΨQ P M xvec N P' AP ΨP Q Q' AQ)
  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) ♯* AP" and "(r  xvec) ♯* AQ"
    and "(r  xvec) ♯* P'" and "(r  xvec) ♯* Q'" and "(r  xvec) ♯* ΨP" and "(r  xvec) ♯* Ψ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, AP, AQ, ΨP, ΨQ, P', Q', C)"])
      (auto simp add: eqvts)

  obtain q::"name prm" where "(q  AQ) ♯* Ψ" and "(q  AQ) ♯* P" and "(q  AQ) ♯* Q" and "(q  AQ) ♯* M"
    and "(q  AQ) ♯* N" and "(q  AQ) ♯* xvec" and "(q  AQ) ♯* Q'" and "(q  AQ) ♯* P'"
    and "(q  AQ) ♯* ΨP" and  "(q  AQ) ♯* AP" and "(q  AQ) ♯* ΨQ" and "(q  AQ) ♯* (r  xvec)"
    and "(q  AQ) ♯* C" and Sq: "(set q)  (set AQ) × (set(q  AQ))"
    by(rule name_list_avoiding[where xvec=AQ and c="(Ψ, P, Q, N, M, xvec, r  xvec, ΨQ, AP, ΨP, Q', P', C)"]) clarsimp

  obtain p::"name prm"  where "(p  AP) ♯* Ψ" and "(p  AP) ♯* P" and "(p  AP) ♯* Q" and "(p  AP) ♯* M"
    and "(p  AP) ♯* N" and "(p  AP) ♯* xvec" and "(p  AP) ♯* Q'" and "(p  AP) ♯* AQ"
    and "(p  AP) ♯* P'" and "(p  AP) ♯* ΨP" and "(p  AP) ♯* ΨQ" and "(p  AP) ♯* (q  AQ)"
    and "(p  AP) ♯* C" and "(p  AP) ♯* (r  xvec)" and Sp: "(set p)  (set AP) × (set(p  AP))"
    by(rule name_list_avoiding[where xvec=AP and c="(Ψ, P, Q, M, N, xvec, r  xvec, AQ, q  AQ, ΨQ, ΨP, Q', P', C)"])
      (auto simp add: eqvts fresh_star_prod)

  from Sp AP ♯* M (p  AP) ♯* M
  have "(p  M) = M"
    by simp
  from Sq AQ ♯* M (q  AQ) ♯* M
  have "(q  M) = M"
    by simp
  from Sr xvec ♯* M (r  xvec) ♯* M
  have "(r  M) = M"
    by simp

  have FrP: "extractFrame P = AP, ΨP" by fact
  have FrQ: "extractFrame Q = AQ, ΨQ" by fact

  from AP ♯* Q FrQ AP ♯* AQ have "AP ♯* ΨQ"
    by(auto dest: extractFrameFreshChain)
  from AQ ♯* P FrP AP ♯* AQ have "AQ ♯* Ψ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 Ψ  ΨQ  P ¡M⦇ν*xvec⦈⟨N  P' Sr (r  xvec) ♯* N (r  xvec) ♯* P'
  have "Ψ  ΨQ  P ¡M⦇ν*(r  xvec)⦈⟨(r  N)  (r  P')" by(simp add: boundOutputChainAlpha'' residualInject)
  then have "(q  (Ψ  ΨQ))  P ¡(q  M)⦇ν*(r  xvec)⦈⟨(r  N)  (r  P')" using Sq AQ ♯* P (q  AQ) ♯* P (r  xvec) ♯* M (r  xvec) ♯* AQ (q  AQ) ♯* (r  xvec)
    by(fastforce intro: broutputPermFrameSubject)
  then have PTrans: "Ψ  (q  ΨQ)  P ¡M⦇ν*(r  xvec)⦈⟨(r  N)  (r  P')" using Sq AQ ♯* Ψ (q  AQ) ♯* Ψ (q  M) = M
    by(simp add: eqvts)

  moreover from extractFrame P = AP, ΨP  Sp (p  AP) ♯* ΨP
  have FrP: "extractFrame P = (p  AP), (p  ΨP)"
    by(simp add: frameChainAlpha)
  moreover from distinct AP have "distinct(p  AP)"  by simp

  moreover from Ψ  ΨP  Q ¿MN  Q' Sr  distinctPerm r xvec ♯* Q (r  xvec) ♯* Q
  have "Ψ  ΨP  Q ¿M(r  N)  (r  Q')" by(rule brinputAlpha)
  then have "(p  (Ψ  ΨP))  Q ¿(p  M)(r  N)  (r  Q')" using Sp AP ♯* Q (p  AP) ♯* Q
    by - (rule brinputPermFrameSubject, (assumption | simp)+)
  then have QTrans: "Ψ  (p  ΨP)  Q ¿M(r  N)  (r  Q')" using Sp AP ♯* Ψ (p  AP) ♯* Ψ (p  M) = M
    by(simp add: eqvts)

  moreover from extractFrame Q = AQ, ΨQ  Sq (q  AQ) ♯* ΨQ
  have FrQ: "extractFrame Q = (q  AQ), (q  ΨQ)"
    by(simp add: frameChainAlpha)
  moreover from distinct AQ have "distinct(q  AQ)"  by simp

  moreover note (p  AP) ♯* Ψ
  moreover from (p  AP) ♯* AQ (p  AP) ♯* (q  AQ) (p  AP) ♯* ΨQ Sq have "(p  AP) ♯* (q  ΨQ)"
    by(simp add: freshChainSimps)
  moreover note (p  AP) ♯* P
  moreover from (p  AP) ♯* xvec (p  AP) ♯* (r  xvec) (p  AP) ♯* N Sr have "(p  AP) ♯* (r  N)"
    by(simp add: freshChainSimps)
  moreover from (p  AP) ♯* xvec (p  AP) ♯* (r  xvec) (p  AP) ♯* P' Sr have "(p  AP) ♯* (r  P')"
    by(simp add: freshChainSimps)
  moreover note  (p  AP) ♯* Q
  moreover from (p  AP) ♯* xvec (p  AP) ♯* (r  xvec) (p  AP) ♯* Q' Sr have "(p  AP) ♯* (r  Q')"
    by(simp add: freshChainSimps)
  moreover note (p  AP) ♯* (q  AQ) (p  AP) ♯* (r  xvec) (q  AQ) ♯* Ψ
  moreover from (q  AQ) ♯* AP (p  AP) ♯* AQ (q  AQ) ♯* ΨP (p  AP) ♯* (q  AQ) Sp Sq have "(q  AQ) ♯* (p  ΨP)"
    by(simp add: freshChainSimps)
  moreover note (q  AQ) ♯* P
  moreover from (q  AQ) ♯* xvec (q  AQ) ♯* (r  xvec) (q  AQ) ♯* N Sr have "(q  AQ) ♯* (r  N)"
    by(simp add: freshChainSimps)
  moreover from (q  AQ) ♯* xvec (q  AQ) ♯* (r  xvec) (q  AQ) ♯* P' Sr have "(q  AQ) ♯* (r  P')"
    by(simp add: freshChainSimps)
  moreover note (q  AQ) ♯* Q
  moreover from (q  AQ) ♯* xvec (q  AQ) ♯* (r  xvec) (q  AQ) ♯* Q' Sr have "(q  AQ) ♯* (r  Q')"
    by(simp add: freshChainSimps)
  moreover note (q  AQ) ♯* (r  xvec) (r  xvec) ♯* Ψ
  moreover from (r  xvec) ♯* AP (p  AP) ♯* (r  xvec) (r  xvec) ♯* ΨP Sp have "(r  xvec) ♯* (p  ΨP)"
    by(simp add: freshChainSimps)
  moreover from (r  xvec) ♯* AQ (q  AQ) ♯* (r  xvec) (r  xvec) ♯* ΨQ Sq have "(r  xvec) ♯* (q  ΨQ)"
    by(simp add: freshChainSimps)
  moreover note (p  AP) ♯* M (q  AQ) ♯* M
  moreover note (r  xvec) ♯* P
  moreover note (r  xvec) ♯* Q (r  xvec) ♯* M
  moreover note (p  AP) ♯* C (q  AQ) ♯* 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 = ⦇νxP y  P have cP_perm: "cP = ⦇νy([(x, y)]  P)" by(simp add: alphaRes)
  from cRs = τ  ⦇νx(⦇ν*xvecP') (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 = ⦇νxP 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(⦇ν*yvecN ≺' 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 = ⦇νxP 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(⦇ν*yvecN ≺' 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 = ⦇νxP y  P have "cP = ⦇νy([(x, y)]  P)" by(simp add: alphaRes)
  moreover from cRs = α  ⦇νxP' 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 (MN⟩.P) = 0"
| "inputLength (MI) = inputLength' I"
| "inputLength (Case C) = 0"
| "inputLength (P  Q) = 0"
| "inputLength (⦇νxP) = 0"
| "inputLength (Ψ) = 0"
| "inputLength (!P) = 0"

| "inputLength' (Trm M P) = 0"
| "inputLength' (ν y I) = 1 + (inputLength' I)"

| "inputLength'' (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 = MN⟩.P; cRs = KN  P; Ψ  M  K  Prop"
  and   rBrOutput: "M K N P. cP = MN⟩.P; cRs = ¡KN  P; Ψ  M  K  Prop"
  and   rCase: "Cs P φ. cP = Cases Cs; Ψ  P cRs; (φ, P)  set Cs; Ψ  φ; guarded P  Prop"
  and   rPar1: "ΨQ P α P' Q AQ. (xvec2 ♯* Ψ; xvec2 ♯* cP; xvec2 ♯* cRs 
                                         cP = P  Q  cRs = α  (P'  Q)  xvec2 = bn α 
                                          Ψ  ΨQ  P α  P'  extractFrame Q = AQ, ΨQ  distinct AQ 
                                          AQ ♯* P  AQ ♯* Q  AQ ♯* Ψ  AQ ♯* α  AQ ♯* P'  AQ ♯* C)  Prop"
  and   rPar2: "ΨP Q α Q' P AP. (xvec3 ♯* Ψ; xvec3 ♯* cP; xvec3 ♯* cRs 
                                          cP = P  Q  cRs = α  (P  Q')  xvec3 = bn α 
                                          Ψ  ΨP  Q α  Q'  extractFrame P = AP, ΨP  distinct AP 
                                          AP ♯* P  AP ♯* Q  AP ♯* Ψ  AP ♯* α  AP ♯* Q'  AP ♯* C)  Prop"
  and   rComm1: "ΨQ P M N P' AP ΨP Q K xvec Q' AQ.
                   cP = P  Q; cRs = τ  ⦇ν*xvecP'  Q';
                    Ψ  ΨQ  P MN  P'; extractFrame P = AP, ΨP; distinct AP;
                    Ψ  ΨP  Q  K⦇ν*xvec⦈⟨N  Q'; extractFrame Q = AQ, ΨQ; distinct AQ;
                    Ψ  ΨP  ΨQ  M  K; AP ♯* Ψ; AP ♯* ΨQ; AP ♯* P; AP ♯* M; AP ♯* N;
                    AP ♯* P'; AP ♯* Q; AP ♯* Q'; AP ♯* AQ; AP ♯* xvec; AQ ♯* Ψ; AQ ♯* ΨP;
                    AQ ♯* P; AQ ♯* K; AQ ♯* N; AQ ♯* P'; AQ ♯* Q; AQ ♯* Q'; AQ ♯* xvec;
                    xvec ♯* Ψ; xvec ♯* ΨP; xvec ♯* ΨQ; xvec ♯* P; xvec ♯* M; xvec ♯* Q;
                    xvec ♯* K; AP ♯* C; AQ ♯* C; xvec ♯* C; distinct xvec  Prop"
  and   rComm2: "ΨQ P M xvec N P' AP ΨP Q K Q' AQ.
                   cP = P  Q; cRs = τ  ⦇ν*xvecP'  Q';
                    Ψ  ΨQ  P M⦇ν*xvec⦈⟨N  P'; extractFrame P = AP, ΨP; distinct AP;
                    Ψ  ΨP  Q  KN  Q'; extractFrame Q = AQ, ΨQ; distinct AQ;
                    Ψ  ΨP  ΨQ  M  K; AP ♯* Ψ; AP ♯* ΨQ; AP ♯* P; AP ♯* M; AP ♯* N;
                    AP ♯* P'; AP ♯* Q; AP ♯* Q'; AP ♯* AQ; AP ♯* xvec; AQ ♯* Ψ; AQ ♯* ΨP;
                    AQ ♯* P; AQ ♯* K; AQ ♯* N; AQ ♯* P'; AQ ♯* Q; AQ ♯* Q'; AQ ♯* xvec;
                    xvec ♯* Ψ; xvec ♯* ΨP; xvec ♯* ΨQ; xvec ♯* P; xvec ♯* M; xvec ♯* Q;
                    xvec ♯* K; AP ♯* C; AQ ♯* C; xvec ♯* C; distinct xvec  Prop"
  and   rBrMerge: "ΨQ P M N P' AP ΨP Q Q' AQ.
                    cP = (P  Q); cRs = ¿MN  (P'  Q');
                    Ψ  ΨQ  P  ¿MN  P'; extractFrame P = AP, ΨP; distinct AP;
                    Ψ  ΨP  Q  ¿MN  Q'; extractFrame Q = AQ, ΨQ; distinct AQ;
                    AP ♯* Ψ; AP ♯* ΨQ; AP ♯* P; AP ♯* N; AP ♯* P';
                    AP ♯* Q; AP ♯* Q'; AP ♯* AQ; AP ♯* M; AQ ♯* M;
                    AQ ♯* Ψ; AQ ♯* ΨP; AQ ♯* P; AQ ♯* N; AQ ♯* P';
                    AQ ♯* Q; AQ ♯* Q'; AP ♯* C; AQ ♯* C  Prop"
  and   rBrComm1: "ΨQ P M N P' AP ΨP Q Q' AQ.
                   (xvec7 ♯* Ψ; xvec7 ♯* cP; xvec7 ♯* cRs 
                    cP = P  Q  cRs = ¡M⦇ν*xvec7⦈⟨N  (P'  Q') 
                    Ψ  ΨQ  P ¿MN  P'  extractFrame P = AP, ΨP  distinct AP 
                    Ψ  ΨP  Q  ¡M⦇ν*xvec7⦈⟨N  Q'  extractFrame Q = AQ, ΨQ  distinct AQ 
                    AP ♯* Ψ  AP ♯* ΨQ  AP ♯* P  AP ♯* N 
                    AP ♯* P' AP ♯* Q  AP ♯* Q'  AP ♯* AQ  AP ♯* xvec7  AQ ♯* Ψ  AQ ♯* ΨP 
                    AQ ♯* P  AQ ♯* N  AQ ♯* P'  AQ ♯* Q  AQ ♯* Q'  AQ ♯* xvec7 
                    xvec7 ♯* Ψ  xvec7 ♯* ΨP  xvec7 ♯* ΨQ  xvec7 ♯* P  xvec7 ♯* Q 
                    AP ♯* M  AQ ♯* M  xvec7 ♯* M 
                    AP ♯* C  AQ ♯* C  distinct xvec7)  Prop"
  and   rBrComm2: "ΨQ P M N P' AP ΨP Q Q' AQ.
                   (xvec8 ♯* Ψ; xvec8 ♯* cP; xvec8 ♯* cRs 
                    cP = P  Q  cRs = ¡M⦇ν*xvec8⦈⟨N  (P'  Q') 
                    Ψ  ΨQ  P ¡M⦇ν*xvec8⦈⟨N  P'  extractFrame P = AP, ΨP  distinct AP 
                    Ψ  ΨP  Q  ¿MN  Q'  extractFrame Q = AQ, ΨQ  distinct AQ 
                    AP ♯* Ψ  AP ♯* ΨQ  AP ♯* P  AP ♯* N 
                    AP ♯* P'  AP ♯* Q  AP ♯* Q'  AP ♯* AQ  AP ♯* xvec8  AQ ♯* Ψ  AQ ♯* ΨP 
                    AQ ♯* P  AQ ♯* N  AQ ♯* P'  AQ ♯* Q  AQ ♯* Q'  AQ ♯* xvec8 
                    xvec8 ♯* Ψ  xvec8 ♯* ΨP  xvec8 ♯* ΨQ  xvec8 ♯* P  xvec8 ♯* Q 
                    AP ♯* M  AQ ♯* M  xvec8 ♯* M 
                    AP ♯* C  AQ ♯* C  distinct xvec8)  Prop"
  and   rBrClose: "P M N xvec P'.
                   (x3  Ψ; x3  cP; x3  cRs 
                    cP = (⦇νx3P)  cRs = τ  ⦇νx3(⦇ν*xvecP') 
                    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 = ⦇νx1P  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 = ⦇νx4P  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 = ⦇νx2P  cRs = α  ⦇νx2P'   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 ΨQ P α P' Q AQ)
  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 AQ ♯* (xvec1, xvec2, xvec3, xvec4, xvec5, xvec6, xvec7, xvec8, xvec9, x1, x2, x3, x4, cP, cRs, C) have "AQ ♯* xvec2" and "AQ ♯* 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 AQ=AQ and ΨQ=Ψ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 Ψ  ΨQ  P α  P' S B C S bn α ♯* xvec2 xvec2 ♯* cRs xvec2 = bn(p  α) bn α ♯* subject α xvec2 ♯* cP
    have "Ψ  ΨQ  P (p  α)  (p  P')"
      by(subst residualAlpha[symmetric]) auto
    moreover note extractFrame Q = AQ, ΨQ distinct AQ AQ ♯* P AQ ♯* Q AQ ♯* Ψ AQ ♯* α
    moreover from AQ ♯* α AQ ♯* xvec2 S xvec2 = bn(p  α) distinctPerm p have "AQ ♯* (p  α)"
      by(subst fresh_star_bij[symmetric, where pi=p]) simp
    moreover from AQ ♯* P' AQ ♯* α AQ ♯* xvec2 S xvec2 = bn(p  α) distinctPerm p have "AQ ♯* (p  P')"
      by(subst fresh_star_bij[symmetric, where pi=p]) simp
    moreover note AQ ♯* C
    ultimately show "cP = P  Q  cRs = (p  α)  (p  P')  Q  xvec2 = bn (p  α) 
      Ψ  ΨQ  P  (p  α)  p  P'  extractFrame Q = AQ, ΨQ  distinct AQ  AQ ♯* P 
      AQ ♯* Q  AQ ♯* Ψ  AQ ♯* (p  α)  AQ ♯* (p  P')  AQ ♯* C" by blast
  qed
next
  case(cPar2 ΨP Q α Q' P AP)
  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 AP ♯* (xvec1, xvec2, xvec3, xvec4, xvec5, xvec6, xvec7, xvec8, xvec9, x1, x2, x3, x4, cP, cRs, C) have "AP ♯* xvec3" and "AP ♯* 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 AP=AP and ΨP=Ψ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 Ψ  ΨP  Q α  Q' S B C S bn α ♯* xvec3 xvec3 ♯* cRs xvec3 = bn(p  α) bn α ♯* subject α xvec3 ♯* cP
    have "Ψ  ΨP  Q (p  α)  (p  Q')"
      by(subst residualAlpha[symmetric]) auto
    moreover note extractFrame P = AP, ΨP distinct AP AP ♯* P AP ♯* Q AP ♯* Ψ AP ♯* α
    moreover from AP ♯* α AP ♯* xvec3 S xvec3 = bn(p  α) distinctPerm p have "AP ♯* (p  α)"
      by(subst fresh_star_bij[symmetric, where pi=p]) simp
    moreover from AP ♯* Q' AP ♯* α AP ♯* xvec3 S xvec3 = bn(p  α) distinctPerm p have "AP ♯* (p  Q')"
      by(subst fresh_star_bij[symmetric, where pi=p]) simp
    moreover note AP ♯* C
    ultimately show "cP = P  Q  cRs = (p  α)  P  (p  Q')  xvec3 = bn (p  α) 
      Ψ  ΨP  Q  (p  α)  p  Q' 
      extractFrame P = AP, ΨP  distinct AP  AP ♯* P  AP ♯* Q  AP ♯* Ψ  AP ♯* (p  α) 
      AP ♯* (p  Q')  AP ♯* C" by blast
  qed
next
  case(cComm1 ΨQ P M N P' AP ΨP Q K xvec Q' AQ)
  then show ?thesis by - (rule rComm1[where P=P and Q=Q], (assumption | simp)+)
next
  case(cComm2 ΨQ P M xvec N P' AP ΨP Q K Q' AQ)
  then show ?thesis by - (rule rComm2[where P=P and Q=Q], (assumption | simp)+)
next
  case(cBrMerge ΨQ P M N P' AP ΨP Q Q' AQ)
  then show ?thesis by - (rule rBrMerge[where P=P and Q=Q], (assumption | simp)+)
next
  case(cBrComm1 ΨQ P M N P' AP ΨP Q xvec Q' AQ)
  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 AP ♯* (xvec1, xvec2, xvec3, xvec4, xvec5, xvec6, xvec7, xvec8, xvec9, x1, x2, x3, x4, cP, cRs, C) have "AP ♯* xvec7" and "AP ♯* C" by simp+
  from AQ ♯* (xvec1, xvec2, xvec3, xvec4, xvec5, xvec6, xvec7, xvec8, xvec9, x1, x2, x3, x4, cP, cRs, C) have "AQ ♯* xvec7" and "AQ ♯* 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 AP=AP and ΨP=ΨP and AQ=AQ and ΨQ=ΨQ and M=M])
    assume "xvec7 ♯* Ψ" and "xvec7 ♯* cP" and "xvec7 ♯* cRs"
    from AQ ♯* xvec7 xvec7 = p  xvec
    have "AQ ♯* (p  xvec)" by simp
    from AP ♯* xvec7 xvec7 = p  xvec
    have "AP ♯* (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 Ψ  ΨQ  P  ¿MN  P' S B distinctPerm p xvec ♯* xvec7 xvec7 = p  xvec xvec ♯* P xvec7 ♯* cP
    have "Ψ  ΨQ  P  ¿M(p  N)  p  P'"
      by(simp add: brinputAlpha)

    moreover from Ψ  ΨP  Q  ¡M⦇ν*xvec⦈⟨N  Q' S B C xvec ♯* xvec7 xvec7 ♯* cRs xvec7 = p  xvec xvec ♯* M xvec7 ♯* cP xvec7 ♯* M
    have "Ψ  ΨP  Q  ¡M⦇ν*xvec7⦈⟨(p  N)  p  Q'"
      by(auto simp add: residualAlpha)

    moreover note
      extractFrame P = AP, ΨP distinct AP extractFrame Q = AQ, ΨQ distinct AQ
      AP ♯* Ψ AP ♯* ΨQ AP ♯* P AP ♯* Q AP ♯* AQ AP ♯* xvec7
      AQ ♯* Ψ AQ ♯* ΨP AQ ♯* P AQ ♯* Q AQ ♯* xvec7

    moreover from S AP ♯* xvec AP ♯* (p  xvec) AP ♯* N
    have "AP ♯* (p  N)"
      by(simp add: freshChainSimps)
    moreover from S AP ♯* xvec AP ♯* (p  xvec) AP ♯* P'
    have "AP ♯* (p  P')"
      by(simp add: freshChainSimps)
    moreover from S AP ♯* xvec AP ♯* (p  xvec) AP ♯* Q'
    have "AP ♯* (p  Q')"
      by(simp add: freshChainSimps)
    moreover from S AQ ♯* xvec AQ ♯* (p  xvec) AQ ♯* N
    have "AQ ♯* (p  N)"
      by(simp add: freshChainSimps)
    moreover from S AQ ♯* xvec AQ ♯* (p  xvec) AQ ♯* P'
    have "AQ ♯* (p  P')"
      by(simp add: freshChainSimps)
    moreover from S AQ ♯* xvec AQ ♯* (p  xvec) AQ ♯* Q'
    have "AQ ♯* (p  Q')"
      by(simp add: freshChainSimps)

    moreover note xvec7 ♯* Ψ

    moreover from xvec7 ♯* cP cP = P  Q extractFrame P = AP, ΨP AP ♯* xvec7
    have "xvec7 ♯* ΨP"
      by simp (metis extractFrameFreshChain freshFrameDest)

    moreover from xvec7 ♯* cP cP = P  Q extractFrame Q = AQ, ΨQ AQ ♯* xvec7
    have "xvec7 ♯* Ψ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 AP ♯* M AQ ♯* M

    moreover note xvec7 ♯* M

    moreover note AP ♯* C AQ ♯* C distinct xvec7

    ultimately show "cP = P  Q  cRs = ¡M⦇ν*xvec7⦈⟨(p  N)  (p  P')  (p  Q') 
      Ψ  ΨQ  P  ¿M(p  N)  p  P'  extractFrame P = AP, ΨP  distinct AP 
      Ψ  ΨP  Q  ¡M⦇ν*xvec7⦈⟨(p  N)  p  Q'  extractFrame Q = AQ, ΨQ  distinct AQ 
      AP ♯* Ψ  AP ♯* ΨQ  AP ♯* P  AP ♯* (p  N)  AP ♯* (p  P')  AP ♯* Q  AP ♯* (p  Q') 
      AP ♯* AQ  AP ♯* xvec7  AQ ♯* Ψ  AQ ♯* ΨP  AQ ♯* P  AQ ♯* (p  N)  AQ ♯* (p  P') 
      AQ ♯* Q  AQ ♯* (p  Q')  AQ ♯* xvec7  xvec7 ♯* Ψ  xvec7 ♯* ΨP  xvec7 ♯* ΨQ  xvec7 ♯* P 
      xvec7 ♯* Q  AP ♯* M  AQ ♯* M  xvec7 ♯* M  AP ♯* C  AQ ♯* C  distinct xvec7" by blast
  qed
next
  case(cBrComm2 ΨQ P M xvec N P' AP ΨP Q Q' AQ)
  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 AP ♯* (xvec1, xvec2, xvec3, xvec4, xvec5, xvec6, xvec7, xvec8, xvec9, x1, x2, x3, x4, cP, cRs, C) have "AP ♯* xvec8" and "AP ♯* C" by simp+
  from AQ ♯* (xvec1, xvec2, xvec3, xvec4, xvec5, xvec6, xvec7, xvec8, xvec9, x1, x2, x3, x4, cP, cRs, C) have "AQ ♯* xvec8" and "AQ ♯* 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 AP=AP and ΨP=ΨP and AQ=AQ and ΨQ=ΨQ and M=M])
    assume "xvec8 ♯* Ψ" and "xvec8 ♯* cP" and "xvec8 ♯* cRs"
    from AQ ♯* xvec8 xvec8 = p  xvec
    have "AQ ♯* (p  xvec)" by simp
    from AP ♯* xvec8 xvec8 = p  xvec
    have "AP ♯* (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 Ψ  ΨP  Q  ¿MN  Q' S B distinctPerm p xvec ♯* xvec8 xvec8 = p  xvec xvec ♯* Q xvec8 ♯* cP
    have "Ψ  ΨP  Q  ¿M(p  N)  p  Q'"
      by(simp add: brinputAlpha)

    moreover from Ψ  ΨQ  P  ¡M⦇ν*xvec⦈⟨N  P' S B C xvec ♯* xvec8 xvec8 ♯* cRs xvec8 = p  xvec xvec ♯* M xvec8 ♯* cP xvec8 ♯* M
    have "Ψ  ΨQ  P  ¡M⦇ν*xvec8⦈⟨(p  N)  p  P'"
      by(auto simp add: residualAlpha)

    moreover note
      extractFrame P = AP, ΨP distinct AP extractFrame Q = AQ, ΨQ distinct AQ
      AP ♯* Ψ AP ♯* ΨQ AP ♯* P AP ♯* Q AP ♯* AQ AP ♯* xvec8
      AQ ♯* Ψ AQ ♯* ΨP AQ ♯* P AQ ♯* Q AQ ♯* xvec8

    moreover from S AP ♯* xvec AP ♯* (p  xvec) AP ♯* N
    have "AP ♯* (p  N)"
      by(simp add: freshChainSimps)
    moreover from S AP ♯* xvec AP ♯* (p  xvec) AP ♯* P'
    have "AP ♯* (p  P')"
      by(simp add: freshChainSimps)
    moreover from S AP ♯* xvec AP ♯* (p  xvec) AP ♯* Q'
    have "AP ♯* (p  Q')"
      by(simp add: freshChainSimps)
    moreover from S AQ ♯* xvec AQ ♯* (p  xvec) AQ ♯* N
    have "AQ ♯* (p  N)"
      by(simp add: freshChainSimps)
    moreover from S AQ ♯* xvec AQ ♯* (p  xvec) AQ ♯* P'
    have "AQ ♯* (p  P')"
      by(simp add: freshChainSimps)
    moreover from S AQ ♯* xvec AQ ♯* (p  xvec) AQ ♯* Q'
    have "AQ ♯* (p  Q')"
      by(simp add: freshChainSimps)

    moreover note xvec8 ♯* Ψ

    moreover from xvec8 ♯* cP cP = P  Q extractFrame P = AP, ΨP AP ♯* xvec8
    have "xvec8 ♯* ΨP"
      by simp (metis extractFrameFreshChain freshFrameDest)

    moreover from xvec8 ♯* cP cP = P  Q extractFrame Q = AQ, ΨQ AQ ♯* xvec8
    have "xvec8 ♯* Ψ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 AP ♯* M AQ ♯* M

    moreover note xvec8 ♯* M

    moreover note AP ♯* C AQ ♯* C distinct xvec8

    ultimately show "cP = P  Q  cRs = ¡M⦇ν*xvec8⦈⟨(p  N)  (p  P')  (p  Q') 
      Ψ  ΨQ  P  ¡M⦇ν*xvec8⦈⟨(p  N)  p  P'  extractFrame P = AP, ΨP  distinct AP 
      Ψ  ΨP  Q  ¿M(p  N)  p  Q'  extractFrame Q = AQ, ΨQ  distinct AQ 
      AP ♯* Ψ  AP ♯* ΨQ  AP ♯* P  AP ♯* (p  N)  AP ♯* (p  P')  AP ♯* Q  AP ♯* (p  Q') 
      AP ♯* AQ  AP ♯* xvec8  AQ ♯* Ψ  AQ ♯* ΨP  AQ ♯* P  AQ ♯* (p  N)  AQ ♯* (p  P') 
      AQ ♯* Q  AQ ♯* (p  Q')  AQ ♯* xvec8  xvec8 ♯* Ψ  xvec8 ♯* ΨP  xvec8 ♯* ΨQ  xvec8 ♯* P 
      xvec8 ♯* Q  AP ♯* M  AQ ♯* M  xvec8 ♯* M  AP ♯* C  AQ ♯* C  distinct xvec8" by blast
  qed
next
  case(cBrClose P M xvec N P' x)
  have B: "cP = ⦇νxP" and C: "cRs = τ  ⦇νx(⦇ν*xvecP')"
    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(⦇ν*xvecP'))" by simp

    then have "x3  (⦇νx(⦇ν*xvecP'))" by simp
    with x  x3 have "x3  (⦇ν*xvecP')" 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 = ⦇νxP 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 = ⦇νxP" 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 = ⦇νxP x  cP x  x1 have "([(x, x1)]  cP) = [(x, x1)]  ⦇νxP"
      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 = ⦇νxP" 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 = ⦇νxP x  cP x  x4 have "([(x, x4)]  cP) = [(x, x4)]  ⦇νxP"
      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 = ⦇νxP" and C: "cRs = α  ⦇νxP'"
    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 = ⦇νxP 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' AQ ΨQ. Ψ  ΨQ  P α  P';  extractFrame Q = AQ, ΨQ; distinct AQ;
                                  AQ ♯* Ψ; AQ ♯* P; AQ ♯* Q; AQ ♯* α; AQ ♯* P'; AQ ♯* C  Prop α (P'  Q)"
  and   rPar2: "Q' AP ΨP. Ψ  ΨP  Q α  Q';  extractFrame P = AP, ΨP; distinct AP;
                                 AP ♯* Ψ; AP ♯* P; AP ♯* Q; AP ♯* α; AP ♯* Q'; AP ♯* C  Prop α (P  Q')"
  and   rComm1: "ΨQ M N P' AP ΨP K xvec Q' AQ.
           Ψ  ΨQ  P MN  P'; extractFrame P = AP, ΨP; distinct AP;
            Ψ  ΨP  Q K⦇ν*xvec⦈⟨N  Q'; extractFrame Q = AQ, ΨQ; distinct AQ;
            Ψ  ΨP  ΨQ  M  K; distinct xvec; α = τ;
            AP ♯* Ψ;  AP ♯* ΨQ;  AP ♯* P;  AP ♯* M;  AP ♯* N;  AP ♯* P';  AP ♯* Q;  AP ♯* xvec;  AP ♯* Q'; AP ♯* AQ;  AP ♯* C;
            AQ ♯* Ψ;  AQ ♯* ΨP; AQ ♯* P;  AQ ♯* K;  AQ ♯* N;  AQ ♯* P';  AQ ♯* Q;  AQ ♯* xvec;  AQ ♯* Q'; AQ ♯* C;
            xvec ♯* Ψ;  xvec ♯* ΨP; xvec ♯* P;  xvec ♯* M;  xvec ♯* K; xvec ♯* Q;  xvec ♯* ΨQ;  xvec ♯* C 
            Prop (τ) (⦇ν*xvec(P'  Q'))"
  and   rComm2: "ΨQ M xvec N P' AP ΨP K Q' AQ.
           Ψ  ΨQ  P M⦇ν*xvec⦈⟨N  P'; extractFrame P = AP, ΨP; distinct AP;
            Ψ  ΨP  Q KN  Q'; extractFrame Q = AQ, ΨQ; distinct AQ;
            Ψ  ΨP  ΨQ  M  K; distinct xvec; α = τ;
            AP ♯* Ψ;  AP ♯* ΨQ;  AP ♯* P;  AP ♯* M;  AP ♯* N;  AP ♯* P';  AP ♯* Q;  AP ♯* xvec;  AP ♯* Q'; AP ♯* AQ;  AP ♯* C;
            AQ ♯* Ψ;  AQ ♯* ΨP; AQ ♯* P;  AQ ♯* K;  AQ ♯* N;  AQ ♯* P';  AQ ♯* Q;  AQ ♯* xvec;  AQ ♯* Q'; AQ ♯* C;
            xvec ♯* Ψ;  xvec ♯* ΨP; xvec ♯* P;  xvec ♯* M;  xvec ♯* K;  xvec ♯* Q;  xvec ♯* ΨQ;  xvec ♯* C 
            Prop (τ) (⦇ν*xvec(P'  Q'))"
  and   rBrMerge: "ΨQ M N P' AP ΨP Q' AQ.
                    Ψ  ΨQ  P  ¿MN  P'; extractFrame P = AP, ΨP; distinct AP;
                    Ψ  ΨP  Q  ¿MN  Q'; extractFrame Q = AQ, ΨQ; distinct AQ;
                    AP ♯* Ψ; AP ♯* ΨQ; AP ♯* P; AP ♯* N; AP ♯* P';
                    AP ♯* Q; AP ♯* Q'; AP ♯* AQ; AP ♯* M; AQ ♯* M;
                    AQ ♯* Ψ; AQ ♯* ΨP; AQ ♯* P; AQ ♯* N; AQ ♯* P';
                    AQ ♯* Q; AQ ♯* Q'; AP ♯* C; AQ ♯* C; α = ¿MN 
                    Prop (¿MN) (P'  Q')"
  and   rBrComm1: "ΨQ M N P' AP ΨP xvec Q' AQ.
           Ψ  ΨQ  P ¿MN  P'; extractFrame P = AP, ΨP; distinct AP;
            Ψ  ΨP  Q ¡M⦇ν*xvec⦈⟨N  Q'; extractFrame Q = AQ, ΨQ; distinct AQ;
            distinct xvec;
            AP ♯* Ψ;  AP ♯* ΨQ;  AP ♯* P; AP ♯* N;  AP ♯* P';  AP ♯* Q;  AP ♯* xvec;  AP ♯* Q'; AP ♯* AQ;  AP ♯* C;
            AQ ♯* Ψ;  AQ ♯* ΨP; AQ ♯* P; AQ ♯* N;  AQ ♯* P';  AQ ♯* Q;  AQ ♯* xvec;  AQ ♯* Q'; AQ ♯* C;
            AP ♯* M; AQ ♯* M; xvec ♯* M; ¡M⦇ν*xvec⦈⟨N = α;
            xvec ♯* Ψ; xvec ♯* ΨP; xvec ♯* P; xvec ♯* Q; xvec ♯* ΨQ 
            Prop (¡M⦇ν*xvec⦈⟨N) (P'  Q')"
  and   rBrComm2: "ΨQ M xvec N P' AP ΨP Q' AQ.
           Ψ  ΨQ  P ¡M⦇ν*xvec⦈⟨N  P'; extractFrame P = AP, ΨP; distinct AP;
            Ψ  ΨP  Q ¿MN  Q'; extractFrame Q = AQ, ΨQ; distinct AQ;
            distinct xvec;
            AP ♯* Ψ;  AP ♯* ΨQ;  AP ♯* P; AP ♯* N;  AP ♯* P';  AP ♯* Q;  AP ♯* xvec;  AP ♯* Q'; AP ♯* AQ;  AP ♯* C;
            AQ ♯* Ψ;  AQ ♯* ΨP; AQ ♯* P; AQ ♯* N;  AQ ♯* P';  AQ ♯* Q;  AQ ♯* xvec;  AQ ♯* Q'; AQ ♯* C;
            AP ♯* M; AQ ♯* M; xvec ♯* M; ¡M⦇ν*xvec⦈⟨N = α;
            xvec ♯* Ψ;  xvec ♯* ΨP; xvec ♯* P; xvec ♯* Q;  xvec ♯* Ψ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 ΨQ P α' P' Q AQ)
    then show ?thesis using assms
      by(force simp add: psi.inject residualInject residualInject' intro: rPar1)
  next
    case (cPar2 ΨP Q α' Q' P AP)
    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 ΨQ P1 M N P' AP ΨP Q1 Q' AQ)
    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' 
      Ψ  ΨQ  P1  ¿MN  P' 
      extractFrame P1 = AP, ΨP 
      distinct AP 
      Ψ  ΨP  Q1  ¡M⦇ν*bn α⦈⟨N  Q' 
      extractFrame Q1 = AQ, ΨQ 
      distinct AQ 
      AP ♯* Ψ 
      AP ♯* ΨQ 
      AP ♯* P1 
      AP ♯* N 
      AP ♯* P' 
      AP ♯* Q1 
      AP ♯* Q' 
      AP ♯* AQ 
      AP ♯* bn α 
      AQ ♯* Ψ 
      AQ ♯* ΨP 
      AQ ♯* P1 
      AQ ♯* N 
      AQ ♯* P' 
      AQ ♯* Q1 
      AQ ♯* Q' 
      AQ ♯* bn α 
      bn α ♯* Ψ 
      bn α ♯* ΨP 
      bn α ♯* ΨQ 
      bn α ♯* P1 
      bn α ♯* Q1 
      AP ♯* M  AQ ♯* M  bn α ♯* M  AP ♯* C  AQ ♯* 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 "Ψ  ΨQ  P1  ¿MN  P'"
      and "extractFrame P1 = AP, ΨP"
      and "distinct AP"
      and "Ψ  ΨP  Q1  ¡M⦇ν*bn α⦈⟨N  Q'"
      and "extractFrame Q1 = AQ, ΨQ"
      and "distinct AQ"
      and "AP ♯* Ψ"
      and "AP ♯* ΨQ"
      and "AP ♯* P1"
      and "AP ♯* N"
      and "AP ♯* P'"
      and "AP ♯* Q1"
      and "AP ♯* Q'"
      and "AP ♯* AQ"
      and "AP ♯* bn α"
      and "AQ ♯* Ψ"
      and "AQ ♯* ΨP"
      and "AQ ♯* P1"
      and "AQ ♯* N"
      and "AQ ♯* P'"
      and "AQ ♯* Q1"
      and "AQ ♯* Q'"
      and "AQ ♯* bn α"
      and "bn α ♯* Ψ"
      and "bn α ♯* ΨP"
      and "bn α ♯* ΨQ"
      and "bn α ♯* P1"
      and "bn α ♯* Q1"
      and "AP ♯* M"
      and "AQ ♯* M"
      and "bn α ♯* M"
      and "AP ♯* C"
      and "AQ ♯* 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 ΨQ P1 M N P' AP ΨP Q1 Q' AQ)
    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' 
      Ψ  ΨQ  P1  ¡M⦇ν*bn α⦈⟨N  P' 
      extractFrame P1 = AP, ΨP 
      distinct AP 
      Ψ  ΨP  Q1  ¿MN  Q' 
      extractFrame Q1 = AQ, ΨQ 
      distinct AQ 
      AP ♯* Ψ 
      AP ♯* ΨQ 
      AP ♯* P1 
      AP ♯* N 
      AP ♯* P' 
      AP ♯* Q1 
      AP ♯* Q' 
      AP ♯* AQ 
      AP ♯* bn α 
      AQ ♯* Ψ 
      AQ ♯* ΨP 
      AQ ♯* P1 
      AQ ♯* N 
      AQ ♯* P' 
      AQ ♯* Q1 
      AQ ♯* Q' 
      AQ ♯* bn α 
      bn α ♯* Ψ 
      bn α ♯* ΨP 
      bn α ♯* ΨQ 
      bn α ♯* P1 
      bn α ♯* Q1 
      AP ♯* M  AQ ♯* M  bn α ♯* M  AP ♯* C  AQ ♯* 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 "Ψ  ΨQ  P1  ¡M⦇ν*bn α⦈⟨N  P'"
      and "extractFrame P1 = AP, ΨP"
      and "distinct AP"
      and "Ψ  ΨP  Q1  ¿MN  Q'"
      and "extractFrame Q1 = AQ, ΨQ"
      and "distinct AQ"
      and "AP ♯* Ψ"
      and "AP ♯* ΨQ"
      and "AP ♯* P1"
      and "AP ♯* N"
      and "AP ♯* P'"
      and "AP ♯* Q1"
      and "AP ♯* Q'"
      and "AP ♯* AQ"
      and "AP ♯* bn α"
      and "AQ ♯* Ψ"
      and "AQ ♯* ΨP"
      and "AQ ♯* P1"
      and "AQ ♯* N"
      and "AQ ♯* P'"
      and "AQ ♯* Q1"
      and "AQ ♯* Q'"
      and "AQ ♯* bn α"
      and "bn α ♯* Ψ"
      and "bn α ♯* ΨP"
      and "bn α ♯* ΨQ"
      and "bn α ♯* P1"
      and "bn α ♯* Q1"
      and "AP ♯* M"
      and "AQ ♯* M"
      and "bn α ♯* M"
      and "AP ♯* C"
      and "AQ ♯* 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 MN  R"
  and   rPar1: "P' AQ ΨQ. Ψ  ΨQ  P MN  P';  extractFrame Q = AQ, ΨQ; distinct AQ;
                       AQ ♯* Ψ; AQ ♯* P; AQ ♯* Q; AQ ♯* M; AQ ♯* N; AQ ♯* C  Prop (P'  Q)"
  and   rPar2: "Q' AP ΨP. Ψ  ΨP  Q MN  Q'; extractFrame P = AP, ΨP; distinct AP;
                       AP ♯* Ψ; AP ♯* P;  AP ♯* Q; AP ♯* M; AP ♯* N; AP ♯* 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 "α = MN" 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 ¿MN  R"
  and   rPar1: "P' AQ ΨQ. Ψ  ΨQ  P ¿MN  P';  extractFrame Q = AQ, ΨQ; distinct AQ;
                       AQ ♯* Ψ; AQ ♯* P; AQ ♯* Q; AQ ♯* M; AQ ♯* N; AQ ♯* C  Prop (P'  Q)"
  and   rPar2: "Q' AP ΨP. Ψ  ΨP  Q ¿MN  Q'; extractFrame P = AP, ΨP; distinct AP;
                       AP ♯* Ψ; AP ♯* P;  AP ♯* Q; AP ♯* M; AP ♯* N; AP ♯* C  Prop (P  Q')"
  and   rBrMerge: "ΨQ P' AP ΨP Q' AQ.
                    Ψ  ΨQ  P  ¿MN  P'; extractFrame P = AP, ΨP; distinct AP;
                    Ψ  ΨP  Q  ¿MN  Q'; extractFrame Q = AQ, ΨQ; distinct AQ;
                    AP ♯* Ψ; AP ♯* ΨQ; AP ♯* P; AP ♯* N; AP ♯* P';
                    AP ♯* Q; AP ♯* Q'; AP ♯* AQ; AP ♯* M; AQ ♯* M;
                    AQ ♯* Ψ; AQ ♯* ΨP; AQ ♯* P; AQ ♯* N; AQ ♯* P';
                    AQ ♯* Q; AQ ♯* Q'; AP ♯* C; AQ ♯* 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 "α = ¿MN" 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' AQ ΨQ. Ψ  ΨQ  P M⦇ν*xvec⦈⟨N  P';  extractFrame Q = AQ, ΨQ; distinct AQ;
                       AQ ♯* Ψ; AQ ♯* P; AQ ♯* Q; AQ ♯* M; AQ ♯* xvec; AQ ♯* N; AQ ♯* C; AQ ♯* xvec; distinct xvec  Prop (P'  Q)"
  and   rPar2: "Q' AP ΨP. Ψ  ΨP  Q M⦇ν*xvec⦈⟨N  Q'; extractFrame P = AP, ΨP; distinct AP;
                       AP ♯* Ψ; AP ♯* P;  AP ♯* Q; AP ♯* M; AP ♯* xvec; AP ♯* N; AP ♯* C; AP ♯* 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' AQ ΨQ. Ψ  ΨQ  P ¡M⦇ν*xvec⦈⟨N  P';  extractFrame Q = AQ, ΨQ; distinct AQ;
                       AQ ♯* Ψ; AQ ♯* P; AQ ♯* Q; AQ ♯* M; AQ ♯* xvec; AQ ♯* N; AQ ♯* C; AQ ♯* xvec; distinct xvec  Prop (P'  Q)"
  and   rPar2: "Q' AP ΨP. Ψ  ΨP  Q ¡M⦇ν*xvec⦈⟨N  Q'; extractFrame P = AP, ΨP; distinct AP;
                       AP ♯* Ψ; AP ♯* P;  AP ♯* Q; AP ♯* M; AP ♯* xvec; AP ♯* N; AP ♯* C; AP ♯* xvec; distinct xvec  Prop (P  Q')"
  and   rBrComm1: "ΨQ P' AP ΨP Q' AQ.
           Ψ  ΨQ  P ¿MN  P'; extractFrame P = AP, ΨP; distinct AP;
            Ψ  ΨP  Q ¡M⦇ν*xvec⦈⟨N  Q'; extractFrame Q = AQ, ΨQ; distinct AQ;
            distinct xvec;
            AP ♯* Ψ;  AP ♯* ΨQ;  AP ♯* P; AP ♯* N;  AP ♯* P';  AP ♯* Q;  AP ♯* xvec;  AP ♯* Q'; AP ♯* AQ;  AP ♯* C;
            AQ ♯* Ψ;  AQ ♯* ΨP; AQ ♯* P; AQ ♯* N;  AQ ♯* P';  AQ ♯* Q;  AQ ♯* xvec;  AQ ♯* Q'; AQ ♯* C;
            AP ♯* M; AQ ♯* M; xvec ♯* M;
            xvec ♯* Ψ; xvec ♯* ΨP; xvec ♯* P; xvec ♯* Q; xvec ♯* ΨQ 
            Prop (P'  Q')"
  and   rBrComm2: "ΨQ P' AP ΨP Q' AQ.
           Ψ  ΨQ  P ¡M⦇ν*xvec⦈⟨N  P'; extractFrame P = AP, ΨP; distinct AP;
            Ψ  ΨP  Q ¿MN  Q'; extractFrame Q = AQ, ΨQ; distinct AQ;
            distinct xvec;
            AP ♯* Ψ;  AP ♯* ΨQ;  AP ♯* P; AP ♯* N;  AP ♯* P';  AP ♯* Q;  AP ♯* xvec;  AP ♯* Q'; AP ♯* AQ;  AP ♯* C;
            AQ ♯* Ψ;  AQ ♯* ΨP; AQ ♯* P; AQ ♯* N;  AQ ♯* P';  AQ ♯* Q;  AQ ♯* xvec;  AQ ♯* Q'; AQ ♯* C;
            AP ♯* M; AQ ♯* M; xvec ♯* M;
            xvec ♯* Ψ;  xvec ♯* ΨP; xvec ♯* P; xvec ♯* Q;  xvec ♯* Ψ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 AP  :: "name list"
    and ΨP :: 'b
    and N   :: 'a
    and K   :: 'a
    and B   :: "name list"

assumes "Ψ  P KN  P'"
  and   "extractFrame P = AP, ΨP"
  and   "distinct AP"
  and   "B ♯* P"
  and   "AP ♯* Ψ"
  and   "AP ♯* B"
  and   "AP ♯* P"
  and   "AP ♯* K"

obtains M where "Ψ  ΨP  K  M" and "B ♯* M"
  using assms
proof(nominal_induct avoiding: B arbitrary: thesis rule: inputFrameInduct)
  case(cAlpha Ψ P K N P' AP ΨP p B)
  then obtain M where subjEq: "Ψ  ΨP  K  M" and "B ♯* M"
    by(auto intro: cAlpha)
  from Ψ  ΨP  K  M
  have "p  (Ψ  ΨP  K  M)" by simp
  with set p  set AP × set (p  AP)
    AP ♯* Ψ (p  AP) ♯* Ψ AP ♯* K (p  AP) ♯* K
  have permEq: "Ψ  (p  ΨP)  K  (p  M)" by(simp add: eqvts)
  from B ♯* M have "p  (B ♯* M)" by simp
  with set p  set AP × set (p  AP)
    AP ♯* B (p  AP) ♯* 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 AP ΨP B)
  then obtain M where "Ψ  ΨP  K  M" and "B ♯* M"
    by - (rule cCase, auto dest: memFreshChain)
  with ΨP  𝟭 show ?case by(blast intro: cCase statEqEnt compositionSym Identity)
next
  case(cPar1 Ψ ΨQ P K N P' AQ Q AP ΨP B)
  then obtain M where "(Ψ  ΨQ)  Ψ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 Ψ ΨP Q K N Q' AP P AQ ΨQ B)
  then obtain M where "(Ψ  ΨP)  Ψ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 AP ΨP B)
  then obtain M where "Ψ  ΨP  K  M" and "B ♯* M"
    by - (rule cScope, auto)
  then show ?case by(auto intro: cScope)
next
  case(cBang Ψ P K N P' AP ΨP B)
  then obtain M where "Ψ  ΨP  𝟭  K  M" and "B ♯* M"
    by - (rule cBang, auto)
  with Ψ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 AP  :: "name list"
    and ΨP :: 'b
    and N   :: 'a
    and K   :: 'a
    and xvec :: "name list"
    and B   :: "name list"

assumes "Ψ  P  ROut K (⦇ν*xvecN ≺' P')"
  and   "extractFrame P = AP, ΨP"
  and   "distinct AP"
  and   "xvec ♯* K"
  and   "distinct xvec"
  and   "B ♯* P"
  and   "AP ♯* Ψ"
  and   "AP ♯* B"
  and   "AP ♯* P"
  and   "AP ♯* K"

obtains M where "Ψ  ΨP  K  M" and "B ♯* M"
  using assms
proof(nominal_induct avoiding: B xvec arbitrary: thesis rule: outputFrameInduct)
  case(cAlpha Ψ P K AP ΨP p α B xvec)
  then obtain M where subjEq: "Ψ  ΨP  K  M" and "B ♯* M"
    by(auto intro: cAlpha)

  from Ψ  ΨP  K  M
  have "p  (Ψ  ΨP  K  M)" by simp

  with set p  set AP × set (p  AP)
    AP ♯* Ψ (p  AP) ♯* Ψ AP ♯* K (p  AP) ♯* K
  have permEq: "Ψ  (p  ΨP)  K  (p  M)" by(simp add: eqvts)

  from B ♯* M have "p  (B ♯* M)" by simp
  with set p  set AP × set (p  AP)
    AP ♯* B (p  AP) ♯* 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 ♯* (MN⟩.P) have "B ♯* M" by simp
  ultimately show ?case by(auto intro: cOutput)
next
  case(cCase Ψ P K P' φ Cs AP ΨP B xvec)
  then obtain M where "Ψ  ΨP  K  M" and "B ♯* M"
    by - (rule cCase, auto dest: memFreshChain)
  with ΨP  𝟭 show ?case by(blast intro: cCase statEqEnt compositionSym Identity)
next
  case(cPar1 Ψ ΨQ P K yvec N P' AQ Q AP ΨP B xvec)
  then obtain M where "(Ψ  ΨQ)  Ψ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 Ψ ΨP Q K yvec N Q' AP P AQ ΨQ B xvec)
  then obtain M where "(Ψ  ΨP)  Ψ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 AP ΨP B xvec)
  then obtain K where "Ψ  Ψ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 AP ΨP B xvec)
  then obtain M where "Ψ  Ψ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' AP ΨP B xvec)
  then obtain M where "Ψ  ΨP  𝟭  K  M" and "B ♯* M"
    by - (rule cBang, auto)
  with Ψ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 AP :: "name list"
    and ΨP :: 'b

assumes "Ψ  P MN  P'"
  and   "extractFrame P = AP, ΨP"
  and   "distinct AP"
  and   "Ψ  ΨP  M  K"
  and   "AP ♯* Ψ"
  and   "AP ♯* P"
  and   "AP ♯* M"
  and   "AP ♯* K"

shows "Ψ  P KN  P'"
  using assms
proof(nominal_induct avoiding: K rule: inputFrameInduct)
  case(cAlpha Ψ P M N P' AP ΨP  p K)
  have S: "set p  set AP × set (p  AP)" by fact
  from Ψ  (p  ΨP)  M  K have "(p  (Ψ  (p  ΨP)))  (p  M)  (p  K)"
    by(rule chanEqClosed)
  with S distinctPerm p AP ♯* Ψ AP ♯* M AP ♯* K (p  AP) ♯* Ψ (p  AP) ♯* M (p  AP) ♯* K
  have "Ψ  ΨP  M  K" by(simp add: eqvts)
  with AP ♯* Ψ AP ♯* P AP ♯* M AP ♯* K
    Ψ  ΨP  M  K; AP ♯* Ψ; AP ♯* P; AP ♯* M; AP ♯* K  Ψ  P KN  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 AP ΨP K)
  from Ψ  𝟭  M  K ΨP  𝟭 have "Ψ  ΨP  M  K"
    by(blast intro: statEqEnt Identity compositionSym AssertionStatEqSym)
  with AP ♯* Ψ AP ♯* P AP ♯* M AP ♯* K
    K. Ψ  ΨP  M  K; AP ♯* Ψ; AP ♯* P; AP ♯* M; AP ♯* K  Ψ  P KN  P'
  have "Ψ  P KN  P'" by force
  then show ?case using (φ, P)  set Cs Ψ  φ guarded P by(rule Case)
next
  case(cPar1 Ψ ΨQ P M N P' AQ Q AP ΨP K)
  from Ψ  ΨP  ΨQ  M  K have "(Ψ  ΨQ)  ΨP  M  K"
    by(metis statEqEnt Associativity Composition AssertionStatEqTrans Commutativity)
  with AP ♯* Ψ AP ♯* ΨQ AP ♯* P AP ♯* M AP ♯* K
    K. (Ψ  ΨQ)  ΨP  M  K; AP ♯* (Ψ  ΨQ); AP ♯* P; AP ♯* M; AP ♯* K  Ψ  ΨQ  P KN  P'
  have "Ψ  ΨQ  P KN  P'" by force
  then show ?case using extractFrame Q = AQ, ΨQ AQ ♯* Ψ AQ ♯* P AQ ♯* K AQ ♯* N
    by(auto intro: Par1)
next
  case(cPar2 Ψ ΨP Q M N Q' AP P AQ ΨQ K)
  from Ψ  ΨP  ΨQ  M  K have "(Ψ  ΨP)  ΨQ  M  K"
    by(rule statEqEnt[OF AssertionStatEqSym[OF Associativity]])
  with AQ ♯* Ψ AQ ♯* ΨP AQ ♯* Q AQ ♯* M AQ ♯* K
    K. (Ψ  ΨP)  ΨQ  M  K; AQ ♯* (Ψ  ΨP); AQ ♯* Q; AQ ♯* M; AQ ♯* K  Ψ  ΨP  Q KN  Q'
  have "Ψ  ΨP  Q KN  Q'" by force
  then show ?case using extractFrame P = AP, ΨP AP ♯* Ψ AP ♯* Q AP ♯* K AP ♯* N
    by(auto intro: Par2)
next
  case(cScope Ψ P M N P' x AP ΨP)
  then have "Ψ  P KN  P'" by force
  with x  Ψ x  K x  N show ?case
    by(auto intro: Scope)
next
  case(cBang Ψ P M N P' AP ΨP K)
  from Ψ  𝟭  M  K ΨP  𝟭 have "Ψ  ΨP  𝟭  M  K"
    by(blast intro: statEqEnt Identity compositionSym AssertionStatEqSym)
  with AP ♯* Ψ AP ♯* P AP ♯* M AP ♯* K
    K. Ψ  ΨP  𝟭  M  K; AP ♯* Ψ; AP ♯* (P  !P); AP ♯* M; AP ♯* K  Ψ  P  !P KN  P'
  have "Ψ  P  !P KN  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 AP   :: "name list"
    and ΨP   :: 'b

assumes "Ψ  P M⦇ν*xvec⦈⟨N  P'"
  and   "extractFrame P = AP, ΨP"
  and   "distinct AP"
  and   "Ψ  ΨP  M  K"
  and   "AP ♯* Ψ"
  and   "AP ♯* P"
  and   "AP ♯* M"
  and   "AP ♯* K"

shows "Ψ  P K⦇ν*xvec⦈⟨N  P'"
  using assms unfolding residualInject
proof(nominal_induct avoiding: K rule: outputFrameInduct)
  case(cAlpha Ψ P M AP ΨP p B K)
  have S: "set p  set AP × set(p  AP)" by fact
  from Ψ  (p  ΨP)  M  K have "(p  (Ψ  (p  ΨP)))  (p  M)  (p  K)"
    by(rule chanEqClosed)
  with S distinctPerm p AP ♯* Ψ AP ♯* M AP ♯* K (p  AP) ♯* Ψ (p  AP) ♯* M (p  AP) ♯* K
  have "Ψ  ΨP  M  K" by(simp add: eqvts)
  with AP ♯* Ψ AP ♯* P AP ♯* M AP ♯* 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 AP ΨP K)
  from Ψ  𝟭  M  K ΨP  𝟭 have "Ψ  ΨP  M  K"
    by(blast intro: statEqEnt Identity compositionSym AssertionStatEqSym)
  with AP ♯* Ψ AP ♯* P AP ♯* M AP ♯* K
    K. Ψ  ΨP  M  K; AP ♯* Ψ; AP ♯* P; AP ♯* M; AP ♯* 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 Ψ ΨQ P M xvec N P' AQ Q AP ΨP K)
  from Ψ  ΨP  ΨQ  M  K have "(Ψ  ΨQ)  ΨP  M  K"
    by(metis statEqEnt Associativity Composition AssertionStatEqTrans Commutativity)
  with AP ♯* Ψ AP ♯* ΨQ AP ♯* P AP ♯* M AP ♯* K
    K. (Ψ  ΨQ)  ΨP  M  K; AP ♯* (Ψ  ΨQ); AP ♯* P; AP ♯* M; AP ♯* K  Ψ  ΨQ  P (ROut K (⦇ν*xvecN ≺' P'))
  have "Ψ  ΨQ  P K⦇ν*xvec⦈⟨N  P'" by(force simp add: residualInject)
  then show ?case using extractFrame Q = AQ, ΨQ xvec ♯* Q AQ ♯* Ψ AQ ♯* P AQ ♯* K  AQ ♯* xvec  AQ ♯* N Par1[where α="K⦇ν*xvec⦈⟨N"]
    by(auto simp add: residualInject)
next
  case(cPar2 Ψ ΨP Q M xvec N Q' AP P AQ ΨQ K)
  from Ψ  ΨP  ΨQ  M  K have "(Ψ  ΨP)  ΨQ  M  K"
    by(rule statEqEnt[OF AssertionStatEqSym[OF Associativity]])
  with AQ ♯* Ψ AQ ♯* ΨP AQ ♯* Q AQ ♯* M AQ ♯* K
    K. (Ψ  ΨP)  ΨQ  M  K; AQ ♯* (Ψ  ΨP); AQ ♯* Q; AQ ♯* M; AQ ♯* K  Ψ  ΨP  Q ROut K (⦇ν*xvecN ≺' Q')
  have "Ψ  ΨP  Q ROut K (⦇ν*xvecN ≺' Q')" by force
  then show ?case using extractFrame P = AP, ΨP xvec ♯* P AP ♯* Ψ AP ♯* Q AP ♯* K AP ♯* xvec AP ♯* N Par2[where α="K⦇ν*xvec⦈⟨N"]
    by(auto simp add: residualInject)
next
  case(cOpen Ψ P M xvec yvec N P' x AP Ψ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 AP Ψ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 AP ΨP K)
  from Ψ  𝟭  M  K ΨP  𝟭 have "Ψ  ΨP  𝟭  M  K"
    by(blast intro: statEqEnt Identity compositionSym AssertionStatEqSym)
  with AP ♯* Ψ AP ♯* P AP ♯* M AP ♯* K
    K. Ψ  ΨP  𝟭  M  K; AP ♯* Ψ; AP ♯* (P  !P); AP ♯* M; AP ♯* 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' AQ ΨQ. Ψ  ΨQ  P α  P';  extractFrame Q = AQ, ΨQ; distinct AQ;
                       AQ ♯* Ψ; AQ ♯* P; AQ ♯* α; AQ ♯* C  Prop α (P'  Q)"
  and   rPar2: "Q' AP ΨP. Ψ  ΨP  Q α  Q'; extractFrame P = AP, ΨP; distinct AP;
                       AP ♯* Ψ; AP ♯* Q; AP ♯* α; AP ♯* C  Prop α (P  Q')"
  and   rComm1: "ΨQ M N P' AP ΨP K xvec Q' AQ.
           Ψ  ΨQ  P MN  P'; extractFrame P = AP, ΨP; distinct AP;
            Ψ  ΨP  Q K⦇ν*xvec⦈⟨N  Q'; extractFrame Q = AQ, ΨQ; distinct AQ;
            Ψ  ΨP  ΨQ  M  K; yvec ♯* M; yvec ♯* K; distinct xvec; α = τ;
            AP ♯* Ψ;  AP ♯* ΨQ;  AP ♯* P;  AP ♯* M;  AP ♯* N;  AP ♯* P';  AP ♯* Q;  AP ♯* xvec;  AP ♯* Q'; AP ♯* AQ;  AP ♯* C;
            AQ ♯* Ψ;  AQ ♯* ΨP; AQ ♯* P;  AQ ♯* K;  AQ ♯* N;  AQ ♯* P';  AQ ♯* Q;  AQ ♯* xvec;  AQ ♯* Q'; AQ ♯* C;
            xvec ♯* Ψ;  xvec ♯* ΨP; xvec ♯* P;  xvec ♯* M;  xvec ♯* K; xvec ♯* Q;  xvec ♯* ΨQ;  xvec ♯* C 
            Prop (τ) (⦇ν*xvec(P'  Q'))"
  and   rComm2: "ΨQ M xvec N P' AP ΨP K Q' AQ.
           Ψ  ΨQ  P M⦇ν*xvec⦈⟨N  P'; extractFrame P = AP, ΨP; distinct AP;
            Ψ  ΨP  Q KN  Q'; extractFrame Q = AQ, ΨQ; distinct AQ;
            Ψ  ΨP  ΨQ  M  K; yvec ♯* M; yvec ♯* K; distinct xvec; α = τ;
            AP ♯* Ψ;  AP ♯* ΨQ;  AP ♯* P;  AP ♯* M;  AP ♯* N;  AP ♯* P';  AP ♯* Q;  AP ♯* xvec;  AP ♯* Q'; AP ♯* AQ;  AP ♯* C;
            AQ ♯* Ψ;  AQ ♯* ΨP; AQ ♯* P;  AQ ♯* K;  AQ ♯* N;  AQ ♯* P';  AQ ♯* Q;  AQ ♯* xvec;  AQ ♯* Q'; AQ ♯* C;
            xvec ♯* Ψ;  xvec ♯* ΨP; xvec ♯* P;  xvec ♯* M;  xvec ♯* K;  xvec ♯* Q;  xvec ♯* ΨQ;  xvec ♯* C 
            Prop (τ) (⦇ν*xvec(P'  Q'))"
  and   rBrMerge: "ΨQ M N P' AP ΨP Q' AQ.
                    Ψ  ΨQ  P  ¿MN  P'; extractFrame P = AP, ΨP; distinct AP;
                    Ψ  ΨP  Q  ¿MN  Q'; extractFrame Q = AQ, ΨQ; distinct AQ;
                    AP ♯* Ψ; AP ♯* ΨQ; AP ♯* P; AP ♯* N; AP ♯* P';
                    AP ♯* Q; AP ♯* Q'; AP ♯* AQ; AP ♯* M; AQ ♯* M;
                    AQ ♯* Ψ; AQ ♯* ΨP; AQ ♯* P; AQ ♯* N; AQ ♯* P';
                    AQ ♯* Q; AQ ♯* Q'; AP ♯* C; AQ ♯* C; α = ¿MN 
                    Prop (¿MN) (P'  Q')"
  and   rBrComm1: "ΨQ M N P' AP ΨP xvec Q' AQ.
           Ψ  ΨQ  P ¿MN  P'; extractFrame P = AP, ΨP; distinct AP;
            Ψ  ΨP  Q ¡M⦇ν*xvec⦈⟨N  Q'; extractFrame Q = AQ, ΨQ; distinct AQ;
            distinct xvec;
            AP ♯* Ψ;  AP ♯* ΨQ;  AP ♯* P; AP ♯* N;  AP ♯* P';  AP ♯* Q;  AP ♯* xvec;  AP ♯* Q'; AP ♯* AQ;  AP ♯* C;
            AQ ♯* Ψ;  AQ ♯* ΨP; AQ ♯* P; AQ ♯* N;  AQ ♯* P';  AQ ♯* Q;  AQ ♯* xvec;  AQ ♯* Q'; AQ ♯* C;
            AP ♯* M; AQ ♯* M; xvec ♯* M; ¡M⦇ν*xvec⦈⟨N = α;
            xvec ♯* Ψ; xvec ♯* ΨP; xvec ♯* P; xvec ♯* Q; xvec ♯* ΨQ 
            Prop (¡M⦇ν*xvec⦈⟨N) (P'  Q')"
  and   rBrComm2: "ΨQ M xvec N P' AP ΨP Q' AQ.
           Ψ  ΨQ  P ¡M⦇ν*xvec⦈⟨N  P'; extractFrame P = AP, ΨP; distinct AP;
            Ψ  ΨP  Q ¿MN  Q'; extractFrame Q = AQ, ΨQ; distinct AQ;
            distinct xvec;
            AP ♯* Ψ;  AP ♯* ΨQ;  AP ♯* P; AP ♯* N;  AP ♯* P';  AP ♯* Q;  AP ♯* xvec;  AP ♯* Q'; AP ♯* AQ;  AP ♯* C;
            AQ ♯* Ψ;  AQ ♯* ΨP; AQ ♯* P; AQ ♯* N;  AQ ♯* P';  AQ ♯* Q;  AQ ♯* xvec;  AQ ♯* Q'; AQ ♯* C;
            AP ♯* M; AQ ♯* M; xvec ♯* M; ¡M⦇ν*xvec⦈⟨N = α;
            xvec ♯* Ψ;  xvec ♯* ΨP; xvec ♯* P; xvec ♯* Q;  xvec ♯* Ψ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' AQ ΨQ)
  then show ?case by(auto intro: rPar1)
next
  case(cPar2 Q' AP ΨP)
  then show ?case by(auto intro: rPar2)
next
  case(cComm1 ΨQ M N P' AP ΨP K xvec Q' AQ)
  from AP ♯* (C, yvec) AQ ♯* (C, yvec) xvec ♯* (C, yvec)
  have "AP ♯* C" and "AQ ♯* C" and "xvec ♯* C" and "AP ♯* yvec" and "AQ ♯* yvec" and "xvec ♯* yvec"
    by simp+

  have FrP: "extractFrame P = AP, ΨP" and FrQ: "extractFrame Q = AQ, ΨQ"
    and MeqK: "Ψ  ΨP  ΨQ  M  K" by fact+

  from Ψ  ΨQ  P MN  P' FrP distinct AP AP ♯* P AQ ♯* P yvec ♯* P AP ♯* Ψ
    AP ♯* AQ AP ♯* yvec AP ♯* xvec AP ♯* P AP ♯* M xvec ♯* P AP ♯* ΨQ
  obtain M' where MeqM': "(Ψ  ΨQ)  ΨP  M  M'" and "xvec ♯* M'" and "yvec ♯* M'" and "AQ ♯* M'"
    by - (rule inputObtainPrefix[where B="xvec@yvec@AQ"], (assumption | force)+)
  from Ψ  ΨP  Q K⦇ν*xvec⦈⟨N  Q' have "Ψ  ΨP  Q  ROut K (⦇ν*xvecN ≺' Q')"
    by(simp add: residualInject)
  with FrQ distinct AQ AP ♯* Q AQ ♯* Q yvec ♯* Q AQ ♯* Ψ
    AP ♯* AQ AQ ♯* yvec AQ ♯* xvec AQ ♯* Q AQ ♯* K xvec ♯* Q AQ ♯* ΨP xvec ♯* K distinct xvec
  obtain K' where KeqK': "(Ψ  ΨP)  ΨQ  K  K'" and "xvec ♯* K'" and "yvec ♯* K'" and "AP ♯* K'"
    by - (rule outputObtainPrefix[where B="xvec@yvec@AP"], (assumption | force | metis freshChainSym)+)

  from MeqK KeqK' have "(Ψ  ΨQ)  ΨP  M  K'"
    by(metis statEqEnt Associativity Commutativity Composition chanEqTrans)
  with Ψ  ΨQ  P MN  P' FrP distinct AP
  have "Ψ  ΨQ  P K'N  P'" using AP ♯* Ψ AP ♯* ΨQ AP ♯* P AP ♯* M AP ♯* K'
    by - (rule inputRenameSubject, (assumption | force)+)
  moreover note FrP distinct AP
  moreover from MeqK MeqM' have "(Ψ  ΨP)  ΨQ  K  M'"
    by(metis statEqEnt Associativity Commutativity Composition chanEqTrans chanEqSym)
  with Ψ  ΨP  Q K⦇ν*xvec⦈⟨N  Q' FrQ distinct AQ
  have "Ψ  ΨP  Q M'⦇ν*xvec⦈⟨N  Q'" using AQ ♯* Ψ AQ ♯* ΨP AQ ♯* Q AQ ♯* K AQ ♯* M'
    by - (rule outputRenameSubject, (assumption | force)+)
  moreover note FrQ distinct AQ
  moreover from MeqM' KeqK' MeqK have "Ψ  ΨP  ΨQ  K'  M'"
    by(metis statEqEnt Associativity Commutativity Composition chanEqTrans chanEqSym)
  moreover note AP ♯* Ψ AP ♯* ΨQ AP ♯* P AP ♯* K' AP ♯* N AP ♯* P' AP ♯* Q AP ♯* xvec AP ♯* Q' AP ♯* AQ AP ♯* C
    AQ ♯* Ψ AQ ♯* ΨP AQ ♯* Q AQ ♯* M' AQ ♯* N AQ ♯* Q' AQ ♯* P AQ ♯* xvec AQ ♯* P' AQ ♯* C α = τ
    xvec ♯* Ψ xvec ♯* ΨP xvec ♯* P xvec ♯* M' xvec ♯* K' xvec ♯* Q xvec ♯* ΨQ xvec ♯* C yvec ♯* M' yvec ♯* K' distinct xvec
  ultimately show ?case
    by(metis rComm1)
next
  case(cComm2 ΨQ M xvec N P' AP ΨP K Q' AQ)
  from AP ♯* (C, yvec) AQ ♯* (C, yvec) xvec ♯* (C, yvec)
  have "AP ♯* C" and "AQ ♯* C" and "xvec ♯* C" and "AP ♯* yvec" and "AQ ♯* yvec" and "xvec ♯* yvec"
    by simp+

  have FrP: "extractFrame P = AP, ΨP" and FrQ: "extractFrame Q = AQ, ΨQ"
    and MeqK: "Ψ  ΨP  ΨQ  M  K" by fact+

  from Ψ  ΨQ  P M⦇ν*xvec⦈⟨N  P' have "Ψ  ΨQ  P  ROut M (⦇ν*xvecN ≺' P')"
    by(simp add: residualInject)
  with FrP distinct AP AP ♯* P AQ ♯* P yvec ♯* P AP ♯* Ψ
    AP ♯* AQ AP ♯* yvec AP ♯* xvec AP ♯* P AP ♯* M xvec ♯* P AP ♯* ΨQ xvec ♯* M distinct xvec
  obtain M' where MeqM': "(Ψ  ΨQ)  ΨP  M  M'" and "xvec ♯* M'" and "yvec ♯* M'" and "AQ ♯* M'"
    by - (rule outputObtainPrefix[where B="xvec@yvec@AQ"], (assumption | force)+)
  from Ψ  ΨP  Q KN  Q' FrQ distinct AQ AP ♯* Q AQ ♯* Q yvec ♯* Q AQ ♯* Ψ
    AP ♯* AQ AQ ♯* yvec AQ ♯* xvec AQ ♯* Q AQ ♯* K xvec ♯* Q AQ ♯* ΨP
  obtain K' where KeqK': "(Ψ  ΨP)  ΨQ  K  K'" and "xvec ♯* K'" and "yvec ♯* K'" and "AP ♯* K'"
    by - (rule inputObtainPrefix[where B="xvec@yvec@AP"], (assumption | force | metis freshChainSym)+)

  from MeqK KeqK' have "(Ψ  ΨQ)  ΨP  M  K'"
    by(metis statEqEnt Associativity Commutativity Composition chanEqTrans)
  with Ψ  ΨQ  P M⦇ν*xvec⦈⟨N  P' FrP distinct AP
  have "Ψ  ΨQ  P K'⦇ν*xvec⦈⟨N  P'" using AP ♯* Ψ AP ♯* ΨQ AP ♯* P AP ♯* M AP ♯* K'
    by - (rule outputRenameSubject, (assumption | force)+)
  moreover note FrP distinct AP
  moreover from MeqK MeqM' have "(Ψ  ΨP)  ΨQ  K  M'"
    by(metis statEqEnt Associativity Commutativity Composition chanEqTrans chanEqSym)
  with Ψ  ΨP  Q KN  Q' FrQ distinct AQ
  have "Ψ  ΨP  Q M'N  Q'" using AQ ♯* Ψ AQ ♯* ΨP AQ ♯* Q AQ ♯* K AQ ♯* M'
    by - (rule inputRenameSubject, (assumption | force)+)
  moreover note FrQ distinct AQ
  moreover from MeqM' KeqK' MeqK have "Ψ  ΨP  ΨQ  K'  M'"
    by(metis statEqEnt Associativity Commutativity Composition chanEqTrans chanEqSym)
  moreover note AP ♯* Ψ AP ♯* ΨQ AP ♯* P AP ♯* K' AP ♯* N AP ♯* P' AP ♯* Q AP ♯* xvec AP ♯* Q' AP ♯* AQ AP ♯* C
    AQ ♯* Ψ AQ ♯* ΨP AQ ♯* Q AQ ♯* M' AQ ♯* N AQ ♯* Q' AQ ♯* P AQ ♯* xvec AQ ♯* P' AQ ♯* C α = τ
    xvec ♯* Ψ xvec ♯* ΨP xvec ♯* P xvec ♯* M' xvec ♯* K' xvec ♯* Q xvec ♯* Ψ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 (KN[xvec::=Tvec]) (P[xvec::=Tvec])"
  and   rBrInput: "K Tvec. Ψ  K  M; set xvec  supp N; length xvec = length Tvec; distinct xvec  Prop (¿KN[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 (KN[xvec::=Tvec]) (P[xvec::=Tvec])"
      and rBrInput: "K Tvec. Ψ  K  M; set xvec  supp N; length xvec = length Tvec; distinct xvec  Prop (¿KN[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 (KN[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 (¿KN[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 "Ψ  MN⟩.P α  P'"
  and   "K. Ψ  M  K  Prop (KN) P"
  and   "K. Ψ  M  K  Prop (¡KN) 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: "Ψ  ⦇νxP α  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 α (⦇νxP')"
  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(⦇ν*xvecP'))"
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(⦇νxP)" and "distinct []"
    by(auto simp add: inputLength_inputLength'_inputLength''.simps)
  moreover have "length [] = inputLength(⦇νxP)" 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: "Ψ  ⦇νxP α  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 α (⦇νxP')"
  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(⦇ν*xvecP'))"

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(⦇νxP)" and "distinct []"
    by(auto simp add: inputLength_inputLength'_inputLength''.simps)
  moreover have "length [] = inputLength(⦇νxP)" 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: "Ψ  ⦇νxP MN  R"
  and   2: "x  Ψ"
  and   "x  (MN)"
  and   4: "x  R"
  and   rScope:  "P'. Ψ  P MN  P'  Prop (⦇νxP')"

shows "Prop R"
proof -
  from Trans obtain α where 1: "Ψ  ⦇νxP α  R" and 5: "bn α ♯* Ψ" and 6: "bn α ♯* P" and 7: "bn α ♯* subject α" and "α = MN" by auto
  from x  (MN) α = (MN) have 3: "x  α" by simp
  show ?thesis using 1 2 3 4 5 6 7 α = MN 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: "Ψ  ⦇νxP ¿MN  R"
  and   2: "x  Ψ"
  and   "x  (¿MN)"
  and   4: "x  R"
  and   rScope:  "P'. Ψ  P ¿MN  P'  Prop (⦇νxP')"

shows "Prop R"
proof -
  from Trans obtain α where 1: "Ψ  ⦇νxP α  R" and 5: "bn α ♯* Ψ" and 6: "bn α ♯* P" and 7: "bn α ♯* subject α" and "α = ¿MN" by auto
  from x  (¿MN) α = (¿MN) have 3: "x  α" by simp
  show ?thesis using 1 2 3 4 5 6 7 α = ¿MN 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: "Ψ  ⦇νxP 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 (⦇νxP')"

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: "Ψ  ⦇νxP α  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: "Ψ  ⦇νxP 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 (⦇νxP')"

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: "Ψ  ⦇νxP α  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 Ψ ΨQ P M N P' AP ΨP Q Q' AQ Ψ')
  then show ?case
    by(clarsimp, intro BrMerge) (blast intro: Composition)+
next
  case(cBrComm1 Ψ ΨQ P M N P' AP ΨP Q xvec Q' AQ Ψ')
  then show ?case
    by(clarsimp, intro BrComm1) (blast intro: Composition)+
next
  case(cBrComm2 Ψ ΨQ P M xvec N P' AP ΨP Q Q' AQ Ψ')
  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  ¿KN  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 Ψ ΨQ P M N P' AQ 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 Ψ ΨP Q M N Q' AP 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 Ψ ΨQ P M N P' AP ΨP Q Q' AQ)
  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 (⦇ν*xvecN ≺' 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 Ψ ΨQ P M xvec N P' AQ 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 Ψ ΨP Q M xvec N Q' AP 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 Ψ ΨQ P M N P' AP ΨP Q xvec Q' AQ)
  then have "((supp M)::name set)  ((supp Q)::name set)"
    by auto
  then show ?case
    by(auto simp add: psi.supp)
next
  case(cBrComm2 Ψ ΨQ P M xvec N P' AP ΨP Q Q' AQ)
  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 AP   :: "name list"
    and ΨP  :: 'b
    and C    :: "'f::fs_name"
    and C'   :: "'g::fs_name"

assumes Trans: "Ψ  P α  P'"
  and   "extractFrame P = AP, ΨP"
  and   "distinct AP"
  and   "bn α ♯* subject α"
  and   "AP ♯* P"
  and   "AP ♯* α"
  and   "AP ♯* C"
  and   "AP ♯* C'"
  and   "bn α ♯* P"
  and   "bn α ♯* C'"
  and   "α  τ"

obtains p Ψ' AP' ΨP' where "set p  set(bn α) × set(bn(p  α))" and "(p  ΨP)  Ψ'  ΨP'" and "distinctPerm p" and
  "extractFrame P' = AP', ΨP'" and "AP' ♯* P'" and "AP' ♯* α" and "AP' ♯* (p  α)" and
  "AP' ♯* C" and "(bn(p  α)) ♯* C'" and "(bn(p  α)) ♯* α" and "(bn(p  α)) ♯* P'" and "distinct AP'"
proof -
  assume A: "p Ψ' ΨP' AP'.
        set p  set(bn α) × set(bn(p  α)); (p  ΨP)  Ψ'  ΨP'; distinctPerm p;
                              extractFrame P' = AP', ΨP'; AP' ♯* P'; AP' ♯* α; AP' ♯* (p  α);
                              AP' ♯* C; (bn(p  α)) ♯* C'; (bn(p  α)) ♯* α; (bn(p  α)) ♯* P'; distinct AP'
         thesis"

  from Trans have "distinct(bn α)" by(auto dest: boundOutputDistinct)

  with Trans bn α ♯* subject α AP ♯* P AP ♯* α have "AP ♯* 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 "AP ♯* V" and "AP ♯* W" and "AP ♯* X" and "AP ♯* Y" and "AP ♯* Z"

    with assms obtain p Ψ' AP' ΨP' where "set p  set(bn α) × set(bn(p  α))" and "(p  ΨP)  Ψ'  ΨP'" and "distinctPerm p"
      and "extractFrame P' = AP', ΨP'" and "AP' ♯* P'" and "AP' ♯* α" and "AP' ♯* (p  α)"
      and "AP' ♯* C" and "(bn(p  α)) ♯* C'" and "(bn(p  α)) ♯* α" and "(bn(p  α)) ♯* P'"
      and "AP' ♯* V" and "AP' ♯* W" and "AP' ♯* X" and "AP' ♯* Y" and "AP' ♯* Z" and "distinct AP'"
      and "(bn(p  α)) ♯* V" and "(bn(p  α)) ♯* W" and "(bn(p  α)) ♯* X" and "(bn(p  α)) ♯* Y" and "(bn(p  α)) ♯* Z"
      using AP ♯* P' distinct(bn α)
    proof(nominal_induct Ψ P Rs=="α  P'" AP ΨP avoiding: C C' α P' V W X Y Z arbitrary: thesis rule: semanticsFrameInduct)
      case(cAlpha Ψ P AP ΨP p C C' α P' V W X Y Z)
      then obtain q Ψ' AP' ΨP' where Sq: "set q  set(bn α) × set(bn(q  α))" and PeqP': "(q  ΨP)  Ψ'  ΨP'" and "distinctPerm q"
        and FrP': "extractFrame P' = AP', ΨP'" and "AP' ♯* P'" and "AP' ♯* α" and "AP' ♯* (q  α)"
        and "AP' ♯* C" and "(bn(q  α)) ♯* C'" and "(bn(q  α)) ♯* α" and "(bn(q  α)) ♯* P'"
        and "AP' ♯* V" and "AP' ♯* W" and "AP' ♯* X" and "AP' ♯* Y" and "AP' ♯* Z" and "distinct AP'"
        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 AP × set (p  AP)" 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 AP ♯* α (p  AP) ♯* α Sp have "set(p  q)  set(bn α) × set(bn((p  q)  α))"
        by(simp add: perm_compose bnEqvt[symmetric])
      moreover from PeqP' have "(p  (q  ΨP)  Ψ')  (p  ΨP')"
        by(simp add: AssertionStatEqClosed)
      then have "((p  q)  p  ΨP)  (p  Ψ')  (p  Ψ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 AP ♯* α (p  AP) ♯* α AP ♯* C' (p  AP) ♯* C' Sp have "bn((p  q)  α) ♯* C'"
        by(simp add: perm_compose bnEqvt[symmetric])
      moreover from FrP' have "(p  extractFrame P') = p  AP', ΨP'" by simp
      with AP ♯* P' (p  AP) ♯* P' Sp have "extractFrame P' = p  AP', p  ΨP'"
        by(simp add: eqvts)
      moreover from AP' ♯* P' have "(p  AP') ♯* (p  P')"
        by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
      with AP ♯* P' (p  AP) ♯* P' Sp have "(p  AP') ♯* P'" by simp
      moreover from AP' ♯* α have "(p  AP') ♯* (p  α)"
        by(simp only: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
      with AP ♯* α (p  AP) ♯* α Sp have "(p  AP') ♯* α" by simp
      moreover from AP' ♯* C have "(p  AP') ♯* (p  C)"
        by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
      with AP ♯* C (p  AP) ♯* C Sp have "(p  AP') ♯* 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 AP ♯* α (p  AP) ♯* α AP ♯* α (p  AP) ♯* α 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 AP ♯* α (p  AP) ♯* α AP ♯* P' (p  AP) ♯* P' Sp have "bn((p  q)  α) ♯* P'"
        by(simp add: perm_compose eqvts)
      moreover from AP' ♯* (q  α) have "(p  AP') ♯* (p  q  α)"
        by(simp only: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
      with Sp AP ♯* α (p  AP) ♯* α have "(p  AP') ♯* ((p  q)  α)"
        by(simp add: perm_compose)
      moreover from AP' ♯* V have "(p  AP') ♯* (p  V)"
        by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
      with AP ♯* V (p  AP) ♯* V Sp have "(p  AP') ♯* V" by simp
      moreover from AP' ♯* W have "(p  AP') ♯* (p  W)"
        by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
      with AP ♯* W (p  AP) ♯* W Sp have "(p  AP') ♯* W" by simp
      moreover from AP' ♯* X have "(p  AP') ♯* (p  X)"
        by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
      with AP ♯* X (p  AP) ♯* X Sp have "(p  AP') ♯* X" by simp
      moreover from AP' ♯* Y have "(p  AP') ♯* (p  Y)"
        by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
      with AP ♯* Y (p  AP) ♯* Y Sp have "(p  AP') ♯* Y" by simp
      moreover from AP' ♯* Z have "(p  AP') ♯* (p  Z)"
        by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
      with AP ♯* Z (p  AP) ♯* Z Sp have "(p  AP') ♯* 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 AP ♯* α (p  AP) ♯* α AP ♯* V (p  AP) ♯* 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 AP ♯* α (p  AP) ♯* α AP ♯* W (p  AP) ♯* 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 AP ♯* α (p  AP) ♯* α AP ♯* X (p  AP) ♯* 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 AP ♯* α (p  AP) ♯* α AP ♯* Y (p  AP) ♯* 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 AP ♯* α (p  AP) ♯* α AP ♯* Z (p  AP) ♯* Z Sp have "bn((p  q)  α) ♯* Z"
        by(simp add: perm_compose eqvts)
      moreover from distinct AP' have "distinct(p  AP')" 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 AP ΨP where "extractFrame(P[xvec::=Tvec]) = AP, ΨP" and "distinct AP"
        and "AP ♯* (C, P[xvec::=Tvec], α, P', V, W, X, Y, Z, N)"
        by(rule freshFrame)
      moreover have "𝟭  ΨP  Ψ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 AP ΨP where "extractFrame(P[xvec::=Tvec]) = AP, ΨP" and "distinct AP"
        and "AP ♯* (C, P[xvec::=Tvec], α, P', V, W, X, Y, Z, N)"
        by(rule freshFrame)
      moreover have "𝟭  ΨP  Ψ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 AP ΨP where "extractFrame P = AP, ΨP" and "distinct AP"
        and "AP ♯* (C, C', P, α, N, P', V, W, X, Y, Z)"
        by(rule freshFrame)
      moreover have "𝟭  ΨP  Ψ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 AP ΨP where "extractFrame P = AP, ΨP" and "distinct AP"
        and "AP ♯* (C, C', P, α, N, P', V, W, X, Y, Z)"
        by(rule freshFrame)
      moreover have "𝟭  ΨP  ΨP"
        by(blast intro: Identity Commutativity AssertionStatEqTrans)
      ultimately show ?case by(simp add: residualInject)
    next
      case(cCase Ψ P φ Cs AP Ψ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 Ψ' AP' ΨP' where S: "set p  set(bn α) × set(bn(p  α))"
        and FrP': "extractFrame P' = AP', ΨP'"
        and PeqP': "(p  ΨP)  Ψ'  ΨP'" and "distinct AP'"
        and "AP' ♯* C" and "AP' ♯* P'" and "AP' ♯* α" and "AP' ♯* (p  α)"
        and "AP' ♯* V" and "AP' ♯* W" and "AP' ♯* X" and "AP' ♯* Y" and "AP' ♯* 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 ΨP  𝟭 have "(p  ΨP)  (p  𝟭)"
        by(simp add: AssertionStatEqClosed)
      then have "(p  ΨP)  𝟭" by(simp add: permBottom)
      with PeqP' have "(𝟭  Ψ')  Ψ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 Ψ ΨQ P α P' AQ Q AP ΨP C C' α' PQ' V W X Y Z)
      have FrP: "extractFrame P = AP, ΨP" and  FrQ: "extractFrame Q = AQ, ΨQ"
        by fact+

      note bn α' ♯* subject α'
      moreover from bn α' ♯* (P  Q) have "bn α' ♯* P" and "bn α' ♯* Q" by simp+
      moreover with FrP FrQ AP ♯* α' AQ ♯* α' have "bn α' ♯* ΨP" and "bn α' ♯* ΨQ"
        by(force dest: extractFrameFreshChain)+

      moreover note bn α' ♯* V bn α' ♯* W
      moreover from bn α' ♯* X AQ ♯* α' have "bn α' ♯* (X@AQ)" by simp
      moreover from bn α' ♯* Y bn α' ♯* ΨQ have "bn α' ♯* (ΨQ#Y)" by simp
      moreover from bn α' ♯* Z bn α' ♯* Q have "bn α' ♯* (Q#Z)" by simp
      moreover note AP ♯* V AP ♯* W
      moreover from AP ♯* X AP ♯* AQ have "AP ♯* (X@AQ)" by simp
      moreover from AP ♯* Y AP ♯* ΨQ have "AP ♯* (ΨQ#Y)" by force
      moreover from AP ♯* Z AP ♯* Q have "AP ♯* (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 AP ♯* PQ' PQ' = P''  Q have "AP ♯* P''" by simp
      ultimately obtain p Ψ' ΨP' AP' where S: "set p  set(bn α') × set (bn(p  α'))" and PeqP': "((p  ΨP)  Ψ')  ΨP'"
        and "distinctPerm p" and "(bn(p  α')) ♯* C'" and FrP': "extractFrame P'' = AP', ΨP'"
        and "AP' ♯* P''" and "AP' ♯* α'" "AP' ♯* (p  α')" and "AP' ♯* C"
        and "(bn(p  α')) ♯* α'" and "(bn(p  α')) ♯* P''" and "distinct AP'"
        and "AP' ♯* V" and "AP' ♯* W" and "AP' ♯* (X @ AQ)" and "AP' ♯* (ΨQ#Y)"
        and "AP' ♯* (Q#Z)" and "(bn(p  α')) ♯* V" and "(bn(p  α')) ♯* W" and "(bn(p  α')) ♯* (X @ AQ)" and "(bn(p  α')) ♯* (ΨQ#Y)"
        and "(bn(p  α')) ♯* (Q#Z)" using cPar1
        by(elim cPar1)

      then have "AP' ♯* Q" and "AP' ♯* Z" and "AP' ♯* AQ" and "AP' ♯* X" and "AP' ♯* ΨQ"  and "AP' ♯* Y"
        and "(bn(p  α')) ♯* AQ" and "(bn(p  α')) ♯* X" and "(bn(p  α')) ♯* Y" and "(bn(p  α')) ♯* Z" and "(bn(p  α')) ♯* ΨQ"
        and "(bn(p  α')) ♯* Q"
        by(simp del: freshChainSimps)+

      from AQ ♯* PQ' PQ' = P''  Q AP' ♯* AQ FrP' have "AQ ♯* ΨP'"
        by(force dest: extractFrameFreshChain)
      note S
      moreover from PeqP' have "((p  (ΨP  ΨQ))  Ψ')  ΨP'  (p  ΨQ)"
        by(simp add: eqvts) (metis Composition Associativity AssertionStatEqTrans AssertionStatEqSym Commutativity)
      with (bn(p  α')) ♯* ΨQ bn α' ♯* ΨQ S have "((p  (ΨP  ΨQ))  Ψ')  ΨP'  ΨQ"
        by simp
      moreover from PQ' = P''  Q AP' ♯* AQ AP' ♯* ΨQ AQ ♯* ΨP' AQ ♯* PQ' FrP' FrQ have "extractFrame PQ' = (AP'@AQ), ΨP'  ΨQ"
        by simp
      moreover note distinctPerm p (bn(p  α')) ♯* C'
      moreover from AP' ♯* P'' AP' ♯* Q PQ' = P''  Q have "AP' ♯* PQ'" by simp
      moreover note AQ ♯* PQ' AP' ♯* α' AQ ♯* α' AP' ♯* C AQ ♯* 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 AP' ♯* α' AQ ♯* α' have "(AP'@AQ) ♯* α'" by simp
      moreover from AP' ♯* C AQ ♯* C have "(AP'@AQ) ♯* C" by simp
      moreover from AP' ♯* V AQ ♯* V have "(AP'@AQ) ♯* V" by simp
      moreover from AP' ♯* W AQ ♯* W have "(AP'@AQ) ♯* W" by simp
      moreover from AP' ♯* X AQ ♯* X have "(AP'@AQ) ♯* X" by simp
      moreover from AP' ♯* Y AQ ♯* Y have "(AP'@AQ) ♯* Y" by simp
      moreover from AP' ♯* Z AQ ♯* Z have "(AP'@AQ) ♯* Z" by simp
      moreover from AP' ♯* PQ' AQ ♯* PQ' have "(AP'@AQ) ♯* PQ'" by simp
      moreover from AP' ♯* α' AQ ♯* α' have "(AP'@AQ) ♯* α'" by simp
      moreover from AQ ♯* α' have "(p  AQ) ♯* (p  α')"
        by(simp only: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst] bnEqvt[symmetric])
      with S AQ ♯* α' bn(p  α') ♯* AQ have "AQ ♯* (p  α')"
        by simp
      with AP' ♯* (p  α') AQ ♯* α' bn(p  α') ♯* AQ S have "(AP'@AQ) ♯* (p  α')"
        by simp
      moreover from AP' ♯* AQ distinct AP' distinct AQ have "distinct(AP'@AQ)" 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 Ψ ΨP Q α Q' AP P AQ ΨQ C C' α' PQ' V W X Y Z)
      have FrP: "extractFrame P = AP, ΨP" and  FrQ: "extractFrame Q = AQ, ΨQ"
        by fact+

      note bn α' ♯* subject α'
      moreover from bn α' ♯* (P  Q) have "bn α' ♯* Q" and "bn α' ♯* P" by simp+
      moreover with FrP FrQ AP ♯* α' AQ ♯* α' have "bn α' ♯* ΨP" and "bn α' ♯* ΨQ"
        by(force dest: extractFrameFreshChain)+

      moreover note bn α' ♯* V bn α' ♯* W
      moreover from bn α' ♯* X AP ♯* α' have "bn α' ♯* (X@AP)" by simp
      moreover from bn α' ♯* Y bn α' ♯* ΨP have "bn α' ♯* (ΨP#Y)" by simp
      moreover from bn α' ♯* Z bn α' ♯* P have "bn α' ♯* (P#Z)" by simp
      moreover note AQ ♯* V AQ ♯* W
      moreover from AQ ♯* X AP ♯* AQ have "AQ ♯* (X@AP)" by simp
      moreover from AQ ♯* Y AQ ♯* ΨP have "AQ ♯* (ΨP#Y)" by force
      moreover from AQ ♯* Z AQ ♯* P have "AQ ♯* (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 AQ ♯* PQ' PQ' = P  Q'' have "AQ ♯* Q''" by simp
      ultimately obtain p Ψ' AQ' ΨQ' where S: "set p  set(bn α') × set (bn(p  α'))" and QeqQ': "((p  ΨQ)  Ψ')  ΨQ'"
        and "distinctPerm p" and "(bn(p  α')) ♯* C'" and FrQ': "extractFrame Q'' = AQ', ΨQ'"
        and "AQ' ♯* Q''" and "AQ' ♯* α'" "AQ' ♯* (p  α')" and "AQ' ♯* C"
        and "(bn(p  α')) ♯* α'" and "(bn(p  α')) ♯* Q''" and "distinct AQ'"
        and "AQ' ♯* V" and "AQ' ♯* W" and "AQ' ♯* (X @ AP)" and "AQ' ♯* (ΨP#Y)"
        and "AQ' ♯* (P#Z)" and "(bn(p  α')) ♯* V" and "(bn(p  α')) ♯* W" and "(bn(p  α')) ♯* (X @ AP)" and "(bn(p  α')) ♯* (ΨP#Y)"
        and "(bn(p  α')) ♯* (P#Z)" using cPar2
        by(elim cPar2)

      then have "AQ' ♯* P" and "AQ' ♯* Z" and "AQ' ♯* AP" and "AQ' ♯* X" and "AQ' ♯* ΨP"  and "AQ' ♯* Y"
        and "(bn(p  α')) ♯* AP" and "(bn(p  α')) ♯* X" and "(bn(p  α')) ♯* Y" and "(bn(p  α')) ♯* Z" and "(bn(p  α')) ♯* ΨP"
        and "(bn(p  α')) ♯* P"
        by(simp del: freshChainSimps)+

      from AP ♯* PQ' PQ' = P  Q'' AQ' ♯* AP FrQ' have "AP ♯* ΨQ'"
        by(force dest: extractFrameFreshChain)
      note S
      moreover from QeqQ' have "((p  (ΨP  ΨQ))  Ψ')  (p  ΨP)  ΨQ'"
        by(simp add: eqvts) (metis Composition Associativity AssertionStatEqTrans AssertionStatEqSym Commutativity)
      with (bn(p  α')) ♯* ΨP bn α' ♯* ΨP S have "((p  (ΨP  ΨQ))  Ψ')  ΨP  ΨQ'"
        by simp
      moreover from PQ' = P  Q'' AQ' ♯* AP AQ' ♯* ΨP AP ♯* ΨQ' AP ♯* PQ' FrQ' FrP have "extractFrame PQ' = (AP@AQ'), ΨP  ΨQ'"
        by simp
      moreover note distinctPerm p (bn(p  α')) ♯* C'
      moreover from AQ' ♯* Q'' AQ' ♯* P PQ' = P  Q'' have "AQ' ♯* PQ'" by simp
      moreover note AQ ♯* PQ' AQ' ♯* α' AQ ♯* α' AQ' ♯* C AQ ♯* 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 AQ' ♯* α' AP ♯* α' have "(AP@AQ') ♯* α'" by simp
      moreover from AQ' ♯* C AP ♯* C have "(AP@AQ') ♯* C" by simp
      moreover from AQ' ♯* V AP ♯* V have "(AP@AQ') ♯* V" by simp
      moreover from AQ' ♯* W AP ♯* W have "(AP@AQ') ♯* W" by simp
      moreover from AQ' ♯* X AP ♯* X have "(AP@AQ') ♯* X" by simp
      moreover from AQ' ♯* Y AP ♯* Y have "(AP@AQ') ♯* Y" by simp
      moreover from AQ' ♯* Z AP ♯* Z have "(AP@AQ') ♯* Z" by simp
      moreover from AQ' ♯* PQ' AP ♯* PQ' have "(AP@AQ') ♯* PQ'" by simp
      moreover from AQ' ♯* α' AP ♯* α' have "(AP@AQ') ♯* α'" by simp
      moreover from AP ♯* α' have "(p  AP) ♯* (p  α')"
        by(simp only: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst] bnEqvt[symmetric])
      with S AP ♯* α' bn(p  α') ♯* AP have "AP ♯* (p  α')"
        by simp
      with AQ' ♯* (p  α') AP ♯* α' bn(p  α') ♯* AP S have "(AP@AQ') ♯* (p  α')"
        by simp
      moreover note (bn(p  α')) ♯* V (bn(p  α')) ♯* W (bn(p  α')) ♯* X (bn(p  α')) ♯* Y (bn(p  α')) ♯* Z
      moreover from AQ' ♯* AP distinct AP distinct AQ' have "distinct(AP@AQ')" 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 Ψ ΨQ P M N P' AP ΨP Q Q' AQ C C' α PQ' V W X Y Z)
      have FrP: "extractFrame P = AP, ΨP" and  FrQ: "extractFrame Q = AQ, ΨQ"
        by fact+

      have "bn (¿MN) ♯* Q" by simp
      have "bn (¿MN) ♯* Q'" by simp
      have "bn (¿MN) ♯* α" by simp
      have "bn (¿MN) ♯* P" by simp
      have "bn (¿MN) ♯* P'" by simp
      have "bn (¿MN) ♯* α" by simp

      from ¿MN  P'  Q' = α  PQ'
      have emptyα: "bn α = ([]::name list)" and
        "α = (¿MN)"
        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 α ♯* ΨP" and "bn α ♯* ΨQ" and "bn α ♯* C'"
        and "bn α ♯* (X@AQ)"
        and "bn α ♯* (ΨQ#Y)" and "bn α ♯* (Q#Z)" by simp+

      moreover note AP ♯* P AP ♯* α AP ♯* C AP ♯* C'
        α  τ

      moreover from AP ♯* X AP ♯* AQ have "AP ♯* (X@AQ)" by simp
      moreover from AP ♯* Y AP ♯* ΨQ have "AP ♯* (ΨQ#Y)" by force
      moreover from AP ♯* Z AP ♯* Q have "AP ♯* (Q#Z)" by simp
      moreover from ¿MN  (P'  Q') = α  PQ' bn (¿MN) ♯* Q' bn α ♯* Q' bn (¿MN) ♯* α
      obtain P'' where A: "¿MN  P' = α  P''" and "PQ' = P''  Q'"
        by(metis actionPar1Dest')
      moreover from AP ♯* PQ' PQ' = P''  Q' have "AP ♯* P''" by simp
      ultimately obtain p PΨ' AP' ΨP' where Sp: "set p  set(bn α) × set (bn(p  α))" and PeqP': "((p  ΨP)  PΨ')  ΨP'"
        and "distinctPerm p" and FrP': "extractFrame P'' = AP', ΨP'"
        and "AP' ♯* P''" and "AP' ♯* α" "AP' ♯* (p  α)" and "AP' ♯* C" and "(bn(p  α)) ♯* C'"
        and "(bn(p  α)) ♯* α" and "(bn(p  α)) ♯* P''" and "distinct AP'"
        and "AP' ♯* V" and "AP' ♯* W" and "AP' ♯* (X @ AQ)" and "AP' ♯* (ΨQ#Y)"
        and "AP' ♯* (Q#Z)" and "(bn(p  α)) ♯* V" and "(bn(p  α)) ♯* W" and "(bn(p  α)) ♯* (X @ AQ)" and "(bn(p  α)) ♯* (ΨQ#Y)"
        and "(bn(p  α)) ♯* (Q#Z)" using cBrMerge
        by(elim cBrMerge)

      then have "AP' ♯* Q" and "AP' ♯* Z" and "AP' ♯* AQ" and "AP' ♯* X" and "AP' ♯* ΨQ"  and "AP' ♯* Y"
        and "(bn(p  α)) ♯* AQ" and "(bn(p  α)) ♯* X" and "(bn(p  α)) ♯* Y" and "(bn(p  α)) ♯* Z" and "(bn(p  α)) ♯* ΨQ"
        and "(bn(p  α)) ♯* Q"
        by(simp del: freshChainSimps)+

      from AQ ♯* PQ' PQ' = (P''  Q') have "AQ ♯* P''" by simp
      with extractFrame P'' = AP', ΨP' AP' ♯* AQ have "AQ ♯* ΨP'"
        by (metis extractFrameFreshChain freshFrameDest)

      from bn α = []
      have "bn α ♯* subject α" and "bn α ♯* Q" and "bn α ♯* P''"
        and "bn α ♯* ΨQ" and "bn α ♯* ΨP'" and "bn α ♯* C'"
        and "bn α ♯* (X@AP')"
        and "bn α ♯* (ΨP'#Y)" and "bn α ♯* (P''#Z)" by simp+

      moreover note AP' ♯* Q AQ ♯* α AQ ♯* C AQ ♯* C'
        α  τ

      moreover from AQ ♯* X AP' ♯* AQ have "AQ ♯* (X@AP')" by simp
      moreover from AQ ♯* Y AQ ♯* ΨP' have "AQ ♯* (ΨP'#Y)" by force
      moreover from AQ ♯* Z AQ ♯* P'' have "AQ ♯* (P''#Z)" by simp
      moreover from ¿MN  (P'  Q') = α  PQ' bn (¿MN) ♯* P' bn α ♯* P' bn (¿MN) ♯* α
      obtain Q'' where A: "¿MN  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 AQ ♯* PQ' PQ' = P''  Q'' have "AQ ♯* Q''" by simp
      ultimately obtain q Ψ' AQ' ΨQ' where Sq: "set q  set(bn α) × set (bn(q  α))" and QeqQ': "((q  ΨQ)  Ψ')  ΨQ'"
        and "distinctPerm q" and FrQ': "extractFrame Q'' = AQ', ΨQ'"
        and "AQ' ♯* Q''" and "AQ' ♯* α" "AQ' ♯* (q  α)" and "AQ' ♯* C" and "(bn(q  α)) ♯* C'"
        and "(bn(q  α)) ♯* α" and "(bn(q  α)) ♯* Q''" and "distinct AQ'"
        and "AQ' ♯* V" and "AQ' ♯* W" and "AQ' ♯* (X @ AP')" and "AQ' ♯* (ΨP'#Y)"
        and "AQ' ♯* (P''#Z)" and "(bn(q  α)) ♯* V" and "(bn(q  α)) ♯* W" and "(bn(q  α)) ♯* (X @ AP')" and "(bn(q  α)) ♯* (ΨP'#Y)"
        and "(bn(q  α)) ♯* (P''#Z)" using cBrMerge
        by(elim cBrMerge(6)[where bb=α]) (rule refl | assumption)+

      then have "AQ' ♯* P''" and "AQ' ♯* Z" and "AQ' ♯* AP'" and "AQ' ♯* X" and "AQ' ♯* ΨP'"  and "AQ' ♯* Y"
        and "(bn(q  α)) ♯* AP'" and "(bn(q  α)) ♯* X" and "(bn(q  α)) ♯* Y" and "(bn(q  α)) ♯* Z" and "(bn(q  α)) ♯* Ψ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 AQ' ♯* P'' AQ' ♯* AP' extractFrame P'' = AP', ΨP' have "AQ' ♯* ΨP'"
        by (metis extractFrameFreshChain freshFrameDest)
      from AP' ♯* α α = ¿MN have "AP' ♯* M" and "AP' ♯* N"
        by simp+

      from Ψ  ΨP  Q  ¿MN  Q' AP' ♯* Q AP' ♯* N
      have "AP' ♯* Q'"
        by(simp add: brinputFreshChainDerivative)
      with PQ' = (P''  Q') PQ' = (P''  Q'') have "AP' ♯* Q''"
        by (simp add: psi.inject)
      with FrQ' AQ' ♯* AP' have "AP' ♯* ΨQ'"
        by(metis extractFrameFreshChain freshFrameDest)
      from AP' ♯* P'' AP' ♯* Q'' PQ' = (P''  Q'')
      have "AP' ♯* PQ'" by simp
      from AQ' ♯* P'' AQ' ♯* Q'' PQ' = (P''  Q'')
      have "AQ' ♯* PQ'" by simp


      from PeqP' have "((p  (ΨP  ΨQ))  PΨ')  ΨP'  (p  ΨQ)"
        by(simp add: eqvts) (metis Composition Associativity AssertionStatEqTrans AssertionStatEqSym Commutativity)

      with p = [] have "((ΨP  ΨQ)  PΨ')  (ΨP'  ΨQ)" by simp
      with QeqQ' have "((q  (ΨQ  ΨP'))  Ψ')  ΨQ'  (q  ΨP')"
        by(simp add: eqvts) (metis Composition Associativity AssertionStatEqTrans AssertionStatEqSym Commutativity)
      with PeqP' p = [] q = [] have "((q  (ΨQ  ((p  ΨP)  PΨ')))  Ψ')  ΨQ'  (q  ΨP')"
        by(simp add: eqvts) (metis AssertionStatEqTrans Composition compositionSym)
      with p = [] have "((q  (ΨQ  (ΨP  PΨ')))  Ψ')  ΨQ'  (q  ΨP')" by simp
      with q = [] have "(q  (ΨP  ΨQ))  (PΨ'  Ψ')  ΨP'  ΨQ'"
        by(simp add: eqvts) (metis AssertionStatEqTrans Commutativity Composition associativitySym)

      moreover note set q  set(bn α) × set (bn (q  α))

      moreover from PQ' = P''  Q'' AQ' ♯* AP' AP' ♯* ΨQ' AQ' ♯* ΨP' AQ' ♯* PQ' FrP' FrQ' have "extractFrame PQ' = (AP'@AQ'), ΨP'  ΨQ'"
        by simp

      moreover note distinctPerm q

      moreover from AP' ♯* PQ' AQ' ♯* PQ'
      have "(AP'@AQ') ♯* PQ'" by simp
      moreover from AP' ♯* α AQ' ♯* α
      have "(AP'@AQ') ♯* α" by simp
      moreover with q = []
      have "(AP'@AQ') ♯* (q  α)" by simp
      moreover from AP' ♯* C AQ' ♯* C
      have "(AP'@AQ') ♯* 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 AP' ♯* V AQ' ♯* V
      have "(AP'@AQ') ♯* V" by simp
      moreover from AP' ♯* W AQ' ♯* W
      have "(AP'@AQ') ♯* W" by simp
      moreover from AP' ♯* X AQ' ♯* X
      have "(AP'@AQ') ♯* X" by simp
      moreover from AP' ♯* Y AQ' ♯* Y
      have "(AP'@AQ') ♯* Y" by simp
      moreover from AP' ♯* Z AQ' ♯* Z
      have "(AP'@AQ') ♯* Z" by simp
      moreover from AQ' ♯* AP' distinct AP' distinct AQ' have "distinct(AP'@AQ')" 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 Ψ ΨQ P M N P' AP ΨP Q xvec Q' AQ C C' α PQ' V W X Y Z)
      have FrP: "extractFrame P = AP, ΨP" and  FrQ: "extractFrame Q = AQ, Ψ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 AQ ♯* PQ' PQ' = (r  P')  Q'' have "AQ ♯* Q''"
        by simp

      from AP ♯* PQ' PQ' = (r  P')  Q'' have "AP ♯* Q''"
        by simp

      from bn α ♯* (P  Q) AP ♯* α
      have "bn α ♯* P" and "AP ♯* bn α" by simp+
      with extractFrame P = AP, ΨP have "bn α ♯* ΨP"
        by (metis extractFrameFreshChain freshFrameDest)

      from bn α ♯* (P  Q) AQ ♯* α
      have "bn α ♯* Q" and "AQ ♯* bn α" by simp+
      with extractFrame Q = AQ, ΨQ have "bn α ♯* ΨQ"
        by (metis extractFrameFreshChain freshFrameDest)

      from rPerm AP ♯* xvec xvec ♯* ΨP AP ♯* bn α bn α ♯* ΨP
      have "r  AP = AP" and "r  ΨP = ΨP" by simp+

      have "¿MN  P' = ¿MN  P'" by simp
      moreover note AP ♯* P AP ♯* C AP ♯* C'
      moreover from AP ♯* M AP ♯* N have "AP ♯* (¿MN)" by simp
      moreover note AP ♯* V
      moreover from AP ♯* W AP ♯* α have "AP ♯* (α#W)" by simp
      moreover from AP ♯* X AP ♯* AQ AP ♯* xvec have "AP ♯* (X@AQ@xvec)" by simp
      moreover from AP ♯* Y AP ♯* ΨQ have "AP ♯* (ΨQ#Y)" by force
      moreover from AP ♯* Z AP ♯* Q AP ♯* Q' AP ♯* Q'' have "AP ♯* (Q#(Q'#(Q''#Z)))" by simp
      moreover note AP ♯* P' AP ♯* Q

      ultimately obtain p PΨ' AP' ΨP' where Sp: "set p  set(bn (¿MN)) × set (bn(p  (¿MN)))" and "((p  ΨP)  PΨ')  ΨP'"
        and "distinctPerm p" and FrP': "extractFrame P' = AP', ΨP'"
        and "AP' ♯* P'" and "AP' ♯* (¿MN)" "AP' ♯* (p  (¿MN))" and "AP' ♯* C" and "(bn(p  (¿MN))) ♯* C'"
        and "(bn(p  (¿MN))) ♯* (¿MN)" and "(bn(p  (¿MN))) ♯* P'" and "distinct AP'"
        and "AP' ♯* V" and "AP' ♯* (α#W)" and "AP' ♯* (X @ AQ @ xvec)" and "AP' ♯* (ΨQ#Y)"
        and "AP' ♯* (Q#(Q'#(Q''#Z)))" and "(bn(p  (¿MN))) ♯* (α#W)" and "(bn(p  (¿MN))) ♯* (X @ AQ @ xvec)" and "(bn(p  (¿MN))) ♯* (ΨQ#Y)"
        and "(bn(p  (¿MN))) ♯* (Q#(Q'#(Q''#Z)))"
        by(elim cBrComm1(4)) (assumption | simp)+

      then have PeqP': "ΨP  PΨ'  ΨP'"
        and "AP' ♯* P'" and "AP' ♯* (¿MN)" and "AP' ♯* C" and "distinct AP'"
        and "AP' ♯* Q" and "AP' ♯* Q'" and "AP' ♯* Q''" and "AP' ♯* Z" and "AP' ♯* AQ" and "AP' ♯* X" and "AP' ♯* ΨQ" and "AP' ♯* Y"
        and "AP' ♯* xvec" and "AP' ♯* α" and "AP' ♯* (subject α)" and "AP' ♯* (bn α)" and "AP' ♯* (object α)"
        and "AP' ♯* V" and "AP' ♯* W"
        by(simp del: freshChainSimps)+

      from rPerm bn α ♯* ΨP xvec ♯* ΨP
      have "(r  ΨP) = ΨP" by simp

      from PeqP' have "r  (ΨP  PΨ'  ΨP')"
        by simp
      with (r  ΨP) = ΨP
      have rPeqP': "ΨP  (r  PΨ')  (r  ΨP')" by(simp add: eqvts)

      from rPerm AP' ♯* (bn α) AP' ♯* xvec
      have "r  AP' = AP'" by simp

      from FrP' have "r  (extractFrame P' = AP', ΨP')"
        by simp

      with r  AP' = AP' have rFrP': "extractFrame (r  P') = AP', (r  ΨP')"
        by(simp add: eqvts)

      from xvec ♯* α have "(bn α) ♯* xvec" by simp

      from (AP @ AQ) ♯* α have "AQ ♯* bn α" by simp

      from AQ ♯* M AQ ♯* xvec AQ ♯* N have "AQ ♯* (¡M⦇ν*xvec⦈⟨N)" by simp
      from Ψ  ΨP  Q  ¡M⦇ν*xvec⦈⟨N  Q' xvec ♯* M distinct xvec bn α ♯* Q xvec ♯* bn α
      have "bn α ♯* Q'" by(simp add: broutputFreshChainDerivative)

      from Ψ  Ψ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 α ♯* ΨP bn α ♯* ΨQ bn α ♯* W
      moreover from bn α ♯* X AP ♯* bn α AP' ♯* bn α bn α ♯* xvec have "bn α ♯* (X@xvec@AP@AP')" by simp
      moreover from bn α ♯* Y bn α ♯* ΨP have "bn α ♯* (ΨP#Y)" by simp
      moreover from bn α ♯* Z bn α ♯* P bn α ♯* Q have "bn α ♯* (P#Q#Z)" by simp
      moreover from AQ ♯* V AQ ♯* N have "AQ ♯* (N#V)" by simp
      moreover note AQ ♯* W
      moreover from AQ ♯* X AP ♯* AQ AP' ♯* AQ AQ ♯* xvec have "AQ ♯* (X@xvec@AP@AP')" by simp
      moreover from AQ ♯* Y AQ ♯* ΨP have "AQ ♯* (ΨP#Y)" by force
      moreover from AQ ♯* Z AQ ♯* P AQ ♯* Q have "AQ ♯* (P#Q#Z)" by simp

      moreover note ¡M⦇ν*xvec⦈⟨N  Q' = α  Q''
      ultimately obtain q QΨ' AQ' ΨQ' where Sq: "set q  set (bn α) × set (bn (q  α))" and QeqQ': "((q  ΨQ)  QΨ')  ΨQ'"
        and "distinctPerm q" and "bn(q  α) ♯* C'" and FrQ': "extractFrame Q'' = AQ', ΨQ'"
        and "AQ' ♯* Q''" and "AQ' ♯* α" and "AQ' ♯* (q  α)" and "AQ' ♯* C"
        and "bn(q  α) ♯* α" and "bn(q  α) ♯* Q''" and "distinct AQ'"
        and "AQ' ♯* (N#V)" and "AQ' ♯* W" and "AQ' ♯* (X @ xvec @ AP @ AP')" and "AQ' ♯* (ΨP#Y)"
        and "AQ' ♯* (P#Q#Z)" and "bn(q  α) ♯* (N#V)" and "bn(q  α) ♯* W" and "bn(q  α) ♯* (X @ xvec @ AP @ AP')" and "bn(q  α) ♯* (ΨP#Y)"
        and "bn(q  α) ♯* (P#Q#Z)"
        using AQ ♯* Q AQ ♯* α AQ ♯* C AQ ♯* C'
          bn α ♯* C' α  τ AQ ♯* Q'' distinct (bn α)
        by(elim cBrComm1(8)[where b=C and ba=C' and bb=α and bc=Q'' and bf="(X @ xvec @ AP @ AP')" and bg="(ΨP#Y)" and bh="(P#Q#Z)"]) (assumption | simp)+
      then have "AQ' ♯* P" and "AQ' ♯* Z" and "AQ' ♯* AP" and "AQ' ♯* AP'" and "AQ' ♯* X" and "AQ' ♯* ΨP" and "AQ' ♯* Y" and "AQ' ♯* Q" and "AQ' ♯* N"
        and "AQ' ♯* xvec" and "AQ' ♯* V" and "AQ' ♯* W"
        and "bn(q  α) ♯* AP" and "bn(q  α) ♯* AP'" and "bn(q  α) ♯* X" and "bn(q  α) ♯* Y" and "bn(q  α) ♯* Z" and "bn(q  α) ♯* Ψ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 AP' ♯* Q'' AQ' ♯* AP' FrQ'
      have "AP' ♯* ΨQ'"
        by(metis extractFrameFreshChain freshFrameDest)

      from Ψ  ΨQ  P  ¿MN  P' AQ' ♯* P AQ' ♯* N
      have "AQ' ♯* P'"
        by(simp add: brinputFreshChainDerivative)
      with FrP' AQ' ♯* AP' have "AQ' ♯* ΨP'"
        by(metis extractFrameFreshChain freshFrameDest)

      have "(ΨP  (r  PΨ'))  ((q  ΨQ)  QΨ')  (r  ΨP')  ΨQ'"
        by(metis Composition' rPeqP' QeqQ')
      then have "(ΨP  ((r  PΨ')  ((q  ΨQ)  QΨ')))  (r  ΨP')  ΨQ'"
        by (metis AssertionStatEqSym AssertionStatEqTrans Associativity)
      then have "(ΨP  (((q  ΨQ)  QΨ')  (r  PΨ')))  (r  ΨP')  ΨQ'"
        by (metis AssertionStatEqSym AssertionStatEqTrans Associativity associativitySym)
      then have "(ΨP  ((q  ΨQ)  (QΨ'  (r  PΨ'))))  (r  ΨP')  ΨQ'"
        by (metis AssertionStatEqSym AssertionStatEqTrans Associativity compositionSym)
      then have "((ΨP  (q  ΨQ))  (QΨ'  (r  PΨ')))  (r  ΨP')  ΨQ'"
        by (metis AssertionStatEqTrans Associativity)
      then have "((ΨP  (q  ΨQ))  ((r  PΨ')  QΨ'))  (r  ΨP')  ΨQ'"
        by (metis AssertionStatEqSym AssertionStatEqTrans Associativity associativitySym)
      with Sq bn α ♯* ΨP bn(q  α) ♯* ΨP have "((q  (ΨP  ΨQ))  ((r  PΨ')  QΨ'))  (r  ΨP')  ΨQ'"
        by(simp add: eqvts)

      from Sq AP' ♯* α bn(q  α) ♯* AP' have "AP' ♯* (q  α)"
        by (metis actionFreshChain freshChainSym freshStarChainSimps fresh_star_set_eq)

      from AQ' ♯* α have "AQ' ♯* bn α" by simp
      from rPerm AP' ♯* P' AP' ♯* xvec AP' ♯* bn α
      have "AP' ♯* (r  P')"
        by (metis freshAlphaPerm freshChainSym name_list_set_fresh permStarFresh)
      from rPerm AQ' ♯* P' AQ' ♯* xvec AQ' ♯* bn α
      have "AQ' ♯* (r  P')"
        by (metis freshAlphaPerm freshChainSym name_list_set_fresh permStarFresh)
      from rPerm AQ' ♯* ΨP' AQ' ♯* xvec AQ' ♯* bn α
      have "AQ' ♯* (r  ΨP')"
        by (metis freshAlphaPerm freshChainSym name_list_set_fresh permStarFresh)

      with AP' ♯* ΨQ' AQ' ♯* AP' rFrP' FrQ'
      have "extractFrame ((r  P')  Q'') = (AP'@AQ'), ((r  ΨP')  ΨQ')"
        by simp
      with PQ' = ((r  P')  Q'')
      have "extractFrame PQ' = (AP'@AQ'), ((r  ΨP')  ΨQ')"
        by simp

      from Ψ  ΨQ  P  ¿MN  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 AP' ♯* Q'' AP' ♯* (r  P') PQ' = (r  P')  Q'' have "AP' ♯* PQ'" by simp
      from AQ' ♯* Q'' AQ' ♯* (r  P') PQ' = (r  P')  Q'' have "AQ' ♯* PQ'" by simp

      note set q  (set (bn α)) × set (bn(q  α))
        ((q  (ΨP  ΨQ))  ((r  PΨ')  QΨ'))  (r  ΨP')  ΨQ' distinctPerm q
        extractFrame PQ' = (AP'@AQ'), ((r  ΨP')  ΨQ')

      moreover from AP' ♯* PQ' AQ' ♯* PQ' have "(AP'@AQ') ♯* PQ'" by simp
      moreover from AP' ♯* α AQ' ♯* α have "(AP'@AQ') ♯* α" by simp
      moreover from AP' ♯* (q  α) AQ' ♯* (q  α) have "(AP'@AQ') ♯* (q  α)" by simp
      moreover from AP' ♯* C AQ' ♯* C have "(AP'@AQ') ♯* 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 AP' ♯* V AQ' ♯* V have "(AP'@AQ') ♯* V" by simp
      moreover from AP' ♯* W AQ' ♯* W have "(AP'@AQ') ♯* W" by simp
      moreover from AP' ♯* X AQ' ♯* X have "(AP'@AQ') ♯* X" by simp
      moreover from AP' ♯* Y AQ' ♯* Y have "(AP'@AQ') ♯* Y" by simp
      moreover from AP' ♯* Z AQ' ♯* Z have "(AP'@AQ') ♯* Z" by simp
      moreover from distinct AP' distinct AQ' AQ' ♯* AP'
      have "distinct(AP'@AQ')" 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 Ψ ΨQ P M xvec N P' AP ΨP Q Q' AQ C C' α PQ' V W X Y Z)
      have FrP: "extractFrame P = AP, ΨP" and  FrQ: "extractFrame Q = AQ, Ψ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 AQ ♯* PQ' PQ' = P''  (r  Q') have "AQ ♯* P''"
        by simp

      from AP ♯* PQ' PQ' = P''  (r  Q') have "AP ♯* P''"
        by simp

      from bn α ♯* (P  Q) AP ♯* α
      have "bn α ♯* P" and "AP ♯* bn α" by simp+
      with extractFrame P = AP, ΨP have "bn α ♯* ΨP"
        by (metis extractFrameFreshChain freshFrameDest)

      from bn α ♯* (P  Q) AQ ♯* α
      have "bn α ♯* Q" and "AQ ♯* bn α" by simp+
      with extractFrame Q = AQ, ΨQ have "bn α ♯* ΨQ"
        by (metis extractFrameFreshChain freshFrameDest)

      from rPerm AQ ♯* xvec xvec ♯* ΨQ AQ ♯* bn α bn α ♯* ΨQ
      have "r  AQ = AQ" and "r  ΨQ = ΨQ" by simp+

      have "¿MN  Q' = ¿MN  Q'" by simp
      moreover note AQ ♯* P AQ ♯* C AQ ♯* C'
      moreover from AQ ♯* M AQ ♯* N have "AQ ♯* (¿MN)" by simp
      moreover note AQ ♯* V
      moreover from AQ ♯* W AQ ♯* α have "AQ ♯* (α#W)" by simp
      moreover from AQ ♯* X AP ♯* AQ AQ ♯* xvec have "AQ ♯* (X@AP@xvec)" by simp
      moreover from AQ ♯* Y AQ ♯* ΨP have "AQ ♯* (ΨP#Y)" by force
      moreover from AQ ♯* Z AQ ♯* P AQ ♯* P' AQ ♯* P'' have "AQ ♯* (P#(P'#(P''#Z)))" by simp
      moreover note AQ ♯* Q' AQ ♯* Q

      ultimately obtain q QΨ' AQ' ΨQ' where Sq: "set q  set(bn (¿MN)) × set (bn(q  (¿MN)))" and "((q  ΨQ)  QΨ')  ΨQ'"
        and "distinctPerm q" and FrQ': "extractFrame Q' = AQ', ΨQ'"
        and "AQ' ♯* Q'" and "AQ' ♯* (¿MN)" "AQ' ♯* (q  (¿MN))" and "AQ' ♯* C" and "(bn(q  (¿MN))) ♯* C'"
        and "(bn(q  (¿MN))) ♯* (¿MN)" and "(bn(q  (¿MN))) ♯* Q'" and "distinct AQ'"
        and "AQ' ♯* V" and "AQ' ♯* (α#W)" and "AQ' ♯* (X @ AP @ xvec)" and "AQ' ♯* (ΨP#Y)"
        and "AQ' ♯* (P#(P'#(P''#Z)))" and "(bn(q  (¿MN))) ♯* (α#W)" and "(bn(q  (¿MN))) ♯* (X @ AP @ xvec)" and "(bn(q  (¿MN))) ♯* (ΨP#Y)"
        and "(bn(q  (¿MN))) ♯* (P#(P'#(P''#Z)))"
        by(elim cBrComm2(8)) (assumption | simp)+

      then have QeqQ': "ΨQ  QΨ'  ΨQ'"
        and "AQ' ♯* Q'" and "AQ' ♯* (¿MN)" and "AQ' ♯* C" and "distinct AQ'"
        and "AQ' ♯* P" and "AQ' ♯* P'" and "AQ' ♯* P''" and "AQ' ♯* Z" and "AQ' ♯* AP" and "AQ' ♯* X" and "AQ' ♯* ΨP" and "AQ' ♯* Y"
        and "AQ' ♯* xvec" and "AQ' ♯* α" and "AQ' ♯* (subject α)" and "AQ' ♯* (bn α)" and "AQ' ♯* (object α)"
        and "AQ' ♯* V" and "AQ' ♯* W"
        by(simp del: freshChainSimps)+

      from rPerm bn α ♯* ΨQ xvec ♯* ΨQ
      have "(r  ΨQ) = ΨQ" by simp

      from QeqQ' have "r  (ΨQ  QΨ'  ΨQ')"
        by simp
      with (r  ΨQ) = ΨQ
      have rQeqQ': "ΨQ  (r  QΨ')  (r  ΨQ')" by(simp add: eqvts)

      from rPerm AQ' ♯* (bn α) AQ' ♯* xvec
      have "r  AQ' = AQ'" by simp

      from FrQ' have "r  (extractFrame Q' = AQ', ΨQ')"
        by simp

      with r  AQ' = AQ' have rFrQ': "extractFrame (r  Q') = AQ', (r  ΨQ')"
        by(simp add: eqvts)

      from xvec ♯* α have "(bn α) ♯* xvec" by simp

      from (AP @ AQ) ♯* α have "AP ♯* bn α" by simp

      from AP ♯* M AP ♯* xvec AP ♯* N have "AP ♯* (¡M⦇ν*xvec⦈⟨N)" by simp
      from Ψ  ΨQ  P  ¡M⦇ν*xvec⦈⟨N  P' xvec ♯* M distinct xvec bn α ♯* P xvec ♯* bn α
      have "bn α ♯* P'" by(simp add: broutputFreshChainDerivative)

      from Ψ  Ψ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 α ♯* ΨP bn α ♯* ΨQ bn α ♯* W
      moreover from bn α ♯* X AQ ♯* bn α AQ' ♯* bn α bn α ♯* xvec have "bn α ♯* (X@xvec@AQ@AQ')" by simp
      moreover from bn α ♯* Y bn α ♯* ΨQ have "bn α ♯* (ΨQ#Y)" by simp
      moreover from bn α ♯* Z bn α ♯* P bn α ♯* Q have "bn α ♯* (P#Q#Z)" by simp
      moreover from AP ♯* V AP ♯* N have "AP ♯* (N#V)" by simp
      moreover note AP ♯* W
      moreover from AP ♯* X AP ♯* AQ AQ' ♯* AP AP ♯* xvec have "AP ♯* (X@xvec@AQ@AQ')" by simp
      moreover from AP ♯* Y AP ♯* ΨQ have "AP ♯* (ΨQ#Y)" by force
      moreover from AP ♯* Z AP ♯* Q AP ♯* P have "AP ♯* (P#Q#Z)" by simp

      moreover note ¡M⦇ν*xvec⦈⟨N  P' = α  P''
      ultimately obtain p PΨ' AP' ΨP' where Sp: "set p  set (bn α) × set (bn (p  α))" and PeqP': "((p  ΨP)  PΨ')  ΨP'"
        and "distinctPerm p" and "bn(p  α) ♯* C'" and FrP': "extractFrame P'' = AP', ΨP'"
        and "AP' ♯* P''" and "AP' ♯* α" and "AP' ♯* (p  α)" and "AP' ♯* C"
        and "bn(p  α) ♯* α" and "bn(p  α) ♯* P''" and "distinct AP'"
        and "AP' ♯* (N#V)" and "AP' ♯* W" and "AP' ♯* (X @ xvec @ AQ @ AQ')" and "AP' ♯* (ΨQ#Y)"
        and "AP' ♯* (P#Q#Z)" and "bn(p  α) ♯* (N#V)" and "bn(p  α) ♯* W" and "bn(p  α) ♯* (X @ xvec @ AQ @ AQ')" and "bn(p  α) ♯* (ΨQ#Y)"
        and "bn(p  α) ♯* (P#Q#Z)"
        using AP ♯* P AP ♯* α AP ♯* C AP ♯* C'
          bn α ♯* C' α  τ AP ♯* P'' distinct (bn α)
        by(elim cBrComm2(4)[where b=C and ba=C' and bb=α and bc=P'' and bf="(X @ xvec @ AQ @ AQ')" and bg="(ΨQ#Y)" and bh="(P#Q#Z)"]) (assumption | simp)+
      then have "AP' ♯* Q" and "AP' ♯* Z" and "AP' ♯* AQ" and "AQ' ♯* AP'" and "AP' ♯* X" and "AP' ♯* ΨQ" and "AP' ♯* Y" and "AP' ♯* P" and "AP' ♯* N"
        and "AP' ♯* xvec" and "AP' ♯* V" and "AP' ♯* W"
        and "bn(p  α) ♯* AQ" and "bn(p  α) ♯* AQ'" and "bn(p  α) ♯* X" and "bn(p  α) ♯* Y" and "bn(p  α) ♯* Z" and "bn(p  α) ♯* Ψ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 AQ' ♯* P'' AQ' ♯* AP' FrP'
      have "AQ' ♯* ΨP'"
        by(metis extractFrameFreshChain freshFrameDest)

      from Ψ  ΨP  Q  ¿MN  Q' AP' ♯* Q AP' ♯* N
      have "AP' ♯* Q'"
        by(simp add: brinputFreshChainDerivative)
      with FrQ' AQ' ♯* AP' have "AP' ♯* ΨQ'"
        by(metis extractFrameFreshChain freshFrameDest)

      have "((p  ΨP)  PΨ')  (ΨQ  (r  QΨ'))  ΨP'  (r  ΨQ')"
        by (metis Composition' PeqP' rQeqQ')
      then have "((p  ΨP)  (PΨ'  (ΨQ  (r  QΨ'))))  ΨP'  (r  ΨQ')"
        by (metis AssertionStatEqSym AssertionStatEqTrans Associativity)
      then have "((p  ΨP)  ((ΨQ  (r  QΨ'))  PΨ'))  ΨP'  (r  ΨQ')"
        by (metis AssertionStatEqSym AssertionStatEqTrans Associativity associativitySym)
      then have "(p  ΨP)  (ΨQ  ((r  QΨ')  PΨ'))  ΨP'  (r  ΨQ')"
        by (metis AssertionStatEqSym AssertionStatEqTrans Associativity compositionSym)
      then have "((p  ΨP)  ΨQ)  ((r  QΨ')  PΨ')  ΨP'  (r  ΨQ')"
        by (metis AssertionStatEqTrans Associativity)
      then have "((p  ΨP)  ΨQ)  (PΨ'  (r  QΨ'))  ΨP'  (r  ΨQ')"
        by (metis AssertionStatEqSym AssertionStatEqTrans Associativity associativitySym)
      with Sp bn α ♯* ΨQ bn (p  α) ♯* ΨQ have "(p  (ΨP  ΨQ))  (PΨ'  (r  QΨ'))  ΨP'  (r  ΨQ')"
        by (simp add: eqvts)

      from Sp AQ' ♯* α bn(p  α) ♯* AQ' have "AQ' ♯* (p  α)"
        by (metis actionFreshChain freshChainSym freshStarChainSimps fresh_star_set_eq)

      from AP' ♯* α have "AP' ♯* bn α" by simp
      from rPerm AQ' ♯* Q' AQ' ♯* xvec AQ' ♯* bn α
      have "AQ' ♯* (r  Q')"
        by (metis freshAlphaPerm freshChainSym name_list_set_fresh permStarFresh)
      from rPerm AP' ♯* Q' AP' ♯* xvec AP' ♯* bn α
      have "AP' ♯* (r  Q')"
        by (metis freshAlphaPerm freshChainSym name_list_set_fresh permStarFresh)
      from rPerm AP' ♯* ΨQ' AP' ♯* xvec AP' ♯* bn α
      have "AP' ♯* (r  ΨQ')"
        by (metis freshAlphaPerm freshChainSym name_list_set_fresh permStarFresh)

      with AQ' ♯* ΨP' AQ' ♯* AP' FrP' rFrQ'
      have "extractFrame (P''  (r  Q')) = (AP'@AQ'), (ΨP'  (r  ΨQ'))"
        by simp
      with PQ' = (P''  (r  Q'))
      have "extractFrame PQ' = (AP'@AQ'), (ΨP'  (r  ΨQ'))"
        by simp

      from Ψ  ΨP  Q  ¿MN  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 AP' ♯* P'' AP' ♯* (r  Q') PQ' = P''  (r  Q') have "AP' ♯* PQ'" by simp
      from AQ' ♯* P'' AQ' ♯* (r  Q') PQ' = P''  (r  Q') have "AQ' ♯* PQ'" by simp

      note set p  (set (bn α)) × set (bn(p  α))
        ((p  (ΨP  ΨQ))  (PΨ'  (r  QΨ')))  ΨP'  (r  ΨQ') distinctPerm p
        extractFrame PQ' = (AP'@AQ'), (ΨP'  (r  ΨQ'))

      moreover from AP' ♯* PQ' AQ' ♯* PQ' have "(AP'@AQ') ♯* PQ'" by simp
      moreover from AP' ♯* α AQ' ♯* α have "(AP'@AQ') ♯* α" by simp
      moreover from AP' ♯* (p  α) AQ' ♯* (p  α) have "(AP'@AQ') ♯* (p  α)" by simp
      moreover from AP' ♯* C AQ' ♯* C have "(AP'@AQ') ♯* 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 AP' ♯* V AQ' ♯* V have "(AP'@AQ') ♯* V" by simp
      moreover from AP' ♯* W AQ' ♯* W have "(AP'@AQ') ♯* W" by simp
      moreover from AP' ♯* X AQ' ♯* X have "(AP'@AQ') ♯* X" by simp
      moreover from AP' ♯* Y AQ' ♯* Y have "(AP'@AQ') ♯* Y" by simp
      moreover from AP' ♯* Z AQ' ♯* Z have "(AP'@AQ') ♯* Z" by simp
      moreover from distinct AP' distinct AQ' AQ' ♯* AP'
      have "distinct(AP'@AQ')" 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 AP Ψ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 α) AP ♯* α 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 "AP ♯* 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 AP ♯* P AP ♯* M
      moreover from Subj yvecEq bn α ♯* subject α have "yvec1 ♯* M" "yvec2 ♯* M" by simp+
      moreover from yvecEq AP ♯* α have "AP ♯* (yvec1@yvec2)" by simp
      moreover note AP ♯* C
      moreover from yvecEq  bn α ♯* ⦇νxP 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 AP ♯* V AP ♯* W
      moreover from AP ♯* X x  AP AP ♯* α yvecEq have "AP ♯* (x#y#X)" by simp
      moreover note AP ♯* Y AP ♯* Z
      moreover from AP ♯* N' AP ♯* P'' x  AP AP ♯* α yvecEq have "AP ♯* ([(x, y)]  N')" and  "AP ♯* ([(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 Ψ' AP' ΨP' where S: "set p  set (yvec1@yvec2) × set (p  (yvec1@yvec2))" and PeqP': "((p  ΨP)  Ψ')  ΨP'"
        and "distinctPerm p" and "(p  (yvec1@yvec2)) ♯* C'" and FrP': "extractFrame([(x, y)]  P'') = AP', ΨP'"
        and "AP' ♯* ([(x, y)]  P'')" and "AP' ♯* ([(x, y)]  N')" and "AP' ♯* C" and "(p  (yvec1@yvec2)) ♯* ([(x, y)]  N')" and "AP' ♯* M" and "(p  (yvec1@yvec2)) ♯* (yvec1@yvec2)" and "(p  (yvec1@yvec2)) ♯* M" and "distinct AP'"
        and "(p  (yvec1@yvec2)) ♯* ([(x, y)]  P'')" and "(yvec1@yvec2) ♯* AP'" and "(p  (yvec1@yvec2)) ♯* AP'"
        and "AP' ♯* V" and "AP' ♯* W" and "AP' ♯* (x#y#X)" and "AP' ♯* Y" and "AP' ♯* 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 AP ♯* 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 AP' ♯* (x#y#X) have "x  AP'" and "y  AP'" and "AP' ♯* 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  ΨP)  Ψ'))  [(y, x)]  ΨP'"
        by(simp add: AssertionStatEqClosed)
      then have "(((y, x)#p)  ΨP)  ([(y, x)]  Ψ')  ([(y, x)]  Ψ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)]  AP', ΨP')"
        by simp
      with x  AP' y  AP' have "extractFrame P'' = AP', ([(y, x)]  Ψ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  AP' (p  (yvec1@yvec2)) ♯* AP' have "(p  (yvec1@x#yvec2)) ♯* AP'"
        by(simp add: eqvts freshChainSimps)
      then have "([(y, x)]  p  (yvec1@x#yvec2)) ♯* ([(y, x)]  AP')"
        by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
      with x  yvec1 x  yvec2 x  y x  AP' y  yvec1 y  yvec2 x  p y  p y  AP' yvecEq
      have "bn(((y, x)#p)  α) ♯* AP'" 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) ♯* AP' y  AP' yvecEq have "bn α ♯* AP'"
        by simp
      moreover from AP' ♯* ([(x, y)]  N') have "([(x, y)]  AP') ♯* ([(x, y)]  [(x, y)]  N')"
        by(simp only: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
      with x  AP' y  AP' have "AP' ♯* N'" by simp
      with AP' ♯* M (yvec1@yvec2) ♯* AP' y  AP' αeq have "AP' ♯* α" by simp
      moreover then have "(((y, x)#p)  AP') ♯* (((y, x)#p)  α)"
        by(simp only: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
      with x  AP' y  AP' S (yvec1@yvec2) ♯* AP' (p  (yvec1@yvec2)) ♯* AP'
      have "AP' ♯* (((y, x)#p)  α)" by(simp add: eqvts)
      moreover from AP' ♯* ([(x, y)]  P'') have "([(x, y)]  AP') ♯* ([(x, y)]  [(x, y)]  P'')"
        by(simp only: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
      with x  AP' y  AP' have "AP' ♯* 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  yvec1y  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 AP' ♯* C AP' ♯* V AP' ♯* W AP' ♯* X AP' ♯* Y AP' ♯* Z distinct AP'

      ultimately show ?case
        by(elim cOpen)
    next
      case(cBrOpen Ψ P M xvec1 xvec2 N P' x AP Ψ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 α) AP ♯* α 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 "AP ♯* 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 AP ♯* P AP ♯* M
      moreover from Subj yvecEq bn α ♯* subject α have "yvec1 ♯* M" "yvec2 ♯* M" by simp+
      moreover from yvecEq AP ♯* α have "AP ♯* (yvec1@yvec2)" by simp
      moreover note AP ♯* C
      moreover from yvecEq  bn α ♯* ⦇νxP 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 AP ♯* V AP ♯* W
      moreover from AP ♯* X x  AP AP ♯* α yvecEq have "AP ♯* (x#y#X)" by simp
      moreover note AP ♯* Y AP ♯* Z
      moreover from AP ♯* N' AP ♯* P'' x  AP AP ♯* α yvecEq have "AP ♯* ([(x, y)]  N')" and  "AP ♯* ([(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 Ψ' AP' ΨP' where S: "set p  set (yvec1@yvec2) × set (p  (yvec1@yvec2))" and PeqP': "((p  ΨP)  Ψ')  ΨP'"
        and "distinctPerm p" and "(p  (yvec1@yvec2)) ♯* C'" and FrP': "extractFrame([(x, y)]  P'') = AP', ΨP'"
        and "AP' ♯* ([(x, y)]  P'')" and "AP' ♯* ([(x, y)]  N')" and "AP' ♯* C" and "(p  (yvec1@yvec2)) ♯* ([(x, y)]  N')" and "AP' ♯* M" and "(p  (yvec1@yvec2)) ♯* (yvec1@yvec2)" and "(p  (yvec1@yvec2)) ♯* M" and "distinct AP'"
        and "(p  (yvec1@yvec2)) ♯* ([(x, y)]  P'')" and "(yvec1@yvec2) ♯* AP'" and "(p  (yvec1@yvec2)) ♯* AP'"
        and "AP' ♯* V" and "AP' ♯* W" and "AP' ♯* (x#y#X)" and "AP' ♯* Y" and "AP' ♯* 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 AP ♯* 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 AP' ♯* (x#y#X) have "x  AP'" and "y  AP'" and "AP' ♯* 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  ΨP)  Ψ'))  [(y, x)]  ΨP'"
        by(simp add: AssertionStatEqClosed)
      then have "(((y, x)#p)  ΨP)  ([(y, x)]  Ψ')  ([(y, x)]  Ψ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)]  AP', ΨP')"
        by simp
      with x  AP' y  AP' have "extractFrame P'' = AP', ([(y, x)]  Ψ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  AP' (p  (yvec1@yvec2)) ♯* AP' have "(p  (yvec1@x#yvec2)) ♯* AP'"
        by(simp add: eqvts freshChainSimps)
      then have "([(y, x)]  p  (yvec1@x#yvec2)) ♯* ([(y, x)]  AP')"
        by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
      with x  yvec1 x  yvec2 x  y x  AP' y  yvec1 y  yvec2 x  p y  p y  AP' yvecEq
      have "bn(((y, x)#p)  α) ♯* AP'" 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) ♯* AP' y  AP' yvecEq have "bn α ♯* AP'"
        by simp
      moreover from AP' ♯* ([(x, y)]  N') have "([(x, y)]  AP') ♯* ([(x, y)]  [(x, y)]  N')"
        by(simp only: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
      with x  AP' y  AP' have "AP' ♯* N'" by simp
      with AP' ♯* M (yvec1@yvec2) ♯* AP' y  AP' αeq have "AP' ♯* α" by simp
      moreover then have "(((y, x)#p)  AP') ♯* (((y, x)#p)  α)"
        by(simp only: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
      with x  AP' y  AP' S (yvec1@yvec2) ♯* AP' (p  (yvec1@yvec2)) ♯* AP'
      have "AP' ♯* (((y, x)#p)  α)" by(simp add: eqvts)
      moreover from AP' ♯* ([(x, y)]  P'') have "([(x, y)]  AP') ♯* ([(x, y)]  [(x, y)]  P'')"
        by(simp only: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
      with x  AP' y  AP' have "AP' ♯* 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  yvec1y  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 AP' ♯* C AP' ♯* V AP' ♯* W AP' ♯* X AP' ♯* Y AP' ♯* Z distinct AP'

      ultimately show ?case
        by(elim cBrOpen)
    next
      case(cScope Ψ P α P' x AP ΨP C C' α' P'' V W X Y Z)
      from α  ⦇νxP' = α'  P'' x  α x  α'
      obtain P''' where "α  P' = α'  P'''" and "P'' = ⦇νxP'''"
        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 Ψ' AP' ΨP' where S: "set p  set(bn α') × set(bn(p  α'))" and PeqP': "((p  ΨP)  Ψ')  ΨP'"
        and "distinctPerm p" and "bn(p  α') ♯* C'" and FrP': "extractFrame P''' = AP', ΨP'"
        and "AP' ♯* P'''" and "AP' ♯* α'" and "AP' ♯* (p  α')" and "AP' ♯* C" and "distinct AP'"
        and "bn(p  α') ♯* P'''" and "AP' ♯* V" and "AP' ♯* W" and "AP' ♯* (x#X)" and "AP' ♯* Y" and "bn(p  α') ♯* α'"
        and "AP' ♯* 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 AP' ♯* (x#X) have "x  AP'" and "AP' ♯* 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'' = ⦇νxP''' have "extractFrame P'' = (x#AP'), ΨP'" by simp
      moreover from AP' ♯* P''' P'' = ⦇νxP''' x  AP' have "(x#AP') ♯* P''" by(simp add: abs_fresh)
      moreover from AP' ♯* α' AP' ♯* C  x  α' x  C have "(x#AP') ♯* α'"  and "(x#AP') ♯* C" by simp+
      moreover note bn(p  α') ♯* α'
      moreover from bn(p  α') ♯* P''' P'' = ⦇νxP''' x  bn(p  α') have "bn(p  α') ♯* P''" by simp
      moreover from AP' ♯* α' x  α' have "(x#AP') ♯* α'" by simp
      moreover from AP' ♯* (p  α') x  α' S x  bn(p  α') have "(x#AP') ♯* (p  α')"
        by(simp add: subjectEqvt[symmetric] bnEqvt[symmetric] okjectEqvt[symmetric] freshChainSimps)
      moreover from AP' ♯* V x  V have "(x#AP') ♯* V" by simp+
      moreover from AP' ♯* W x  W have "(x#AP') ♯* W" by simp+
      moreover from AP' ♯* X x  X have "(x#AP') ♯* X" by simp+
      moreover from AP' ♯* Y x  Y have "(x#AP') ♯* Y" by simp+
      moreover from AP' ♯* Z x  Z have "(x#AP') ♯* Z" by simp+
      moreover note bn(p  α') ♯* V bn(p  α') ♯* W bn(p  α') ♯* X bn(p  α') ♯* Y bn(p  α') ♯* Z
      moreover from distinct AP' x  AP' have "distinct(x#AP')" by simp
      ultimately show ?case by(elim cScope)
    next
      case(cBang Ψ P AP ΨP C C' α P' V W X Y Z)
      then obtain p Ψ' AP' ΨP' where S: "set p  set(bn α) × set(bn(p  α))"
        and FrP': "extractFrame P' = AP', ΨP'"
        and PeqP': "(p  (ΨP  𝟭))  Ψ'  ΨP'"
        and "AP' ♯* C" and "AP' ♯* P'" and "AP' ♯* α" and "AP' ♯* (p  α)"
        and "AP' ♯* V" and "AP' ♯* W" and "AP' ♯* X" and "AP' ♯* Y" and "AP' ♯* Z" and "distinct AP'"
        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 ΨP  𝟭 have "(p  ΨP)  (p  𝟭)"
        by(simp add: AssertionStatEqClosed)
      then have "(p  ΨP)  𝟭" by(simp add: permBottom)
      with PeqP' have "(𝟭  Ψ')  Ψ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  "AP ♯* ([]::'a list)" and "AP ♯* ([]::('a action) list)" and "AP ♯* ([]::name list)" and "AP ♯* ([]::'b list)" and "AP ♯* ([]::('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 AP   :: "name list"
    and ΨP   :: 'b
    and C    :: "'f::fs_name"

assumes "Ψ  P τ  P'"
  and   "extractFrame P = AP, ΨP"
  and   "distinct AP"
  and   "AP ♯* P"
  and   "AP ♯* C"

obtains Ψ' AP' ΨP' where "extractFrame P' = AP', ΨP'"  and "ΨP  Ψ'  ΨP'" and "AP' ♯* C" and "AP' ♯* P'" and "distinct AP'"
proof -
  assume A: "AP' ΨP' Ψ'.
        extractFrame P' = AP', ΨP'; ΨP  Ψ'  ΨP'; AP' ♯* C; AP' ♯* P'; distinct AP'
         thesis"

  from Ψ  P τ P' AP ♯* P have "AP ♯* P'" by(rule tauFreshChainDerivative)

  {
    fix X :: "name list"
      and Y :: "'b list"
      and Z :: "('a, 'b, 'c) psi list"

    assume "AP ♯* X"
      and  "AP ♯* Y"
      and  "AP ♯* Z"

    with assms AP ♯* P' obtain Ψ' AP' ΨP' where "extractFrame P' = AP', ΨP'" and "ΨP  Ψ'  ΨP'" and "AP' ♯* C"
      and "AP' ♯* P'" and "AP' ♯* X" and "AP' ♯* Y" and "AP' ♯* Z" and "distinct AP'"
    proof(nominal_induct avoiding: C X Y Z arbitrary: thesis rule: tauFrameInduct)
      case(cAlpha Ψ P P' AP ΨP p C X Y Z)
      then obtain Ψ' AP' ΨP' where FrP': "extractFrame P' = AP', ΨP'" and "ΨP  Ψ'  ΨP'" and "distinct AP'"
        and "AP' ♯* C" and "AP' ♯* P'" and "AP' ♯* X" and "AP' ♯* Y" and "AP' ♯* Z"
        by metis

      have S: "set p  set AP × set(p  AP)" by fact

      from FrP' have "(p  extractFrame P') = p  AP', ΨP'" by simp
      with AP ♯* P' (p  AP) ♯* P' S have "extractFrame P' = (p  AP'), (p  ΨP')" by(simp add: eqvts)
      moreover from ΨP  Ψ'  ΨP' have "(p  (ΨP  Ψ'))  (p  ΨP')" by(rule AssertionStatEqClosed)
      then have "(p  ΨP)  (p  Ψ')  (p  ΨP')" by(simp add: eqvts)
      moreover from AP' ♯* C have "(p  AP') ♯* (p  C)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
      with AP ♯* C (p  AP) ♯* C S have "(p  AP') ♯* C" by simp
      moreover from AP' ♯* P' have "(p  AP') ♯* (p  P')" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
      with AP ♯* P' (p  AP) ♯* P' S have "(p  AP') ♯* P'" by simp
      moreover from AP' ♯* X have "(p  AP') ♯* (p  X)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
      with AP ♯* X (p  AP) ♯* X S have "(p  AP') ♯* X" by simp
      moreover from AP' ♯* Y have "(p  AP') ♯* (p  Y)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
      with AP ♯* Y (p  AP) ♯* Y S have "(p  AP') ♯* Y" by simp
      moreover from AP' ♯* Z have "(p  AP') ♯* (p  Z)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
      with AP ♯* Z (p  AP) ♯* Z S have "(p  AP') ♯* Z" by simp
      moreover from distinct AP' have "distinct(p  AP')" by simp
      ultimately show ?case by(rule cAlpha)
    next
      case(cCase Ψ P P' φ Cs AP ΨP C B Y Z thesis)
      then obtain Ψ' AP' ΨP' where FrP': "extractFrame P' = AP', ΨP'"
        and "ΨP  Ψ'  ΨP'" and "distinct AP'"
        and "AP' ♯* C" and "AP' ♯* P'"
        and "AP' ♯* B" and "AP' ♯* Y" and "AP' ♯* Z"
        apply -
        by(rule cCase) (assumption | simp (no_asm_use))+
      with ΨP  𝟭 ΨP  Ψ'  ΨP' have "𝟭  Ψ'  ΨP'"
        by(metis Identity AssertionStatEqTrans composition' Commutativity Associativity AssertionStatEqSym)
      then show ?case using FrP' AP' ♯* P' AP' ♯* C AP' ♯* B AP' ♯* Y AP' ♯* Z distinct AP' using cCase
        by force
    next
      case(cPar1 Ψ ΨQ P P' AQ Q AP ΨP C X Y Z)
      moreover from AP ♯* X AP ♯* AQ AP ♯* Y AP ♯* ΨQ AP ♯* Q AP ♯* Z
      have "AP ♯* (X@AQ)" and  "AP ♯* (ΨQ#Y)" and  "AP ♯* (Q#Z)"
        by simp+
      ultimately obtain Ψ' AP' ΨP' where FrP': "extractFrame P' = AP', ΨP'" and "distinct AP'"
        and "ΨP  Ψ'  ΨP'" and "AP ♯* P" and "AP ♯* C" and "AP' ♯* C" and "AP' ♯* P'"
        and "AP' ♯* (X@AQ)" and "AP' ♯* (ΨQ#Y)" and "AP' ♯* (Q#Z)"
        by metis

      then have "AP' ♯* X" and "AP' ♯* AQ" and "AP' ♯* Y" and "AP' ♯* ΨQ" and "AP' ♯* Q" and "AP' ♯* Z"
        by simp+

      from AP' ♯* AQ AQ ♯* P' FrP' have "AQ ♯* ΨP'" by(force dest: extractFrameFreshChain)
      with AP' ♯* ΨQ AP' ♯* AQ extractFrame Q = AQ, ΨQ FrP'
      have "extractFrame(P'  Q) = (AP'@AQ), ΨP'  ΨQ" by simp

      moreover from ΨP  Ψ'  ΨP'have "(ΨP  ΨQ)  Ψ'  ΨP'  ΨQ"
        by(metis Associativity Commutativity Composition AssertionStatEqTrans AssertionStatEqSym)

      moreover from AP' ♯* C AQ ♯* C have "(AP'@AQ) ♯* C" by simp
      moreover from AP' ♯* P' AQ ♯* P' AP' ♯* Q AQ ♯* Q have "(AP'@AQ) ♯* (P'  Q)" by simp
      moreover from AP' ♯* X AQ ♯* X have "(AP'@AQ) ♯* X" by simp
      moreover from AP' ♯* Y AQ ♯* Y have "(AP'@AQ) ♯* Y" by simp
      moreover from AP' ♯* Z AQ ♯* Z have "(AP'@AQ) ♯* Z" by simp
      moreover from AP' ♯* AQ distinct AP' distinct AQ have "distinct(AP'@AQ)" by simp
      ultimately show ?case by(rule cPar1)
    next
      case(cPar2 Ψ ΨP Q Q' AP P AQ ΨQ C X Y Z)
      moreover from AQ ♯* X AP ♯* AQ AQ ♯* Y AQ ♯* ΨP AQ ♯* P AQ ♯* Z
      have "AQ ♯* (X@AP)" and  "AQ ♯* (ΨP#Y)" and  "AQ ♯* (P#Z)"
        by(simp add: freshChainSimps)+
      ultimately obtain Ψ' AQ' ΨQ' where FrQ': "extractFrame Q' = AQ', ΨQ'" and "distinct AQ'"
        and "ΨQ  Ψ'  ΨQ'"and "AQ' ♯* C" and "AQ' ♯* Q'"
        and "AQ' ♯* (X@AP)" and "AQ' ♯* (ΨP#Y)" and "AQ' ♯* (P#Z)"
        by metis

      then have "AQ' ♯* X" and "AQ' ♯* AP" and "AQ' ♯* Y" and "AQ' ♯* ΨP" and "AQ' ♯* P" and "AQ' ♯* Z"
        by simp+

      from AQ' ♯* AP AP ♯* Q' FrQ' have "AP ♯* ΨQ'" by(force dest: extractFrameFreshChain)
      with AQ' ♯* ΨP AQ' ♯* AP extractFrame P = AP, ΨP FrQ'
      have "extractFrame(P  Q') = (AP@AQ'), ΨP  ΨQ'" by simp

      moreover from ΨQ  Ψ'  ΨQ'have "(ΨP  ΨQ)  Ψ'  ΨP  ΨQ'"
        by(metis Associativity Commutativity Composition AssertionStatEqTrans AssertionStatEqSym)
      moreover from AP ♯* C AQ' ♯* C have "(AP@AQ') ♯* C" by simp
      moreover from AP ♯* P AQ' ♯* P AP ♯* Q' AQ' ♯* Q' have "(AP@AQ') ♯* (P  Q')" by simp
      moreover from AP ♯* X AQ' ♯* X have "(AP@AQ') ♯* X" by simp
      moreover from AP ♯* Y AQ' ♯* Y have "(AP@AQ') ♯* Y" by simp
      moreover from AP ♯* Z AQ' ♯* Z have "(AP@AQ') ♯* Z" by simp
      moreover from AQ' ♯* AP distinct AP distinct AQ' have "distinct(AP@AQ')" by simp
      ultimately show ?case by(rule cPar2)
    next
      case(cComm1 Ψ ΨQ P M N P' AP ΨP Q K xvec Q' AQ C X Y Z)
      have PTrans: "Ψ  ΨQ  P MN  P'" and QTrans: "Ψ  ΨP  Q K⦇ν*xvec⦈⟨N  Q'" by fact+
      from PTrans extractFrame P = AP, ΨP distinct AP AP ♯* P AP ♯* N AP ♯* M
        AP ♯* Q' AP ♯* C AP ♯* X AP ♯* Y AP ♯* Z AP ♯* AQ AP ♯* xvec
      obtain Ψ' AP' ΨP' where FrP': "extractFrame P' = AP', ΨP'" and PeqP': "ΨP  Ψ'  ΨP'"
        and "AP' ♯* Q'" and "AP' ♯* C" and "AP' ♯* X" and "AP' ♯* Y" and "distinct AP'"
        and "AP' ♯* Z" and "AP' ♯* AQ" and "AP' ♯* xvec" and "AP' ♯* P'"
        by(elim expandNonTauFrame[where C="(Q', C, X, Y, Z, AQ, xvec)" and C'="(Q', C, X, Y, Z, AQ, xvec)"]) auto
      moreover from QTrans  have "distinct xvec" by(auto dest: boundOutputDistinct)
      from QTrans extractFrame Q = AQ, ΨQ distinct AQ AQ ♯* P AQ ♯* xvec AQ ♯* K xvec ♯* K distinct xvec
        AP' ♯* AQ AP' ♯* xvec AQ ♯* Q xvec ♯* Q AQ ♯* ΨP xvec ♯* ΨP AQ ♯* N
        AQ ♯* C AQ ♯* X AQ ♯* Y AQ ♯* Z xvec ♯* P xvec ♯* C xvec ♯* X xvec ♯* Y xvec ♯* Z
      obtain p Ψ'' AQ' ΨQ' where S: "set p  set xvec × set(p  xvec)" and QeqQ': "(p  ΨQ)  Ψ''  ΨQ'" and FrQ': "extractFrame Q' = AQ', ΨQ'"
        and "AQ' ♯* AP'" and "AQ' ♯* C" and "AQ' ♯* X" and "AQ' ♯* Y" and "AQ' ♯* Z" and "AQ' ♯* P" and "AQ' ♯* N" and "distinct AQ'"
        and "(p  xvec) ♯* AP'" 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) ♯* ΨP" and "(p  xvec) ♯* AQ'"and "(p  xvec) ♯* Q'"
        and "distinctPerm p" and "AQ' ♯* xvec" and "AQ' ♯* Q'"
        by(elim expandNonTauFrame[where C="(P, C, X, Y, Z, AP', ΨP)" and C'="(P, C, X, Y, Z, AP', ΨP)"]) (assumption | simp)+

      from PTrans AQ' ♯* P AQ' ♯* N (p  xvec) ♯* P (p  xvec) ♯* N
      have "AQ' ♯* P'" and "(p  xvec) ♯* P'" by(force dest: inputFreshChainDerivative)+
      with FrP' AQ' ♯* AP' (p  xvec) ♯* AP' have "AQ' ♯* ΨP'" and "(p  xvec) ♯* ΨP'" by(force dest: extractFrameFreshChain)+
      from FrQ' AQ' ♯* AP' AP' ♯* Q' (p  xvec) ♯* AQ' (p  xvec) ♯* Q' have "AP' ♯* ΨQ'" and "(p  xvec) ♯* ΨQ'"
        by(force dest: extractFrameFreshChain)+

      have "extractFrame(⦇ν*xvec(P'  Q')) = ((p  xvec)@AP'@AQ'), (p  ΨP')  (p  ΨQ')"
      proof -
        from FrP' FrQ' AP' ♯* ΨQ' AQ' ♯* AP' AQ' ♯* ΨP' have "extractFrame(P'  Q') = (AP'@AQ'), ΨP'  ΨQ'"
          by simp
        then have "extractFrame(⦇ν*xvec(P'  Q')) = (xvec@AP'@AQ'), ΨP'  ΨQ'"
          by(induct xvec) auto
        moreover from (p  xvec) ♯* ΨP' (p  xvec) ♯* ΨQ' S
        have "⦇ν*xvec(⦇ν*(AP'@AQ')(FAssert (ΨP'  ΨQ'))) = ⦇ν*(p  xvec)(p  ⦇ν*(AP'@AQ')(FAssert(ΨP'  ΨQ')))"
          by(intro frameChainAlpha) (auto simp add: fresh_star_def frameResChainFresh)
        then have "⦇ν*xvec(⦇ν*(AP'@AQ')(FAssert (ΨP'  ΨQ'))) = ⦇ν*(p  xvec)(⦇ν*(AP'@AQ')(FAssert((p  ΨP')  (p  ΨQ'))))"
          using AP' ♯* xvec (p  xvec) ♯* AP' AQ' ♯* xvec (p  xvec) ♯* AQ' S
          by(auto simp add: eqvts)
        ultimately show ?thesis
          by(simp add: frameChainAppend)
      qed

      moreover have "(ΨP  ΨQ)  ((p  Ψ')  (p  Ψ''))  (p  ΨP')  (p  ΨQ')"
      proof -
        have "(ΨP  (p  ΨQ))  (Ψ'  Ψ'')  (ΨP  Ψ')  ((p  ΨQ)  Ψ'')"
          by(metis Associativity Commutativity Composition AssertionStatEqTrans)
        moreover from PeqP' QeqQ' have "(ΨP  Ψ')  ((p  ΨQ)  Ψ'')  ΨP'  ΨQ'"
          by(metis Associativity Commutativity Composition AssertionStatEqTrans)
        ultimately have "(ΨP  (p  ΨQ))  (Ψ'  Ψ'')  ΨP'  ΨQ'"
          by(metis AssertionStatEqTrans)
        then have "(p  ((ΨP  (p  ΨQ))  (Ψ'  Ψ'')))  (p  (ΨP'  ΨQ'))"
          by(rule AssertionStatEqClosed)
        with xvec ♯* ΨP (p  xvec) ♯* ΨP S distinctPerm p show ?thesis
          by(simp add: eqvts)
      qed

      moreover from (p  xvec) ♯* C AP' ♯* C AQ' ♯* C have "((p  xvec)@AP'@AQ') ♯* C" by simp
      moreover from (p  xvec) ♯* X AP' ♯* X AQ' ♯* X have "((p  xvec)@AP'@AQ') ♯* X" by simp
      moreover from (p  xvec) ♯* Y AP' ♯* Y AQ' ♯* Y have "((p  xvec)@AP'@AQ') ♯* Y" by simp
      moreover from (p  xvec) ♯* Z AP' ♯* Z AQ' ♯* Z have "((p  xvec)@AP'@AQ') ♯* Z" by simp
      moreover from (p  xvec) ♯* P' (p  xvec) ♯* Q' AP' ♯* P' AP' ♯* Q'AQ' ♯* P' AQ' ♯* Q'
      have "((p  xvec)@AP'@AQ') ♯* (⦇ν*xvec(P'  Q'))" by(auto simp add: resChainFresh fresh_star_def)
      moreover from (p  xvec) ♯* AP' (p  xvec) ♯* AQ' AQ' ♯* AP' distinct xvec distinct AP' distinct AQ'
      have "distinct((p  xvec)@AP'@AQ')"
        by auto (simp add: name_list_supp fresh_star_def fresh_def)+

      ultimately show ?case using cComm1
        by metis
    next
      case(cComm2 Ψ ΨQ P M xvec N P' AP ΨP Q K Q' AQ C X Y Z)
      have PTrans: "Ψ  ΨQ  P M⦇ν*xvec⦈⟨N  P'" and QTrans: "Ψ  ΨP  Q KN  Q'" by fact+
      from PTrans have "distinct xvec" by(auto dest: boundOutputDistinct)
      from PTrans extractFrame P = AP, ΨP distinct AP AP ♯* P AP ♯* Q xvec ♯* Q distinct xvec xvec ♯* M
        AP ♯* C AP ♯* X AP ♯* Y AP ♯* Z AP ♯* AQ AP ♯* xvec AP ♯* ΨQ AP ♯* M AP ♯* N
        xvec ♯* P xvec ♯* C xvec ♯* X xvec ♯* Y xvec ♯* Z AQ ♯* xvec xvec ♯* ΨQ
      obtain p Ψ' AP' ΨP' where S: "set p  set xvec × set(p  xvec)"
        and FrP': "extractFrame P' = AP', ΨP'" and PeqP': "(p  ΨP)  Ψ'  ΨP'" and "distinct AP'"
        and "AP' ♯* C" and "AP' ♯* X" and "AP' ♯* Y" and "AP' ♯* N" and "AP' ♯* Q" and "(p  xvec) ♯* Q"
        and "AP' ♯* Z" and "AP' ♯* AQ" and "AP' ♯* xvec" and "AP' ♯* P'" and "(p  xvec) ♯* N" and "(p  xvec) ♯* ΨQ"
        and "(p  xvec) ♯* AP'" and "(p  xvec) ♯* C" and "(p  xvec) ♯* X" and "(p  xvec) ♯* AQ"
        and "(p  xvec) ♯* Y" and "(p  xvec) ♯* Z" and "(p  xvec) ♯* P'" and "distinctPerm p"
        by(elim expandNonTauFrame[where C="(C, X, Y, Z, AQ, Q, ΨQ)" and C'="(C, X, Y, Z, AQ, Q, ΨQ)"]) (assumption | simp)+

      from QTrans extractFrame Q = AQ, ΨQ distinct AQ AQ ♯* Q AQ ♯* xvec AQ ♯* P' (p  xvec) ♯* AQ
        AP' ♯* AQ AQ ♯* P AQ ♯* ΨP AQ ♯* C AQ ♯* X AQ ♯* Y AQ ♯* Z AQ ♯* N  AQ ♯* K
      obtain Ψ'' AQ' ΨQ' where QeqQ': "ΨQ  Ψ''  ΨQ'" and FrQ': "extractFrame Q' = AQ', ΨQ'" and "distinct AQ'"
        and "AQ' ♯* xvec" and "AQ' ♯* Q'" and "AQ' ♯* xvec" and "AQ' ♯* P'" and "AQ' ♯* (p  xvec)"
        and "AQ' ♯* AP'" and "AQ' ♯* C" and "AQ' ♯* X" and "AQ' ♯* Y" and "AQ' ♯* Z" and "AQ' ♯* P"
        by(elim expandNonTauFrame[where C="(P, C, P', X, Y, Z, AP', xvec, (p  xvec), ΨP)" and C'="(P, C, P', X, Y, Z, AP', xvec, (p  xvec), ΨP)"]) (assumption | simp)+

      from QTrans AP' ♯* Q AP' ♯* N (p  xvec) ♯* Q (p  xvec) ♯* N
      have "AP' ♯* Q'" and "(p  xvec) ♯* Q'" by(force dest: inputFreshChainDerivative)+
      with FrQ' AQ' ♯* AP' AQ' ♯* (p  xvec) have "AP' ♯* ΨQ'" and "(p  xvec) ♯* ΨQ'" by(force dest: extractFrameFreshChain)+
      from FrP' AQ' ♯* AP' AQ' ♯* P' (p  xvec) ♯* AP' (p  xvec) ♯* P' have "AQ' ♯* ΨP'" and "(p  xvec) ♯* ΨP'"
        by(force dest: extractFrameFreshChain)+

      have "extractFrame(⦇ν*xvec(P'  Q')) = ((p  xvec)@AP'@AQ'), (p  ΨP')  (p  ΨQ')"
      proof -
        from FrP' FrQ' AP' ♯* ΨQ' AQ' ♯* AP' AQ' ♯* ΨP' have "extractFrame(P'  Q') = (AP'@AQ'), ΨP'  ΨQ'"
          by simp
        then have "extractFrame(⦇ν*xvec(P'  Q')) = (xvec@AP'@AQ'), ΨP'  ΨQ'"
          by(induct xvec) auto
        moreover from (p  xvec) ♯* ΨP' (p  xvec) ♯* ΨQ' S
        have "⦇ν*xvec(⦇ν*(AP'@AQ')(FAssert (ΨP'  ΨQ'))) = ⦇ν*(p  xvec)(p  ⦇ν*(AP'@AQ')(FAssert(ΨP'  ΨQ')))"
          by(intro frameChainAlpha) (auto simp add: fresh_star_def frameResChainFresh)
        then have "⦇ν*xvec(⦇ν*(AP'@AQ')(FAssert (ΨP'  ΨQ'))) = ⦇ν*(p  xvec)(⦇ν*(AP'@AQ')(FAssert((p  ΨP')  (p  ΨQ'))))"
          using AP' ♯* xvec (p  xvec) ♯* AP' AQ' ♯* xvec AQ' ♯* (p  xvec) S
          by(auto simp add: eqvts)
        ultimately show ?thesis
          by(simp add: frameChainAppend)
      qed

      moreover have "(ΨP  ΨQ)  ((p  Ψ')  (p  Ψ''))  (p  ΨP')  (p  ΨQ')"
      proof -
        have "((p  ΨP)  ΨQ)  (Ψ'  Ψ'')  ((p  ΨP)  Ψ')  (ΨQ  Ψ'')"
          by(metis Associativity Commutativity Composition AssertionStatEqTrans)
        moreover from PeqP' QeqQ' have "((p  ΨP)  Ψ')  (ΨQ  Ψ'')  ΨP'  ΨQ'"
          by(metis Associativity Commutativity Composition AssertionStatEqTrans)
        ultimately have "((p  ΨP)  ΨQ)  (Ψ'  Ψ'')  ΨP'  ΨQ'"
          by(metis AssertionStatEqTrans)
        then have "(p  ((p  ΨP)  ΨQ)  (Ψ'  Ψ''))  (p  (ΨP'  ΨQ'))"
          by(rule AssertionStatEqClosed)
        with xvec ♯* ΨQ (p  xvec) ♯* ΨQ S distinctPerm p show ?thesis
          by(simp add: eqvts)
      qed

      moreover from (p  xvec) ♯* C AP' ♯* C AQ' ♯* C have "((p  xvec)@AP'@AQ') ♯* C" by simp
      moreover from (p  xvec) ♯* X AP' ♯* X AQ' ♯* X have "((p  xvec)@AP'@AQ') ♯* X" by simp
      moreover from (p  xvec) ♯* Y AP' ♯* Y AQ' ♯* Y have "((p  xvec)@AP'@AQ') ♯* Y" by simp
      moreover from (p  xvec) ♯* Z AP' ♯* Z AQ' ♯* Z have "((p  xvec)@AP'@AQ') ♯* Z" by simp
      moreover from (p  xvec) ♯* P' (p  xvec) ♯* Q' AP' ♯* P' AP' ♯* Q'AQ' ♯* P' AQ' ♯* Q'
      have "((p  xvec)@AP'@AQ') ♯* (⦇ν*xvec(P'  Q'))" by(auto simp add: resChainFresh fresh_star_def)
      moreover from (p  xvec) ♯* AP' AQ' ♯* (p  xvec) AQ' ♯* AP' distinct xvec distinct AP' distinct AQ'
      have "distinct((p  xvec)@AP'@AQ')"
        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' AP ΨP x C X Y Z)
      from (x # AP) ♯* C have "AP ♯* C" by simp
      from (x # AP) ♯* X x  AP have "AP ♯* (x#X)" by simp
      from (x # AP) ♯* Y have "AP ♯* Y" by simp
      from (x # AP) ♯* Z have "AP ♯* 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 = AP, ΨP distinct AP xvec ♯* M AP ♯* P AP ♯* M AP ♯* xvec AP ♯* N AP ♯* C
        AP ♯* (x#X) AP ♯* Y AP ♯* Z xvec ♯* P xvec ♯* C xvec ♯* (x#X) xvec ♯* Y xvec ♯* Z

      obtain p Ψ' AP' ΨP' where S: "set p  set xvec × set(p  xvec)"
        and FrP': "extractFrame P' = AP', ΨP'" and PeqP': "(p  ΨP)  Ψ'  ΨP'" and "distinct AP'"
        and "AP' ♯* C" and "AP' ♯* (x#X)" and "AP' ♯* Y" and "AP' ♯* N"
        and "AP' ♯* Z" and "AP' ♯* xvec" and "AP' ♯* P'" and "(p  xvec) ♯* xvec" and "(p  xvec) ♯* N"
        and "(p  xvec) ♯* AP'" 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 "AP' ♯* X" and "x  AP'" and "(p  xvec) ♯* X" and "x  (p  xvec)" by simp+
      from FrP' (p  xvec) ♯* AP' (p  xvec) ♯* P' have "(p  xvec) ♯* ΨP'"
        by (metis extractFrameFreshChain freshFrameDest)
      from S AP' ♯* xvec (p  xvec) ♯* AP' have "p  AP' = AP'" by simp

      from (p  xvec) ♯* ΨP' S
      have "⦇ν*xvec(AP', ΨP') = ⦇ν*(p  xvec)(p  AP', ΨP')"
        by(intro frameChainAlpha) (auto simp add: fresh_star_def frameResChainFresh)
      then have "(xvec@AP'), ΨP' = (p  (xvec@AP'), ΨP')"
        by (metis frameChainAppend frameResChainEqvt)
      with p  AP' = AP' have "(xvec@AP'), ΨP' = ((p  xvec)@AP'), (p  ΨP')"
        by(simp add: eqvts)
      then have "(x#(xvec@AP')), ΨP' = (x#((p  xvec)@AP')), (p  ΨP')"
        by simp

      moreover from FrP' have "extractFrame (⦇νx(⦇ν*xvecP')) = (x#(xvec@AP')), ΨP'"
        by (metis extractFrameResChain extractFrame_extractFrame'_extractFrame''.simps frameChainAppend frameResChain.step)

      ultimately have FrP'2: "extractFrame (⦇νx(⦇ν*xvecP')) = (x#((p  xvec)@AP')), (p  ΨP')"
        by simp

      from PeqP' have "(p  ((p  ΨP)  Ψ'))  (p  ΨP')"
        by (metis AssertionStatEqClosed)
      with distinctPerm p have PeqP'2: "ΨP  (p  Ψ')  (p  ΨP')"
        by(simp add: eqvts)

      note FrP'2 PeqP'2

      moreover from x  C (p  xvec) ♯* C AP' ♯* C
      have "(x#((p  xvec)@AP')) ♯* C" by simp
      moreover from (x # AP) ♯* ⦇νx(⦇ν*xvecP') AP' ♯* (x#X) AP' ♯* P' AP' ♯* xvec
        x  (p  xvec) (p  xvec) ♯* xvec (p  xvec) ♯* P'
      have "(x#((p  xvec)@AP')) ♯* (⦇νx(⦇ν*xvecP'))"
        by simp
      moreover from x  X (p  xvec) ♯* X AP' ♯* X
      have "(x#((p  xvec)@AP')) ♯* X" by simp
      moreover from x  Y (p  xvec) ♯* Y AP' ♯* Y
      have "(x#((p  xvec)@AP')) ♯* Y" by simp
      moreover from x  Z (p  xvec) ♯* Z AP' ♯* Z
      have "(x#((p  xvec)@AP')) ♯* Z" by simp
      moreover from x  (p  xvec) x  AP' (p  xvec) ♯* AP'
        distinct AP' distinct xvec distinctPerm p
      have "distinct (x#((p  xvec)@AP'))" by simp
      ultimately show ?case
        by(rule cBrClose)
    next
      case(cScope Ψ P P' x AP ΨP C X Y Z)
      then obtain Ψ' AP' ΨP' where FrP': "extractFrame P' = AP', ΨP'" and "distinct AP'"
        and "ΨP  Ψ'  ΨP'" and "AP' ♯* C" and "AP' ♯* P'"
        and "AP' ♯* (x#X)" and "AP' ♯* Y" and "AP' ♯* Z"
        using cScope(4)[where ba="x#X"] by (metis freshStarAtom fresh_star_list_cons(1))
      from AP' ♯* (x#X) have "x  AP'" and "AP' ♯* X" by simp+
      moreover from FrP' have "extractFrame(⦇νxP') = (x#AP'), ΨP'" by simp
      moreover note ΨP  Ψ'  ΨP'
      moreover from x  C AP' ♯* C have "(x#AP') ♯* C" by simp
      moreover from  AP' ♯* P' have "(x#AP') ♯* (⦇νxP')" by(simp add: abs_fresh fresh_star_def)
      moreover from x  X AP' ♯* X have "(x#AP') ♯* X" by simp
      moreover from x  Y AP' ♯* Y have "(x#AP') ♯* Y" by simp
      moreover from x  Z AP' ♯* Z have "(x#AP') ♯* Z" by simp
      moreover from x  AP' distinct AP' have "distinct(x#AP')" by simp
      ultimately show ?case by(elim cScope)
    next
      case(cBang Ψ P P' AP ΨP C B Y Z)
      then obtain Ψ' AP' ΨP' where FrP': "extractFrame P' = AP', ΨP'"
        and "(ΨP  𝟭)  Ψ'  ΨP'"
        and "AP' ♯* C" and "AP' ♯* P'" and "distinct AP'"
        and "AP' ♯* B" and "AP' ♯* Y" and "AP' ♯* Z"
        using cBang by (metis psiFreshVec(4) psiFreshVec(7))
      with ΨP  𝟭 (ΨP  𝟭)  Ψ'  ΨP' have "𝟭  Ψ'  ΨP'"
        by(metis Identity AssertionStatEqTrans composition' Commutativity Associativity AssertionStatEqSym)
      then show ?case using FrP' AP' ♯* P' AP' ♯* C AP' ♯* B AP' ♯* Y AP' ♯* Z distinct AP'
        by(elim cBang)
    qed
    with A have ?thesis by simp
  }
  moreover have "AP ♯* ([]::name list)" and "AP ♯* ([]::'b list)" and "AP ♯* ([]::('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 AP   :: "name list"
    and ΨP  :: 'b
    and C    :: "'f::fs_name"
    and C'   :: "'g::fs_name"

assumes Trans: "Ψ  P α  P'"
  and   "extractFrame P = AP, ΨP"
  and   "distinct AP"
  and   "bn α ♯* subject α"
  and   "distinct(bn α)"
  and   "AP ♯* α"
  and   "AP ♯* P"
  and   "AP ♯* C"
  and   "AP ♯* C'"
  and   "bn α ♯* P"
  and   "bn α ♯* C'"

obtains p Ψ' AP' ΨP' where "set p  set(bn α) × set(bn(p  α))" and "(p  ΨP)  Ψ'  ΨP'" and "distinctPerm p" and
  "extractFrame P' = AP', ΨP'" and "AP' ♯* P'" and "AP' ♯* α" and "AP' ♯* (p  α)" and
  "AP' ♯* C" and "(bn(p  α)) ♯* C'" and "(bn(p  α)) ♯* α" and "(bn(p  α)) ♯* P'" and "distinct AP'"
  using assms
  apply(cases "α=τ")
  by(auto intro: expandTauFrame[where C=C] expandNonTauFrame[where C=C and C'=C'])

abbreviation
  frameImpJudge ("_ F _" [80, 80] 80)
  where "F 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 F G"
  and   "G F H"
  and   "H F I"

shows "F F I"
  using assms
  by(auto simp add: FrameStatEq_def) (blast intro: FrameStatImpTrans)

lemma transferNonTauFrame:
  fixes ΨF  :: 'b
    and P    :: "('a, 'b, 'c) psi"
    and α    :: "'a action"
    and P'   :: "('a, 'b, 'c) psi"
    and AF   :: "name list"
    and AG   :: "name list"
    and ΨG   :: 'b

assumes "ΨF  P α  P'"
  and   "extractFrame P = AP, ΨP"
  and   "distinct AP"
  and   "distinct(bn α)"
  and   "AF, ΨF  ΨP F AG, ΨG  ΨP"
  and   "AF ♯* P"
  and   "AG ♯* P"
  and   "AF ♯* subject α"
  and   "AG ♯* subject α"
  and   "AP ♯* AF"
  and   "AP ♯* AG"
  and   "AP ♯* ΨF"
  and   "AP ♯* ΨG"
  and   "α  τ"

shows "ΨG  P α  P'"
  using assms
proof(nominal_induct ΨF P Rs=="α  P'" AP ΨP avoiding: α P' ΨG AF AG rule: semanticsFrameInduct)
  case(cAlpha ΨF P AP ΨP p α P' ΨG AF AG)
  from AF, ΨF  (p  ΨP) F AG, ΨG  (p  ΨP)
  have "(p  (AF, ΨF  (p  ΨP))) F (p  (AG, ΨG  (p  ΨP)))"
    by(rule FrameStatImpClosed)
  with AP ♯* AF (p  AP) ♯* AF AP ♯* ΨF (p  AP) ♯* ΨF AP ♯* AG (p  AP) ♯* AG AP ♯* ΨG (p  AP) ♯* ΨG
    distinctPerm p set p  set AP × set (p  AP) have "AF, ΨF  ΨP F AG, ΨG  ΨP"
    by(simp add: eqvts)
  with cAlpha show ?case by force
next
  case(cInput ΨF M K xvec N Tvec P α P' ΨG AF AG)
  from cInput have "AF ♯* K" and "AG ♯* K" by(auto simp add: residualInject)

  from AF ♯* (M⦇λ*xvec N⦈.P) AG ♯* (M⦇λ*xvec N⦈.P) have "AF ♯* M" and "AG ♯* M" by simp+
  from ΨF  M  K
  have "ΨF  𝟭  M  K"
    by(blast intro: statEqEnt Identity AssertionStatEqSym)
  with AF ♯* M AF ♯* K
  have "(AF, ΨF  𝟭) F M  K"
    by(force intro: frameImpI)
  with AF, ΨF  𝟭 F AG, ΨG  𝟭
  have "(AG, ΨG  𝟭) F M  K"
    by(simp add: FrameStatEq_def FrameStatImp_def)
  with AG ♯* M AG ♯* K
  have "ΨG  𝟭  M  K" by(force dest: frameImpE)
  then have "Ψ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 ΨF M K xvec N Tvec P α P' ΨG AF AG)
  from cBrInput have "AF ♯* K" and "AG ♯* K" by(auto simp add: residualInject)

  from AF ♯* (M⦇λ*xvec N⦈.P) AG ♯* (M⦇λ*xvec N⦈.P) have "AF ♯* M" and "AG ♯* M" by simp+
  from ΨF  K  M
  have "ΨF  𝟭  K  M"
    by(blast intro: statEqEnt Identity AssertionStatEqSym)
  with AF ♯* M AF ♯* K
  have "(AF, ΨF  𝟭) F K  M"
    by(force intro: frameImpI)
  with AF, ΨF  𝟭 F AG, ΨG  𝟭
  have "(AG, ΨG  𝟭) F K  M"
    by(simp add: FrameStatEq_def FrameStatImp_def)
  with AG ♯* M AG ♯* K
  have "ΨG  𝟭  K  M" by(force dest: frameImpE)
  then have "Ψ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 ΨF M K N P α P' ΨG AF AG)
  from cOutput have "AF ♯* K" and "AG ♯* K" by(auto simp add: residualInject)

  from AF ♯* (MN⟩.P) AG ♯* (MN⟩.P) have "AF ♯* M" and "AG ♯* M" by simp+
  from ΨF  M  K
  have "ΨF  𝟭  M  K"
    by(blast intro: statEqEnt Identity AssertionStatEqSym)
  with AF ♯* M AF ♯* K
  have "(AF, ΨF  𝟭) F M  K"
    by(force intro: frameImpI)
  with AF, ΨF  𝟭 F AG, ΨG  𝟭
  have "(AG, ΨG  𝟭) F M  K"
    by(simp add: FrameStatImp_def)
  with AG ♯* M AG ♯* K
  have "ΨG  𝟭  M  K" by(force dest: frameImpE)
  then have "ΨG  M  K" by(blast intro: statEqEnt Identity)
  then show ?case using cOutput Output by(force simp add: residualInject)
next
  case(cBrOutput ΨF M K N P α P' ΨG AF AG)
  from cBrOutput have "AF ♯* K" and "AG ♯* K" by(auto simp add: residualInject)

  from AF ♯* (MN⟩.P) AG ♯* (MN⟩.P) have "AF ♯* M" and "AG ♯* M" by simp+
  from ΨF  M  K
  have "ΨF  𝟭  M  K"
    by(blast intro: statEqEnt Identity AssertionStatEqSym)
  with AF ♯* M AF ♯* K
  have "(AF, ΨF  𝟭) F M  K"
    by(force intro: frameImpI)
  with AF, ΨF  𝟭 F AG, ΨG  𝟭
  have "(AG, ΨG  𝟭) F M  K"
    by(simp add: FrameStatImp_def)
  with AG ♯* M AG ♯* K
  have "ΨG  𝟭  M  K" by(force dest: frameImpE)
  then have "ΨG  M  K" by(blast intro: statEqEnt Identity)
  then show ?case using cBrOutput BrOutput by(force simp add: residualInject)
next
  case(cCase ΨF P φ Cs AP ΨP α P' ΨG AF AG)
  from ΨP  𝟭 have "AF, ΨF  ΨP F AF, ΨF  𝟭"
    by(metis frameIntCompositionSym Identity AssertionStatEqTrans)
  moreover note AF, ΨF  𝟭 F AG, ΨG  𝟭
  moreover from ΨP  𝟭 have "AG, ΨG  𝟭 F AG, ΨG  ΨP"
    by(metis frameIntCompositionSym Identity AssertionStatEqTrans AssertionStatEqSym)
  ultimately have "AF, ΨF  ΨP F AG, ΨG  ΨP"
    by(rule FrameStatEqImpCompose)
  with cCase have "ΨG  P α  P'"
    by (metis freshStarPair(2) memFreshChain(1) psiCasesFreshChain(1) psiFreshVec(3))
  moreover note (φ, P)  set Cs
  moreover from AF ♯* (Cases Cs)AG ♯* (Cases Cs) (φ, P)  set Cs have "AF ♯* φ" and "AG ♯* φ"
    by(auto dest: memFreshChain)
  from ΨF  φ have "ΨF  𝟭  φ" by(blast intro: statEqEnt Identity AssertionStatEqSym)
  with AF ♯* φ have "(AF, ΨF  𝟭) F φ" by(force intro: frameImpI)
  with AF, ΨF  𝟭 F AG, ΨG  𝟭 have "(AG, ΨG  𝟭) F φ"
    by(simp add: FrameStatImp_def)
  with AG ♯* φ have "ΨG  𝟭  φ" by(force dest: frameImpE)
  then have "ΨG  φ" by(blast intro: statEqEnt Identity)
  ultimately show ?case using guarded P cCase Case by(force intro: residualInject)
next
  case(cPar1 ΨF ΨQ P α P' AQ Q AP ΨP α' PQ' ΨG AF AG)
  from AF ♯* (P  Q) AG ♯* (P  Q) have "AF ♯* P" and "AG ♯* P" and "AF ♯* Q" and "AG ♯* Q"
    by simp+
  have FrQ: "extractFrame Q = AQ, ΨQ" by fact
  have "AF, (ΨF  ΨQ)  ΨP F AF, ΨF  ΨP  ΨQ"
    by(metis Associativity Composition AssertionStatEqSym AssertionStatEqTrans Commutativity frameResChainPres frameNilStatEq)
  moreover note AF, ΨF  ΨP  ΨQ F AG, ΨG  ΨP  ΨQ
  moreover have "AG, ΨG  ΨP  ΨQ F AG, (ΨG  ΨQ)  ΨP "
    by(metis Associativity Composition AssertionStatEqSym AssertionStatEqTrans Commutativity frameResChainPres frameNilStatEq)
  ultimately have "AF, (ΨF  ΨQ)  ΨP F AG, (ΨG  ΨQ)  ΨP"
    by(rule FrameStatEqImpCompose)
  moreover from AF ♯* Q AG ♯* Q FrQ AQ ♯* AF AQ ♯* AG have "AF ♯* ΨQ" and  "AG ♯* ΨQ"
    by(force dest: extractFrameFreshChain)+
  moreover note AF ♯* P AG ♯* P AF ♯* subject α' AG ♯* subject α' AP ♯* AF AP ♯* AG AP ♯* ΨF AP ♯* ΨQ AP ♯* AG AP ♯* Ψ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  "ΨG  ΨQ  P α  P'" using α'  τ by(force intro: cPar1)
  then show ?case using FrQ (bn α) ♯* Q AQ ♯* ΨG AQ ♯* P AQ ♯* α using cPar1
    by(metis Par1)
next
  case(cPar2 ΨF ΨP Q α Q' AP P AQ ΨQ α' PQ' ΨG AF AG)
  from AF ♯* (P  Q) AG ♯* (P  Q) have "AF ♯* P" and "AG ♯* P" and "AF ♯* Q" and "AG ♯* Q"
    by simp+
  have FrP: "extractFrame P = AP, ΨP" by fact
  have "AF, (ΨF  ΨP)  ΨQ F AF, ΨF  ΨP  ΨQ"
    by(metis Associativity frameResChainPres frameNilStatEq)
  moreover note AF, ΨF  ΨP  ΨQ F AG, ΨG  ΨP  ΨQ
  moreover have "AG, ΨG  ΨP  ΨQ F AG, (ΨG  ΨP)  ΨQ "
    by(metis Associativity AssertionStatEqSym frameResChainPres frameNilStatEq)
  ultimately have "AF, (ΨF  ΨP)  ΨQ F AG, (ΨG  ΨP)  ΨQ"
    by(rule FrameStatEqImpCompose)
  moreover from AF ♯* P AG ♯* P FrP AP ♯* AF AP ♯* AG have "AF ♯* ΨP" and  "AG ♯* ΨP"
    by(force dest: extractFrameFreshChain)+
  moreover note AF ♯* Q AG ♯* Q AF ♯* subject α' AG ♯* subject α' AQ ♯* AF AQ ♯* AG AQ ♯* ΨF AQ ♯* ΨP AQ ♯* AG AQ ♯* Ψ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  "ΨG  ΨP  Q α  Q'" using α'  τ by(force intro: cPar2)
  then show ?case using FrP (bn α) ♯* P AP ♯* ΨG AP ♯* Q AP ♯* α 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 ΨF ΨQ P M N P' AP ΨP Q Q' AQ α PQ' ΨG AF AG)
  from AF, ΨF  ΨP  ΨQ F AG, ΨG  ΨP  ΨQ
  have "AF, (ΨF  ΨQ)  ΨP F AG, ΨG  ΨP  ΨQ"
    by (metis AssertionStatEqTrans AssertionStatEq_def Associativity FrameStatImpTrans associativitySym frameImpNilStatEq frameImpResChainPres)
  moreover have "AG, ΨG  ΨP  ΨQ F AG, (ΨG  ΨQ)  ΨP"
    by (metis AssertionStatEqTrans AssertionStatEq_def Associativity associativitySym frameImpNilStatEq frameImpResChainPres)
  ultimately have FimpP: "AF, (ΨF  ΨQ)  ΨP F AG, (ΨG  ΨQ)  ΨP"
    by (metis FrameStatImpTrans)

  from ¿MN  P'  Q' = α  PQ' have "α = ¿MN" and "(P'  Q') = PQ'"
    by(simp add: residualInject)+

  moreover note FimpP AF ♯* (P  Q) AG ♯* (P  Q) AF ♯* subject α AG ♯* subject α
    AP ♯* AF AP ♯* AG AP ♯* ΨF AP ♯* ΨG AP ♯* ΨQ
  ultimately have TransP: "ΨG  ΨQ  P  ¿MN  P'"
    by(intro cBrMerge(2)[where bd="AG"]) auto

  from AF, ΨF  ΨP  ΨQ F AG, ΨG  ΨP  ΨQ
  have "AF, (ΨF  ΨP)  ΨQ F AG, ΨG  ΨP  ΨQ"
    by (metis FrameStatEq_def FrameStatImpTrans frameIntAssociativity)
  moreover have "AG, ΨG  ΨP  ΨQ F AG, (ΨG  ΨP)  ΨQ"
    by (metis FrameStatEq_def frameIntAssociativity)
  ultimately have FimpQ: "AF, (ΨF  ΨP)  ΨQ F AG, (ΨG  ΨP)  ΨQ"
    by (metis FrameStatImpTrans)

  note α = ¿MN (P'  Q') = PQ' FimpQ AF ♯* (P  Q) AG ♯* (P  Q) AF ♯* subject α AG ♯* subject α
    AQ ♯* AF AQ ♯* AG AQ ♯* ΨF AQ ♯* ΨG AQ ♯* ΨP
  then have TransQ: "ΨG  ΨP  Q  ¿MN  Q'"
    by(intro cBrMerge(6)[where bd="AG"]) auto

  have "ΨG  P  Q  ¿MN  P'  Q'" using TransP TransQ cBrMerge
    by(intro BrMerge)
  with ¿MN  P'  Q' = α  PQ'
  show ?case by simp
next
  case(cBrComm1 ΨF ΨQ P M N P' AP ΨP Q xvec Q' AQ α PQ' ΨG AF AG)
  from ¡M⦇ν*xvec⦈⟨N  P'  Q' = α  PQ' AF ♯* subject α AG ♯* subject α
  have "AF ♯* M" and "AG ♯* M"
    by(auto simp add: residualInject)

  from AF, ΨF  ΨP  ΨQ F AG, ΨG  ΨP  ΨQ
  have "AF, (ΨF  ΨQ)  ΨP F AG, ΨG  ΨP  ΨQ"
    by (metis AssertionStatEqTrans AssertionStatEq_def Associativity FrameStatImpTrans associativitySym frameImpNilStatEq frameImpResChainPres)
  moreover have "AG, ΨG  ΨP  ΨQ F AG, (ΨG  ΨQ)  ΨP"
    by (metis AssertionStatEqTrans AssertionStatEq_def Associativity associativitySym frameImpNilStatEq frameImpResChainPres)
  ultimately have FimpP: "AF, (ΨF  ΨQ)  ΨP F AG, (ΨG  ΨQ)  ΨP"
    by (metis FrameStatImpTrans)

  note FimpP AF ♯* (P  Q) AG ♯* (P  Q) AF ♯* M AG ♯* M
    AP ♯* AF AP ♯* AG AP ♯* ΨF AP ♯* ΨG AP ♯* ΨQ
  then have TransP: "ΨG  ΨQ  P  ¿MN  P'"
    by (intro cBrComm1(4)[where bc="AF" and bd="AG"]) auto
  from AF, ΨF  ΨP  ΨQ F AG, ΨG  ΨP  ΨQ
  have "AF, (ΨF  ΨP)  ΨQ F AG, ΨG  ΨP  ΨQ"
    by (metis FrameStatEq_def FrameStatImpTrans frameIntAssociativity)
  moreover have "AG, ΨG  ΨP  ΨQ F AG, (ΨG  ΨP)  ΨQ"
    by (metis FrameStatEq_def frameIntAssociativity)
  ultimately have FimpQ: "AF, (ΨF  ΨP)  ΨQ F AG, (ΨG  ΨP)  ΨQ"
    by (metis FrameStatImpTrans)
  note distinct xvec FimpQ AF ♯* (P  Q) AG ♯* (P  Q) AF ♯* M AG ♯* M
    AQ ♯* AF AQ ♯* AG AQ ♯* ΨF AQ ♯* ΨG AQ ♯* ΨP
  then have TransQ: "ΨG  ΨP  Q  ¡M⦇ν*xvec⦈⟨N  Q'"
    by (simp add: cBrComm1.hyps(8) freshCompChain(1))

  from TransP TransQ cBrComm1
  have "Ψ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 ΨF ΨQ P M xvec N P' AP ΨP Q Q' AQ α PQ' ΨG AF AG)
  from ¡M⦇ν*xvec⦈⟨N  P'  Q' = α  PQ' AF ♯* subject α AG ♯* subject α
  have "AF ♯* M" and "AG ♯* M"
    by(auto simp add: residualInject)

  from AF, ΨF  ΨP  ΨQ F AG, ΨG  ΨP  ΨQ
  have "AF, (ΨF  ΨP)  ΨQ F AG, ΨG  ΨP  ΨQ"
    by (metis FrameStatEq_def FrameStatImpTrans frameIntAssociativity)
  moreover have "AG, ΨG  ΨP  ΨQ F AG, (ΨG  ΨP)  ΨQ"
    by (metis FrameStatEq_def frameIntAssociativity)
  ultimately have FimpQ: "AF, (ΨF  ΨP)  ΨQ F AG, (ΨG  ΨP)  ΨQ"
    by (metis FrameStatImpTrans)

  note FimpQ AF ♯* (P  Q) AG ♯* (P  Q) AF ♯* M AG ♯* M
    AQ ♯* AF AQ ♯* AG AQ ♯* ΨF AQ ♯* ΨG AQ ♯* ΨP
  then have TransQ: "ΨG  ΨP  Q  ¿MN  Q'"
    by(intro cBrComm2(8)[where bc="AF" and bd="AG"]) auto

  from AF, ΨF  ΨP  ΨQ F AG, ΨG  ΨP  ΨQ
  have "AF, (ΨF  ΨQ)  ΨP F AG, ΨG  ΨP  ΨQ"
    by (metis AssertionStatEqTrans AssertionStatEq_def Associativity FrameStatImpTrans associativitySym frameImpNilStatEq frameImpResChainPres)
  moreover have "AG, ΨG  ΨP  ΨQ F AG, (ΨG  ΨQ)  ΨP"
    by (metis AssertionStatEqTrans AssertionStatEq_def Associativity associativitySym frameImpNilStatEq frameImpResChainPres)
  ultimately have FimpP: "AF, (ΨF  ΨQ)  ΨP F AG, (ΨG  ΨQ)  ΨP"
    by (metis FrameStatImpTrans)

  note distinct xvec FimpP AF ♯* (P  Q) AG ♯* (P  Q) AF ♯* M AG ♯* M
    AP ♯* AF AP ♯* AG AP ♯* ΨF AP ♯* ΨG AP ♯* ΨQ
  then have TransP: "ΨG  ΨQ  P  ¡M⦇ν*xvec⦈⟨N  P'"
    by(intro cBrComm2(4)[where bc="AF" and bd="AG"]) auto

  from TransP TransQ cBrComm2
  have "Ψ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 ΨF P M xvec yvec N P' x AP ΨP α P'' ΨG AF AG)
  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 "ΨG  P M⦇ν*(xvec@yvec)⦈⟨N  P'" using cOpen
    by(intro cOpen(4)[where bc="AF" and bd="AG"]) auto
  with x  supp N x  ΨG x  M x  xvec x  yvec
  have "ΨG  ⦇νxP 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 ΨF P M xvec yvec N P' x AP ΨP α P'' ΨG AF AG)
  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 "ΨG  P ¡M⦇ν*(xvec@yvec)⦈⟨N  P'" using cBrOpen
    by(intro cBrOpen) (assumption | simp)+
  with x  supp N x  ΨG x  M x  xvec x  yvec
  have "ΨG  ⦇νxP ¡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 ΨF P α P' x AP ΨP α' xP ΨG AF AG)
  from α  ⦇νxP' = α'  xP x  α x  α' obtain P'' where "xP = ⦇νxP''" and "α  P' = α'  P''"
    by(drule_tac sym) (force intro: actionScopeDest)
  then have "ΨG  P α  P'" using cScope by auto
  with x  ΨG x  α' α  P' = α'  P'' xP = ⦇νxP'' show ?case
    by(metis Scope)
next
  case(cBang ΨF P AP ΨP α P' ΨG AF AG)
  from ΨP  𝟭 have "AF, ΨF  ΨP  𝟭 F AF, ΨF  𝟭"
    by(metis frameIntCompositionSym Identity AssertionStatEqTrans)
  moreover note AF, ΨF  𝟭 F AG, ΨG  𝟭
  moreover from ΨP  𝟭 have "AG, ΨG  𝟭 F AG, ΨG  ΨP  𝟭"
    by(metis frameIntCompositionSym Identity AssertionStatEqTrans AssertionStatEqSym)
  ultimately have "AF, ΨF  ΨP  𝟭 F AG, ΨG  ΨP  𝟭"
    by(rule FrameStatEqImpCompose)
  with cBang have "ΨG  P  !P α  P'" by force
  then show ?case using guarded P using cBang by(metis Bang)
qed

lemma transferTauFrame:
  fixes ΨF  :: 'b
    and P  :: "('a, 'b, 'c) psi"
    and P'   :: "('a, 'b, 'c) psi"
    and AF   :: "name list"
    and AG   :: "name list"
    and ΨG   :: 'b

assumes "ΨF  P τ  P'"
  and   "extractFrame P = AP, ΨP"
  and   "distinct AP"
  and   "AF, ΨF  ΨP F AG, ΨG  ΨP"
  and   "AF ♯* P"
  and   "AG ♯* P"
  and   "AP ♯* AF"
  and   "AP ♯* AG"
  and   "AP ♯* ΨF"
  and   "AP ♯* ΨG"

shows "ΨG  P τ  P'"
  using assms
proof(nominal_induct avoiding: ΨG AF AG rule: tauFrameInduct)
  case(cAlpha ΨF P P' AP ΨP p ΨG AF AG)
  from AF, ΨF  (p  ΨP) F AG, ΨG  (p  ΨP)
  have "(p  (AF, ΨF  (p  ΨP))) F (p  (AG, ΨG  (p  ΨP)))"
    by(rule FrameStatImpClosed)
  with AP ♯* AF (p  AP) ♯* AF AP ♯* ΨF (p  AP) ♯* ΨF AP ♯* AG (p  AP) ♯* AG AP ♯* ΨG (p  AP) ♯* ΨG
    distinctPerm p set p  set AP × set (p  AP) have "AF, ΨF  ΨP F AG, ΨG  ΨP"
    by(simp add: eqvts)
  with cAlpha show ?case by blast
next
  case(cCase ΨF P P' φ Cs AP ΨP ΨG AF AG)
  from ΨP  𝟭 have "AF, ΨF  ΨP F AF, ΨF  𝟭"
    by(metis frameIntCompositionSym Identity AssertionStatEqTrans)
  moreover note AF, ΨF  𝟭 F AG, ΨG  𝟭
  moreover from ΨP  𝟭 have "AG, ΨG  𝟭 F AG, ΨG  ΨP"
    by(metis frameIntCompositionSym Identity AssertionStatEqTrans AssertionStatEqSym)
  ultimately have "AF, ΨF  ΨP F AG, ΨG  ΨP"
    by(rule FrameStatEqImpCompose)
  with cCase have "ΨG  P τ  P'"
    by (metis freshStarPair(2) memFreshChain(1) psiCasesFreshChain(1) psiFreshVec(3))
  moreover note (φ, P)  set Cs
  moreover from AF ♯* (Cases Cs)AG ♯* (Cases Cs) (φ, P)  set Cs have "AF ♯* φ" and "AG ♯* φ"
    by(auto dest: memFreshChain)
  from ΨF  φ have "ΨF  𝟭  φ" by(blast intro: statEqEnt Identity AssertionStatEqSym)
  with AF ♯* φ have "(AF, ΨF  𝟭) F φ" by(force intro: frameImpI)
  with AF, ΨF  𝟭 F AG, ΨG  𝟭 have "(AG, ΨG  𝟭) F φ"
    by(simp add: FrameStatImp_def)
  with AG ♯* φ have "ΨG  𝟭  φ" by(force dest: frameImpE)
  then have "ΨG  φ" by(blast intro: statEqEnt Identity)
  ultimately show ?case using guarded P by(rule Case)
next
  case(cPar1 ΨF ΨQ P P' AQ Q AP ΨP ΨG AF AG)
  from AF ♯* (P  Q) AG ♯* (P  Q) have "AF ♯* P" and "AG ♯* P" and "AF ♯* Q" and "AG ♯* Q"
    by simp+
  have IH: "Ψ AF AG. AF, (ΨF  ΨQ)  ΨP F AG, Ψ  ΨP; AF ♯* P; AG ♯* P;
                           AP ♯* AF; AP ♯* AG; AP ♯* (ΨF  ΨQ); AP ♯* Ψ  Ψ  P τ  P'"
    by fact
  have FrQ: "extractFrame Q = AQ, ΨQ" by fact
  have "AF, (ΨF  ΨQ)  ΨP F AF, ΨF  ΨP  ΨQ"
    by(metis Associativity Composition AssertionStatEqSym AssertionStatEqTrans Commutativity frameResChainPres frameNilStatEq)
  moreover note AF, ΨF  ΨP  ΨQ F AG, ΨG  ΨP  ΨQ
  moreover have "AG, ΨG  ΨP  ΨQ F AG, (ΨG  ΨQ)  ΨP "
    by(metis Associativity Composition AssertionStatEqSym AssertionStatEqTrans Commutativity frameResChainPres frameNilStatEq)
  ultimately have "AF, (ΨF  ΨQ)  ΨP F AG, (ΨG  ΨQ)  ΨP"
    by(rule FrameStatEqImpCompose)
  moreover from AF ♯* Q AG ♯* Q FrQ AQ ♯* AF AQ ♯* AG have "AF ♯* ΨQ" and  "AG ♯* ΨQ"
    by(force dest: extractFrameFreshChain)+
  moreover note AF ♯* P AG ♯* P AP ♯* AF AP ♯* AG AP ♯* ΨF AP ♯* ΨQ AP ♯* AG AP ♯* ΨG
  ultimately have  "ΨG  ΨQ  P τ  P'"
    using IH by blast
  then show ?case using FrQ AQ ♯* ΨG AQ ♯* P
    by(intro Par1) auto
next
  case(cPar2 ΨF ΨP Q Q' AP P AQ ΨQ ΨG AF AG)
  from AF ♯* (P  Q) AG ♯* (P  Q) have "AF ♯* P" and "AG ♯* P" and "AF ♯* Q" and "AG ♯* Q"
    by simp+
  have IH: "Ψ AF AG. AF, (ΨF  ΨP)  ΨQ F AG, Ψ  ΨQ; AF ♯* Q; AG ♯* Q;
                           AQ ♯* AF; AQ ♯* AG; AQ ♯* (ΨF  ΨP); AQ ♯* Ψ  Ψ  Q τ  Q'"
    by fact
  have FrP: "extractFrame P = AP, ΨP" by fact
  have "AF, (ΨF  ΨP)  ΨQ F AF, ΨF  ΨP  ΨQ"
    by(metis Associativity frameResChainPres frameNilStatEq)
  moreover note AF, ΨF  ΨP  ΨQ F AG, ΨG  ΨP  ΨQ
  moreover have "AG, ΨG  ΨP  ΨQ F AG, (ΨG  ΨP)  ΨQ "
    by(metis Associativity AssertionStatEqSym frameResChainPres frameNilStatEq)
  ultimately have "AF, (ΨF  ΨP)  ΨQ F AG, (ΨG  ΨP)  ΨQ"
    by(rule FrameStatEqImpCompose)
  moreover from AF ♯* P AG ♯* P FrP AP ♯* AF AP ♯* AG have "AF ♯* ΨP" and  "AG ♯* ΨP"
    by(force dest: extractFrameFreshChain)+
  moreover note AF ♯* Q AG ♯* Q AQ ♯* AF AQ ♯* AG AQ ♯* ΨF AQ ♯* ΨP AQ ♯* AG AQ ♯* ΨG
  ultimately have  "ΨG  ΨP  Q τ  Q'"
    using IH by blast
  then show ?case using FrP AP ♯* ΨG AP ♯* Q
    by(intro Par2) auto
next
  case(cComm1 ΨF ΨQ P M N P' AP ΨP Q K xvec Q' AQ ΨG AF AG)
  have FimpG: "AF, ΨF  ΨP  ΨQ F AG, ΨG  ΨP  ΨQ" by fact
  from AF ♯* (P  Q) AG ♯* (P  Q) have "AF ♯* P" and "AG ♯* P" and "AF ♯* Q" and "AG ♯* Q"
    by simp+
  have FrP: "extractFrame P = AP, ΨP" by fact
  have FrQ: "extractFrame Q = AQ, ΨQ" by fact
  from AF ♯* P AG ♯* P FrP AP ♯* AF AP ♯* AG have "AF ♯* ΨP" and  "AG ♯* ΨP"
    by(force dest: extractFrameFreshChain)+
  from AF ♯* Q AG ♯* Q FrQ AQ ♯* AF AQ ♯* AG have "AF ♯* ΨQ" and  "AG ♯* ΨQ"
    by(force dest: extractFrameFreshChain)+

  from ΨF  ΨP  Q K⦇ν*xvec⦈⟨N  Q' have "ΨF  ΨP  Q  ROut K (⦇ν*xvecN ≺' Q')"
    by(simp add: residualInject)
  with FrQ distinct AQ
  obtain K' where KeqK': "(ΨF  ΨP)  ΨQ  K  K'" and "AP ♯* K'" and "AF ♯* K'" and "AG ♯* K'"
    using AP ♯* Q AQ ♯* AF AQ ♯* AG AF ♯* Q AG ♯* Q AQ ♯* ΨF AQ ♯* ΨP AP ♯* AQ AQ ♯* Q AQ ♯* K xvec ♯* K distinct xvec
    by(elim outputObtainPrefix[where B="AP@AF@AG"]) force+
  have "ΨG  ΨQ  P K'N  P'"
  proof -
    from KeqK' have "ΨF  (ΨP  ΨQ)  K  K'" by(rule statEqEnt[OF Associativity])
    with ΨF  (ΨP  ΨQ)  M  K have "ΨF  (ΨP  ΨQ)  M  K'"
      by(rule chanEqTrans)
    then have "(ΨF  ΨQ)  ΨP  M  K'"
      by(metis statEqEnt AssertionStatEqSym Associativity AssertionStatEqTrans compositionSym Commutativity)
    with ΨF  ΨQ  P MN  P' FrP distinct AP
    have "ΨF  ΨQ  P K'N  P'" using AP ♯* ΨF AP ♯* ΨQ AP ♯* P AP ♯* M AP ♯* K'
      by(force intro: inputRenameSubject)
    moreover have "AF, (ΨF  ΨQ)  ΨP F AG, (ΨG  ΨQ)  ΨP"
    proof -
      have "AF, (ΨF  ΨQ)  ΨP F AF, ΨF  ΨP  ΨQ"
        by(metis Associativity Composition AssertionStatEqSym AssertionStatEqTrans Commutativity frameResChainPres frameNilStatEq)
      moreover have "AG, ΨG  ΨP  ΨQ F AG, (ΨG  ΨQ)  ΨP "
        by(metis Associativity Composition AssertionStatEqSym AssertionStatEqTrans Commutativity frameResChainPres frameNilStatEq)
      ultimately show ?thesis using FimpG
        by(elim FrameStatEqImpCompose)
    qed
    ultimately show ?thesis using AF ♯* P AG ♯* P AF ♯* K'
        AG ♯* K' AP ♯* AF AP ♯* AG AP ♯* ΨF AP ♯* ΨG AP ♯* ΨQ FrP distinct AP
      by(auto intro: transferNonTauFrame)
  qed

  moreover from FrP ΨF  ΨQ  P MN  P' distinct AP
  obtain M' where MeqM': "(ΨF  ΨQ)  ΨP  M  M'" and "AQ ♯* M'" and "AF ♯* M'" and "AG ♯* M'"
    using AQ ♯* P AP ♯* AF AP ♯* AG AF ♯* P AG ♯* P AP ♯* ΨF AP ♯* ΨQ AP ♯* AQ AP ♯* P AP ♯* M
    by(elim inputObtainPrefix[where B="AQ@AF@AG"]) force+

  have "ΨG  ΨP  Q M'⦇ν*xvec⦈⟨N  Q'"
  proof -
    from MeqM' have "ΨF  (ΨQ  ΨP)  M  M'"
      by(rule statEqEnt[OF Associativity])
    with ΨF  (ΨP  ΨQ)  M  K have "ΨF  (ΨQ  ΨP)  K  M'"
      by(blast intro: chanEqTrans chanEqSym compositionSym Commutativity statEqEnt)
    then have "(ΨF  ΨP)  ΨQ  K  M'"
      by(blast intro: statEqEnt AssertionStatEqSym Associativity
          AssertionStatEqTrans compositionSym Commutativity)
    with ΨF  ΨP  Q K⦇ν*xvec⦈⟨N  Q' FrQ distinct AQ
    have "ΨF  ΨP  Q M'⦇ν*xvec⦈⟨N  Q'" using AQ ♯* ΨF AQ ♯* ΨP AQ ♯* Q AQ ♯* K AQ ♯* M'
      by(force intro: outputRenameSubject)
    moreover have "AF, (ΨF  ΨP)  ΨQ F AG, (ΨG  ΨP)  ΨQ"
    proof -
      have "AF, (ΨF  ΨP)  ΨQ F AF, ΨF  ΨP  ΨQ"
        by(metis Associativity frameResChainPres frameNilStatEq)
      moreover have "AG, ΨG  ΨP  ΨQ F AG, (ΨG  ΨP)  ΨQ "
        by(metis Associativity AssertionStatEqSym frameResChainPres frameNilStatEq)
      ultimately show ?thesis using FimpG
        by(elim FrameStatEqImpCompose)
    qed

    ultimately show ?thesis using AF ♯* Q AG ♯* Q AF ♯* M' AG ♯* M'
        AQ ♯* AF AQ ♯* AG AQ ♯* ΨF AQ ♯* ΨG AQ ♯* ΨP FrQ distinct AQ distinct xvec
      by(auto intro: transferNonTauFrame)
  qed

  moreover have "ΨG  ΨP  ΨQ  K'  M'"
  proof -
    from MeqM' have "ΨF  ΨP  ΨQ  M'  M"
      by(blast intro: chanEqSym Associativity statEqEnt Commutativity compositionSym)
    moreover from KeqK' have "ΨF  ΨP  ΨQ  K  K'"
      by(blast intro: chanEqSym Associativity statEqEnt Commutativity compositionSym)
    ultimately have "ΨF  ΨP  ΨQ  K'  M'" using ΨF  ΨP  ΨQ  M  K
      by(blast intro: chanEqSym chanEqTrans)
    then show ?thesis using AF ♯* M' AF ♯* K' AG ♯* M' AG ♯* 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 AP ♯* ΨG AP ♯* P AP ♯* Q AP ♯* AQ AP ♯* K' AQ ♯* ΨG AQ ♯* P AQ ♯* Q AQ ♯* M' xvec ♯* P FrP FrQ
    by(intro Comm1) (assumption | simp)+
next
  case(cComm2 ΨF ΨQ P M xvec N P' AP ΨP Q K Q' AQ ΨG AF AG)
  have FimpG: "AF, ΨF  ΨP  ΨQ F AG, ΨG  ΨP  ΨQ" by fact
  from AF ♯* (P  Q) AG ♯* (P  Q) have "AF ♯* P" and "AG ♯* P" and "AF ♯* Q" and "AG ♯* Q"
    by simp+
  have FrP: "extractFrame P = AP, ΨP" by fact
  have FrQ: "extractFrame Q = AQ, ΨQ" by fact
  from AF ♯* P AG ♯* P FrP AP ♯* AF AP ♯* AG have "AF ♯* ΨP" and  "AG ♯* ΨP"
    by(force dest: extractFrameFreshChain)+
  from AF ♯* Q AG ♯* Q FrQ AQ ♯* AF AQ ♯* AG have "AF ♯* ΨQ" and  "AG ♯* ΨQ"
    by(force dest: extractFrameFreshChain)+
  from ΨF  ΨP  Q KN  Q' FrQ distinct AQ
  obtain K' where KeqK': "(ΨF  ΨP)  ΨQ  K  K'" and "AP ♯* K'" and "AF ♯* K'" and "AG ♯* K'"
    using AP ♯* Q AQ ♯* AF AQ ♯* AG AF ♯* Q AG ♯* Q AQ ♯* ΨF AQ ♯* ΨP AP ♯* AQ AQ ♯* Q AQ ♯* K
    by(elim inputObtainPrefix[where B="AP@AF@AG"]) force+
  have "ΨG  ΨQ  P K'⦇ν*xvec⦈⟨N  P'"
  proof -
    from KeqK' have "ΨF  (ΨP  ΨQ)  K  K'"
      by(rule statEqEnt[OF Associativity])
    with ΨF  (ΨP  ΨQ)  M  K have "ΨF  (ΨP  ΨQ)  M  K'"
      by(rule chanEqTrans)
    then have "(ΨF  ΨQ)  ΨP  M  K'"
      by(metis statEqEnt AssertionStatEqSym Associativity AssertionStatEqTrans compositionSym Commutativity)
    with ΨF  ΨQ  P M⦇ν*xvec⦈⟨N  P' FrP distinct AP
    have "ΨF  ΨQ  P K'⦇ν*xvec⦈⟨N  P'" using AP ♯* ΨF AP ♯* ΨQ AP ♯* P AP ♯* M AP ♯* K'
      by(force intro: outputRenameSubject)
    moreover have "AF, (ΨF  ΨQ)  ΨP F AG, (ΨG  ΨQ)  ΨP"
    proof -
      have "AF, (ΨF  ΨQ)  ΨP F AF, ΨF  ΨP  ΨQ"
        by(metis Associativity Composition AssertionStatEqSym AssertionStatEqTrans Commutativity frameResChainPres frameNilStatEq)
      moreover have "AG, ΨG  ΨP  ΨQ F AG, (ΨG  ΨQ)  ΨP "
        by(metis Associativity Composition AssertionStatEqSym AssertionStatEqTrans Commutativity frameResChainPres frameNilStatEq)
      ultimately show ?thesis using FimpG
        by(elim FrameStatEqImpCompose)
    qed
    ultimately show ?thesis using  AF ♯* P AG ♯* P AF ♯* K'
        AG ♯* K' AP ♯* AF AP ♯* AG AP ♯* ΨF AP ♯* ΨG AP ♯* ΨQ FrP distinct AP
        distinct xvec
      by(auto intro: transferNonTauFrame)
  qed

  moreover from ΨF  ΨQ  P M⦇ν*xvec⦈⟨N  P' have "ΨF  ΨQ  P  ROut M (⦇ν*xvecN ≺' P')"
    by(simp add: residualInject)
  moreover with FrP distinct AP
  obtain M' where MeqM': "(ΨF  ΨQ)  ΨP  M  M'" and "AQ ♯* M'" and "AF ♯* M'" and "AG ♯* M'"
    using AQ ♯* P AP ♯* AF AP ♯* AG AF ♯* P AG ♯* P AP ♯* ΨF AP ♯* ΨQ AP ♯* AQ AP ♯* P AP ♯* M xvec ♯* M distinct xvec
    by(elim outputObtainPrefix[where B="AQ@AF@AG"]) force+

  have "ΨG  ΨP  Q M'N  Q'"
  proof -
    from MeqM' have "ΨF  (ΨQ  ΨP)  M  M'" by(rule statEqEnt[OF Associativity])
    with ΨF  (ΨP  ΨQ)  M  K have "ΨF  (ΨQ  ΨP)  K  M'"
      by(blast intro: chanEqTrans chanEqSym compositionSym Commutativity statEqEnt)
    then have "(ΨF  ΨP)  ΨQ  K  M'"
      by(blast intro: statEqEnt AssertionStatEqSym Associativity
          AssertionStatEqTrans compositionSym Commutativity)
    with ΨF  ΨP  Q KN  Q' FrQ distinct AQ
    have "ΨF  ΨP  Q M'N  Q'" using AQ ♯* ΨF AQ ♯* ΨP AQ ♯* Q AQ ♯* K AQ ♯* M'
      by(force intro: inputRenameSubject)
    moreover have "AF, (ΨF  ΨP)  ΨQ F AG, (ΨG  ΨP)  ΨQ"
    proof -
      have "AF, (ΨF  ΨP)  ΨQ F AF, ΨF  ΨP  ΨQ"
        by(metis Associativity frameResChainPres frameNilStatEq)
      moreover have "AG, ΨG  ΨP  ΨQ F AG, (ΨG  ΨP)  ΨQ "
        by(metis Associativity AssertionStatEqSym frameResChainPres frameNilStatEq)
      ultimately show ?thesis using FimpG
        by(elim FrameStatEqImpCompose)
    qed

    ultimately show ?thesis using AF ♯* Q AG ♯* Q AF ♯* M' AG ♯* M'
        AQ ♯* AF AQ ♯* AG AQ ♯* ΨF AQ ♯* ΨG AQ ♯* ΨP FrQ distinct AQ distinct xvec
      by(auto intro: transferNonTauFrame)
  qed

  moreover have "ΨG  ΨP  ΨQ  K'  M'"
  proof -
    from MeqM' have "ΨF  ΨP  ΨQ  M'  M"
      by(blast intro: chanEqSym Associativity statEqEnt Commutativity compositionSym)
    moreover from KeqK' have "ΨF  ΨP  ΨQ  K  K'"
      by(blast intro: chanEqSym Associativity statEqEnt Commutativity compositionSym)
    ultimately have "ΨF  ΨP  ΨQ  K'  M'" using ΨF  ΨP  ΨQ  M  K
      by(blast intro: chanEqSym chanEqTrans)
    then show ?thesis using AF ♯* M' AF ♯* K' AG ♯* M' AG ♯* 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 AP ♯* ΨG AP ♯* P AP ♯* Q AP ♯* AQ AP ♯* K' AQ ♯* ΨG AQ ♯* P AQ ♯* Q AQ ♯* M' xvec ♯* Q FrP FrQ
    by(intro Comm2) (assumption | simp)+
next
  case(cBrClose ΨF P M xvec N P' AP ΨP x ΨG AF AG)
  from ΨF  P  ¡M⦇ν*xvec⦈⟨N  P'
  have suppM: "((supp M)::name set)  ((supp P)::name set)"
    by(simp add: residualInject brOutputTermSupp)

  note ΨF  P  ¡M⦇ν*xvec⦈⟨N  P' extractFrame P = AP, ΨP
    distinct AP distinct xvec AF, ΨF  ΨP F AG, ΨG  ΨP

  moreover from x  AF AF ♯* ⦇νxP x  AG AG ♯* ⦇νxP
  have "AF ♯* P" and "AG ♯* P" by simp+
  moreover with suppM
  have "AF ♯* M" and "AG ♯* M"
    by(auto simp add: fresh_star_def fresh_def)
  moreover note AP ♯* AF AP ♯* AG AP ♯* ΨF AP ♯* ΨG
  ultimately have "ΨG  P  ¡M⦇ν*xvec⦈⟨N  P'"
    by(simp add: transferNonTauFrame)
  with x  supp M x  ΨG
  show ?case
    by(simp add: BrClose)
next
  case(cScope ΨF P P' x AP ΨP ΨG AF AG)
  then have "ΨG  P τ  P'" by auto
  with x  ΨG show ?case
    by(intro Scope) auto
next
  case(cBang ΨF P P' AP ΨP ΨG AF AG)
  from ΨP  𝟭 have "AF, ΨF  ΨP  𝟭 F AF, ΨF  𝟭"
    by(metis frameIntCompositionSym Identity AssertionStatEqTrans)
  moreover note AF, ΨF  𝟭 F AG, ΨG  𝟭
  moreover from ΨP  𝟭 have "AG, ΨG  𝟭 F AG, ΨG  ΨP  𝟭"
    by(metis frameIntCompositionSym Identity AssertionStatEqTrans AssertionStatEqSym)
  ultimately have "AF, ΨF  ΨP  𝟭 F AG, ΨG  ΨP  𝟭"
    by(rule FrameStatEqImpCompose)
  with cBang have "ΨG  P  !P τ  P'" by force
  then show ?case using guarded P by(rule Bang)
qed

lemma transferFrame:
  fixes ΨF  :: 'b
    and P    :: "('a, 'b, 'c) psi"
    and α    :: "'a action"
    and P'   :: "('a, 'b, 'c) psi"
    and AF   :: "name list"
    and AG   :: "name list"
    and ΨG   :: 'b

assumes "ΨF  P α  P'"
  and   "extractFrame P = AP, ΨP"
  and   "distinct AP"
  and   "AF, ΨF  ΨP F AG, ΨG  ΨP"
  and   "AF ♯* P"
  and   "AG ♯* P"
  and   "AF ♯* subject α"
  and   "AG ♯* subject α"
  and   "AP ♯* AF"
  and   "AP ♯* AG"
  and   "AP ♯* ΨF"
  and   "AP ♯* ΨG"

shows "ΨG  P α  P'"
  using assms
proof -
  from Ψ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 MN  T"
  and   "extractFrame(P  Q) = APQ, ΨPQ"
  and   "distinct APQ"
  and   "APQ ♯* Ψ"
  and   "APQ ♯* P"
  and   "APQ ♯* Q"
  and   "APQ ♯* M"
  and   rPar1: "P' AP ΨP AQ ΨQ. Ψ  ΨQ  P MN  P'; extractFrame P = AP, ΨP; extractFrame Q = AQ, ΨQ;
                                      distinct AP; distinct AQ; AP ♯* Ψ; AP ♯* P; AP ♯* Q; AP ♯* M;  AQ ♯* Ψ; AQ ♯* P; AQ ♯* Q; AQ ♯* M;
                                      AP ♯* ΨQ; AQ ♯* ΨP; AP ♯* AQ; APQ = AP@AQ; ΨPQ = ΨP  ΨQ  Prop (P'  Q)"
  and   rPar2: "Q' AP ΨP AQ ΨQ. Ψ  ΨP  Q MN  Q'; extractFrame P = AP, ΨP; extractFrame Q = AQ, ΨQ;
                                      distinct AP; distinct AQ; AP ♯* Ψ; AP ♯* P; AP ♯* Q; AP ♯* M;  AQ ♯* Ψ; AQ ♯* P; AQ ♯* Q; AQ ♯* M;
                                      AP ♯* ΨQ; AQ ♯* ΨP; AP ♯* AQ; APQ = AP@AQ; ΨPQ = ΨP  ΨQ  Prop (P  Q')"
shows "Prop T"
  using Trans
proof(induct rule: parInputCases[of _ _ _ _ _ _ "(APQ, ΨPQ)"])
  case(cPar1 P' AQ ΨQ)
  from AQ ♯* (APQ, ΨPQ) have "AQ ♯* APQ" and "AQ ♯* ΨPQ" by simp+
  obtain AP ΨP where FrP: "extractFrame P = AP, ΨP" and "distinct AP"
    "AP ♯* (P, Q, Ψ, M, AQ, APQ, ΨQ)"
    by(rule freshFrame)
  then have "AP ♯* P" and "AP ♯* Q" and "AP ♯* Ψ" and "AP ♯* M" and "AP ♯* AQ" and "AP ♯* APQ" and "AP ♯* ΨQ"
    by simp+

  have FrQ: "extractFrame Q = AQ, ΨQ" by fact

  from AQ ♯* P AP ♯* AQ FrP have "AQ ♯* ΨP"
    by(force dest: extractFrameFreshChain)

  from extractFrame(P  Q) = APQ, ΨPQ FrP FrQ AP ♯* AQ AP ♯* ΨQ AQ ♯* ΨP
  have "(AP@AQ), ΨP  ΨQ = APQ, ΨPQ" by simp
  moreover from distinct AP distinct AQ AP ♯* AQ have "distinct(AP@AQ)"
    by(auto simp add: fresh_star_def fresh_def name_list_supp)
  ultimately obtain p where S: "set p  set(AP@AQ) × set((p  AP)@(p  AQ))"  and "distinctPerm p"
    and Ψeq: "ΨPQ = (p  ΨP)  (p  ΨQ)" and Aeq: "APQ = (p  AP)@(p  AQ)"
    using AP ♯* APQ AQ ♯* APQ distinct APQ
    by(elim frameChainEq') (assumption | simp add: eqvts)+

  from Ψ  ΨQ  P MN  P' S APQ ♯* P AP ♯* P AQ ♯* P APQ ♯* M AP ♯* M AQ ♯* M Aeq
  have "(p  (Ψ  ΨQ))  P MN  P'"
    by(elim inputPermFrame) auto
  with S APQ ♯* Ψ AP ♯* Ψ AQ ♯* Ψ Aeq have "Ψ  (p  ΨQ)  P MN  P'"
    by(simp add: eqvts)
  moreover from FrP have "(p  extractFrame P) = p  AP, ΨP" by simp
  with S APQ ♯* P AP ♯* P AQ ♯* P Aeq have "extractFrame P = (p  AP), p  ΨP"
    by(simp add: eqvts)
  moreover from FrQ have "(p  extractFrame Q) = p  AQ, ΨQ" by simp
  with S APQ ♯* Q AP ♯* Q AQ ♯* Q Aeq have "extractFrame Q = (p  AQ), p  ΨQ"
    by(simp add: eqvts)
  moreover from distinct AP distinct AQ have "distinct(p  AP)" and "distinct(p  AQ)"
    by simp+
  moreover from AP ♯* AQ have "(p  AP) ♯* (p  AQ)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
  moreover from AP ♯* ΨQ have "(p  AP) ♯* (p  ΨQ)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
  moreover from AQ ♯* ΨP have "(p  AQ) ♯* (p  ΨP)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
  ultimately show ?case using APQ ♯* Ψ APQ ♯* P APQ ♯* Q APQ ♯* M Aeq Ψeq
    by(intro rPar1) (assumption | simp)+
next
  case(cPar2 Q' AP ΨP)
  from AP ♯* (APQ, ΨPQ) have "AP ♯* APQ" and "AP ♯* ΨPQ" by simp+
  obtain AQ ΨQ where FrQ: "extractFrame Q = AQ, ΨQ" and "distinct AQ"
    "AQ ♯* (P, Q, Ψ, M, AP, APQ, ΨP)"
    by(rule freshFrame)
  then have "AQ ♯* P" and "AQ ♯* Q" and "AQ ♯* Ψ" and "AQ ♯* M" and "AQ ♯* AP" and "AQ ♯* APQ" and "AQ ♯* ΨP"
    by simp+

  have FrP: "extractFrame P = AP, ΨP" by fact

  from AP ♯* Q AQ ♯* AP FrQ have "AP ♯* ΨQ"
    by(force dest: extractFrameFreshChain)

  from extractFrame(P  Q) = APQ, ΨPQ FrP FrQ AQ ♯* AP AP ♯* ΨQ AQ ♯* ΨP
  have "(AP@AQ), ΨP  ΨQ = APQ, ΨPQ" by simp
  moreover from distinct AP distinct AQ AQ ♯* AP have "distinct(AP@AQ)"
    by(auto simp add: fresh_star_def fresh_def name_list_supp)
  ultimately obtain p where S: "set p  set(AP@AQ) × set((p  AP)@(p  AQ))"  and "distinctPerm p"
    and Ψeq: "ΨPQ = (p  ΨP)  (p  ΨQ)" and Aeq: "APQ = (p  AP)@(p  AQ)"
    using AP ♯* APQ AQ ♯* APQ distinct APQ
    by(elim frameChainEq') (assumption | simp add: eqvts)+

  from Ψ  ΨP  Q MN  Q' S APQ ♯* Q AP ♯* Q AQ ♯* Q APQ ♯* M AP ♯* M AQ ♯* M Aeq
  have "(p  (Ψ  ΨP))  Q MN  Q'"
    by(elim inputPermFrame) auto
  with S APQ ♯* Ψ AP ♯* Ψ AQ ♯* Ψ Aeq have "Ψ  (p  ΨP)  Q MN  Q'"
    by(simp add: eqvts)
  moreover from FrP have "(p  extractFrame P) = p  AP, ΨP" by simp
  with S APQ ♯* P AP ♯* P AQ ♯* P Aeq have "extractFrame P = (p  AP), p  ΨP"
    by(simp add: eqvts)
  moreover from FrQ have "(p  extractFrame Q) = p  AQ, ΨQ" by simp
  with S APQ ♯* Q AP ♯* Q AQ ♯* Q Aeq have "extractFrame Q = (p  AQ), p  ΨQ"
    by(simp add: eqvts)
  moreover from distinct AP distinct AQ have "distinct(p  AP)" and "distinct(p  AQ)"
    by simp+
  moreover from AQ ♯* AP have "(p  AP) ♯* (p  AQ)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
  moreover from AP ♯* ΨQ have "(p  AP) ♯* (p  ΨQ)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
  moreover from AQ ♯* ΨP have "(p  AQ) ♯* (p  ΨP)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
  ultimately show ?case using APQ ♯* Ψ APQ ♯* P APQ ♯* Q APQ ♯* 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 ¿MN  T"
  and   "extractFrame(P  Q) = APQ, ΨPQ"
  and   "distinct APQ"
  and   "APQ ♯* Ψ"
  and   "APQ ♯* P"
  and   "APQ ♯* Q"
  and   "APQ ♯* M"
  and   rPar1: "P' AP ΨP AQ ΨQ. Ψ  ΨQ  P ¿MN  P'; extractFrame P = AP, ΨP; extractFrame Q = AQ, ΨQ;
                                      distinct AP; distinct AQ; AP ♯* Ψ; AP ♯* P; AP ♯* Q; AP ♯* M;  AQ ♯* Ψ; AQ ♯* P; AQ ♯* Q; AQ ♯* M;
                                      AP ♯* ΨQ; AQ ♯* ΨP; AP ♯* AQ; APQ = AP@AQ; ΨPQ = ΨP  ΨQ  Prop (P'  Q)"
  and   rPar2: "Q' AP ΨP AQ ΨQ. Ψ  ΨP  Q ¿MN  Q'; extractFrame P = AP, ΨP; extractFrame Q = AQ, ΨQ;
                                      distinct AP; distinct AQ; AP ♯* Ψ; AP ♯* P; AP ♯* Q; AP ♯* M;  AQ ♯* Ψ; AQ ♯* P; AQ ♯* Q; AQ ♯* M;
                                      AP ♯* ΨQ; AQ ♯* ΨP; AP ♯* AQ; APQ = AP@AQ; ΨPQ = ΨP  ΨQ  Prop (P  Q')"
  and   rBrMerge: "ΨQ P' AP ΨP Q' AQ.
                    Ψ  ΨQ  P  ¿MN  P'; extractFrame P = AP, ΨP; distinct AP;
                    Ψ  ΨP  Q  ¿MN  Q'; extractFrame Q = AQ, ΨQ; distinct AQ;
                    AP ♯* Ψ; AP ♯* ΨQ; AP ♯* P;
                    AP ♯* Q; AP ♯* AQ;
                    AQ ♯* Ψ; AQ ♯* ΨP; AQ ♯* P;
                    AQ ♯* Q;
                    APQ = AP@AQ; ΨPQ = ΨP  ΨQ 
                    Prop (P'  Q')"
shows "Prop T"
  using Trans
proof(induct rule: parBrInputCases[of _ _ _ _ _ _ "(APQ, ΨPQ)"])
  case(cPar1 P' AQ ΨQ)
  from AQ ♯* (APQ, ΨPQ) have "AQ ♯* APQ" and "AQ ♯* ΨPQ" by simp+
  obtain AP ΨP where FrP: "extractFrame P = AP, ΨP" and "distinct AP"
    "AP ♯* (P, Q, Ψ, M, AQ, APQ, ΨQ)"
    by(rule freshFrame)
  then have "AP ♯* P" and "AP ♯* Q" and "AP ♯* Ψ" and "AP ♯* M" and "AP ♯* AQ" and "AP ♯* APQ" and "AP ♯* ΨQ"
    by simp+

  have FrQ: "extractFrame Q = AQ, ΨQ" by fact

  from AQ ♯* P AP ♯* AQ FrP have "AQ ♯* ΨP"
    by(force dest: extractFrameFreshChain)

  from extractFrame(P  Q) = APQ, ΨPQ FrP FrQ AP ♯* AQ AP ♯* ΨQ AQ ♯* ΨP
  have "(AP@AQ), ΨP  ΨQ = APQ, ΨPQ" by simp
  moreover from distinct AP distinct AQ AP ♯* AQ have "distinct(AP@AQ)"
    by(auto simp add: fresh_star_def fresh_def name_list_supp)
  ultimately obtain p where S: "set p  set(AP@AQ) × set((p  AP)@(p  AQ))"  and "distinctPerm p"
    and Ψeq: "ΨPQ = (p  ΨP)  (p  ΨQ)" and Aeq: "APQ = (p  AP)@(p  AQ)"
    using AP ♯* APQ AQ ♯* APQ distinct APQ
    by(elim frameChainEq') (assumption | simp add: eqvts)+

  from Ψ  ΨQ  P ¿MN  P' S APQ ♯* P AP ♯* P AQ ♯* P APQ ♯* M AP ♯* M AQ ♯* M Aeq
  have "(p  (Ψ  ΨQ))  P ¿MN  P'"
    by(elim brinputPermFrame) auto
  with S APQ ♯* Ψ AP ♯* Ψ AQ ♯* Ψ Aeq have "Ψ  (p  ΨQ)  P ¿MN  P'"
    by(simp add: eqvts)
  moreover from FrP have "(p  extractFrame P) = p  AP, ΨP" by simp
  with S APQ ♯* P AP ♯* P AQ ♯* P Aeq have "extractFrame P = (p  AP), p  ΨP"
    by(simp add: eqvts)
  moreover from FrQ have "(p  extractFrame Q) = p  AQ, ΨQ" by simp
  with S APQ ♯* Q AP ♯* Q AQ ♯* Q Aeq have "extractFrame Q = (p  AQ), p  ΨQ"
    by(simp add: eqvts)
  moreover from distinct AP distinct AQ have "distinct(p  AP)" and "distinct(p  AQ)"
    by simp+
  moreover from AP ♯* AQ have "(p  AP) ♯* (p  AQ)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
  moreover from AP ♯* ΨQ have "(p  AP) ♯* (p  ΨQ)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
  moreover from AQ ♯* ΨP have "(p  AQ) ♯* (p  ΨP)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
  ultimately show ?case using APQ ♯* Ψ APQ ♯* P APQ ♯* Q APQ ♯* M Aeq Ψeq
    by(intro rPar1) (assumption | simp)+
next
  case(cPar2 Q' AP ΨP)
  from AP ♯* (APQ, ΨPQ) have "AP ♯* APQ" and "AP ♯* ΨPQ" by simp+
  obtain AQ ΨQ where FrQ: "extractFrame Q = AQ, ΨQ" and "distinct AQ"
    "AQ ♯* (P, Q, Ψ, M, AP, APQ, ΨP)"
    by(rule freshFrame)
  then have "AQ ♯* P" and "AQ ♯* Q" and "AQ ♯* Ψ" and "AQ ♯* M" and "AQ ♯* AP" and "AQ ♯* APQ" and "AQ ♯* ΨP"
    by simp+

  have FrP: "extractFrame P = AP, ΨP" by fact

  from AP ♯* Q AQ ♯* AP FrQ have "AP ♯* ΨQ"
    by(force dest: extractFrameFreshChain)

  from extractFrame(P  Q) = APQ, ΨPQ FrP FrQ AQ ♯* AP AP ♯* ΨQ AQ ♯* ΨP
  have "(AP@AQ), ΨP  ΨQ = APQ, ΨPQ" by simp
  moreover from distinct AP distinct AQ AQ ♯* AP have "distinct(AP@AQ)"
    by(auto simp add: fresh_star_def fresh_def name_list_supp)
  ultimately obtain p where S: "set p  set(AP@AQ) × set((p  AP)@(p  AQ))"  and "distinctPerm p"
    and Ψeq: "ΨPQ = (p  ΨP)  (p  ΨQ)" and Aeq: "APQ = (p  AP)@(p  AQ)"
    using AP ♯* APQ AQ ♯* APQ distinct APQ
    by(elim frameChainEq') (assumption | simp add: eqvts)+

  from Ψ  ΨP  Q ¿MN  Q' S APQ ♯* Q AP ♯* Q AQ ♯* Q APQ ♯* M AP ♯* M AQ ♯* M Aeq
  have "(p  (Ψ  ΨP))  Q ¿MN  Q'"
    by(elim brinputPermFrame) auto
  with S APQ ♯* Ψ AP ♯* Ψ AQ ♯* Ψ Aeq have "Ψ  (p  ΨP)  Q ¿MN  Q'"
    by(simp add: eqvts)
  moreover from FrP have "(p  extractFrame P) = p  AP, ΨP" by simp
  with S APQ ♯* P AP ♯* P AQ ♯* P Aeq have "extractFrame P = (p  AP), p  ΨP"
    by(simp add: eqvts)
  moreover from FrQ have "(p  extractFrame Q) = p  AQ, ΨQ" by simp
  with S APQ ♯* Q AP ♯* Q AQ ♯* Q Aeq have "extractFrame Q = (p  AQ), p  ΨQ"
    by(simp add: eqvts)
  moreover from distinct AP distinct AQ have "distinct(p  AP)" and "distinct(p  AQ)"
    by simp+
  moreover from AQ ♯* AP have "(p  AP) ♯* (p  AQ)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
  moreover from AP ♯* ΨQ have "(p  AP) ♯* (p  ΨQ)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
  moreover from AQ ♯* ΨP have "(p  AQ) ♯* (p  ΨP)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
  ultimately show ?case using APQ ♯* Ψ APQ ♯* P APQ ♯* Q APQ ♯* M Aeq Ψeq
    by(intro rPar2) (assumption | simp)+
next
  case(cBrMerge ΨQ P' AP ΨP Q' AQ)
  then have FrP: "extractFrame P = AP, ΨP" and FrQ: "extractFrame Q = AQ, ΨQ"
    and "AP ♯* APQ" and "AQ ♯* APQ"
    by simp+

  from extractFrame(P  Q) = APQ, ΨPQ FrP FrQ AP ♯* AQ AP ♯* ΨQ AQ ♯* ΨP
  have "(AP@AQ), ΨP  ΨQ = APQ, ΨPQ" by simp
  moreover from distinct AP distinct AQ AP ♯* AQ have "distinct(AP@AQ)"
    by(auto simp add: fresh_star_def fresh_def name_list_supp)
  ultimately obtain p where S: "set p  set(AP@AQ) × set((p  AP)@(p  AQ))"  and "distinctPerm p"
    and Ψeq: "ΨPQ = (p  ΨP)  (p  ΨQ)" and Aeq: "APQ = (p  AP)@(p  AQ)"
    using AP ♯* APQ AQ ♯* APQ distinct APQ
    by(elim frameChainEq') (assumption | simp add: eqvts)+

  from Ψ  ΨQ  P ¿MN  P' S APQ ♯* P AP ♯* P AQ ♯* P APQ ♯* M AP ♯* M AQ ♯* M Aeq
  have "(p  (Ψ  ΨQ))  P ¿MN  P'"
    by(elim brinputPermFrame) auto
  with S APQ ♯* Ψ AP ♯* Ψ AQ ♯* Ψ Aeq have "Ψ  (p  ΨQ)  P ¿MN  P'"
    by(simp add: eqvts)

  from Ψ  ΨP  Q ¿MN  Q' S APQ ♯* Q AQ ♯* Q AP ♯* Q APQ ♯* M AQ ♯* M AP ♯* M Aeq
  have "(p  (Ψ  ΨP))  Q ¿MN  Q'"
    by(elim brinputPermFrame) auto
  with S APQ ♯* Ψ AP ♯* Ψ AQ ♯* Ψ Aeq have "Ψ  (p  ΨP)  Q ¿MN  Q'"
    by(simp add: eqvts)

  note Ψ  (p  ΨQ)  P ¿MN  P' Ψ  (p  ΨP)  Q ¿MN  Q'
  moreover from FrP have "(p  extractFrame P) = p  AP, ΨP" by simp
  with S APQ ♯* P AP ♯* P AQ ♯* P Aeq have "extractFrame P = (p  AP), p  ΨP"
    by(simp add: eqvts)
  moreover from FrQ have "(p  extractFrame Q) = p  AQ, ΨQ" by simp
  with S APQ ♯* Q AP ♯* Q AQ ♯* Q Aeq have "extractFrame Q = (p  AQ), p  ΨQ"
    by(simp add: eqvts)
  moreover from distinct AP distinct AQ have "distinct(p  AP)" and "distinct(p  AQ)"
    by simp+
  moreover from AP ♯* AQ have "(p  AP) ♯* (p  AQ)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
  moreover from AP ♯* ΨQ have "(p  AP) ♯* (p  ΨQ)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
  moreover from AQ ♯* ΨP have "(p  AQ) ♯* (p  ΨP)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
  ultimately show ?case using APQ ♯* Ψ APQ ♯* P APQ ♯* Q APQ ♯* 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) = APQ, ΨPQ"
  and   "distinct APQ"
  and   "APQ ♯* Ψ"
  and   "APQ ♯* P"
  and   "APQ ♯* Q"
  and   "APQ ♯* M"
  and   rPar1: "P' AP ΨP AQ ΨQ. Ψ  ΨQ  P M⦇ν*xvec⦈⟨N  P'; extractFrame P = AP, ΨP; extractFrame Q = AQ, ΨQ;
                                      distinct AP; distinct AQ; AP ♯* Ψ; AP ♯* P; AP ♯* Q; AP ♯* M;  AQ ♯* Ψ; AQ ♯* P; AQ ♯* Q; AQ ♯* M;
                                      AP ♯* ΨQ; AQ ♯* ΨP; AP ♯* AQ; APQ = AP@AQ; ΨPQ = ΨP  ΨQ  Prop (P'  Q)"
  and   rPar2: "Q' AP ΨP AQ ΨQ. Ψ  ΨP  Q M⦇ν*xvec⦈⟨N  Q'; extractFrame P = AP, ΨP; extractFrame Q = AQ, ΨQ;
                                      distinct AP; distinct AQ; AP ♯* Ψ; AP ♯* P; AP ♯* Q; AP ♯* M;  AQ ♯* Ψ; AQ ♯* P; AQ ♯* Q; AQ ♯* M;
                                      AP ♯* ΨQ; AQ ♯* ΨP; AP ♯* AQ; APQ = AP@AQ; ΨPQ = ΨP  ΨQ  Prop (P  Q')"
shows "Prop T"
  using Trans xvec ♯* Ψ xvec ♯* P xvec ♯* Q xvec ♯* M
proof(induct rule: parOutputCases[of _ _ _ _ _ _ _ "(APQ, ΨPQ)"])
  case(cPar1 P' AQ ΨQ)
  from AQ ♯* (APQ, ΨPQ) have "AQ ♯* APQ" and "AQ ♯* ΨPQ" by simp+
  obtain AP ΨP where FrP: "extractFrame P = AP, ΨP" and "distinct AP"
    "AP ♯* (P, Q, Ψ, M, AQ, APQ, ΨQ)"
    by(rule freshFrame)
  then have "AP ♯* P" and "AP ♯* Q" and "AP ♯* Ψ" and "AP ♯* M" and "AP ♯* AQ" and "AP ♯* APQ" and "AP ♯* ΨQ"
    by simp+

  have FrQ: "extractFrame Q = AQ, ΨQ" by fact

  from AQ ♯* P AP ♯* AQ FrP have "AQ ♯* ΨP"
    by(force dest: extractFrameFreshChain)

  from extractFrame(P  Q) = APQ, ΨPQ FrP FrQ AP ♯* AQ AP ♯* ΨQ AQ ♯* ΨP
  have "(AP@AQ), ΨP  ΨQ = APQ, ΨPQ" by simp
  moreover from distinct AP distinct AQ AP ♯* AQ have "distinct(AP@AQ)"
    by(auto simp add: fresh_star_def fresh_def name_list_supp)
  ultimately obtain p where S: "set p  set(AP@AQ) × set((p  AP)@(p  AQ))"  and "distinctPerm p"
    and Ψeq: "ΨPQ = (p  ΨP)  (p  ΨQ)" and Aeq: "APQ = (p  AP)@(p  AQ)"
    using AP ♯* APQ AQ ♯* APQ distinct APQ
    by(elim frameChainEq') (assumption | simp add: eqvts)+

  from Ψ  ΨQ  P M⦇ν*xvec⦈⟨N  P' S APQ ♯* P AP ♯* P AQ ♯* P APQ ♯* M AP ♯* M AQ ♯* M Aeq
  have "(p  (Ψ  ΨQ))  P M⦇ν*xvec⦈⟨N  P'"
    by(elim outputPermFrame) (assumption | simp)+

  with S APQ ♯* Ψ AP ♯* Ψ AQ ♯* Ψ Aeq have "Ψ  (p  ΨQ)  P M⦇ν*xvec⦈⟨N  P'"
    by(simp add: eqvts)
  moreover from FrP have "(p  extractFrame P) = p  AP, ΨP" by simp
  with S APQ ♯* P AP ♯* P AQ ♯* P Aeq have "extractFrame P = (p  AP), p  ΨP"
    by(simp add: eqvts)
  moreover from FrQ have "(p  extractFrame Q) = p  AQ, ΨQ" by simp
  with S APQ ♯* Q AP ♯* Q AQ ♯* Q Aeq have "extractFrame Q = (p  AQ), p  ΨQ"
    by(simp add: eqvts)
  moreover from distinct AP distinct AQ have "distinct(p  AP)" and "distinct(p  AQ)"
    by simp+
  moreover from AP ♯* AQ have "(p  AP) ♯* (p  AQ)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
  moreover from AP ♯* ΨQ have "(p  AP) ♯* (p  ΨQ)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
  moreover from AQ ♯* ΨP have "(p  AQ) ♯* (p  ΨP)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
  ultimately show ?case using APQ ♯* Ψ APQ ♯* P APQ ♯* Q APQ ♯* M Aeq Ψeq
    by(intro rPar1) (assumption | simp)+
next
  case(cPar2 Q' AP ΨP)
  from AP ♯* (APQ, ΨPQ) have "AP ♯* APQ" and "AP ♯* ΨPQ" by simp+
  obtain AQ ΨQ where FrQ: "extractFrame Q = AQ, ΨQ" and "distinct AQ"
    "AQ ♯* (P, Q, Ψ, M, AP, APQ, ΨP)"
    by(rule freshFrame)
  then have "AQ ♯* P" and "AQ ♯* Q" and "AQ ♯* Ψ" and "AQ ♯* M" and "AQ ♯* AP" and "AQ ♯* APQ" and "AQ ♯* ΨP"
    by simp+

  have FrP: "extractFrame P = AP, ΨP" by fact

  from AP ♯* Q AQ ♯* AP FrQ have "AP ♯* ΨQ"
    by(force dest: extractFrameFreshChain)

  from extractFrame(P  Q) = APQ, ΨPQ FrP FrQ AQ ♯* AP AP ♯* ΨQ AQ ♯* ΨP
  have "(AP@AQ), ΨP  ΨQ = APQ, ΨPQ" by simp
  moreover from distinct AP distinct AQ AQ ♯* AP have "distinct(AP@AQ)"
    by(auto simp add: fresh_star_def fresh_def name_list_supp)
  ultimately obtain p where S: "set p  set(AP@AQ) × set((p  AP)@(p  AQ))"  and "distinctPerm p"
    and Ψeq: "ΨPQ = (p  ΨP)  (p  ΨQ)" and Aeq: "APQ = (p  AP)@(p  AQ)"
    using AP ♯* APQ AQ ♯* APQ distinct APQ
    by(elim frameChainEq') (assumption | simp add: eqvts)+

  from Ψ  ΨP  Q M⦇ν*xvec⦈⟨N  Q' S APQ ♯* Q AP ♯* Q AQ ♯* Q APQ ♯* M AP ♯* M AQ ♯* M Aeq
  have "(p  (Ψ  ΨP))  Q M⦇ν*xvec⦈⟨N  Q'"
    by(elim outputPermFrame) (assumption | simp)+
  with S APQ ♯* Ψ AP ♯* Ψ AQ ♯* Ψ Aeq have "Ψ  (p  ΨP)  Q M⦇ν*xvec⦈⟨N  Q'"
    by(simp add: eqvts)
  moreover from FrP have "(p  extractFrame P) = p  AP, ΨP" by simp
  with S APQ ♯* P AP ♯* P AQ ♯* P Aeq have "extractFrame P = (p  AP), p  ΨP"
    by(simp add: eqvts)
  moreover from FrQ have "(p  extractFrame Q) = p  AQ, ΨQ" by simp
  with S APQ ♯* Q AP ♯* Q AQ ♯* Q Aeq have "extractFrame Q = (p  AQ), p  ΨQ"
    by(simp add: eqvts)
  moreover from distinct AP distinct AQ have "distinct(p  AP)" and "distinct(p  AQ)"
    by simp+
  moreover from AQ ♯* AP have "(p  AP) ♯* (p  AQ)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
  moreover from AP ♯* ΨQ have "(p  AP) ♯* (p  ΨQ)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
  moreover from AQ ♯* ΨP have "(p  AQ) ♯* (p  ΨP)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
  ultimately show ?case using APQ ♯* Ψ APQ ♯* P APQ ♯* Q APQ ♯* 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) = APQ, ΨPQ"
  and   "distinct APQ"
  and   "APQ ♯* Ψ"
  and   "APQ ♯* P"
  and   "APQ ♯* Q"
  and   "APQ ♯* M"
  and   rPar1: "P' AP ΨP AQ ΨQ. Ψ  ΨQ  P ¡M⦇ν*xvec⦈⟨N  P'; extractFrame P = AP, ΨP; extractFrame Q = AQ, ΨQ;
                                      distinct AP; distinct AQ; AP ♯* Ψ; AP ♯* P; AP ♯* Q; AP ♯* M;  AQ ♯* Ψ; AQ ♯* P; AQ ♯* Q; AQ ♯* M;
                                      AP ♯* ΨQ; AQ ♯* ΨP; AP ♯* AQ; APQ = AP@AQ; ΨPQ = ΨP  ΨQ  Prop (P'  Q)"
  and   rPar2: "Q' AP ΨP AQ ΨQ. Ψ  ΨP  Q ¡M⦇ν*xvec⦈⟨N  Q'; extractFrame P = AP, ΨP; extractFrame Q = AQ, ΨQ;
                                      distinct AP; distinct AQ; AP ♯* Ψ; AP ♯* P; AP ♯* Q; AP ♯* M;  AQ ♯* Ψ; AQ ♯* P; AQ ♯* Q; AQ ♯* M;
                                      AP ♯* ΨQ; AQ ♯* ΨP; AP ♯* AQ; APQ = AP@AQ; ΨPQ = ΨP  ΨQ  Prop (P  Q')"
  and   rBrComm1: "ΨQ P' AP ΨP Q' AQ.
           Ψ  ΨQ  P ¿MN  P'; extractFrame P = AP, ΨP; distinct AP;
            Ψ  ΨP  Q ¡M⦇ν*xvec⦈⟨N  Q'; extractFrame Q = AQ, ΨQ; distinct AQ;
            distinct xvec;
            AP ♯* Ψ; AP ♯* ΨQ; AP ♯* P; AP ♯* Q; AP ♯* AQ;
            AQ ♯* Ψ; AQ ♯* ΨP; AQ ♯* P; AQ ♯* Q;
            AP ♯* M; AQ ♯* M; xvec ♯* M;
            xvec ♯* Ψ; xvec ♯* P; xvec ♯* Q;
            APQ = AP@AQ; ΨPQ = ΨP  ΨQ 
            Prop (P'  Q')"
  and   rBrComm2: "ΨQ P' AP ΨP Q' AQ.
           Ψ  ΨQ  P ¡M⦇ν*xvec⦈⟨N  P'; extractFrame P = AP, ΨP; distinct AP;
            Ψ  ΨP  Q ¿MN  Q'; extractFrame Q = AQ, ΨQ; distinct AQ;
            distinct xvec;
            AP ♯* Ψ; AP ♯* ΨQ; AP ♯* P; AP ♯* Q; AP ♯* AQ;
            AQ ♯* Ψ; AQ ♯* ΨP; AQ ♯* P; AQ ♯* Q;
            AP ♯* M; AQ ♯* M; xvec ♯* M;
            xvec ♯* Ψ; xvec ♯* P; xvec ♯* Q;
            APQ = AP@AQ; ΨPQ = ΨP  ΨQ 
            Prop (P'  Q')"
shows "Prop T"
  using Trans xvec ♯* Ψ xvec ♯* P xvec ♯* Q xvec ♯* M
proof(induct rule: parBrOutputCases[of _ _ _ _ _ _ _ "(APQ, ΨPQ)"])
  case(cPar1 P' AQ ΨQ)
  from AQ ♯* (APQ, ΨPQ) have "AQ ♯* APQ" and "AQ ♯* ΨPQ" by simp+
  obtain AP ΨP where FrP: "extractFrame P = AP, ΨP" and "distinct AP"
    "AP ♯* (P, Q, Ψ, M, AQ, APQ, ΨQ)"
    by(rule freshFrame)
  then have "AP ♯* P" and "AP ♯* Q" and "AP ♯* Ψ" and "AP ♯* M" and "AP ♯* AQ" and "AP ♯* APQ" and "AP ♯* ΨQ"
    by simp+

  have FrQ: "extractFrame Q = AQ, ΨQ" by fact

  from AQ ♯* P AP ♯* AQ FrP have "AQ ♯* ΨP"
    by(force dest: extractFrameFreshChain)

  from extractFrame(P  Q) = APQ, ΨPQ FrP FrQ AP ♯* AQ AP ♯* ΨQ AQ ♯* ΨP
  have "(AP@AQ), ΨP  ΨQ = APQ, ΨPQ" by simp
  moreover from distinct AP distinct AQ AP ♯* AQ have "distinct(AP@AQ)"
    by(auto simp add: fresh_star_def fresh_def name_list_supp)
  ultimately obtain p where S: "set p  set(AP@AQ) × set((p  AP)@(p  AQ))"  and "distinctPerm p"
    and Ψeq: "ΨPQ = (p  ΨP)  (p  ΨQ)" and Aeq: "APQ = (p  AP)@(p  AQ)"
    using AP ♯* APQ AQ ♯* APQ distinct APQ
    by(elim frameChainEq') (assumption | simp add: eqvts)+

  from Ψ  ΨQ  P ¡M⦇ν*xvec⦈⟨N  P' S APQ ♯* P AP ♯* P AQ ♯* P APQ ♯* M AP ♯* M AQ ♯* M Aeq
  have "(p  (Ψ  ΨQ))  P ¡M⦇ν*xvec⦈⟨N  P'"
    by(elim broutputPermFrame) (assumption | simp)+

  with S APQ ♯* Ψ AP ♯* Ψ AQ ♯* Ψ Aeq have "Ψ  (p  ΨQ)  P ¡M⦇ν*xvec⦈⟨N  P'"
    by(simp add: eqvts)
  moreover from FrP have "(p  extractFrame P) = p  AP, ΨP" by simp
  with S APQ ♯* P AP ♯* P AQ ♯* P Aeq have "extractFrame P = (p  AP), p  ΨP"
    by(simp add: eqvts)
  moreover from FrQ have "(p  extractFrame Q) = p  AQ, ΨQ" by simp
  with S APQ ♯* Q AP ♯* Q AQ ♯* Q Aeq have "extractFrame Q = (p  AQ), p  ΨQ"
    by(simp add: eqvts)
  moreover from distinct AP distinct AQ have "distinct(p  AP)" and "distinct(p  AQ)"
    by simp+
  moreover from AP ♯* AQ have "(p  AP) ♯* (p  AQ)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
  moreover from AP ♯* ΨQ have "(p  AP) ♯* (p  ΨQ)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
  moreover from AQ ♯* ΨP have "(p  AQ) ♯* (p  ΨP)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
  ultimately show ?case using APQ ♯* Ψ APQ ♯* P APQ ♯* Q APQ ♯* M Aeq Ψeq
    by(intro rPar1) (assumption | simp)+
next
  case(cPar2 Q' AP ΨP)
  from AP ♯* (APQ, ΨPQ) have "AP ♯* APQ" and "AP ♯* ΨPQ" by simp+
  obtain AQ ΨQ where FrQ: "extractFrame Q = AQ, ΨQ" and "distinct AQ"
    "AQ ♯* (P, Q, Ψ, M, AP, APQ, ΨP)"
    by(rule freshFrame)
  then have "AQ ♯* P" and "AQ ♯* Q" and "AQ ♯* Ψ" and "AQ ♯* M" and "AQ ♯* AP" and "AQ ♯* APQ" and "AQ ♯* ΨP"
    by simp+

  have FrP: "extractFrame P = AP, ΨP" by fact

  from AP ♯* Q AQ ♯* AP FrQ have "AP ♯* ΨQ"
    by(force dest: extractFrameFreshChain)

  from extractFrame(P  Q) = APQ, ΨPQ FrP FrQ AQ ♯* AP AP ♯* ΨQ AQ ♯* ΨP
  have "(AP@AQ), ΨP  ΨQ = APQ, ΨPQ" by simp
  moreover from distinct AP distinct AQ AQ ♯* AP have "distinct(AP@AQ)"
    by(auto simp add: fresh_star_def fresh_def name_list_supp)
  ultimately obtain p where S: "set p  set(AP@AQ) × set((p  AP)@(p  AQ))"  and "distinctPerm p"
    and Ψeq: "ΨPQ = (p  ΨP)  (p  ΨQ)" and Aeq: "APQ = (p  AP)@(p  AQ)"
    using AP ♯* APQ AQ ♯* APQ distinct APQ
    by(elim frameChainEq') (assumption | simp add: eqvts)+

  from Ψ  ΨP  Q ¡M⦇ν*xvec⦈⟨N  Q' S APQ ♯* Q AP ♯* Q AQ ♯* Q APQ ♯* M AP ♯* M AQ ♯* M Aeq
  have "(p  (Ψ  ΨP))  Q ¡M⦇ν*xvec⦈⟨N  Q'"
    by(elim broutputPermFrame) (assumption | simp)+
  with S APQ ♯* Ψ AP ♯* Ψ AQ ♯* Ψ Aeq have "Ψ  (p  ΨP)  Q ¡M⦇ν*xvec⦈⟨N  Q'"
    by(simp add: eqvts)
  moreover from FrP have "(p  extractFrame P) = p  AP, ΨP" by simp
  with S APQ ♯* P AP ♯* P AQ ♯* P Aeq have "extractFrame P = (p  AP), p  ΨP"
    by(simp add: eqvts)
  moreover from FrQ have "(p  extractFrame Q) = p  AQ, ΨQ" by simp
  with S APQ ♯* Q AP ♯* Q AQ ♯* Q Aeq have "extractFrame Q = (p  AQ), p  ΨQ"
    by(simp add: eqvts)
  moreover from distinct AP distinct AQ have "distinct(p  AP)" and "distinct(p  AQ)"
    by simp+
  moreover from AQ ♯* AP have "(p  AP) ♯* (p  AQ)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
  moreover from AP ♯* ΨQ have "(p  AP) ♯* (p  ΨQ)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
  moreover from AQ ♯* ΨP have "(p  AQ) ♯* (p  ΨP)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
  ultimately show ?case using APQ ♯* Ψ APQ ♯* P APQ ♯* Q APQ ♯* M Aeq Ψeq
    by(intro rPar2) (assumption | simp)+
next
  case(cBrComm1 ΨQ P' AP ΨP Q' AQ)
  then have FrP: "extractFrame P = AP, ΨP" and FrQ: "extractFrame Q = AQ, ΨQ"
    and "AP ♯* APQ" and "AQ ♯* APQ"
    by simp+

  from extractFrame(P  Q) = APQ, ΨPQ FrP FrQ AP ♯* AQ AP ♯* ΨQ AQ ♯* ΨP
  have "(AP@AQ), ΨP  ΨQ = APQ, ΨPQ" by simp
  moreover from distinct AP distinct AQ AP ♯* AQ have "distinct(AP@AQ)"
    by(auto simp add: fresh_star_def fresh_def name_list_supp)
  ultimately obtain p where S: "set p  set(AP@AQ) × set((p  AP)@(p  AQ))"  and "distinctPerm p"
    and Ψeq: "ΨPQ = (p  ΨP)  (p  ΨQ)" and Aeq: "APQ = (p  AP)@(p  AQ)"
    using AP ♯* APQ AQ ♯* APQ distinct APQ
    by(elim frameChainEq') (assumption | simp add: eqvts)+

  from Ψ  ΨQ  P ¿MN  P' S APQ ♯* P AP ♯* P AQ ♯* P APQ ♯* M AP ♯* M AQ ♯* M Aeq
  have "(p  (Ψ  ΨQ))  P ¿MN  P'"
    by(elim brinputPermFrame) (assumption | simp)+

  from Ψ  ΨP  Q ¡M⦇ν*xvec⦈⟨N  Q' S APQ ♯* Q AQ ♯* Q AP ♯* Q APQ ♯* M AP ♯* M AQ ♯* M Aeq
  have "(p  (Ψ  ΨP))  Q ¡M⦇ν*xvec⦈⟨N  Q'"
    by(elim broutputPermFrame) (assumption | simp)+

  from (p  (Ψ  ΨQ))  P ¿MN  P' S APQ ♯* Ψ AP ♯* Ψ AQ ♯* Ψ Aeq have "Ψ  (p  ΨQ)  P ¿MN  P'"
    by(simp add: eqvts)
  moreover from (p  (Ψ  ΨP))  Q ¡M⦇ν*xvec⦈⟨N  Q' S APQ ♯* Ψ AP ♯* Ψ AQ ♯* Ψ Aeq have "Ψ  (p  ΨP)  Q ¡M⦇ν*xvec⦈⟨N  Q'"
    by(simp add: eqvts)

  moreover from FrP have "(p  extractFrame P) = p  AP, ΨP" by simp
  with S APQ ♯* P AP ♯* P AQ ♯* P Aeq have "extractFrame P = (p  AP), p  ΨP"
    by(simp add: eqvts)
  moreover from FrQ have "(p  extractFrame Q) = p  AQ, ΨQ" by simp
  with S APQ ♯* Q AP ♯* Q AQ ♯* Q Aeq have "extractFrame Q = (p  AQ), p  ΨQ"
    by(simp add: eqvts)
  moreover from distinct AP distinct AQ have "distinct(p  AP)" and "distinct(p  AQ)"
    by simp+
  moreover from AP ♯* AQ have "(p  AP) ♯* (p  AQ)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
  moreover from AP ♯* ΨQ have "(p  AP) ♯* (p  ΨQ)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
  moreover from AQ ♯* ΨP have "(p  AQ) ♯* (p  Ψ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 APQ ♯* Ψ APQ ♯* P APQ ♯* Q APQ ♯* M Aeq Ψeq
    by(intro rBrComm1) (assumption | simp)+
next
  case(cBrComm2 ΨQ P' AP ΨP Q' AQ)
  then have FrP: "extractFrame P = AP, ΨP" and FrQ: "extractFrame Q = AQ, ΨQ"
    and "AP ♯* APQ" and "AQ ♯* APQ"
    by simp+

  from extractFrame(P  Q) = APQ, ΨPQ FrP FrQ AP ♯* AQ AP ♯* ΨQ AQ ♯* ΨP
  have "(AP@AQ), ΨP  ΨQ = APQ, ΨPQ" by simp
  moreover from distinct AP distinct AQ AP ♯* AQ have "distinct(AP@AQ)"
    by(auto simp add: fresh_star_def fresh_def name_list_supp)
  ultimately obtain p where S: "set p  set(AP@AQ) × set((p  AP)@(p  AQ))"  and "distinctPerm p"
    and Ψeq: "ΨPQ = (p  ΨP)  (p  ΨQ)" and Aeq: "APQ = (p  AP)@(p  AQ)"
    using AP ♯* APQ AQ ♯* APQ distinct APQ
    by(elim frameChainEq') (assumption | simp add: eqvts)+

  from Ψ  ΨQ  P ¡M⦇ν*xvec⦈⟨N  P' S APQ ♯* P AP ♯* P AQ ♯* P APQ ♯* M AP ♯* M AQ ♯* M Aeq
  have "(p  (Ψ  ΨQ))  P ¡M⦇ν*xvec⦈⟨N  P'"
    by(elim broutputPermFrame) (assumption | simp)+

  from Ψ  ΨP  Q ¿MN  Q' S APQ ♯* Q AQ ♯* Q AP ♯* Q APQ ♯* M AP ♯* M AQ ♯* M Aeq
  have "(p  (Ψ  ΨP))  Q ¿MN  Q'"
    by(elim brinputPermFrame) (assumption | simp)+

  from (p  (Ψ  ΨQ))  P ¡M⦇ν*xvec⦈⟨N  P' S APQ ♯* Ψ AP ♯* Ψ AQ ♯* Ψ Aeq have "Ψ  (p  ΨQ)  P ¡M⦇ν*xvec⦈⟨N   P'"
    by(simp add: eqvts)
  moreover from (p  (Ψ  ΨP))  Q ¿MN  Q' S APQ ♯* Ψ AP ♯* Ψ AQ ♯* Ψ Aeq have "Ψ  (p  ΨP)  Q ¿MN  Q'"
    by(simp add: eqvts)

  moreover from FrP have "(p  extractFrame P) = p  AP, ΨP" by simp
  with S APQ ♯* P AP ♯* P AQ ♯* P Aeq have "extractFrame P = (p  AP), p  ΨP"
    by(simp add: eqvts)
  moreover from FrQ have "(p  extractFrame Q) = p  AQ, ΨQ" by simp
  with S APQ ♯* Q AP ♯* Q AQ ♯* Q Aeq have "extractFrame Q = (p  AQ), p  ΨQ"
    by(simp add: eqvts)
  moreover from distinct AP distinct AQ have "distinct(p  AP)" and "distinct(p  AQ)"
    by simp+
  moreover from AP ♯* AQ have "(p  AP) ♯* (p  AQ)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
  moreover from AP ♯* ΨQ have "(p  AP) ♯* (p  ΨQ)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
  moreover from AQ ♯* ΨP have "(p  AQ) ♯* (p  Ψ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 APQ ♯* Ψ APQ ♯* P APQ ♯* Q APQ ♯* 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 MN  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 KN  P''; C. Prop C Ψ (!P) (KN  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 ¿MN  P'; Ψ  !P ¿MN  P''; C. Prop C Ψ (!P) (¿MN  P'') 
                                         Prop C Ψ (P  !P) (¿MN  (P'  P''))"
  and   rBrComm1: "M N P' xvec P'' C. Ψ  P ¿MN  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 ¿MN  P''; C. Prop C Ψ (!P) (¿MN  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 (MN⟩.Q)")
    next
      case(BrOutput Ψ' M K N Q Ψ C)
      then show ?case by - (ind_cases "bangPred P (MN⟩.Q)")
    next
      case(Case Ψ' Q Rs φ Cs Ψ C)
      then show ?case by - (ind_cases "bangPred P (Cases Cs)")
    next
      case(cPar1 Ψ' ΨR Q α P' R AR Ψ 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 = AR, ΨR have "AR = []" and "ΨR = 𝟭" by auto
      from Ψ'  ΨR  Q α  P' Q = P Ψ  Ψ' Ψ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 Ψ' ΨP R α P' Q AP Ψ 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 = AP, ΨP guarded P have "ΨP  𝟭" and "supp ΨP = ({}::name set)"
        by(blast dest: guardedStatEq)+
      from Ψ'  ΨP  R α  P' R = !P Ψ  Ψ' ΨP  𝟭 have "Ψ  !P α  P'"
        by(metis statEqTransition Identity Composition Commutativity AssertionStatEqSym)
      moreover
      {
        fix C
        have "bangPred P (!P)" by(rule aux1)
        moreover from Ψ  Ψ' ΨP  𝟭 have "Ψ  Ψ'  Ψ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 Ψ' ΨR Q M N P' AP ΨP R K xvec P'' AR Ψ C)
      have rComm1: "M N P' K xvec P'' C. Ψ  P MN  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 = AR, ΨR have "AR = []" and "ΨR = 𝟭" by auto
      from Ψ'  ΨR  Q MN  P' Q = P Ψ  Ψ' ΨR = 𝟭 have "Ψ  P MN  P'"
        by(metis statEqTransition Identity AssertionStatEqSym)
      moreover from Q = P extractFrame Q = AP, ΨP guarded P have "ΨP  𝟭" and "supp ΨP = ({}::name set)"
        by(blast dest: guardedStatEq)+
      moreover from Ψ'  ΨP  R K⦇ν*xvec⦈⟨N  P'' R = !P Ψ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 Ψ  Ψ' ΨP  𝟭 have "Ψ  Ψ'  Ψ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 Ψ'  ΨP  ΨR  M  K ΨP  𝟭 Ψ  Ψ' Ψ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 Ψ' ΨR Q M xvec N P' AP ΨP R K P'' AR Ψ C)
      have rComm2: "M xvec N P' K P'' C. Ψ  P M⦇ν*xvec⦈⟨N  P'; Ψ  !P KN  P''; C. Prop C Ψ (!P) (KN  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 = AR, ΨR have "AR = []" and "ΨR = 𝟭" by auto
      from Ψ'  ΨR  Q M⦇ν*xvec⦈⟨N  P' Q = P Ψ  Ψ' ΨR = 𝟭 have "Ψ  P M⦇ν*xvec⦈⟨N  P'"
        by(metis statEqTransition Identity AssertionStatEqSym)
      moreover from Q = P extractFrame Q = AP, ΨP guarded P have "ΨP  𝟭" and "supp ΨP = ({}::name set)"
        by(blast dest: guardedStatEq)+
      moreover from Ψ'  ΨP  R KN  P'' R = !P ΨP  𝟭 Ψ  Ψ' have "Ψ  !P KN  P''"
        by(metis statEqTransition Identity Composition Commutativity AssertionStatEqSym)
      moreover
      {
        fix C
        have "bangPred P (!P)" by(rule aux1)
        moreover from Ψ  Ψ' ΨP  𝟭 have "Ψ  Ψ'  ΨP" by(metis Composition Identity Commutativity AssertionStatEqSym AssertionStatEqTrans)
        ultimately have "C. Prop C Ψ (!P) (KN  P'')" using cComm2 R = !P guarded P by simp
      }
      moreover from Ψ'  ΨP  ΨR  M  K ΨP  𝟭 Ψ  Ψ' Ψ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 Ψ' ΨR Q M N P' AP ΨP R P'' AR Ψ C)
      have rBrMerge: "M N P' P'' C. Ψ  P ¿MN  P'; Ψ  !P ¿MN  P''; C. Prop C Ψ (!P) (¿MN  P'') 
                                         Prop C Ψ (P  !P) (¿MN  (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 = AR, ΨR have "AR = []" and "ΨR = 𝟭" by auto
      from Ψ'  ΨR  Q ¿MN  P' Q = P Ψ  Ψ' ΨR = 𝟭 have "Ψ  P ¿MN  P'"
        by(metis statEqTransition Identity AssertionStatEqSym)
      moreover from Q = P extractFrame Q = AP, ΨP guarded P have "ΨP  𝟭" and "supp ΨP = ({}::name set)"
        by(blast dest: guardedStatEq)+
      from Ψ'  ΨP  R ¿MN  P'' R = !P Ψ  Ψ' ΨP  𝟭 have "Ψ  !P ¿MN  P''"
        by (metis AssertionStatEqSym Identity compositionSym statEqTransition)
      moreover
      {
        fix C
        have "bangPred P (!P)" by(rule aux1)
        moreover from Ψ  Ψ' ΨP  𝟭 have "Ψ  Ψ'  ΨP" by(metis Composition Identity Commutativity AssertionStatEqSym AssertionStatEqTrans)
        ultimately have "C. Prop C Ψ (!P) (¿MN  P'')" using cBrMerge R = !P guarded P by simp
      }
      ultimately have "Prop C Ψ (P  !P) (¿MN  (P'  P''))"
        by(elim rBrMerge) auto
      with R = !P Q = P show ?case by simp
    next
      case(cBrComm1 Ψ' ΨR Q M N P' AP ΨP R xvec P'' AR Ψ C)
      have rBrComm1: "M N P' xvec P'' C. Ψ  P ¿MN  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 = AR, ΨR have "AR = []" and "ΨR = 𝟭" by auto
      from Ψ'  ΨR  Q ¿MN  P' Q = P Ψ  Ψ' ΨR = 𝟭 have "Ψ  P ¿MN  P'"
        by(metis statEqTransition Identity AssertionStatEqSym)
      moreover from Q = P extractFrame Q = AP, ΨP guarded P have "ΨP  𝟭" and "supp ΨP = ({}::name set)"
        by(blast dest: guardedStatEq)+
      moreover from Ψ'  ΨP  R ¡M⦇ν*xvec⦈⟨N  P'' R = !P Ψ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 Ψ  Ψ' ΨP  𝟭 have "Ψ  Ψ'  Ψ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 Ψ' ΨR Q M xvec N P' AP ΨP R P'' AR Ψ C)
      have rBrComm2: "M N P' xvec P'' C. Ψ  P ¡M⦇ν*xvec⦈⟨N  P'; Ψ  !P ¿MN  P''; C. Prop C Ψ (!P) (¿MN  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 = AR, ΨR have "AR = []" and "ΨR = 𝟭" by auto
      from Ψ'  ΨR  Q ¡M⦇ν*xvec⦈⟨N  P' Q = P Ψ  Ψ' ΨR = 𝟭 have "Ψ  P ¡M⦇ν*xvec⦈⟨N  P'"
        by(metis statEqTransition Identity AssertionStatEqSym)
      moreover from Q = P extractFrame Q = AP, ΨP guarded P have "ΨP  𝟭" and "supp ΨP = ({}::name set)"
        by(blast dest: guardedStatEq)+
      moreover from Ψ'  ΨP  R ¿MN  P'' R = !P ΨP  𝟭 Ψ  Ψ' have "Ψ  !P ¿MN  P''"
        by(metis statEqTransition Identity Composition Commutativity AssertionStatEqSym)
      moreover
      {
        fix C
        have "bangPred P (!P)" by(rule aux1)
        moreover from Ψ  Ψ' ΨP  𝟭 have "Ψ  Ψ'  ΨP" by(metis Composition Identity Commutativity AssertionStatEqSym AssertionStatEqTrans)
        ultimately have "C. Prop C Ψ (!P) (¿MN  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 (⦇νxQ)")
    next
      case(cOpen Ψ Q M xvec yvec N P' x C)
      then show ?case by - (ind_cases "bangPred P (⦇νxQ)")
    next
      case(cBrOpen Ψ Q M xvec yvec N P' x C)
      then show ?case by - (ind_cases "bangPred P (⦇νxQ)")
    next
      case(cScope Ψ Q α P' x C)
      then show ?case by - (ind_cases "bangPred P (⦇νxQ)")
    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 MN  P'"
  and   rPar1: "P'. Ψ  P MN  P'  Prop Ψ (P  !P) M N (P'  !P)"
  and   rPar2: "P'. Ψ  !P MN  P'; Prop Ψ (!P) M N P'  Prop Ψ (P  !P) M N (P  P')"
  and   rBang: "P'. Ψ  P  !P MN  P'; Prop Ψ (P  !P) M N P'; guarded P  Prop Ψ (!P) M N P'"
shows "Prop Ψ (!P) M N P'"
  using Ψ  !P MN  P'
  by(nominal_induct Ψ P Rs=="MN  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 ¿MN  P'"
  and   rPar1: "P'. Ψ  P ¿MN  P'  Prop Ψ (P  !P) M N (P'  !P)"
  and   rPar2: "P'. Ψ  !P ¿MN  P'; Prop Ψ (!P) M N P'  Prop Ψ (P  !P) M N (P  P')"
  and   rBrMerge: "P' P'' C. Ψ  P ¿MN  P'; Ψ  !P ¿MN  P''; C. Prop Ψ (!P) M N P'' 
                                         Prop Ψ (P  !P) M N (P'  P'')"
  and   rBang: "P'. Ψ  P  !P ¿MN  P'; Prop Ψ (P  !P) M N P'; guarded P  Prop Ψ (!P) M N P'"
shows "Prop Ψ (!P) M N P'"
  using Ψ  !P ¿MN  P'
  by(nominal_induct Ψ P Rs=="¿MN  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 (⦇ν*xvecN ≺' (P'  !P))"
  and   rPar2: "xvec N P' C. Ψ  !P M⦇ν*xvec⦈⟨N  P'; C. Prop C Ψ (!P) M (⦇ν*xvecN ≺' P'); xvec ♯* Ψ; xvec ♯* P; xvec ♯* M; xvec ♯* C; distinct xvec 
                                  Prop C Ψ (P  !P) M (⦇ν*xvecN ≺' (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 (⦇ν*xvecN ≺' P')"
  using Ψ  !P M⦇ν*xvec⦈⟨N  P'
  apply(simp add: residualInject)
  by(nominal_induct Ψ P Rs=="ROut M (⦇ν*xvecN ≺' 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 MN  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 KN  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 MN  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 KN  P''; C. Prop C Ψ (!P) (KN) 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 ¿MN  P'; Ψ  !P ¿MN  P''; C. Prop C Ψ (!P) (¿MN) P'' 
                                         Prop C Ψ (P  !P) (¿MN) (P'  P'')"
  and   rBrComm1: "M N P' xvec P'' C. Ψ  P ¿MN  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 ¿MN  P''; C. Prop C Ψ (!P) (¿MN) 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) (¿MN) (P'  P'')"
      by(elim rBrMerge) (assumption | simp)+
    then show ?case using ¿MN  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 ΨQ   :: 'b
    and R    :: "('a, 'b, 'c) psi"
    and M    :: 'a
    and N    :: 'a
    and R'   :: "('a, 'b, 'c) psi"
    and AR   :: "name list"
    and ΨR   :: 'b
    and AP   :: "name list"
    and ΨP   :: 'b
    and AQ   :: "name list"

assumes RTrans: "Ψ  ΨQ  R ¿MN  R'"
  and   FrR: "extractFrame R = AR, ΨR"
  and   "distinct AR"
  and   QimpP: "AQ, (Ψ  ΨQ)  ΨR F AP, (Ψ  ΨP)  ΨR"
  and   "AR ♯* AP"
  and   "AR ♯* AQ"
  and   "AR ♯* Ψ"
  and   "AR ♯* ΨP"
  and   "AR ♯* ΨQ"
  and   "AR ♯* (Ψ  ΨQ)"
  and   "AR ♯* R"
  and   "AR ♯* M"
  and   "AP ♯* R"
  and   "AP ♯* M"
  and   "AQ ♯* R"
  and   "AQ ♯* M"

shows "Ψ  ΨP  R ¿MN  R'"
  using assms
proof(nominal_induct avoiding: AP ΨP AQ ΨQ Ψ rule: brinputFrameInduct)
  case(cAlpha Ψ' P M N P' AR ΨR p AP ΨP AQ ΨQ Ψ)
  have S: "set p  set AR × set (p  AR)" by fact
  from AQ, Ψ'  (p  ΨR) F AP, (Ψ  ΨP)  (p  ΨR)
  have "(p  AQ, Ψ'  (p  ΨR)) F (p  AP, (Ψ  ΨP)  (p  ΨR))"
    by(rule FrameStatImpClosed)
  with AR ♯* AP (p  AR) ♯* AP AR ♯* Ψ' (p  AR) ♯* Ψ' AR ♯* ΨP (p  AR) ♯* ΨP AR ♯* AQ
    (p  AR) ♯* AQ AR ♯* Ψ (p  AR) ♯* Ψ S distinctPerm p
  have "AQ, Ψ'  ΨR F AP, (Ψ  ΨP)  ΨR" by(simp add: eqvts)
  moreover note AR ♯* Ψ' AR ♯* Ψ AR ♯* P AR ♯* ΨP AR ♯* ΨQ AR ♯* M AP ♯* P
    AQ ♯* P AP ♯* M AQ ♯* M AR ♯* AP AR ♯* AQ
  ultimately show ?case
    by(elim cAlpha)
next
  case(cBrInput Ψ' M K xvec N Tvec P AP ΨP AQ ΨQ Ψ)
  from AP ♯* (M⦇λ*xvec N⦈.P) AQ ♯* (M⦇λ*xvec N⦈.P)
  have "AP ♯* M" and "AQ ♯* M" by simp+
  from Ψ'  K  M
  have "Ψ'  𝟭  K  M"
    by(blast intro: statEqEnt Identity AssertionStatEqSym)
  with AQ ♯* K AQ ♯* M
  have "(AQ, Ψ'  𝟭) F K  M"
    by(force intro: frameImpI)
  with AQ, Ψ'  𝟭 F AP, (Ψ  ΨP)  𝟭
  have "(AP, (Ψ  ΨP)  𝟭) F K  M"
    by(simp add: FrameStatImp_def)
  with AP ♯* K AP ♯* M have "(Ψ  ΨP)  𝟭  K  M"
    by(force dest: frameImpE)
  then have "Ψ  Ψ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 AR ΨR AQ ΨQ AP ΨP Ψ)
  from AP ♯* (Cases Cs)  AQ ♯* (Cases Cs) (φ, P)  set Cs
  have "AP ♯* P" and "AQ ♯* P" and "AP ♯* φ" and "AQ ♯* φ"
    by(auto dest: memFreshChain)

  from Ψ'  φ
  have "Ψ'  𝟭  φ"
    by(blast intro: statEqEnt Identity AssertionStatEqSym)
  with AP ♯* φ
  have "(AP, Ψ'  𝟭) F φ"
    by(force intro: frameImpI)
  with AP, Ψ'  𝟭 F AQ, (Ψ  ΨQ)  𝟭
  have "(AQ, (Ψ  ΨQ)  𝟭) F φ"
    by(simp add: FrameStatImp_def)
  with AQ ♯* φ have "(Ψ  ΨQ)  𝟭  φ"
    by(force dest: frameImpE)
  then have "Ψ  ΨQ  φ" by(blast intro: statEqEnt Identity)

  have "AP, Ψ'  ΨR F AQ, (Ψ  ΨQ)  ΨR"
  proof -
    from ΨR  𝟭 have "AP,  Ψ'  ΨR F AP, Ψ'  𝟭"
      by(metis Identity Commutativity AssertionStatEqSym Composition frameResChainPres frameNilStatEq AssertionStatEqTrans)
    moreover note AP, Ψ'  𝟭 F AQ, (Ψ  ΨQ)  𝟭
    moreover from ΨR  𝟭 have "AQ, (Ψ  ΨQ)  𝟭 F AQ, (Ψ  ΨQ)  ΨR"
      by(metis Identity Commutativity AssertionStatEqSym Composition frameResChainPres frameNilStatEq AssertionStatEqTrans)
    ultimately show ?thesis by(rule FrameStatEqImpCompose)
  qed
  moreover note AR ♯* Ψ' AR ♯* Ψ AR ♯* P AR ♯* ΨQ AR ♯* M
    AQ ♯* P AP ♯* P AQ ♯* M AP ♯* M AR ♯* AQ AR ♯* AP
  ultimately have "Ψ  ΨQ  P  ¿MN  P'"
    by(elim cCase(4))
  moreover note (φ, P)  set Cs Ψ  ΨQ  φ guarded P
  ultimately show ?case
    by(rule Case)
next
  case(cPar1 Ψ' ΨQ P M N P' AQ Q AP ΨP AP' ΨP' AQ' ΨQ' Ψ)
  from extractFrame Q = AQ, ΨQ AQ ♯* AP' AP' ♯* (P  Q) have "AP' ♯* ΨQ"
    by(force dest: extractFrameFreshChain)

  have "AP', (Ψ  ΨP')  ΨP  ΨQ F AP', (Ψ  ΨP')  ΨQ  ΨP"
    by (metis Commutativity FrameStatEq_def frameIntCompositionSym)
  moreover have "AP', (Ψ  ΨP')  ΨQ  ΨP F AP', ((Ψ  ΨP')  ΨQ)  ΨP"
    by (metis FrameStatEq_def frameIntAssociativity)
  moreover have "AP', ((Ψ  ΨP')  ΨQ)  ΨP F AP', ((Ψ  ΨQ)  ΨP')  ΨP"
    by (metis FrameStatEq_def associativitySym frameIntComposition)
  ultimately have right: "AP', (Ψ  ΨP')  ΨP  ΨQ F AP', ((Ψ  ΨQ)  ΨP')  ΨP"
    by (metis FrameStatImpTrans)
  have "AQ', Ψ'  ΨP  ΨQ F AQ', Ψ'  ΨQ  ΨP"
    by (metis AssertionStatEq_def Commutativity compositionSym frameImpNilStatEq frameImpResChainPres)
  moreover have "AQ', Ψ'  ΨQ  ΨP F AQ', (Ψ'  ΨQ)  ΨP"
    by (metis FrameStatEq_def frameIntAssociativity)
  ultimately have left: "AQ', Ψ'  ΨP  ΨQ F AQ', (Ψ'  ΨQ)  ΨP"
    by (metis FrameStatImpTrans)
  from AQ', Ψ'  ΨP  ΨQ F AP', (Ψ  ΨP')  ΨP  ΨQ left right
  have "AQ', (Ψ'  ΨQ)  ΨP F AP', ((Ψ  ΨQ)  ΨP')  ΨP"
    by (metis AssertionStatEqSym AssertionStatEqTrans AssertionStatEq_def Associativity FrameStatImpTrans associativitySym frameImpNilStatEq frameImpResChainPres)
  moreover note AP ♯* Ψ' AP ♯* ΨQ AP ♯* Ψ AP ♯* ΨQ AP ♯* P AP ♯* M
    AP ♯* ΨP' AP' ♯* (P  Q) AQ' ♯* (P  Q) AP' ♯* ΨQ AP' ♯* M AQ' ♯* M AP ♯* AP' AP ♯* AQ'
  ultimately have "(Ψ  ΨQ)  ΨP'  P  ¿MN  P'"
    by(elim cPar1(6)) (simp | force)+
  then have "(Ψ  ΨP')  ΨQ  P  ¿MN  P'"
    by (metis associativitySym statEqTransition)

  moreover note extractFrame Q = AQ, ΨQ AQ ♯* Ψ AQ ♯* ΨP'
    AQ ♯* P AQ ♯* M AQ ♯* N
  ultimately show ?case
    by(elim Par1) (simp | force)+
next
  case(cPar2 Ψ' ΨP Q M N Q' AP P AQ ΨQ AP' ΨP' AQ' ΨQ' Ψ)
  from extractFrame P = AP, ΨP AP ♯* AP' AP' ♯* (P  Q) have "AP' ♯* ΨP"
    by(force dest: extractFrameFreshChain)
  have "AP', (Ψ  ΨP')  ΨP  ΨQ F AP', ((Ψ  ΨP')  ΨP)  ΨQ"
    by (metis FrameStatEq_def frameIntAssociativity)
  moreover have "AP', ((Ψ  ΨP')  ΨP)  ΨQ F AP', ((Ψ  ΨP)  ΨP')  ΨQ"
    by (metis FrameStatEq_def associativitySym frameIntComposition)
  ultimately have right: "AP', (Ψ  ΨP')  ΨP  ΨQ F AP', ((Ψ  ΨP)  ΨP')  ΨQ"
    by (metis FrameStatImpTrans)
  moreover have left: "AQ', Ψ'  ΨP  ΨQ F AQ', (Ψ'  ΨP)  ΨQ"
    by (metis FrameStatEq_def frameIntAssociativity)
  from AQ', Ψ'  ΨP  ΨQ F AP', (Ψ  ΨP')  ΨP  ΨQ left right
  have "AQ', (Ψ'  ΨP)  ΨQ F AP', ((Ψ  ΨP)  ΨP')  ΨQ"
    by (metis FrameStatEq_def FrameStatImpTrans frameIntAssociativity)
  moreover note AQ ♯* Ψ' AQ ♯* ΨP AQ ♯* Ψ AQ ♯* ΨP AQ ♯* Q AQ ♯* M
    AQ ♯* ΨP' AP' ♯* (P  Q) AQ' ♯* (P  Q) AP' ♯* ΨP AP' ♯* M AQ' ♯* M AQ ♯* AQ' AQ ♯* AP'
  ultimately have "(Ψ  ΨP)  ΨP'  Q  ¿MN  Q'"
    by(elim cPar2(6)) (simp | force)+
  then have "(Ψ  ΨP')  ΨP  Q  ¿MN  Q'"
    by (metis associativitySym statEqTransition)
  moreover note extractFrame P = AP, ΨP AP ♯* Ψ AP ♯* ΨP'
    AP ♯* Q AP ♯* M AP ♯* N
  ultimately show ?case
    by(elim Par2) (simp | force)+
next
  case(cBrMerge Ψ' ΨQ P M N P' AP ΨP Q Q' AQ AP' ΨP' AQ' ΨQ' Ψ)
  from extractFrame Q = AQ, ΨQ AQ ♯* AP' AP' ♯* (P  Q) have "AP' ♯* ΨQ"
    by(force dest: extractFrameFreshChain)
  from extractFrame P = AP, ΨP AP ♯* AP' AP' ♯* (P  Q) have "AP' ♯* ΨP"
    by(force dest: extractFrameFreshChain)
  have "AP', (Ψ  ΨP')  ΨP  ΨQ F AP', (Ψ  ΨP')  ΨQ  ΨP"
    by (metis Commutativity FrameStatEq_def frameIntCompositionSym)
  moreover have "AP', (Ψ  ΨP')  ΨQ  ΨP F AP', ((Ψ  ΨP')  ΨQ)  ΨP"
    by (metis FrameStatEq_def frameIntAssociativity)
  moreover have "AP', ((Ψ  ΨP')  ΨQ)  ΨP F AP', ((Ψ  ΨQ)  ΨP')  ΨP"
    by (metis FrameStatEq_def associativitySym frameIntComposition)
  ultimately have right: "AP', (Ψ  ΨP')  ΨP  ΨQ F AP', ((Ψ  ΨQ)  ΨP')  ΨP"
    by (metis FrameStatImpTrans)
  have "AQ', Ψ'  ΨP  ΨQ F AQ', Ψ'  ΨQ  ΨP"
    by (metis AssertionStatEq_def Commutativity compositionSym frameImpNilStatEq frameImpResChainPres)
  moreover have "AQ', Ψ'  ΨQ  ΨP F AQ', (Ψ'  ΨQ)  ΨP"
    by (metis FrameStatEq_def frameIntAssociativity)
  ultimately have left: "AQ', Ψ'  ΨP  ΨQ F AQ', (Ψ'  ΨQ)  ΨP"
    by (metis FrameStatImpTrans)
  from AQ', Ψ'  ΨP  ΨQ F AP', (Ψ  ΨP')  ΨP  ΨQ left right
  have "AQ', (Ψ'  ΨQ)  ΨP F AP', ((Ψ  ΨQ)  ΨP')  ΨP"
    by (metis AssertionStatEqSym AssertionStatEqTrans AssertionStatEq_def Associativity FrameStatImpTrans associativitySym frameImpNilStatEq frameImpResChainPres)
  moreover note AP ♯* Ψ' AP ♯* ΨQ AP ♯* Ψ AP ♯* ΨQ AP ♯* P AP ♯* M
    AP ♯* ΨP' AP' ♯* (P  Q) AQ' ♯* (P  Q) AP' ♯* ΨQ AP' ♯* M AQ' ♯* M AP ♯* AP' AP ♯* AQ'
  ultimately have "(Ψ  ΨQ)  ΨP'  P  ¿MN  P'"
    by(elim cBrMerge(2)) (simp | force)+
  then have Ptrans: "(Ψ  ΨP')  ΨQ  P  ¿MN  P'"
    by (metis associativitySym statEqTransition)

  have left2: "AQ', Ψ'  ΨP  ΨQ F AQ', (Ψ'  ΨP)  ΨQ"
    by (metis FrameStatEq_def frameIntAssociativity)
  have "AP', (Ψ  ΨP')  ΨP  ΨQ F AP', ((Ψ  ΨP')  ΨP)  ΨQ"
    by (metis FrameStatEq_def frameIntAssociativity)
  moreover have "AP', ((Ψ  ΨP')  ΨP)  ΨQ F AP', ((Ψ  ΨP)  ΨP')  ΨQ"
    by (metis FrameStatEq_def associativitySym frameIntComposition)
  ultimately have right2: "AP', (Ψ  ΨP')  ΨP  ΨQ F AP', ((Ψ  ΨP)  ΨP')  ΨQ"
    by (metis FrameStatImpTrans)
  from AQ', Ψ'  ΨP  ΨQ F AP', (Ψ  ΨP')  ΨP  ΨQ left2 right2
  have "AQ', (Ψ'  ΨP)  ΨQ F AP', ((Ψ  ΨP)  ΨP')  ΨQ"
    by (metis FrameStatEq_def FrameStatImpTrans frameIntAssociativity)
  moreover note AQ ♯* Ψ' AQ ♯* ΨP AQ ♯* Ψ AQ ♯* ΨP AQ ♯* Q AQ ♯* M
    AQ ♯* ΨP' AP' ♯* (P  Q) AQ' ♯* (P  Q) AP' ♯* ΨP AP' ♯* M AQ' ♯* M AQ ♯* AQ' AQ ♯* AP'
  ultimately have "(Ψ  ΨP)  ΨP'  Q  ¿MN  Q'"
    by(elim cBrMerge(6)) (simp | force)+
  then have Qtrans: "(Ψ  ΨP')  ΨP  Q  ¿MN  Q'"
    by (metis associativitySym statEqTransition)

  from Ptrans extractFrame P = AP, ΨP Qtrans extractFrame Q = AQ, ΨQ
    AP ♯* Ψ AP ♯* ΨP' AP ♯* P AP ♯* Q AP ♯* M AP ♯* AQ AQ ♯* Ψ AQ ♯* ΨP' AQ ♯* P
    AQ ♯* Q AQ ♯* M
  show ?case
    by(elim BrMerge) (simp | force)+
next
  case(cScope Ψ' P M N P' x AP ΨP AP' ΨP' AQ ΨQ Ψ)
  then have "Ψ  ΨP'  P  ¿MN  P'"
    by simp
  with x  Ψ x  ΨP' x  M x  N
  show ?case
    by(elim Scope) (simp | force)+
next
  case(cBang Ψ' P M N P' AR ΨR AP ΨP AQ ΨQ Ψ)
  from AR ♯* P have "AR ♯* (P  !P)"
    by simp
  from AP ♯* !P have "AP ♯* (P  !P)"
    by simp
  from AQ ♯* !P have "AQ ♯* (P  !P)"
    by simp

  have "AQ, Ψ'  ΨR  𝟭 F AP, (Ψ  ΨP)  ΨR  𝟭"
  proof -
    from ΨR  𝟭 have "AQ, Ψ'  ΨR  𝟭 F AQ, Ψ'  𝟭"
      by(metis Identity Commutativity AssertionStatEqSym Composition frameResChainPres frameNilStatEq AssertionStatEqTrans)
    moreover note AQ, Ψ'  𝟭 F AP, (Ψ  ΨP)  𝟭
    moreover  from ΨR  𝟭 have "AP, (Ψ  ΨP)  𝟭 F AP, (Ψ  ΨP)  ΨR  𝟭"
      by(metis Identity Commutativity AssertionStatEqSym Composition frameResChainPres frameNilStatEq AssertionStatEqTrans)
    ultimately show ?thesis by(rule FrameStatEqImpCompose)
  qed
  then have "Ψ  ΨP  P  !P  ¿MN  P'"
    using AR ♯* Ψ' AR ♯* Ψ AR ♯* (P  !P) AR ♯* ΨP
      AP ♯* (P  !P) AQ ♯* (P  !P) AP ♯* M AQ ♯* M AR ♯* M AR ♯* AP AR ♯* AQ
    by(elim cBang(5))
  then show ?case using guarded P
    by(rule Bang)
qed

lemma brCommInAux:
  fixes Ψ    :: 'b
    and ΨQ   :: 'b
    and R    :: "('a, 'b, 'c) psi"
    and M    :: 'a
    and N    :: 'a
    and R'   :: "('a, 'b, 'c) psi"
    and AR   :: "name list"
    and ΨR   :: 'b
    and AP   :: "name list"
    and ΨP   :: 'b
    and AQ   :: "name list"

assumes RTrans: "Ψ  ΨQ  R ¿MN  R'"
  and   FrR: "extractFrame R = AR, ΨR"
  and   "distinct AR"
  and   QimpP: "AQ, (Ψ  ΨQ)  ΨR F AP, (Ψ  ΨP)  ΨR"
  and   "AR ♯* AP"
  and   "AR ♯* AQ"
  and   "AR ♯* Ψ"
  and   "AR ♯* ΨP"
  and   "AR ♯* ΨQ"
  and   "AR ♯* R"
  and   "AR ♯* M"
  and   "AP ♯* R"
  and   "AP ♯* M"
  and   "AQ ♯* R"
  and   "AQ ♯* M"

shows "Ψ  ΨP  R ¿MN  R'"
proof -
  from AR ♯* Ψ AR ♯* ΨQ
  have "AR ♯* (Ψ  ΨQ)"
    by auto
  with assms
  show ?thesis
    by(simp add: brCommInAuxTooMuch)
qed

lemma brCommOutAuxTooMuch:
  fixes Ψ    :: 'b
    and Ψ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 AR   :: "name list"
    and ΨR   :: 'b
    and AP   :: "name list"
    and ΨP   :: 'b
    and AQ   :: "name list"

assumes RTrans: "Ψ  ΨQ  R  RBrOut M (⦇ν*xvecN ≺' R')"
  and   FrR: "extractFrame R = AR, ΨR"
  and   "distinct AR"
  and   QimpP: "AQ, (Ψ  ΨQ)  ΨR F AP, (Ψ  ΨP)  ΨR"
  and   "AR ♯* AP"
  and   "AR ♯* AQ"
  and   "AR ♯* Ψ"
  and   "AR ♯* ΨP"
  and   "AR ♯* ΨQ"
  and   "AR ♯* (Ψ  ΨQ)"
  and   "AR ♯* R"
  and   "AR ♯* M"
  and   "AP ♯* R"
  and   "AP ♯* M"
  and   "AQ ♯* R"
  and   "AQ ♯* M"

shows "Ψ  ΨP  R  ¡M⦇ν*xvec⦈⟨N  R'"
  using assms
proof(nominal_induct R M B=="⦇ν*xvecN ≺' R'" AR ΨR avoiding: Ψ AP ΨP AQ ΨQ N R' xvec rule: broutputFrameInduct)
  case (cAlpha Ψ' P M AR ΨR p Ψ AP ΨP AQ ΨQ N P' xvec)
  have S: "set p  set AR × set (p  AR)" by fact
  from AQ, Ψ'  (p  ΨR) F AP, (Ψ  ΨP)  (p  ΨR)
  have "(p  AQ, Ψ'  (p  ΨR)) F (p  AP, (Ψ  ΨP)  (p  ΨR))"
    by(rule FrameStatImpClosed)
  with AR ♯* AP (p  AR) ♯* AP AR ♯* Ψ' (p  AR) ♯* Ψ' AR ♯* ΨP (p  AR) ♯* ΨP AR ♯* AQ
    (p  AR) ♯* AQ AR ♯* Ψ (p  AR) ♯* Ψ S distinctPerm p
  have "AQ, Ψ'  ΨR F AP, (Ψ  ΨP)  ΨR" by(simp add: eqvts)
  moreover note AR ♯* Ψ' AR ♯* Ψ AR ♯* P AR ♯* ΨP AR ♯* ΨQ AR ♯* M AP ♯* P
    AQ ♯* P AP ♯* M AQ ♯* M AR ♯* AP AR ♯* AQ
  ultimately show ?case
    by(elim cAlpha)
next
  case(cBrOutput Ψ' M K N P Ψ AP ΨP AQ ΨQ N' R' xvec)
  from AP ♯* (MN⟩.P) AQ ♯* (MN⟩.P)
  have "AP ♯* M" and "AQ ♯* M" by simp+
  from Ψ'  M  K
  have "Ψ'  𝟭  M  K"
    by(blast intro: statEqEnt Identity AssertionStatEqSym)
  with AQ ♯* K AQ ♯* M
  have "(AQ, Ψ'  𝟭) F M  K"
    by(force intro: frameImpI)
  with AQ, Ψ'  𝟭 F AP, (Ψ  ΨP)  𝟭
  have "(AP, (Ψ  ΨP)  𝟭) F M  K"
    by(simp add: FrameStatImp_def)
  with AP ♯* K AP ♯* M have "(Ψ  ΨP)  𝟭  M  K"
    by(force dest: frameImpE)
  then have "Ψ  ΨP  M  K" by(blast intro: statEqEnt Identity)
  then have "Ψ  ΨP  MN⟩.P  ¡KN  P"
    by(rule BrOutput)
  with N ≺' P = ⦇ν*xvecN' ≺' R'
  show ?case
    by(simp add: residualInject)
next
  case(cCase Ψ' R M φ Cs AR ΨR Ψ AP ΨP AQ ΨQ N R' xvec)
  from AP ♯* (Cases Cs)  AQ ♯* (Cases Cs) (φ, R)  set Cs
  have "AP ♯* R" and "AQ ♯* R" and "AP ♯* φ" and "AQ ♯* φ"
    by(auto dest: memFreshChain)

  from Ψ'  φ have "Ψ'  𝟭  φ" by(blast intro: statEqEnt Identity AssertionStatEqSym)
  with AQ ♯* φ have "(AQ, Ψ'  𝟭) F φ" by(force intro: frameImpI)
  with AQ, Ψ'  𝟭 F AP, (Ψ  ΨP)  𝟭 have "(AP, (Ψ  ΨP)  𝟭) F φ"
    by(simp add: FrameStatImp_def)
  with AP ♯* φ have "(Ψ  ΨP)  𝟭  φ"  by(force dest: frameImpE)
  then have "Ψ  ΨP  φ" by(blast intro: statEqEnt Identity)

  have "AQ, Ψ'  ΨR F AP, (Ψ  ΨP)  ΨR"
  proof -
    from ΨR  𝟭 have "AQ,  Ψ'  ΨR F AQ, Ψ'  𝟭"
      by(metis Identity Commutativity AssertionStatEqSym Composition frameResChainPres frameNilStatEq AssertionStatEqTrans)
    moreover note AQ, Ψ'  𝟭 F AP, (Ψ  ΨP)  𝟭
    moreover from ΨR  𝟭 have "AP, (Ψ  ΨP)  𝟭 F AP, (Ψ  ΨP)  ΨR"
      by(metis Identity Commutativity AssertionStatEqSym Composition frameResChainPres frameNilStatEq AssertionStatEqTrans)
    ultimately show ?thesis by(rule FrameStatEqImpCompose)
  qed
  with AP ♯* R AQ ♯* R
  have "Ψ  ΨP  R  ¡M⦇ν*xvec⦈⟨N  R'" using cCase
    by(intro cCase(4)) simp+
  then show ?case using (φ, R)  set Cs Ψ  ΨP  φ guarded R
    by(rule Case)
next
  case(cPar1 Ψ' ΨQ P M xvec N P' AQ Q AP ΨP Ψ AP' ΨP' AQ' ΨQ' N' R' yvec)
  from extractFrame Q = AQ, ΨQ AQ ♯* AP' AP' ♯* (P  Q) have "AP' ♯* ΨQ"
    by(force dest: extractFrameFreshChain)

  have "AP', (Ψ  ΨP')  ΨP  ΨQ F AP', (Ψ  ΨP')  ΨQ  ΨP"
    by (metis Commutativity FrameStatEq_def frameIntCompositionSym)
  moreover have "AP', (Ψ  ΨP')  ΨQ  ΨP F AP', ((Ψ  ΨP')  ΨQ)  ΨP"
    by (metis FrameStatEq_def frameIntAssociativity)
  moreover have "AP', ((Ψ  ΨP')  ΨQ)  ΨP F AP', ((Ψ  ΨQ)  ΨP')  ΨP"
    by (metis FrameStatEq_def associativitySym frameIntComposition)
  ultimately have right: "AP', (Ψ  ΨP')  ΨP  ΨQ F AP', ((Ψ  ΨQ)  ΨP')  ΨP"
    by (metis FrameStatImpTrans)
  have "AQ', Ψ'  ΨP  ΨQ F AQ', Ψ'  ΨQ  ΨP"
    by (metis AssertionStatEq_def Commutativity compositionSym frameImpNilStatEq frameImpResChainPres)
  moreover have "AQ', Ψ'  ΨQ  ΨP F AQ', (Ψ'  ΨQ)  ΨP"
    by (metis FrameStatEq_def frameIntAssociativity)
  ultimately have left: "AQ', Ψ'  ΨP  ΨQ F AQ', (Ψ'  ΨQ)  ΨP"
    by (metis FrameStatImpTrans)
  from AQ', Ψ'  ΨP  ΨQ F AP', (Ψ  ΨP')  ΨP  ΨQ left right
  have "AQ', (Ψ'  ΨQ)  ΨP F AP', ((Ψ  ΨQ)  ΨP')  ΨP"
    by (metis AssertionStatEqSym AssertionStatEqTrans AssertionStatEq_def Associativity FrameStatImpTrans associativitySym frameImpNilStatEq frameImpResChainPres)
  moreover note AP ♯* Ψ' AP ♯* ΨQ AP ♯* Ψ AP ♯* ΨQ AP ♯* P AP ♯* M
    AP ♯* ΨP' AP' ♯* (P  Q) AQ' ♯* (P  Q) AP' ♯* ΨQ AP' ♯* M AQ' ♯* M AP ♯* AP' AP ♯* AQ'
  ultimately have "(Ψ  ΨQ)  ΨP'  P ¡M⦇ν*xvec⦈⟨N  P'"
    by(intro cPar1(6)) (simp | force)+
  then have "(Ψ  ΨP')  ΨQ  P ¡M⦇ν*xvec⦈⟨N  P'"
    by (metis associativitySym statEqTransition)
  moreover note extractFrame Q = AQ, ΨQ xvec ♯* Q AQ ♯* Ψ AQ ♯* ΨP'
    AQ ♯* P AQ ♯* M AQ ♯* xvec AQ ♯* N
  ultimately have "Ψ  ΨP'  P  Q ¡M⦇ν*xvec⦈⟨N  P'  Q"
    by(elim Par1) (simp | force)+
  then show ?case using ⦇ν*xvecN ≺' (P'  Q) = ⦇ν*yvecN' ≺' R'
    by(simp add: residualInject)
next
  case(cPar2 Ψ' ΨP Q M xvec N Q' AP P AQ ΨQ Ψ AP' ΨP' AQ' ΨQ' N' R' yvec)
  from extractFrame P = AP, ΨP AP ♯* AP' AP' ♯* (P  Q) have "AP' ♯* ΨP"
    by(force dest: extractFrameFreshChain)
  have "AP', (Ψ  ΨP')  ΨP  ΨQ F AP', ((Ψ  ΨP')  ΨP)  ΨQ"
    by (metis FrameStatEq_def frameIntAssociativity)
  moreover have "AP', ((Ψ  ΨP')  ΨP)  ΨQ F AP', ((Ψ  ΨP)  ΨP')  ΨQ"
    by (metis FrameStatEq_def associativitySym frameIntComposition)
  ultimately have right: "AP', (Ψ  ΨP')  ΨP  ΨQ F AP', ((Ψ  ΨP)  ΨP')  ΨQ"
    by (metis FrameStatImpTrans)
  moreover have left: "AQ', Ψ'  ΨP  ΨQ F AQ', (Ψ'  ΨP)  ΨQ"
    by (metis FrameStatEq_def frameIntAssociativity)
  from AQ', Ψ'  ΨP  ΨQ F AP', (Ψ  ΨP')  ΨP  ΨQ left right
  have "AQ', (Ψ'  ΨP)  ΨQ F AP', ((Ψ  ΨP)  ΨP')  ΨQ"
    by (metis FrameStatEq_def FrameStatImpTrans frameIntAssociativity)
  moreover note AQ ♯* Ψ' AQ ♯* ΨP AQ ♯* Ψ AQ ♯* ΨP AQ ♯* Q AQ ♯* M
    AQ ♯* ΨP' AP' ♯* (P  Q) AQ' ♯* (P  Q) AP' ♯* ΨP AP' ♯* M AQ' ♯* M AQ ♯* AQ' AQ ♯* AP'
  ultimately have "(Ψ  ΨP)  ΨP'  Q ¡M⦇ν*xvec⦈⟨N  Q'"
    by(intro cPar2(6)) (simp | force)+
  then have "(Ψ  ΨP')  ΨP  Q ¡M⦇ν*xvec⦈⟨N  Q'"
    by (metis associativitySym statEqTransition)
  moreover note extractFrame P = AP, ΨP xvec ♯* P AP ♯* Ψ AP ♯* ΨP'
    AP ♯* Q AP ♯* M AP ♯* xvec AP ♯* N
  ultimately have "Ψ  ΨP'  P  Q ¡M⦇ν*xvec⦈⟨N  P  Q'"
    by(elim Par2) (simp | force)+
  then show ?case using ⦇ν*xvecN ≺' (P  Q') = ⦇ν*yvecN' ≺' R'
    by(simp add: residualInject)
next
  case(cBrComm1 Ψ' ΨQ P M N P' AP ΨP Q xvec Q' AQ Ψ AP' ΨP' AQ' ΨQ' N' R' yvec)
  have right: "AP', (Ψ  ΨP')  ΨP  ΨQ F AP', (ΨQ  (Ψ  ΨP'))  ΨP"
    by (metis AssertionStatEqTrans AssertionStatEq_def Associativity Commutativity FrameStatImpTrans frameImpNilStatEq frameImpResChainPres)
  with AQ', Ψ'  ΨP  ΨQ F AP', (Ψ  ΨP')  ΨP  ΨQ
  have "AQ', (ΨQ  Ψ')  ΨP F AP', (ΨQ  (Ψ  ΨP'))  ΨP"
    by (metis AssertionStatEqTrans AssertionStatEq_def Associativity Commutativity FrameStatImpTrans frameImpNilStatEq frameImpResChainPres)
  moreover from Ψ'  ΨQ  P  ¿MN  P'
  have "ΨQ  Ψ'  P  ¿MN  P'"
    by (metis Commutativity statEqTransition)
  moreover note extractFrame P = AP, ΨP distinct AP AP ♯* AP'
    AP ♯* AQ' AP ♯* ΨQ AP ♯* Ψ AP ♯* ΨP' AP ♯* Ψ' AP ♯* P AP ♯* M AP' ♯* (P  Q)
    AP' ♯* M AQ' ♯* (P  Q) AQ' ♯* M
  ultimately have "ΨQ  (Ψ  ΨP')  P  ¿MN  P'"
    by(elim brCommInAux) (simp | force)+
  then have Ptrans: "(Ψ  ΨP')  ΨQ  P  ¿MN  P'"
    by (metis Commutativity statEqTransition)

  have right2: "AP', (Ψ  ΨP')  ΨP  ΨQ F AP', ((Ψ  ΨP')  ΨP)  ΨQ"
    by (metis FrameStatEq_def frameIntAssociativity)
  with AQ', Ψ'  ΨP  ΨQ F AP', (Ψ  ΨP')  ΨP  ΨQ
  have "AQ', (Ψ'  ΨP)  ΨQ F AP', ((Ψ  ΨP')  ΨP)  ΨQ"
    by (metis FrameStatEq_def FrameStatImpTrans frameIntAssociativity)
  then have Qtrans: "(Ψ  ΨP')  ΨP  Q  ¡M⦇ν*xvec⦈⟨N  Q'"
    using AQ ♯* AP' AQ ♯* AQ' AQ ♯* Ψ AQ ♯* ΨP'
      AQ ♯* ΨP AQ ♯* Ψ' AQ ♯* ΨP AQ ♯* Q AQ ♯* M AP' ♯* (P  Q) AP' ♯* M
      AQ' ♯* (P  Q) AQ' ♯* M
    by(intro cBrComm1(8)) (simp | force)+
  from Ptrans extractFrame P = AP, ΨP Qtrans extractFrame Q = AQ, ΨQ
    AP ♯* Ψ AP ♯* ΨP' AP ♯* P AP ♯* Q
    AP ♯* M AP ♯* AQ AQ ♯* Ψ AQ ♯* ΨP' AQ ♯* P AQ ♯* Q AQ ♯* M xvec ♯* P
  have "Ψ  ΨP'  P  Q  ¡M⦇ν*xvec⦈⟨N  P'  Q'"
    by(elim BrComm1) (simp | force)+
  then show ?case using ⦇ν*xvecN ≺' (P'  Q') = ⦇ν*yvecN' ≺' R'
    by(simp add: residualInject)
next
  case(cBrComm2 Ψ' ΨQ P M xvec N P' AP ΨP Q Q' AQ Ψ AP' ΨP' AQ' ΨQ' N' R' yvec)
  have "AP', (Ψ  ΨP')  ΨP  ΨQ F AP', (ΨP  (Ψ  ΨP'))  ΨQ"
    by (metis AssertionStatEqTrans AssertionStatEq_def Commutativity associativitySym frameImpNilStatEq frameImpResChainPres)
  with AQ', Ψ'  ΨP  ΨQ F AP', (Ψ  ΨP')  ΨP  ΨQ
  have "AQ', (ΨP  Ψ')  ΨQ F AP', (ΨP  (Ψ  ΨP'))  ΨQ"
    by (metis AssertionStatEqTrans AssertionStatEq_def Commutativity FrameStatImpTrans associativitySym frameImpNilStatEq frameImpResChainPres)
  moreover from Ψ'  ΨP  Q  ¿MN  Q'
  have "ΨP  Ψ'  Q  ¿MN  Q'"
    by (metis Commutativity statEqTransition)
  moreover note extractFrame Q = AQ, ΨQ distinct AQ
    AQ ♯* AP' AQ ♯* AQ' AQ ♯* ΨP AQ ♯* Ψ AQ ♯* ΨP'
    AQ ♯* Ψ' AQ ♯* Q AQ ♯* M AP' ♯* (P  Q) AP' ♯* M AQ' ♯* (P  Q) AQ' ♯* M
  ultimately have "ΨP  (Ψ  ΨP')  Q  ¿MN  Q'"
    by (elim brCommInAux) (simp | force)+
  then have Qtrans: "(Ψ  ΨP')  ΨP  Q  ¿MN  Q'"
    by (metis Commutativity statEqTransition)

  have "AP', (Ψ  ΨP')  ΨP  ΨQ F AP', ((Ψ  ΨP')  ΨQ)  ΨP"
    by (metis AssertionStatEqTrans AssertionStatEq_def Associativity associativitySym frameImpNilStatEq frameImpResChainPres)
  with AQ', Ψ'  ΨP  ΨQ F AP', (Ψ  ΨP')  ΨP  ΨQ
  have "AQ', (Ψ'  ΨQ)  ΨP F AP', ((Ψ  ΨP')  ΨQ)  ΨP"
    by (metis AssertionStatEqTrans AssertionStatEq_def Associativity FrameStatImpTrans associativitySym frameImpNilStatEq frameImpResChainPres)
  with AP ♯* AP' AP ♯* AQ' AP ♯* Ψ AP ♯* ΨP' AP ♯* ΨQ
    AP ♯* Ψ' AP ♯* ΨQ AP ♯* P AP ♯* M AP' ♯* (P  Q) AP' ♯* M
    AQ' ♯* (P  Q) AQ' ♯* M
  have Ptrans: "(Ψ  ΨP')  ΨQ  P  ¡M⦇ν*xvec⦈⟨N  P'"
    by(intro cBrComm2(4)) (simp | force)+
  from Ptrans extractFrame P = AP, ΨP Qtrans extractFrame Q = AQ, ΨQ
    AP ♯* Ψ AP ♯* ΨP' AP ♯* P AP ♯* Q AP ♯* M AP ♯* AQ
    AQ ♯* Ψ AQ ♯* ΨP' AQ ♯* P AQ ♯* Q AQ ♯* M xvec ♯* Q
  have "Ψ  ΨP'  P  Q  ¡M⦇ν*xvec⦈⟨N  P'  Q'"
    by(elim BrComm2) (simp | force)+
  then show ?case using ⦇ν*xvecN ≺' (P'  Q') = ⦇ν*yvecN' ≺' R'
    by(simp add: residualInject)
next
  case(cBrOpen Ψ' R M' xvec yvec N' R' x AR ΨR Ψ AP ΨP AQ ΨQ N R'' zvec)
  have "Ψ  ΨP  R  ¡M'⦇ν*(xvec@yvec)⦈⟨N'  R'" using cBrOpen
    by(intro cBrOpen(4)) (assumption | simp)+

  then have "Ψ  ΨP  ⦇νxR  ¡M'⦇ν*(xvec@x#yvec)⦈⟨N'  R'"
    using x  supp N' x  Ψ x  ΨP x  M' x  xvec x  yvec
    by(elim BrOpen) (simp | force)+
  with ⦇ν*(xvec @ x # yvec)N' ≺' R' = ⦇ν*zvecN ≺' R''
  show ?case
    by(simp add: residualInject)
next
  case(cScope Ψ' R M' xvec N' R' x AR ΨR Ψ AP ΨP AQ ΨQ N R'' yvec)
  then have "Ψ  ΨP  R  ¡M'⦇ν*xvec⦈⟨N'  R'"
    by(intro cScope(4)) (simp | force)+
  with x  Ψ x  ΨP x  M' x  xvec x  N'
  have "Ψ  ΨP  ⦇νxR  ¡M'⦇ν*xvec⦈⟨N'  ⦇νxR'"
    by(elim Scope) (simp | force)+
  then show ?case using ⦇ν*xvecN' ≺' ⦇νxR' = ⦇ν*yvecN ≺' R''
    by(simp add: residualInject)
next
  case(cBang Ψ' R M' AR ΨR Ψ AP ΨP AQ ΨQ N R' xvec)
  have "AQ, Ψ'  ΨR  𝟭 F AP, (Ψ  ΨP)  ΨR  𝟭"
  proof -
    from ΨR  𝟭 have "AQ,  Ψ'  ΨR  𝟭 F AQ, Ψ'  𝟭"
      by(metis Identity Commutativity AssertionStatEqSym Composition frameResChainPres frameNilStatEq AssertionStatEqTrans)
    moreover note AQ, Ψ'  𝟭 F AP, (Ψ  ΨP)  𝟭
    moreover from ΨR  𝟭 have "AP, (Ψ  ΨP)  𝟭 F AP, (Ψ  ΨP)  ΨR  𝟭"
      by(metis Identity Commutativity AssertionStatEqSym Composition frameResChainPres frameNilStatEq AssertionStatEqTrans)
    ultimately show ?thesis by(rule FrameStatEqImpCompose)
  qed
  then have "Ψ  Ψ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 Ψ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 AR   :: "name list"
    and ΨR   :: 'b
    and AP   :: "name list"
    and ΨP   :: 'b
    and AQ   :: "name list"

assumes RTrans: "Ψ  ΨQ  R  ¡M⦇ν*xvec⦈⟨N  R'"
  and   FrR: "extractFrame R = AR, ΨR"
  and   "distinct AR"
  and   QimpP: "AQ, (Ψ  ΨQ)  ΨR F AP, (Ψ  ΨP)  ΨR"
  and   "AR ♯* AP"
  and   "AR ♯* AQ"
  and   "AR ♯* Ψ"
  and   "AR ♯* ΨP"
  and   "AR ♯* ΨQ"
  and   "AR ♯* R"
  and   "AR ♯* M"
  and   "AP ♯* R"
  and   "AP ♯* M"
  and   "AQ ♯* R"
  and   "AQ ♯* M"

shows "Ψ  ΨP  R  ¡M⦇ν*xvec⦈⟨N  R'"
proof -
  from RTrans have "Ψ  ΨQ  R  RBrOut M (⦇ν*xvecN ≺' R')"
    by(simp add: residualInject)
  moreover from AR ♯* Ψ AR ♯* ΨQ have "AR ♯* (Ψ  ΨQ)"
    by force
  ultimately show ?thesis using assms
    by(elim brCommOutAuxTooMuch)
qed

lemma comm1Aux:
  fixes Ψ    :: 'b
    and Ψ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 AR   :: "name list"
    and ΨR   :: 'b
    and P    :: "('a, 'b, 'c) psi"
    and M    :: 'a
    and L    :: 'a
    and P'   :: "('a, 'b, 'c) psi"
    and AP   :: "name list"
    and ΨP   :: 'b
    and AQ   :: "name list"

assumes RTrans: "Ψ  ΨQ  R K⦇ν*xvec⦈⟨N  R'"
  and   FrR: "extractFrame R = AR, ΨR"
  and   PTrans: "Ψ  ΨR  P ML  P'"
  and   MeqK: "Ψ  ΨQ  ΨR  M  K"
  and   PeqQ: "AQ, (Ψ  ΨQ)  ΨR F AP, (Ψ  ΨP)  ΨR"
  and   FrP: "extractFrame P = AP, ΨP"
  and   FrQ: "extractFrame Q = AQ, ΨQ"
  and   "distinct AP"
  and   "distinct AR"
  and   "AR ♯* AP"
  and   "AR ♯* AQ"
  and   "AR ♯* Ψ"
  and   "AR ♯* P"
  and   "AR ♯* Q"
  and   "AR ♯* R"
  and   "AR ♯* K"
  and   "AP ♯* Ψ"
  and   "AP ♯* R"
  and   "AP ♯* P"
  and   "AP ♯* M"
  and   "AQ ♯* R"
  and   "AQ ♯* M"

obtains K' where "Ψ  ΨP  R K'⦇ν*xvec⦈⟨N  R'" and "Ψ  ΨP  ΨR  M  K'" and "AR ♯* K'"
proof -
  from AR ♯* P AR ♯* Q AR ♯* AP AR ♯* AQ FrP FrQ have "AR ♯* ΨP" and "AR ♯* ΨQ"
    by(force dest: extractFrameFreshChain)+
  assume Assumptions: "K'. Ψ  ΨP  R K'⦇ν*xvec⦈⟨N  R'; Ψ  ΨP  ΨR  M  K'; AR ♯* K'  thesis"
  {
    fix Ψ'   ::'b
      and zvec ::"name list"

    assume A: "Ψ  ΨQ  Ψ'"
    assume "Ψ'  R K⦇ν*xvec⦈⟨N  R'"
    then have "Ψ'  R ROut K (⦇ν*xvecN ≺' R')" by(simp add: residualInject)
    moreover note FrR distinct AR PTrans
    moreover from Ψ'  R K⦇ν*xvec⦈⟨N  R' have "distinct xvec" by(auto dest: boundOutputDistinct)
    moreover assume "Ψ'  ΨR  M  K" and "AQ, Ψ'  ΨR F AP, (Ψ  ΨP)  ΨR"
      and "AR ♯* zvec" and "AP ♯* zvec" and "zvec ♯* R" and "zvec ♯* P"
      and "AR ♯* Ψ'"
    ultimately have "K'. Ψ  ΨP  R ROut K' (⦇ν*xvecN ≺' R')  Ψ  ΨP  ΨR  M  K'  zvec ♯* K'  AR ♯* K'"
      using FrP AP ♯* R AQ ♯* R AP ♯* Ψ AP ♯* P distinct AP  AR ♯* Ψ AR ♯* P AR ♯* R
        AP ♯* M AQ ♯* M AR ♯* K AR ♯* AP AR ♯* AQ AR ♯* ΨP
    proof(nominal_induct Ψ' R K B=="⦇ν*xvecN ≺' R'" AR ΨR avoiding: Ψ P AP ΨP AQ zvec xvec N R' arbitrary: M rule: outputFrameInduct)
      case(cAlpha Ψ' R K AR ΨR p Ψ P AP ΨP AQ zvec xvec N R' M)
      have S: "set p  set AR × set (p  AR)" by fact
      from Ψ'  (p  ΨR)  M  K have "(p  (Ψ'  (p  ΨR)))  (p  M)  (p  K)"
        by(rule chanEqClosed)
      with AR ♯* Ψ' (p  AR) ♯* Ψ' AR ♯* K (p  AR) ♯* K S distinctPerm p
      have "Ψ'  ΨR  (p  M)  K" by(simp add: eqvts)
      moreover from Ψ  (p  ΨR)  P ML  P' S AR ♯* P (p  AR) ♯* P have "(p  (Ψ  (p  ΨR)))  P (p  M)L  P'"
        by(elim inputPermFrameSubject) auto
      with AR ♯* Ψ (p  AR) ♯* Ψ S distinctPerm p have "Ψ  ΨR  P (p  M)L  P'"
        by(simp add: eqvts)
      moreover from AP ♯* M have "(p  AP) ♯* (p  M)"
        by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
      with AR ♯* AP (p  AR) ♯* AP S have "AP ♯* (p  M)" by simp
      moreover from AQ ♯* M have "(p  AQ) ♯* (p  M)"
        by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
      with AR ♯* AQ (p  AR) ♯* AQ S have "AQ ♯* (p  M)" by simp
      moreover from AQ, Ψ'  (p  ΨR) F AP, (Ψ  ΨP)  (p  ΨR)
      have "(p  AQ, Ψ'  (p  ΨR)) F (p  AP, (Ψ  ΨP)  (p  ΨR))"
        by(rule FrameStatImpClosed)
      with AR ♯* AP (p  AR) ♯* AP AR ♯* Ψ' (p  AR) ♯* Ψ' AR ♯* ΨP (p  AR) ♯* ΨP AR ♯* AQ
        (p  AR) ♯* AQ AR ♯* Ψ (p  AR) ♯* Ψ S distinctPerm p
      have "AQ, Ψ'  ΨR F AP, (Ψ  ΨP)  ΨR" by(simp add: eqvts)
      ultimately obtain K' where "Ψ  ΨP  R ROut K' (⦇ν*xvecN ≺' R')" and "Ψ  ΨP  ΨR  (p  M)  K'" and "zvec ♯* K'" and "AR ♯* K'"
        using cAlpha
        by metis
      from Ψ  ΨP  R ROut K' (⦇ν*xvecN ≺' R') S AR ♯* R (p  AR) ♯* R
      have "(p  (Ψ  ΨP))  R (ROut (p  K') (⦇ν*xvecN ≺' R'))" using outputPermFrameSubject
        by(auto simp add: residualInject)
      with S AR ♯* Ψ (p  AR) ♯* Ψ AR ♯* ΨP (p  AR) ♯* ΨP have "Ψ  ΨP  R (ROut (p  K') (⦇ν*xvecN ≺' R'))"
        by(simp add: eqvts)
      moreover from Ψ  ΨP  ΨR  (p  M)  K' have "(p  (Ψ  ΨP  ΨR)) (p  p  M)  (p  K')"
        by(rule chanEqClosed)
      with S AR ♯* Ψ (p  AR) ♯* Ψ AR ♯* ΨP (p  AR) ♯* ΨP distinctPerm p have "Ψ  ΨP  (p  Ψ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 AR ♯* zvec (p  AR) ♯* zvec S have "zvec ♯* (p  K')" by simp
      moreover from AR ♯* K' have "(p  AR) ♯* (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 AP ΨP AQ zvec xvec N' R' M)
      from AP ♯* (M'N⟩.R) AQ ♯* (M'N⟩.R) zvec ♯* (M'N⟩.R)
      have "AP ♯* M'" and "AQ ♯* 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 AQ ♯* M' have "(AQ, Ψ'  𝟭) F M'  M'" by(force intro: frameImpI)

      with AQ, Ψ'  𝟭 F AP, (Ψ  ΨP)  𝟭 have "(AP, (Ψ  ΨP)  𝟭) F M'  M'"
        by(simp add: FrameStatImp_def)
      with AP ♯* M' have "(Ψ  ΨP)  𝟭  M'  M'"  by(force dest: frameImpE)
      then have "Ψ  ΨP  M'  M'" by(blast intro: statEqEnt Identity) then have "Ψ  ΨP  M'N⟩.R M'N  R"
        by(rule Output)

      moreover from Ψ'  𝟭  M  K Ψ'  𝟭  M'  K
      have "Ψ'  𝟭  M  M'" by(metis chanEqSym chanEqTrans)
      with AQ ♯* M AQ ♯* M'
      have "(AQ, Ψ'  𝟭) F M  M'"
        by(force intro: frameImpI)
      with AQ, Ψ'  𝟭 F AP, (Ψ  ΨP)  𝟭
      have "(AP, (Ψ  ΨP)  𝟭) F M  M'"
        by(simp add: FrameStatImp_def)
      with AP ♯* M AP ♯* M' have "(Ψ  ΨP)  𝟭  M  M'"
        by(force dest: frameImpE)
      then have "Ψ  ΨP  𝟭  M  M'"
        by(metis statEqEnt Associativity)
      ultimately show ?case using cOutput by(auto simp add: residualInject)
    next
      case(cCase Ψ' R M' φ Cs AR ΨR Ψ P AP ΨP AQ zvec xvec N R' M)
      from guarded R extractFrame R = AR, ΨR have "ΨR  𝟭"
        by(metis guardedStatEq)
      with Ψ'  𝟭  M  M' have "Ψ'  ΨR  M  M'"
        by(metis Identity Commutativity statEqEnt AssertionStatEqSym Composition)
      moreover have "AQ, Ψ'  ΨR F AP, (Ψ  ΨP)  ΨR"
      proof -
        from ΨR  𝟭 have "AQ,  Ψ'  ΨR F AQ, Ψ'  𝟭"
          by(metis Identity Commutativity AssertionStatEqSym Composition frameResChainPres frameNilStatEq AssertionStatEqTrans)
        moreover note AQ, Ψ'  𝟭 F AP, (Ψ  ΨP)  𝟭
        moreover from ΨR  𝟭 have "AP, (Ψ  ΨP)  𝟭 F AP, (Ψ  ΨP)  ΨR"
          by(metis Identity Commutativity AssertionStatEqSym Composition frameResChainPres frameNilStatEq AssertionStatEqTrans)
        ultimately show ?thesis by(rule FrameStatEqImpCompose)
      qed
      moreover from Ψ  𝟭  P ML  P' ΨR  𝟭
      have "Ψ  ΨR   P ML  P'" by(metis statEqTransition Identity Commutativity AssertionStatEqSym Composition)
      moreover from zvec ♯* (Cases Cs) AP ♯* (Cases Cs)  AQ ♯* (Cases Cs) (φ, R)  set Cs
      have "AP ♯* R" and "AQ ♯* R" and "zvec ♯* R" and "AP ♯* φ" and "AQ ♯* φ"
        by(auto dest: memFreshChain)
      ultimately have "K'. Ψ  ΨP  R ROut K' (⦇ν*xvecN ≺' R')  Ψ  ΨP  ΨR  M  K'  zvec ♯* K'  AR ♯* K'"  using cCase
        by(intro cCase) (assumption | simp)+
      then obtain K' where RTrans: "Ψ  ΨP  R ROut K' (⦇ν*xvecN ≺' R')"
        and MeqK': "Ψ  ΨP  ΨR  M  K'" and "zvec ♯* K'" and "AR ♯* K'"
        by metis
      note RTrans (φ, R)  set Cs
      moreover from Ψ'  φ have "Ψ'  𝟭  φ" by(blast intro: statEqEnt Identity AssertionStatEqSym)
      with AQ ♯* φ have "(AQ, Ψ'  𝟭) F φ" by(force intro: frameImpI)
      with AQ, Ψ'  𝟭 F AP, (Ψ  ΨP)  𝟭 have "(AP, (Ψ  ΨP)  𝟭) F φ"
        by(simp add: FrameStatImp_def)
      with AP ♯* φ have "(Ψ  ΨP)  𝟭  φ"  by(force dest: frameImpE)
      then have "Ψ  ΨP  φ" by(blast intro: statEqEnt Identity)
      ultimately have "Ψ  ΨP  Cases Cs ROut K' (⦇ν*xvecN ≺' R')" using guarded R by(rule Case)
      moreover from MeqK' ΨR  𝟭 have "Ψ  ΨP  𝟭  M  K'"
        by(metis Identity Commutativity statEqEnt AssertionStatEqSym Composition AssertionStatEqTrans)
      ultimately show ?case using zvec ♯* K'
        by auto
    next
      case(cPar1 Ψ' ΨR2 R1 M' xvec N' R1' AR2 R2 AR1 ΨR1 Ψ P AP ΨP AQ zvec yvec N R' M)
      have FrR2: "extractFrame R2 = AR2, ΨR2" by fact
      from Ψ'  ΨR1  ΨR2  M  M' have "(Ψ'  ΨR2)  ΨR1  M  M'"
        by(metis statEqEnt Associativity Composition Commutativity)
      moreover have "AQ, (Ψ'  ΨR2)  ΨR1 F AP, ((Ψ  ΨR2)  ΨP)  ΨR1"
      proof -
        have "AQ, (Ψ'  ΨR2)  ΨR1 F AQ, Ψ'  ΨR1  ΨR2"
          by(metis Associativity Composition Commutativity AssertionStatEqTrans AssertionStatEqSym frameNilStatEq frameResChainPres)
        moreover note AQ, Ψ'  ΨR1  ΨR2 F AP, (Ψ  ΨP)  ΨR1  ΨR2
        moreover have  "AP, (Ψ  ΨP)  ΨR1  ΨR2 F AP, ((Ψ  ΨR2)  ΨP)  ΨR1"
          by(metis Associativity Composition Commutativity AssertionStatEqTrans AssertionStatEqSym frameNilStatEq frameResChainPres)
        ultimately show ?thesis by(rule FrameStatEqImpCompose)
      qed
      moreover from Ψ  ΨR1  ΨR2  P ML  P' have "(Ψ  ΨR2)  ΨR1  P ML  P'"
        by(metis statEqTransition Associativity Composition Commutativity)
      moreover from AR1 ♯* AP AR2 ♯* AP AP ♯* (R1  R2) extractFrame R1 = AR1, ΨR1 FrR2 have "AP ♯* ΨR1" and "AP ♯* ΨR2"
        by(force dest: extractFrameFreshChain)+

      moreover from ⦇ν*xvecN' ≺' (R1'  R2) = ⦇ν*yvecN ≺' R' xvec ♯* yvec
      obtain p T where "⦇ν*xvecN' ≺' R1' = ⦇ν*yvecN ≺' T" and "R' = T  (p  R2)" and "set p  set yvec × set xvec"
        apply(drule_tac sym)
        by(elim boundOutputPar1Dest') (assumption | simp | blast dest: sym)+
      ultimately have "K'. (Ψ  ΨR2)  ΨP  R1 ROut K' (⦇ν*yvecN ≺' T)  (Ψ  ΨR2)  ΨP  ΨR1  M  K'  (AR2@zvec) ♯* K'  AR1 ♯* K'" using cPar1
        by(elim cPar1(6)[where ba=P and bb=AP and bd=AQ]) auto
      then  obtain K' where RTrans: "(Ψ  ΨR2)  ΨP  R1 K'⦇ν*xvec⦈⟨N'  R1'"
        and MeqK': "(Ψ  ΨR2)  ΨP  ΨR1  M  K'"  and "AR2 ♯* K'" and "AR1 ♯* K'" and "zvec ♯* K'"
        using ⦇ν*xvecN' ≺' R1' = ⦇ν*yvecN ≺' T by(auto simp add: residualInject)

      from RTrans have "(Ψ  ΨP)  ΨR2  R1 K'⦇ν*xvec⦈⟨N'  R1'"
        by(metis statEqTransition Associativity Composition Commutativity)
      then have "Ψ  ΨP  (R1  R2) K'⦇ν*xvec⦈⟨N'  (R1'  R2)" using FrR2 xvec ♯* R2 AR2 ♯* Ψ AR2 ♯* ΨP AR2 ♯* K' AR2 ♯* R1 AR2 ♯* xvec AR2 ♯* N'
        by(force intro: Par1)
      moreover from MeqK' have "Ψ  ΨP  ΨR1  ΨR2  M  K'"
        by(metis statEqEnt Associativity Composition Commutativity)
      ultimately show ?case using zvec ♯* K' AR1 ♯* K' AR2 ♯* K'  ⦇ν*xvecN' ≺' (R1'  R2) = ⦇ν*yvecN ≺' R'
        by(auto simp add: residualInject)
    next
      case(cPar2 Ψ' ΨR1 R2 M' xvec N' R2' AR1 R1 AR2 ΨR2 Ψ P AP ΨP AQ zvec yvec N R' M)
      have FrR1: "extractFrame R1 = AR1, ΨR1" by fact
      from Ψ'  ΨR1  ΨR2  M  M' have "(Ψ'  ΨR1)  ΨR2  M  M'"
        by(metis statEqEnt Associativity Composition Commutativity)
      moreover have "AQ, (Ψ'  ΨR1)  ΨR2 F AP, ((Ψ  ΨR1)  ΨP)  ΨR2"
      proof -
        have "AQ, (Ψ'  ΨR1)  ΨR2 F AQ, Ψ'  ΨR1  ΨR2"
          by(metis Associativity Composition Commutativity AssertionStatEqTrans AssertionStatEqSym frameNilStatEq frameResChainPres)
        moreover note AQ, Ψ'  ΨR1  ΨR2 F AP, (Ψ  ΨP)  ΨR1  ΨR2
        moreover have "AP, (Ψ  ΨP)  ΨR1  ΨR2 F AP, ((Ψ  ΨR1)  ΨP)  ΨR2"
          by(metis Associativity Composition Commutativity AssertionStatEqTrans AssertionStatEqSym frameNilStatEq frameResChainPres)
        ultimately show ?thesis by(rule FrameStatEqImpCompose)
      qed
      moreover from Ψ  ΨR1  ΨR2  P ML  P' have "(Ψ  ΨR1)  ΨR2  P ML  P'"
        by(metis statEqTransition Associativity Composition Commutativity)
      moreover from AR1 ♯* AP AR2 ♯* AP AP ♯* (R1  R2) FrR1 extractFrame R2 = AR2, ΨR2 have "AP ♯* ΨR1" and "AP ♯* ΨR2"
        by(force dest: extractFrameFreshChain)+
      moreover from ⦇ν*xvecN' ≺' (R1  R2') = ⦇ν*yvecN ≺' R' xvec ♯* yvec
      obtain p T where "⦇ν*xvecN' ≺' R2' = ⦇ν*yvecN ≺' T" and "R' = (p  R1)  T" and "set p  set yvec × set xvec"
        apply(drule_tac sym)
        by(elim boundOutputPar2Dest') (assumption | simp | blast dest: sym)+
      ultimately have "K'. (Ψ  ΨR1)  ΨP  R2 ROut K' (⦇ν*yvecN ≺' T)  (Ψ  ΨR1)  ΨP  ΨR2  M  K'  (AR1@zvec) ♯* K'  AR2 ♯* K'" using cPar2
        by(elim cPar2(6)) (assumption | simp | auto)+
      then obtain K' where RTrans: "(Ψ  ΨR1)  ΨP  R2 K'⦇ν*xvec⦈⟨N'  R2'"
        and MeqK': "(Ψ  ΨR1)  ΨP  ΨR2  M  K'"  and "AR1 ♯* K'" and "zvec ♯* K'" and "AR2 ♯* K'"
        using ⦇ν*xvecN' ≺' R2' = ⦇ν*yvecN ≺' T by(auto simp add: residualInject)

      from RTrans have "(Ψ  ΨP)  ΨR1  R2 K'⦇ν*xvec⦈⟨N'  R2'"
        by(metis statEqTransition Associativity Composition Commutativity)
      then have "Ψ  ΨP  (R1  R2) K'⦇ν*xvec⦈⟨N'  (R1  R2')" using FrR1 xvec ♯* R1 AR1 ♯* Ψ AR1 ♯* ΨP AR1 ♯* K'AR1 ♯* xvec AR1 ♯* N' AR1 ♯* R2
        by(force intro: Par2)
      moreover from MeqK' have "Ψ  ΨP  ΨR1  ΨR2  M  K'"
        by(metis statEqEnt Associativity Composition Commutativity)
      ultimately show ?case using zvec ♯* K' AR1 ♯* K' AR2 ♯* K' ⦇ν*xvecN' ≺' (R1  R2') = ⦇ν*yvecN ≺' R'
        by(auto simp add: residualInject)
    next
      case(cOpen Ψ' R M' xvec yvec N' R' x AR ΨR Ψ P AP ΨP AQ zvec zvec2 N R'' M)
      from ⦇ν*(xvec @ x # yvec)N' ≺' R' = ⦇ν*zvec2N ≺' 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'. Ψ  ΨP  R ROut K' (⦇ν*(xvec'@yvec')([(x, x')]  N) ≺' ([(x, x')]  R''))  Ψ  ΨP  ΨR  M  K'  (zvec) ♯* K'  AR ♯* K'" using cOpen
        by(elim cOpen(4)) (assumption | simp)+
      then  obtain K' where RTrans: "Ψ  ΨP  R K'⦇ν*(xvec@yvec)⦈⟨N'  R'"
        and MeqK': "Ψ  ΨP  ΨR  M  K'" and "zvec ♯* K'" and "AR ♯* K'"
        using A by(auto simp add: residualInject)
      from AR ♯* AP AP ♯* (⦇νxR) x  AP extractFrame R = AR, ΨR have "AP ♯* ΨR"
        by(force dest: extractFrameFreshChain)+

      from Ψ  ΨR  P ML  P' extractFrame P = AP, ΨP distinct AP zvec ♯* P AP ♯* ΨR x  AP AP ♯* M AP ♯* P AP ♯* zvec AP ♯* Ψ AP ♯* zvec AR ♯* P AR ♯* AP x  AP x  P

      obtain K'' where MeqK'': "(Ψ  ΨR)  ΨP  M  K''" and "AR ♯* K''" and "zvec ♯* K''" and "x  K''"
        by(elim inputObtainPrefix[where B="(x#AR@zvec)"]) (assumption | simp | force)+

      from MeqK'' MeqK' have KeqK'': "(Ψ  ΨP)  ΨR  K'  K''"
        by(metis statEqEnt Associativity Composition Commutativity chanEqSym chanEqTrans)
      with RTrans extractFrame R = AR, ΨR distinct AR AR ♯* Ψ AR ♯* ΨP AR ♯* K' AR ♯* K'' AR ♯* R
      have "Ψ  ΨP  R K''⦇ν*(xvec@yvec)⦈⟨N'  R'"
        by(elim outputRenameSubject) (assumption | force)+
      then have "Ψ  ΨP  ⦇νxR K''⦇ν*(xvec@x#yvec)⦈⟨N'  R'"
        using x  Ψ x  ΨP x  K'' x  xvec x  yvec x  supp N' xvec ♯* Ψ xvec ♯* ΨP xvec ♯* R x  K''
        by(elim Open) (assumption | force)+
      moreover from MeqK'' have "Ψ  ΨP  ΨR  M  K''"
        by(metis statEqEnt Associativity Composition Commutativity)
      ultimately show ?case using zvec ♯* K'' x  K'' AR ♯* K'' B ⦇ν*(xvec @ x # yvec)N' ≺' R' = ⦇ν*zvec2N ≺' R''
        by(auto simp add: residualInject)
    next
      case(cScope Ψ' R M' xvec N' R' x AR ΨR Ψ P AP ΨP AQ zvec yvec N R'' M)
      from ⦇ν*xvecN' ≺' ⦇νxR' = ⦇ν*yvecN ≺' R'' x  xvec x  yvec
      obtain R''' where "R'' = ⦇νxR'''" and "⦇ν*xvecN' ≺' R' = ⦇ν*yvecN ≺' R'''"
        apply(drule_tac sym)
        by(metis boundOutputScopeDest)
      then have "K'. Ψ  ΨP  R ROut K' (⦇ν*yvecN ≺' R''')  Ψ  ΨP  ΨR  M  K'  zvec ♯* K'  AR ♯* K'"  using cScope
        by(elim cScope(4)) (assumption | simp)+
      then obtain K' where RTrans: "Ψ  ΨP  R K'⦇ν*xvec⦈⟨N'  R'"
        and MeqK': "Ψ  ΨP  ΨR  M  K'" and "zvec ♯* K'" and "AR ♯* K'"
        using ⦇ν*xvecN' ≺' R' = ⦇ν*yvecN ≺' R'''
        by(auto simp add: residualInject)
      from AR ♯* AP AP ♯* (⦇νxR) x  AP extractFrame R = AR, ΨR have "AP ♯* ΨR"
        by(force dest: extractFrameFreshChain)+
      from Ψ  ΨR  P ML  P' extractFrame P = AP, ΨP distinct AP x  P zvec ♯* P AP ♯* ΨR x  AP AP ♯* M AP ♯* P AP ♯* zvec  AP ♯* Ψ AP ♯* zvec AR ♯* P AR ♯* AP
      obtain K'' where MeqK'': "(Ψ  ΨR)  ΨP  M  K''" and "x  K''" and "AR ♯* K''" and "zvec ♯* K''"
        by(elim inputObtainPrefix[where B="(x#AR@zvec)"]) (assumption | force)+

      from MeqK'' MeqK' have KeqK'': "(Ψ  ΨP)  ΨR  K'  K''"
        by(metis statEqEnt Associativity Composition Commutativity chanEqSym chanEqTrans)
      with RTrans extractFrame R = AR, ΨR distinct AR AR ♯* Ψ AR ♯* ΨP AR ♯* K' AR ♯* K'' AR ♯* R
      have "Ψ  ΨP  R K''⦇ν*xvec⦈⟨N'  R'"
        by(elim outputRenameSubject) (assumption | force)+
      then have "Ψ  ΨP  ⦇νxR K''⦇ν*xvec⦈⟨N'  ⦇νxR'" using x  Ψ x  ΨP x  K'' x  xvec x  N' xvec ♯* Ψ xvec ♯* ΨP xvec ♯* R x  K''
        by(elim Scope) (assumption | force)+
      moreover from MeqK'' have "Ψ  ΨP  ΨR  M  K''"
        by(metis statEqEnt Associativity Composition Commutativity)
      ultimately show ?case using zvec ♯* K'' x  K'' AR ♯* K'' ⦇ν*xvecN' ≺' ⦇νxR' = ⦇ν*yvecN ≺' R''
        by(auto simp add: residualInject)
    next
      case(cBang Ψ' R M' AR ΨR Ψ P AP ΨP AQ zvec xvec N R' M)
      from guarded R extractFrame R = AR, ΨR have "ΨR  𝟭"
        by(metis guardedStatEq)
      with Ψ'  𝟭  M  M' have "Ψ'  ΨR  𝟭  M  M'"
        by(metis Identity Commutativity statEqEnt AssertionStatEqSym Composition)
      moreover have "AQ, Ψ'  ΨR  𝟭 F AP, (Ψ  ΨP)  ΨR  𝟭"
      proof -
        from ΨR  𝟭 have "AQ,  Ψ'  ΨR  𝟭 F AQ, Ψ'  𝟭"
          by(metis Identity Commutativity AssertionStatEqSym Composition frameResChainPres frameNilStatEq AssertionStatEqTrans)
        moreover note AQ, Ψ'  𝟭 F AP, (Ψ  ΨP)  𝟭
        moreover from ΨR  𝟭 have "AP, (Ψ  ΨP)  𝟭 F AP, (Ψ  ΨP)  ΨR  𝟭"
          by(metis Identity Commutativity AssertionStatEqSym Composition frameResChainPres frameNilStatEq AssertionStatEqTrans)
        ultimately show ?thesis by(rule FrameStatEqImpCompose)
      qed
      moreover from Ψ  𝟭  P ML  P' ΨR  𝟭
      have "Ψ  ΨR  𝟭  P ML  P'" by(metis statEqTransition Identity Commutativity AssertionStatEqSym Composition)
      ultimately have "K'. Ψ  ΨP  R  !R ROut K' (⦇ν*xvecN ≺' R')  Ψ  ΨP  ΨR  𝟭  M  K'  zvec ♯* K'  AR ♯* K'"  using cBang
        by(intro cBang(5)) (assumption |simp)+
      then  obtain K' where RTrans: "Ψ  ΨP  R  !R ROut K' (⦇ν*xvecN ≺' R')"
        and MeqK': "Ψ  ΨP  ΨR  𝟭  M  K'" and "zvec ♯* K'" and "AR ♯* K'"
        by metis
      from RTrans guarded R have "Ψ  ΨP  !R ROut K' (⦇ν*xvecN ≺' R')" by(rule Bang)
      moreover from MeqK' ΨR  𝟭 have "Ψ  ΨP  𝟭  M  K'"
        by(metis Identity Commutativity statEqEnt AssertionStatEqSym Composition AssertionStatEqTrans)
      ultimately show ?case using zvec ♯* K'
        by force
    qed
  }
  note Goal = this
  have "Ψ  ΨQ  Ψ  ΨQ" by simp
  moreover note RTrans
  moreover from MeqK have "(Ψ  ΨQ)  ΨR  M  K"
    by(metis statEqEnt Associativity Commutativity)
  moreover note PeqQ AR ♯* Ψ AR ♯* ΨQ
  ultimately have "K'. Ψ  ΨP  R ROut K' (⦇ν*xvecN ≺' R')  Ψ  ΨP  ΨR  M  K'  ([]::name list) ♯* K'  AR ♯* K'"
    by(elim Goal) (assumption | force simp add: residualInject)+
  with Assumptions show ?thesis
    by(force simp add: residualInject)
qed

lemma comm2Aux:
  fixes Ψ    :: 'b
    and ΨQ   :: 'b
    and R    :: "('a, 'b, 'c) psi"
    and K    :: 'a
    and N    :: 'a
    and R'   :: "('a, 'b, 'c) psi"
    and AR   :: "name list"
    and ΨR   :: 'b
    and P    :: "('a, 'b, 'c) psi"
    and M    :: 'a
    and xvec :: "name list"
    and P'   :: "('a, 'b, 'c) psi"
    and AP   :: "name list"
    and ΨP   :: 'b
    and AQ   :: "name list"

assumes RTrans: "Ψ  ΨQ  R KN  R'"
  and   FrR: "extractFrame R = AR, ΨR"
  and   PTrans: "Ψ  ΨR  P M⦇ν*xvec⦈⟨N  P'"
  and   MeqK: "Ψ  ΨQ  ΨR  M  K"
  and   QimpP: "AQ, (Ψ  ΨQ)  ΨR F AP, (Ψ  ΨP)  ΨR"
  and   FrP: "extractFrame P = AP, ΨP"
  and   FrQ: "extractFrame Q = AQ, ΨQ"
  and   "distinct AP"
  and   "distinct AR"
  and   "AR ♯* AP"
  and   "AR ♯* AQ"
  and   "AR ♯* Ψ"
  and   "AR ♯* P"
  and   "AR ♯* Q"
  and   "AR ♯* R"
  and   "AR ♯* K"
  and   "AP ♯* Ψ"
  and   "AP ♯* R"
  and   "AP ♯* P"
  and   "AP ♯* M"
  and   "AQ ♯* R"
  and   "AQ ♯* M"
  and   "AR ♯* xvec"
  and   "xvec ♯* M"

obtains  K' where "Ψ  ΨP  R K'N  R'" and "Ψ  ΨP  ΨR  M  K'" and "AR ♯* K'"
proof -
  from AR ♯* P AR ♯* Q AR ♯* AP AR ♯* AQ FrP FrQ have "AR ♯* ΨP" and "AR ♯* ΨQ"
    by(force dest: extractFrameFreshChain)+
  assume Assumptions: "K'. Ψ  ΨP  R K'N  R'; Ψ  ΨP  ΨR  M  K'; AR ♯* K'  thesis"
  {
    fix Ψ'::'b
    fix zvec::"name list"
    assume "AR ♯* Ψ'"
    assume "AR ♯* zvec"
    assume "AP ♯* zvec"
    assume "zvec ♯* R"
    assume "zvec ♯* P"

    assume A: "Ψ  ΨQ  Ψ'"
    with RTrans have "Ψ'  R KN  R'"
      by(rule statEqTransition)
    moreover note FrR distinct AR
    moreover from Ψ  ΨQ  ΨR  M  K have "(Ψ  ΨQ)  ΨR  M  K"
      by(blast intro: statEqEnt Associativity AssertionStatEqSym)
    with A have "Ψ'  ΨR  M  K" by(rule statEqEnt[OF Composition])
    moreover have "AQ, Ψ'  ΨR F AQ, (Ψ  ΨQ)  ΨR" using A
      by(blast dest: frameIntComposition FrameStatEqTrans FrameStatEqSym)
    with QimpP have "AQ, Ψ'  ΨR F AP, (Ψ  ΨP)  ΨR"
      by(force intro: FrameStatEqImpCompose)
    moreover from PTrans have "distinct xvec" by(auto dest: boundOutputDistinct)
    ultimately have "K'. Ψ  ΨP  R K'N  R'  Ψ  ΨP  ΨR  M  K'  zvec ♯* K'  AR ♯* K'"
      using PTrans FrP AR ♯* K AR ♯* Ψ' AR ♯* Ψ AR ♯* P AR ♯* R AR ♯* ΨP AP ♯* R AQ ♯* R AP ♯* Ψ
        AP ♯* P AP ♯* M AP ♯* zvec AQ ♯* M AR ♯* zvec zvec ♯* R zvec ♯* P distinct AP
        AR ♯* AP AR ♯* AQ AR ♯* xvec xvec ♯* M
    proof(nominal_induct avoiding: Ψ P AP ΨP AQ zvec xvec arbitrary: M rule: inputFrameInduct)
      case(cAlpha Ψ' R K N R' AR ΨR p Ψ P AP ΨP AQ zvec xvec M)
      have S: "set p  set AR × set (p  AR)" by fact
      from Ψ'  (p  ΨR)  M  K have "(p  (Ψ'  (p  ΨR)))  (p  M)  (p  K)"
        by(rule chanEqClosed)
      with AR ♯* Ψ' (p  AR) ♯* Ψ' AR ♯* K (p  AR) ♯* K S distinctPerm p
      have "Ψ'  ΨR  (p  M)  K" by(simp add: eqvts)
      moreover from Ψ  (p  ΨR)  P M⦇ν*xvec⦈⟨N  P' S AR ♯* P (p  AR) ♯* P AR ♯* xvec (p  AR) ♯* xvec xvec ♯* M
      have "(p  (Ψ  (p  ΨR)))  P (p  M)⦇ν*xvec⦈⟨N  P'"
        using outputPermFrameSubject by(auto simp add: residualInject)
      with AR ♯* Ψ (p  AR) ♯* Ψ S distinctPerm p have "Ψ  ΨR  P (p  M)⦇ν*xvec⦈⟨N  P'"
        by(simp add: eqvts)
      moreover from AP ♯* M have "(p  AP) ♯* (p  M)"
        by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
      with AR ♯* AP (p  AR) ♯* AP S have "AP ♯* (p  M)" by simp
      moreover from AQ ♯* M have "(p  AQ) ♯* (p  M)"
        by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst])
      with AR ♯* AQ (p  AR) ♯* AQ S have "AQ ♯* (p  M)" by simp

      moreover from AQ, Ψ'  (p  ΨR) F AP, (Ψ  ΨP)  (p  ΨR)
      have "(p  AQ, Ψ'  (p  ΨR)) F (p  AP, (Ψ  ΨP)  (p  ΨR))"
        by(rule FrameStatImpClosed)
      with AR ♯* AP (p  AR) ♯* AP AR ♯* Ψ' (p  AR) ♯* Ψ' AR ♯* ΨP (p  AR) ♯* ΨP AR ♯* AQ
        (p  AR) ♯* AQ AR ♯* Ψ (p  AR) ♯* Ψ S distinctPerm p
      have "AQ, Ψ'  ΨR F AP, (Ψ  ΨP)  Ψ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 AR ♯* xvec (p  AR) ♯* xvec have "xvec ♯* (p  M)" by simp
      ultimately obtain K' where "Ψ  ΨP  R K'N  R'" and "Ψ  ΨP  ΨR  (p  M)  K'" and "zvec ♯* K'" and "AR ♯* K'"
        using cAlpha
        by(auto simp del: freshChainSimps)
      from Ψ  ΨP  R K'N  R' S AR ♯* R (p  AR) ♯* R have "(p  (Ψ  ΨP))  R (p  K')N  R'"
        by(elim inputPermFrameSubject) auto
      with S AR ♯* Ψ (p  AR) ♯* Ψ AR ♯* ΨP (p  AR) ♯* ΨP have "Ψ  ΨP  R (p  K')N  R'"
        by(simp add: eqvts)
      moreover from Ψ  ΨP  ΨR  (p  M)  K' have "(p  (Ψ  ΨP  ΨR)) (p  p  M)  (p  K')"
        by(rule chanEqClosed)
      with S AR ♯* Ψ (p  AR) ♯* Ψ AR ♯* ΨP (p  AR) ♯* ΨP distinctPerm p have "Ψ  ΨP  (p  Ψ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 AR ♯* zvec (p  AR) ♯* zvec S have "zvec ♯* (p  K')" by simp
      moreover from AR ♯* K' have "(p  AR) ♯* (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 AP ΨP AQ zvec yvec M)
      from AP ♯* (M'⦇λ*xvec N⦈.R) AQ ♯* (M'⦇λ*xvec N⦈.R) zvec ♯* (M'⦇λ*xvec N⦈.R)
      have "AP ♯* M'" and "AQ ♯* 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 AQ ♯* M'
      have "(AQ, Ψ'  𝟭) F M'  M'"
        by(force intro: frameImpI)

      with AQ, Ψ'  𝟭 F AP, (Ψ  ΨP)  𝟭
      have "(AP, (Ψ  ΨP)  𝟭) F M'  M'"
        by(simp add: FrameStatImp_def)
      with AP ♯* M' have "(Ψ  ΨP)  𝟭  M'  M'"
        by(force dest: frameImpE)
      then have "Ψ  ΨP  M'  M'" by(blast intro: statEqEnt Identity)
      then have "Ψ  Ψ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 AQ ♯* M AQ ♯* M'
      have "(AQ, Ψ'  𝟭) F M  M'"
        by(force intro: frameImpI)
      with AQ, Ψ'  𝟭 F AP, (Ψ  ΨP)  𝟭
      have "(AP, (Ψ  ΨP)  𝟭) F M  M'"
        by(simp add: FrameStatImp_def)
      with AP ♯* M AP ♯* M' have "(Ψ  ΨP)  𝟭  M  M'"
        by(force dest: frameImpE)
      then have "Ψ  ΨP  𝟭  M  M'"
        by(metis statEqEnt Associativity)
      ultimately show ?case using zvec ♯* M'
        by force
    next
      case(cCase Ψ' R M' N R' φ Cs AR ΨR Ψ P AP ΨP AQ zvec xvec M)
      from guarded R extractFrame R = AR, ΨR have "ΨR  𝟭"
        by(metis guardedStatEq)
      with Ψ'  𝟭  M  M' have "Ψ'  ΨR  M  M'"
        by(metis Identity Commutativity statEqEnt AssertionStatEqSym Composition)
      moreover have "AQ, Ψ'  ΨR F AP, (Ψ  ΨP)  ΨR"
      proof -
        from ΨR  𝟭 have "AQ,  Ψ'  ΨR F AQ, Ψ'  𝟭"
          by(metis Identity Commutativity AssertionStatEqSym Composition frameResChainPres frameNilStatEq AssertionStatEqTrans)
        moreover note AQ, Ψ'  𝟭 F AP, (Ψ  ΨP)  𝟭
        moreover from ΨR  𝟭 have "AP, (Ψ  ΨP)  𝟭 F AP, (Ψ  ΨP)  ΨR"
          by(metis Identity Commutativity AssertionStatEqSym Composition frameResChainPres frameNilStatEq AssertionStatEqTrans)
        ultimately show ?thesis by(rule FrameStatEqImpCompose)
      qed
      moreover from Ψ  𝟭  P M⦇ν*xvec⦈⟨N  P' ΨR  𝟭
      have "Ψ  ΨR   P M⦇ν*xvec⦈⟨N  P'" by(metis statEqTransition Identity Commutativity AssertionStatEqSym Composition)
      moreover from zvec ♯* (Cases Cs) AP ♯* (Cases Cs)  AQ ♯* (Cases Cs) (φ, R)  set Cs
      have "AP ♯* R" and "AQ ♯* R" and "zvec ♯* R" and "AP ♯* φ" and "AQ ♯* φ"
        by(auto dest: memFreshChain)
      ultimately have "K'. Ψ  ΨP  R K'N  R' Ψ  ΨP  ΨR  M  K'  zvec ♯* K'  AR ♯* K'"  using cCase
        by(elim cCase) (assumption |simp)+
      then  obtain K' where RTrans: "Ψ  ΨP  R K'N  R'"
        and MeqK': "Ψ  ΨP  ΨR  M  K'" and "zvec ♯* K'" and "AR ♯* K'"
        by metis
      note RTrans (φ, R)  set Cs
      moreover from Ψ'  φ have "Ψ'  𝟭  φ" by(blast intro: statEqEnt Identity AssertionStatEqSym)
      with AQ ♯* φ have "(AQ, Ψ'  𝟭) F φ" by(force intro: frameImpI)
      with AQ, Ψ'  𝟭 F AP, (Ψ  ΨP)  𝟭 have "(AP, (Ψ  ΨP)  𝟭) F φ"
        by(simp add: FrameStatImp_def)
      with AP ♯* φ have "(Ψ  ΨP)  𝟭  φ"  by(force dest: frameImpE)
      then have "Ψ  ΨP  φ" by(blast intro: statEqEnt Identity)
      ultimately have "Ψ  ΨP  Cases Cs K'N  R'"  using guarded R by(rule Case)
      moreover from MeqK' ΨR  𝟭 have "Ψ  ΨP  𝟭  M  K'"
        by(metis Identity Commutativity statEqEnt AssertionStatEqSym Composition AssertionStatEqTrans)
      ultimately show ?case using zvec ♯* K'
        by force
    next
      case(cPar1 Ψ' ΨR2 R1 M' N R1' AR2 R2 AR1 ΨR1 Ψ P AP ΨP AQ zvec xvec M)
      have FrR2: "extractFrame R2 = AR2, ΨR2" by fact
      from Ψ'  ΨR1  ΨR2  M  M' have "(Ψ'  ΨR2)  ΨR1  M  M'"
        by(metis statEqEnt Associativity Composition Commutativity)
      moreover have "AQ, (Ψ'  ΨR2)  ΨR1 F AP, ((Ψ  ΨR2)  ΨP)  ΨR1"
      proof -
        have "AQ, (Ψ'  ΨR2)  ΨR1 F AQ, Ψ'  ΨR1  ΨR2"
          by(metis Associativity Composition Commutativity AssertionStatEqTrans AssertionStatEqSym frameNilStatEq frameResChainPres)
        moreover note AQ, Ψ'  ΨR1  ΨR2 F AP, (Ψ  ΨP)  ΨR1  ΨR2
        moreover have  "AP, (Ψ  ΨP)  ΨR1  ΨR2 F AP, ((Ψ  ΨR2)  ΨP)  ΨR1"
          by(metis Associativity Composition Commutativity AssertionStatEqTrans AssertionStatEqSym frameNilStatEq frameResChainPres)
        ultimately show ?thesis by(rule FrameStatEqImpCompose)
      qed
      moreover from Ψ  ΨR1  ΨR2  P M⦇ν*xvec⦈⟨N  P' have "(Ψ  ΨR2)  ΨR1  P M⦇ν*xvec⦈⟨N  P'"
        by(metis statEqTransition Associativity Composition Commutativity)
      moreover from AR1 ♯* AP AR2 ♯* AP AP ♯* (R1  R2) extractFrame R1 = AR1, ΨR1 FrR2 have "AP ♯* ΨR1" and "AP ♯* ΨR2"
        by(force dest: extractFrameFreshChain)+
      moreover note distinct xvec

      ultimately have "K'. (Ψ  ΨR2)  ΨP  R1 K'N  R1'  (Ψ  ΨR2)  ΨP  ΨR1  M  K'  (AR2@zvec) ♯* K'  AR1 ♯* K'" using cPar1
        by(elim cPar1(6)[where ba=P and bb=AP and bd=AQ and bf=xvec]) auto
      then  obtain K' where RTrans: "(Ψ  ΨR2)  ΨP  R1 K'N  R1'"
        and MeqK': "(Ψ  ΨR2)  ΨP  ΨR1  M  K'"  and "AR2 ♯* K'" and "zvec ♯* K'" and "AR1 ♯* K'"
        by force

      from RTrans have "(Ψ  ΨP)  ΨR2  R1 K'N  R1'"
        by(metis statEqTransition Associativity Composition Commutativity)
      then have "Ψ  ΨP  (R1  R2) K'N  (R1'  R2)" using FrR2 AR2 ♯* Ψ AR2 ♯* ΨP AR2 ♯* K' AR2 ♯* R1 AR2 ♯* N
        by(force intro: Par1)
      moreover from MeqK' have "Ψ  ΨP  ΨR1  ΨR2  M  K'"
        by(metis statEqEnt Associativity Composition Commutativity)
      ultimately show ?case using zvec ♯* K' AR1 ♯* K' AR2 ♯* K'
        by force
    next
      case(cPar2 Ψ' ΨR1 R2 M' N R2' AR1 R1 AR2 ΨR2 Ψ P AP ΨP AQ zvec xvec M)
      have FrR1: "extractFrame R1 = AR1, ΨR1" by fact
      from Ψ'  ΨR1  ΨR2  M  M' have "(Ψ'  ΨR1)  ΨR2  M  M'"
        by(metis statEqEnt Associativity Composition Commutativity)
      moreover have "AQ, (Ψ'  ΨR1)  ΨR2 F AP, ((Ψ  ΨR1)  ΨP)  ΨR2"
      proof -
        have "AQ, (Ψ'  ΨR1)  ΨR2 F AQ, Ψ'  ΨR1  ΨR2"
          by(metis Associativity Composition Commutativity AssertionStatEqTrans AssertionStatEqSym frameNilStatEq frameResChainPres)
        moreover note AQ, Ψ'  ΨR1  ΨR2 F AP, (Ψ  ΨP)  ΨR1  ΨR2
        moreover have "AP, (Ψ  ΨP)  ΨR1  ΨR2 F AP, ((Ψ  ΨR1)  ΨP)  ΨR2"
          by(metis Associativity Composition Commutativity AssertionStatEqTrans AssertionStatEqSym frameNilStatEq frameResChainPres)
        ultimately show ?thesis by(rule FrameStatEqImpCompose)
      qed
      moreover from Ψ  ΨR1  ΨR2  P M⦇ν*xvec⦈⟨N  P' have "(Ψ  ΨR1)  ΨR2  P M⦇ν*xvec⦈⟨N  P'"
        by(metis statEqTransition Associativity Composition Commutativity)
      moreover from AR1 ♯* AP AR2 ♯* AP AP ♯* (R1  R2) FrR1 extractFrame R2 = AR2, ΨR2 have "AP ♯* ΨR1" and "AP ♯* ΨR2"
        by(force dest: extractFrameFreshChain)+
      ultimately have "K'. (Ψ  ΨR1)  ΨP  R2 K'N  R2'  (Ψ  ΨR1)  ΨP  ΨR2  M  K'  (AR1@zvec) ♯* K'  AR2 ♯* K'" using distinct xvec cPar2
        by(elim cPar2(6)[where ba=P and bb=AP and bd=AQ and bf=xvec]) auto
      then  obtain K' where RTrans: "(Ψ  ΨR1)  ΨP  R2 K'N  R2'"
        and MeqK': "(Ψ  ΨR1)  ΨP  ΨR2  M  K'"  and "AR1 ♯* K'" and "zvec ♯* K'" and "AR2 ♯* K'"
        by force

      from RTrans have "(Ψ  ΨP)  ΨR1  R2 K'N  R2'"
        by(metis statEqTransition Associativity Composition Commutativity)
      then have "Ψ  ΨP  (R1  R2) K'N  (R1  R2')" using FrR1 AR1 ♯* Ψ AR1 ♯* ΨP AR1 ♯* K' AR1 ♯* R2 AR1 ♯* N
        by(force intro: Par2)
      moreover from MeqK' have "Ψ  ΨP  ΨR1  ΨR2  M  K'"
        by(metis statEqEnt Associativity Composition Commutativity)
      ultimately show ?case using zvec ♯* K' AR1 ♯* K' AR2 ♯* K'
        by force
    next
      case(cScope Ψ' R M' N R' x AR ΨR Ψ P AP ΨP AQ zvec xvec M)
      then have "K'. Ψ  ΨP  R K'N  R'  Ψ  ΨP  ΨR  M  K'  zvec ♯* K'  AR ♯* K'"
        by(elim cScope(4)) (assumption | simp del: freshChainSimps)+
      then  obtain K' where RTrans: "Ψ  ΨP  R K'N  R'"
        and MeqK': "Ψ  ΨP  ΨR  M  K'" and "zvec ♯* K'" and "AR ♯* K'"
        by metis
      from AR ♯* AP AP ♯* (⦇νxR) x  AP extractFrame R = AR, ΨR have "AP ♯* ΨR"
        by(force dest: extractFrameFreshChain)+

      from Ψ  ΨR  P M⦇ν*xvec⦈⟨N  P' have "Ψ  ΨR  P  ROut M (⦇ν*xvecN ≺' P')" by(simp add: residualInject)
      with extractFrame P = AP, ΨP distinct AP x  P zvec ♯* P AP ♯* ΨR x  AP AP ♯* M AP ♯* P AP ♯* zvec  AP ♯* Ψ AP ♯* zvec AR ♯* P AR ♯* AP xvec ♯* M distinct xvec
      obtain K'' where MeqK'': "(Ψ  ΨR)  ΨP  M  K''" and "x  K''" and "AR ♯* K''" and "zvec ♯* K''"
        by(elim outputObtainPrefix[where B="(x#AR@zvec)"]) (assumption | simp | force | metis freshChainSym)+

      from MeqK'' MeqK' have KeqK'': "(Ψ  ΨP)  ΨR  K'  K''"
        by(metis statEqEnt Associativity Composition Commutativity chanEqSym chanEqTrans)
      with RTrans extractFrame R = AR, ΨR distinct AR AR ♯* Ψ AR ♯* ΨP AR ♯* K' AR ♯* K'' AR ♯* R
      have "Ψ  ΨP  R K''N  R'"
        by(elim inputRenameSubject) (assumption | force)+
      then have "Ψ  ΨP  ⦇νxR K''N  ⦇νxR'" using x  Ψ x  ΨP x  K'' x  N
        by(elim Scope) (assumption | force)+
      moreover from MeqK'' have "Ψ  ΨP  ΨR  M  K''"
        by(metis statEqEnt Associativity Composition Commutativity)
      ultimately show ?case using zvec ♯* K'' x  K'' AR ♯* K''
        by force
    next
      case(cBang Ψ' R M' N R' AR ΨR Ψ P AP ΨP AQ zvec xvec M)
      from guarded R extractFrame R = AR, ΨR have "ΨR  𝟭"
        by(metis guardedStatEq)
      with Ψ'  𝟭  M  M' have "Ψ'  ΨR  𝟭  M  M'"
        by(metis Identity Commutativity statEqEnt AssertionStatEqSym Composition)
      moreover have "AQ, Ψ'  ΨR  𝟭 F AP, (Ψ  ΨP)  ΨR  𝟭"
      proof -
        from ΨR  𝟭 have "AQ, Ψ'  ΨR  𝟭 F AQ, Ψ'  𝟭"
          by(metis Identity Commutativity AssertionStatEqSym Composition frameResChainPres frameNilStatEq AssertionStatEqTrans)
        moreover note AQ, Ψ'  𝟭 F AP, (Ψ  ΨP)  𝟭
        moreover  from ΨR  𝟭 have "AP, (Ψ  ΨP)  𝟭 F AP, (Ψ  ΨP)  ΨR  𝟭"
          by(metis Identity Commutativity AssertionStatEqSym Composition frameResChainPres frameNilStatEq AssertionStatEqTrans)
        ultimately show ?thesis by(rule FrameStatEqImpCompose)
      qed
      moreover from Ψ  𝟭  P M⦇ν*xvec⦈⟨N  P' ΨR  𝟭
      have "Ψ  ΨR  𝟭  P M⦇ν*xvec⦈⟨N  P'" by(metis statEqTransition Identity Commutativity AssertionStatEqSym Composition)
      ultimately have "K'. Ψ  ΨP  R  !R K'N  R' Ψ  ΨP  ΨR  𝟭  M  K'  zvec ♯* K'  AR ♯* K'"  using cBang
        by(elim cBang(5)) (assumption | simp del: freshChainSimps)+
      then  obtain K' where RTrans: "Ψ  ΨP  R  !R K'N  R'"
        and MeqK': "Ψ  ΨP  ΨR  𝟭  M  K'" and "zvec ♯* K'" and "AR ♯* K'"
        by metis
      from RTrans guarded R have "Ψ  ΨP  !R K'N  R'" by(rule Bang)
      moreover from MeqK' ΨR  𝟭 have "Ψ  ΨP  𝟭  M  K'"
        by(metis Identity Commutativity statEqEnt AssertionStatEqSym Composition AssertionStatEqTrans)
      ultimately show ?case using zvec ♯* K'
        by force
    qed
  }
  note Goal = this
  have "Ψ  ΨQ  Ψ  ΨQ" by(simp add: AssertionStatEqRefl)
  moreover from AR ♯* Ψ AR ♯* ΨQ have "AR ♯* (Ψ  ΨQ)" by force
  ultimately have "K'. Ψ  ΨP  R K'N  R'  Ψ  ΨP  ΨR  M  K'  ([]::name list) ♯* K'  AR ♯* K'"
    by(intro Goal) (assumption | force)+
  with Assumptions show ?thesis
    by blast
qed

end

end